src/HOL/Tools/inductive_codegen.ML
author bulwahn
Mon, 18 Jul 2011 10:34:21 +0200
changeset 44748 abd1f074cb98
parent 44500 3803869014aa
child 44749 eeb10fdd9535
permissions -rw-r--r--
removing generator registration
haftmann@24584
     1
(*  Title:      HOL/Tools/inductive_codegen.ML
wenzelm@11539
     2
    Author:     Stefan Berghofer, TU Muenchen
berghofe@11537
     3
wenzelm@11539
     4
Code generator for inductive predicates.
berghofe@11537
     5
*)
berghofe@11537
     6
berghofe@11537
     7
signature INDUCTIVE_CODEGEN =
berghofe@11537
     8
sig
wenzelm@43306
     9
  val add: string option -> int option -> attribute
wenzelm@43306
    10
  val poke_test_fn: (int * int * int -> term list option) -> unit
wenzelm@43306
    11
  val test_term: Proof.context -> (term * term list) list -> int list ->
wenzelm@43306
    12
    term list option * Quickcheck.report option
wenzelm@43306
    13
  val setup: theory -> theory
wenzelm@43306
    14
  val quickcheck_setup: theory -> theory
berghofe@11537
    15
end;
berghofe@11537
    16
haftmann@37365
    17
structure Inductive_Codegen : INDUCTIVE_CODEGEN =
berghofe@11537
    18
struct
berghofe@11537
    19
berghofe@12557
    20
(**** theory data ****)
berghofe@11537
    21
wenzelm@18930
    22
fun merge_rules tabs =
haftmann@22642
    23
  Symtab.join (fn _ => AList.merge (Thm.eq_thm_prop) (K true)) tabs;
wenzelm@18930
    24
wenzelm@33522
    25
structure CodegenData = Theory_Data
wenzelm@22846
    26
(
berghofe@15031
    27
  type T =
berghofe@22271
    28
    {intros : (thm * (string * int)) list Symtab.table,
berghofe@15031
    29
     graph : unit Graph.T,
berghofe@16645
    30
     eqns : (thm * string) list Symtab.table};
berghofe@15031
    31
  val empty =
berghofe@15031
    32
    {intros = Symtab.empty, graph = Graph.empty, eqns = Symtab.empty};
wenzelm@16424
    33
  val extend = I;
wenzelm@41720
    34
  fun merge
wenzelm@41720
    35
    ({intros = intros1, graph = graph1, eqns = eqns1},
wenzelm@41720
    36
      {intros = intros2, graph = graph2, eqns = eqns2}) : T =
wenzelm@18930
    37
    {intros = merge_rules (intros1, intros2),
berghofe@15031
    38
     graph = Graph.merge (K true) (graph1, graph2),
wenzelm@18930
    39
     eqns = merge_rules (eqns1, eqns2)};
wenzelm@22846
    40
);
berghofe@12557
    41
berghofe@12557
    42
wenzelm@43307
    43
fun warn thy thm =
wenzelm@43307
    44
  warning ("Inductive_Codegen: Not a proper clause:\n" ^
wenzelm@43307
    45
    Display.string_of_thm_global thy thm);
berghofe@12557
    46
wenzelm@33260
    47
fun add_node x g = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
berghofe@14162
    48
berghofe@22271
    49
fun add optmod optnparms = Thm.declaration_attribute (fn thm => Context.mapping (fn thy =>
berghofe@16645
    50
  let
berghofe@16645
    51
    val {intros, graph, eqns} = CodegenData.get thy;
berghofe@16645
    52
    fun thyname_of s = (case optmod of
haftmann@27398
    53
      NONE => Codegen.thyname_of_const thy s | SOME s => s);
wenzelm@43307
    54
  in
wenzelm@43307
    55
    (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
haftmann@39093
    56
      SOME (Const (@{const_name HOL.eq}, _), [t, _]) =>
wenzelm@35364
    57
        (case head_of t of
wenzelm@35364
    58
          Const (s, _) =>
wenzelm@35364
    59
            CodegenData.put {intros = intros, graph = graph,
wenzelm@35364
    60
               eqns = eqns |> Symtab.map_default (s, [])
wenzelm@35364
    61
                 (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
wenzelm@43307
    62
        | _ => (warn thy thm; thy))
berghofe@22271
    63
    | SOME (Const (s, _), _) =>
berghofe@22271
    64
        let
wenzelm@29287
    65
          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
berghofe@22271
    66
          val rules = Symtab.lookup_list intros s;
wenzelm@41737
    67
          val nparms =
wenzelm@41737
    68
            (case optnparms of
wenzelm@41737
    69
              SOME k => k
wenzelm@41737
    70
            | NONE =>
wenzelm@41737
    71
                (case rules of
wenzelm@41737
    72
                  [] =>
wenzelm@43232
    73
                    (case try (Inductive.the_inductive (Proof_Context.init_global thy)) s of
wenzelm@41737
    74
                      SOME (_, {raw_induct, ...}) =>
wenzelm@41737
    75
                        length (Inductive.params_of raw_induct)
wenzelm@41737
    76
                    | NONE => 0)
wenzelm@41737
    77
                | xs => snd (snd (List.last xs))))
berghofe@22271
    78
        in CodegenData.put
berghofe@22271
    79
          {intros = intros |>
haftmann@22642
    80
           Symtab.update (s, (AList.update Thm.eq_thm_prop
haftmann@22642
    81
             (thm, (thyname_of s, nparms)) rules)),
wenzelm@33345
    82
           graph = fold_rev (Graph.add_edge o pair s) cs (fold add_node (s :: cs) graph),
berghofe@22271
    83
           eqns = eqns} thy
berghofe@22271
    84
        end
wenzelm@43307
    85
    | _ => (warn thy thm; thy))
wenzelm@20897
    86
  end) I);
berghofe@12557
    87
berghofe@12557
    88
fun get_clauses thy s =
wenzelm@43307
    89
  let val {intros, graph, ...} = CodegenData.get thy in
wenzelm@43307
    90
    (case Symtab.lookup intros s of
wenzelm@43307
    91
      NONE =>
wenzelm@43307
    92
        (case try (Inductive.the_inductive (Proof_Context.init_global thy)) s of
wenzelm@43307
    93
          NONE => NONE
wenzelm@43307
    94
        | SOME ({names, ...}, {intrs, raw_induct, ...}) =>
wenzelm@43307
    95
            SOME (names, Codegen.thyname_of_const thy s,
wenzelm@43307
    96
              length (Inductive.params_of raw_induct),
wenzelm@43307
    97
              Codegen.preprocess thy intrs))
skalberg@15531
    98
    | SOME _ =>
berghofe@16645
    99
        let
berghofe@16645
   100
          val SOME names = find_first
haftmann@22642
   101
            (fn xs => member (op =) xs s) (Graph.strong_conn graph);
haftmann@22642
   102
          val intrs as (_, (thyname, nparms)) :: _ =
haftmann@22642
   103
            maps (the o Symtab.lookup intros) names;
wenzelm@43307
   104
        in SOME (names, thyname, nparms, Codegen.preprocess thy (map fst (rev intrs))) end)
berghofe@14162
   105
  end;
berghofe@12557
   106
berghofe@12557
   107
berghofe@11537
   108
(**** check if a term contains only constructor functions ****)
berghofe@11537
   109
berghofe@11537
   110
fun is_constrt thy =
berghofe@11537
   111
  let
haftmann@31784
   112
    val cnstrs = flat (maps
berghofe@11537
   113
      (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
haftmann@33962
   114
      (Symtab.dest (Datatype_Data.get_all thy)));
wenzelm@43307
   115
    fun check t =
wenzelm@43307
   116
      (case strip_comb t of
berghofe@11537
   117
        (Var _, []) => true
wenzelm@43307
   118
      | (Const (s, _), ts) =>
wenzelm@43307
   119
          (case AList.lookup (op =) cnstrs s of
skalberg@15531
   120
            NONE => false
skalberg@15531
   121
          | SOME i => length ts = i andalso forall check ts)
wenzelm@43307
   122
      | _ => false);
berghofe@11537
   123
  in check end;
berghofe@11537
   124
wenzelm@43307
   125
berghofe@11537
   126
(**** check if a type is an equality type (i.e. doesn't contain fun) ****)
berghofe@11537
   127
berghofe@11537
   128
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
berghofe@11537
   129
  | is_eqT _ = true;
berghofe@11537
   130
wenzelm@43307
   131
berghofe@11537
   132
(**** mode inference ****)
berghofe@11537
   133
berghofe@15204
   134
fun string_of_mode (iss, is) = space_implode " -> " (map
skalberg@15531
   135
  (fn NONE => "X"
skalberg@15531
   136
    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
skalberg@15531
   137
       (iss @ [SOME is]));
berghofe@15204
   138
wenzelm@41704
   139
fun print_modes modes = Codegen.message ("Inferred modes:\n" ^
wenzelm@26931
   140
  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
berghofe@33765
   141
    (fn (m, rnd) => string_of_mode m ^
berghofe@33765
   142
       (if rnd then " (random)" else "")) ms)) modes));
berghofe@15204
   143
wenzelm@29265
   144
val term_vs = map (fst o fst o dest_Var) o OldTerm.term_vars;
wenzelm@32952
   145
val terms_vs = distinct (op =) o maps term_vs;
berghofe@11537
   146
berghofe@11537
   147
(** collect all Vars in a term (with duplicates!) **)
wenzelm@16861
   148
fun term_vTs tm =
wenzelm@16861
   149
  fold_aterms (fn Var ((x, _), T) => cons (x, T) | _ => I) tm [];
berghofe@11537
   150
berghofe@11537
   151
fun get_args _ _ [] = ([], [])
haftmann@36677
   152
  | get_args is i (x::xs) = (if member (op =) is i then apfst else apsnd) (cons x)
berghofe@11537
   153
      (get_args is (i+1) xs);
berghofe@11537
   154
berghofe@11537
   155
fun merge xs [] = xs
berghofe@11537
   156
  | merge [] ys = ys
berghofe@11537
   157
  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
berghofe@11537
   158
      else y::merge (x::xs) ys;
berghofe@11537
   159
berghofe@11537
   160
fun subsets i j = if i <= j then
berghofe@11537
   161
       let val is = subsets (i+1) j
berghofe@11537
   162
       in merge (map (fn ks => i::ks) is) is end
berghofe@11537
   163
     else [[]];
berghofe@11537
   164
berghofe@12557
   165
fun cprod ([], ys) = []
berghofe@12557
   166
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
berghofe@12557
   167
wenzelm@30193
   168
fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
berghofe@12557
   169
berghofe@33765
   170
datatype mode = Mode of ((int list option list * int list) * bool) * int list * mode option list;
berghofe@33765
   171
berghofe@33765
   172
fun needs_random (Mode ((_, b), _, ms)) =
berghofe@33765
   173
  b orelse exists (fn NONE => false | SOME m => needs_random m) ms;
berghofe@12557
   174
berghofe@12557
   175
fun modes_of modes t =
berghofe@12557
   176
  let
berghofe@22271
   177
    val ks = 1 upto length (binder_types (fastype_of t));
berghofe@33765
   178
    val default = [Mode ((([], ks), false), ks, [])];
wenzelm@32952
   179
    fun mk_modes name args = Option.map
berghofe@33765
   180
     (maps (fn (m as ((iss, is), _)) =>
berghofe@22271
   181
        let
berghofe@22271
   182
          val (args1, args2) =
berghofe@22271
   183
            if length args < length iss then
berghofe@22271
   184
              error ("Too few arguments for inductive predicate " ^ name)
berghofe@22271
   185
            else chop (length iss) args;
berghofe@22271
   186
          val k = length args2;
berghofe@22271
   187
          val prfx = 1 upto k
berghofe@22271
   188
        in
berghofe@22271
   189
          if not (is_prefix op = prfx is) then [] else
berghofe@22271
   190
          let val is' = map (fn i => i - k) (List.drop (is, k))
berghofe@22271
   191
          in map (fn x => Mode (m, is', x)) (cprods (map
berghofe@22271
   192
            (fn (NONE, _) => [NONE]
wenzelm@33325
   193
              | (SOME js, arg) => map SOME (filter
berghofe@22271
   194
                  (fn Mode (_, js', _) => js=js') (modes_of modes arg)))
berghofe@22271
   195
                    (iss ~~ args1)))
berghofe@22271
   196
          end
berghofe@22271
   197
        end)) (AList.lookup op = modes name)
berghofe@12557
   198
wenzelm@43307
   199
  in
wenzelm@43307
   200
    (case strip_comb t of
haftmann@39093
   201
      (Const (@{const_name HOL.eq}, Type (_, [T, _])), _) =>
berghofe@33765
   202
        [Mode ((([], [1]), false), [1], []), Mode ((([], [2]), false), [2], [])] @
berghofe@33765
   203
        (if is_eqT T then [Mode ((([], [1, 2]), false), [1, 2], [])] else [])
berghofe@22271
   204
    | (Const (name, _), args) => the_default default (mk_modes name args)
berghofe@22271
   205
    | (Var ((name, _), _), args) => the (mk_modes name args)
berghofe@22271
   206
    | (Free (name, _), args) => the (mk_modes name args)
berghofe@22271
   207
    | _ => default)
berghofe@12557
   208
  end;
berghofe@12557
   209
berghofe@26806
   210
datatype indprem = Prem of term list * term * bool | Sidecond of term;
berghofe@12557
   211
berghofe@33765
   212
fun missing_vars vs ts = subtract (fn (x, ((y, _), _)) => x = y) vs
berghofe@33765
   213
  (fold Term.add_vars ts []);
berghofe@33765
   214
berghofe@33765
   215
fun monomorphic_vars vs = null (fold (Term.add_tvarsT o snd) vs []);
berghofe@33765
   216
berghofe@33765
   217
fun mode_ord p = int_ord (pairself (fn (Mode ((_, rnd), _, _), vs) =>
berghofe@33765
   218
  length vs + (if null vs then 0 else 1) + (if rnd then 1 else 0)) p);
berghofe@33765
   219
berghofe@11537
   220
fun select_mode_prem thy modes vs ps =
berghofe@33765
   221
  sort (mode_ord o pairself (hd o snd))
berghofe@33765
   222
    (filter_out (null o snd) (ps ~~ map
berghofe@33765
   223
      (fn Prem (us, t, is_set) => sort mode_ord
wenzelm@43308
   224
          (map_filter (fn m as Mode (_, is, _) =>
berghofe@33765
   225
            let
berghofe@33765
   226
              val (in_ts, out_ts) = get_args is 1 us;
berghofe@33765
   227
              val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
berghofe@33765
   228
              val vTs = maps term_vTs out_ts';
berghofe@33765
   229
              val dupTs = map snd (duplicates (op =) vTs) @
berghofe@33765
   230
                map_filter (AList.lookup (op =) vTs) vs;
berghofe@33765
   231
              val missing_vs = missing_vars vs (t :: in_ts @ in_ts')
berghofe@33765
   232
            in
berghofe@33765
   233
              if forall (is_eqT o fastype_of) in_ts' andalso forall is_eqT dupTs
berghofe@33765
   234
                andalso monomorphic_vars missing_vs
berghofe@33765
   235
              then SOME (m, missing_vs)
berghofe@33765
   236
              else NONE
berghofe@33765
   237
            end)
berghofe@33765
   238
              (if is_set then [Mode ((([], []), false), [], [])]
berghofe@33765
   239
               else modes_of modes t handle Option =>
berghofe@33765
   240
                 error ("Bad predicate: " ^ Syntax.string_of_term_global thy t)))
berghofe@33765
   241
        | Sidecond t =>
berghofe@33765
   242
            let val missing_vs = missing_vars vs [t]
berghofe@33765
   243
            in
berghofe@33765
   244
              if monomorphic_vars missing_vs
berghofe@33765
   245
              then [(Mode ((([], []), false), [], []), missing_vs)]
berghofe@33765
   246
              else []
berghofe@33765
   247
            end)
berghofe@33765
   248
              ps));
berghofe@11537
   249
wenzelm@43292
   250
fun use_random codegen_mode = member (op =) codegen_mode "random_ind";
berghofe@33765
   251
wenzelm@43292
   252
fun check_mode_clause thy codegen_mode arg_vs modes ((iss, is), rnd) (ts, ps) =
berghofe@11537
   253
  let
wenzelm@32952
   254
    val modes' = modes @ map_filter
berghofe@33765
   255
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [(([], js), false)]))
berghofe@12557
   256
        (arg_vs ~~ iss);
berghofe@33765
   257
    fun check_mode_prems vs rnd [] = SOME (vs, rnd)
wenzelm@43307
   258
      | check_mode_prems vs rnd ps =
wenzelm@43307
   259
          (case select_mode_prem thy modes' vs ps of
wenzelm@43307
   260
            (x, (m, []) :: _) :: _ =>
wenzelm@43307
   261
              check_mode_prems
wenzelm@43307
   262
                (case x of Prem (us, _, _) => union (op =) vs (terms_vs us) | _ => vs)
wenzelm@43307
   263
                (rnd orelse needs_random m)
wenzelm@43307
   264
                (filter_out (equal x) ps)
wenzelm@43307
   265
          | (_, (_, vs') :: _) :: _ =>
wenzelm@43307
   266
              if use_random codegen_mode then
wenzelm@43307
   267
                check_mode_prems (union (op =) vs (map (fst o fst) vs')) true ps
wenzelm@43307
   268
              else NONE
wenzelm@43307
   269
          | _ => NONE);
skalberg@15570
   270
    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is 1 ts));
berghofe@11537
   271
    val in_vs = terms_vs in_ts;
berghofe@11537
   272
  in
berghofe@33765
   273
    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
berghofe@33765
   274
      forall (is_eqT o fastype_of) in_ts'
wenzelm@43307
   275
    then
wenzelm@43307
   276
      (case check_mode_prems (union (op =) arg_vs in_vs) rnd ps of
wenzelm@43307
   277
        NONE => NONE
wenzelm@43307
   278
      | SOME (vs, rnd') =>
wenzelm@43307
   279
          let val missing_vs = missing_vars vs ts
wenzelm@43307
   280
          in
wenzelm@43307
   281
            if null missing_vs orelse
wenzelm@43307
   282
              use_random codegen_mode andalso monomorphic_vars missing_vs
wenzelm@43307
   283
            then SOME (rnd' orelse not (null missing_vs))
wenzelm@43307
   284
            else NONE
wenzelm@43307
   285
          end)
berghofe@33765
   286
    else NONE
berghofe@11537
   287
  end;
berghofe@11537
   288
wenzelm@43292
   289
fun check_modes_pred thy codegen_mode arg_vs preds modes (p, ms) =
wenzelm@43307
   290
  let val SOME rs = AList.lookup (op =) preds p in
wenzelm@43308
   291
    (p, map_filter (fn m as (m', _) =>
wenzelm@43307
   292
      let val xs = map (check_mode_clause thy codegen_mode arg_vs modes m) rs in
wenzelm@43307
   293
        (case find_index is_none xs of
wenzelm@43307
   294
          ~1 => SOME (m', exists (fn SOME b => b) xs)
wenzelm@43307
   295
        | i => (Codegen.message ("Clause " ^ string_of_int (i+1) ^ " of " ^
wenzelm@43307
   296
          p ^ " violates mode " ^ string_of_mode m'); NONE))
wenzelm@43307
   297
      end) ms)
berghofe@15204
   298
  end;
berghofe@11537
   299
berghofe@33765
   300
fun fixp f (x : (string * ((int list option list * int list) * bool) list) list) =
berghofe@11537
   301
  let val y = f x
berghofe@11537
   302
  in if x = y then x else fixp f y end;
berghofe@11537
   303
wenzelm@43292
   304
fun infer_modes thy codegen_mode extra_modes arities arg_vs preds = fixp (fn modes =>
wenzelm@43292
   305
  map (check_modes_pred thy codegen_mode arg_vs preds (modes @ extra_modes)) modes)
berghofe@33765
   306
    (map (fn (s, (ks, k)) => (s, map (rpair false) (cprod (cprods (map
skalberg@15531
   307
      (fn NONE => [NONE]
berghofe@22271
   308
        | SOME k' => map SOME (subsets 1 k')) ks),
berghofe@33765
   309
      subsets 1 k)))) arities);
berghofe@11537
   310
wenzelm@43307
   311
berghofe@11537
   312
(**** code generation ****)
berghofe@11537
   313
berghofe@11537
   314
fun mk_eq (x::xs) =
wenzelm@43307
   315
  let
wenzelm@43307
   316
    fun mk_eqs _ [] = []
wenzelm@43307
   317
      | mk_eqs a (b :: cs) = Codegen.str (a ^ " = " ^ b) :: mk_eqs b cs;
berghofe@11537
   318
  in mk_eqs x xs end;
berghofe@11537
   319
wenzelm@43307
   320
fun mk_tuple xs =
wenzelm@43307
   321
  Pretty.block (Codegen.str "(" ::
wenzelm@43307
   322
    flat (separate [Codegen.str ",", Pretty.brk 1] (map single xs)) @
wenzelm@43307
   323
    [Codegen.str ")"]);
berghofe@11537
   324
wenzelm@33260
   325
fun mk_v s (names, vs) =
wenzelm@33260
   326
  (case AList.lookup (op =) vs s of
wenzelm@33260
   327
    NONE => (s, (names, (s, [s])::vs))
wenzelm@33260
   328
  | SOME xs =>
wenzelm@44206
   329
      let val s' = singleton (Name.variant_list names) s
wenzelm@33260
   330
      in (s', (s'::names, AList.update (op =) (s, s'::xs) vs)) end);
berghofe@11537
   331
wenzelm@33260
   332
fun distinct_v (Var ((s, 0), T)) nvs =
wenzelm@33260
   333
      let val (s', nvs') = mk_v s nvs
wenzelm@33260
   334
      in (Var ((s', 0), T), nvs') end
wenzelm@33260
   335
  | distinct_v (t $ u) nvs =
berghofe@11537
   336
      let
wenzelm@33260
   337
        val (t', nvs') = distinct_v t nvs;
wenzelm@33260
   338
        val (u', nvs'') = distinct_v u nvs';
wenzelm@33260
   339
      in (t' $ u', nvs'') end
wenzelm@33260
   340
  | distinct_v t nvs = (t, nvs);
berghofe@11537
   341
berghofe@15061
   342
fun is_exhaustive (Var _) = true
haftmann@37365
   343
  | is_exhaustive (Const (@{const_name Pair}, _) $ t $ u) =
berghofe@15061
   344
      is_exhaustive t andalso is_exhaustive u
berghofe@15061
   345
  | is_exhaustive _ = false;
berghofe@15061
   346
berghofe@15061
   347
fun compile_match nvs eq_ps out_ps success_p can_fail =
wenzelm@41704
   348
  let val eqs = flat (separate [Codegen.str " andalso", Pretty.brk 1]
wenzelm@32952
   349
    (map single (maps (mk_eq o snd) nvs @ eq_ps)));
berghofe@11537
   350
  in
berghofe@11537
   351
    Pretty.block
wenzelm@41704
   352
     ([Codegen.str "(fn ", mk_tuple out_ps, Codegen.str " =>", Pretty.brk 1] @
wenzelm@41704
   353
      (Pretty.block ((if null eqs then [] else Codegen.str "if " ::
wenzelm@41704
   354
         [Pretty.block eqs, Pretty.brk 1, Codegen.str "then "]) @
berghofe@11537
   355
         (success_p ::
wenzelm@41704
   356
          (if null eqs then [] else [Pretty.brk 1, Codegen.str "else DSeq.empty"]))) ::
berghofe@15061
   357
       (if can_fail then
wenzelm@41704
   358
          [Pretty.brk 1, Codegen.str "| _ => DSeq.empty)"]
wenzelm@41704
   359
        else [Codegen.str ")"])))
berghofe@11537
   360
  end;
berghofe@11537
   361
berghofe@17144
   362
fun modename module s (iss, is) gr =
haftmann@39093
   363
  let val (id, gr') = if s = @{const_name HOL.eq} then (("", "equal"), gr)
wenzelm@41704
   364
    else Codegen.mk_const_id module s gr
haftmann@28537
   365
  in (space_implode "__"
wenzelm@41704
   366
    (Codegen.mk_qual_id module id ::
wenzelm@32952
   367
      map (space_implode "_" o map string_of_int) (map_filter I iss @ [is])), gr')
berghofe@17144
   368
  end;
berghofe@16645
   369
wenzelm@41704
   370
fun mk_funcomp brack s k p = (if brack then Codegen.parens else I)
wenzelm@41704
   371
  (Pretty.block [Pretty.block ((if k = 0 then [] else [Codegen.str "("]) @
wenzelm@41704
   372
    separate (Pretty.brk 1) (Codegen.str s :: replicate k (Codegen.str "|> ???")) @
wenzelm@41704
   373
    (if k = 0 then [] else [Codegen.str ")"])), Pretty.brk 1, p]);
berghofe@22271
   374
wenzelm@43292
   375
fun compile_expr thy codegen_mode defs dep module brack modes (NONE, t) gr =
wenzelm@43292
   376
      apfst single (Codegen.invoke_codegen thy codegen_mode defs dep module brack t gr)
wenzelm@43292
   377
  | compile_expr _ _ _ _ _ _ _ (SOME _, Var ((name, _), _)) gr =
wenzelm@41704
   378
      ([Codegen.str name], gr)
wenzelm@43307
   379
  | compile_expr thy codegen_mode
wenzelm@43307
   380
        defs dep module brack modes (SOME (Mode ((mode, _), _, ms)), t) gr =
berghofe@22271
   381
      (case strip_comb t of
wenzelm@43307
   382
        (Const (name, _), args) =>
wenzelm@43307
   383
          if name = @{const_name HOL.eq} orelse AList.defined op = modes name then
wenzelm@43307
   384
            let
wenzelm@43307
   385
              val (args1, args2) = chop (length ms) args;
wenzelm@43307
   386
              val ((ps, mode_id), gr') =
wenzelm@43307
   387
                gr |> fold_map
wenzelm@43307
   388
                  (compile_expr thy codegen_mode defs dep module true modes) (ms ~~ args1)
wenzelm@43307
   389
                ||>> modename module name mode;
wenzelm@43307
   390
               val (ps', gr'') =
wenzelm@43307
   391
                (case mode of
wenzelm@41704
   392
                   ([], []) => ([Codegen.str "()"], gr')
haftmann@28537
   393
                 | _ => fold_map
wenzelm@43307
   394
                     (Codegen.invoke_codegen thy codegen_mode defs dep module true) args2 gr');
wenzelm@43307
   395
             in
wenzelm@43307
   396
              ((if brack andalso not (null ps andalso null ps') then
wenzelm@43307
   397
                single o Codegen.parens o Pretty.block else I)
wenzelm@43307
   398
                  (flat (separate [Pretty.brk 1]
wenzelm@43307
   399
                    ([Codegen.str mode_id] :: ps @ map single ps'))), gr')
berghofe@22271
   400
             end
wenzelm@43307
   401
          else
wenzelm@43307
   402
            apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
wenzelm@43307
   403
              (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr)
wenzelm@43307
   404
      | _ =>
wenzelm@43307
   405
        apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
wenzelm@43307
   406
          (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr));
berghofe@12557
   407
wenzelm@43292
   408
fun compile_clause thy codegen_mode defs dep module all_vs arg_vs modes (iss, is) (ts, ps) inp gr =
berghofe@11537
   409
  let
wenzelm@32952
   410
    val modes' = modes @ map_filter
berghofe@33765
   411
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [(([], js), false)]))
berghofe@12557
   412
        (arg_vs ~~ iss);
berghofe@12557
   413
wenzelm@33260
   414
    fun check_constrt t (names, eqs) =
wenzelm@33260
   415
      if is_constrt thy t then (t, (names, eqs))
wenzelm@33260
   416
      else
wenzelm@44206
   417
        let val s = singleton (Name.variant_list names) "x";
wenzelm@33260
   418
        in (Var ((s, 0), fastype_of t), (s::names, (s, t)::eqs)) end;
berghofe@11537
   419
haftmann@28537
   420
    fun compile_eq (s, t) gr =
wenzelm@41704
   421
      apfst (Pretty.block o cons (Codegen.str (s ^ " = ")) o single)
wenzelm@43292
   422
        (Codegen.invoke_codegen thy codegen_mode defs dep module false t gr);
berghofe@15031
   423
berghofe@12557
   424
    val (in_ts, out_ts) = get_args is 1 ts;
wenzelm@33260
   425
    val (in_ts', (all_vs', eqs)) = fold_map check_constrt in_ts (all_vs, []);
berghofe@11537
   426
haftmann@28537
   427
    fun compile_prems out_ts' vs names [] gr =
berghofe@11537
   428
          let
wenzelm@33260
   429
            val (out_ps, gr2) =
wenzelm@43307
   430
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
wenzelm@43307
   431
                out_ts gr;
haftmann@28537
   432
            val (eq_ps, gr3) = fold_map compile_eq eqs gr2;
wenzelm@33260
   433
            val (out_ts'', (names', eqs')) = fold_map check_constrt out_ts' (names, []);
wenzelm@33260
   434
            val (out_ts''', nvs) =
wenzelm@33260
   435
              fold_map distinct_v out_ts'' (names', map (fn x => (x, [x])) vs);
wenzelm@33260
   436
            val (out_ps', gr4) =
wenzelm@43307
   437
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
wenzelm@43307
   438
                out_ts''' gr3;
haftmann@28537
   439
            val (eq_ps', gr5) = fold_map compile_eq eqs' gr4;
berghofe@33765
   440
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts'));
berghofe@33765
   441
            val missing_vs = missing_vars vs' out_ts;
berghofe@33765
   442
            val final_p = Pretty.block
wenzelm@41704
   443
              [Codegen.str "DSeq.single", Pretty.brk 1, mk_tuple out_ps]
berghofe@11537
   444
          in
berghofe@33765
   445
            if null missing_vs then
berghofe@33765
   446
              (compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'
berghofe@33765
   447
                 final_p (exists (not o is_exhaustive) out_ts'''), gr5)
berghofe@33765
   448
            else
berghofe@33765
   449
              let
wenzelm@43307
   450
                val (pat_p, gr6) =
wenzelm@43307
   451
                  Codegen.invoke_codegen thy codegen_mode defs dep module true
wenzelm@43307
   452
                    (HOLogic.mk_tuple (map Var missing_vs)) gr5;
wenzelm@41704
   453
                val gen_p =
wenzelm@41704
   454
                  Codegen.mk_gen gr6 module true [] ""
wenzelm@43307
   455
                    (HOLogic.mk_tupleT (map snd missing_vs));
berghofe@33765
   456
              in
berghofe@33765
   457
                (compile_match (snd nvs) eq_ps' out_ps'
wenzelm@43307
   458
                  (Pretty.block [Codegen.str "DSeq.generator ", gen_p,
wenzelm@43307
   459
                    Codegen.str " :->", Pretty.brk 1,
wenzelm@43307
   460
                    compile_match [] eq_ps [pat_p] final_p false])
wenzelm@43307
   461
                  (exists (not o is_exhaustive) out_ts'''),
berghofe@33765
   462
                 gr6)
berghofe@33765
   463
              end
berghofe@11537
   464
          end
haftmann@28537
   465
      | compile_prems out_ts vs names ps gr =
berghofe@11537
   466
          let
haftmann@31852
   467
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
wenzelm@33260
   468
            val (out_ts', (names', eqs)) = fold_map check_constrt out_ts (names, []);
wenzelm@43307
   469
            val (out_ts'', nvs) =
wenzelm@43307
   470
              fold_map distinct_v out_ts' (names', map (fn x => (x, [x])) vs);
wenzelm@41704
   471
            val (out_ps, gr0) =
wenzelm@43307
   472
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module false)
wenzelm@43307
   473
                out_ts'' gr;
haftmann@28537
   474
            val (eq_ps, gr1) = fold_map compile_eq eqs gr0;
berghofe@11537
   475
          in
berghofe@33765
   476
            (case hd (select_mode_prem thy modes' vs' ps) of
wenzelm@43307
   477
              (p as Prem (us, t, is_set), (mode as Mode (_, js, _), []) :: _) =>
wenzelm@43307
   478
                let
wenzelm@43307
   479
                  val ps' = filter_out (equal p) ps;
wenzelm@43307
   480
                  val (in_ts, out_ts''') = get_args js 1 us;
wenzelm@43307
   481
                  val (in_ps, gr2) =
wenzelm@43307
   482
                    fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
wenzelm@43307
   483
                      in_ts gr1;
wenzelm@43307
   484
                  val (ps, gr3) =
wenzelm@43307
   485
                    if not is_set then
wenzelm@43307
   486
                      apfst (fn ps => ps @
wenzelm@43307
   487
                          (if null in_ps then [] else [Pretty.brk 1]) @
wenzelm@43307
   488
                          separate (Pretty.brk 1) in_ps)
wenzelm@43307
   489
                        (compile_expr thy codegen_mode defs dep module false modes
wenzelm@43307
   490
                          (SOME mode, t) gr2)
wenzelm@43307
   491
                    else
wenzelm@43307
   492
                      apfst (fn p =>
wenzelm@43307
   493
                        Pretty.breaks [Codegen.str "DSeq.of_list", Codegen.str "(case", p,
wenzelm@43307
   494
                        Codegen.str "of", Codegen.str "Set", Codegen.str "xs", Codegen.str "=>",
wenzelm@43307
   495
                        Codegen.str "xs)"])
wenzelm@43307
   496
                        (*this is a very strong assumption about the generated code!*)
wenzelm@43307
   497
                        (Codegen.invoke_codegen thy codegen_mode defs dep module true t gr2);
haftmann@28537
   498
                   val (rest, gr4) = compile_prems out_ts''' vs' (fst nvs) ps' gr3;
berghofe@11537
   499
                 in
haftmann@28537
   500
                   (compile_match (snd nvs) eq_ps out_ps
wenzelm@43307
   501
                     (Pretty.block (ps @
wenzelm@43307
   502
                       [Codegen.str " :->", Pretty.brk 1, rest]))
wenzelm@43307
   503
                       (exists (not o is_exhaustive) out_ts''), gr4)
berghofe@11537
   504
                 end
wenzelm@43307
   505
            | (p as Sidecond t, [(_, [])]) =>
wenzelm@43307
   506
                let
wenzelm@43307
   507
                  val ps' = filter_out (equal p) ps;
wenzelm@43307
   508
                  val (side_p, gr2) =
wenzelm@43292
   509
                    Codegen.invoke_codegen thy codegen_mode defs dep module true t gr1;
wenzelm@43307
   510
                  val (rest, gr3) = compile_prems [] vs' (fst nvs) ps' gr2;
wenzelm@43307
   511
                in
wenzelm@43307
   512
                  (compile_match (snd nvs) eq_ps out_ps
wenzelm@43307
   513
                    (Pretty.block [Codegen.str "?? ", side_p,
wenzelm@43307
   514
                      Codegen.str " :->", Pretty.brk 1, rest])
wenzelm@43307
   515
                    (exists (not o is_exhaustive) out_ts''), gr3)
wenzelm@43307
   516
                end
wenzelm@43307
   517
            | (_, (_, missing_vs) :: _) =>
wenzelm@43307
   518
                let
wenzelm@43307
   519
                  val T = HOLogic.mk_tupleT (map snd missing_vs);
wenzelm@43307
   520
                  val (_, gr2) =
wenzelm@43292
   521
                    Codegen.invoke_tycodegen thy codegen_mode defs dep module false T gr1;
wenzelm@43307
   522
                  val gen_p = Codegen.mk_gen gr2 module true [] "" T;
wenzelm@43307
   523
                  val (rest, gr3) = compile_prems
wenzelm@43307
   524
                    [HOLogic.mk_tuple (map Var missing_vs)] vs' (fst nvs) ps gr2;
wenzelm@43307
   525
                in
wenzelm@43307
   526
                  (compile_match (snd nvs) eq_ps out_ps
wenzelm@43307
   527
                    (Pretty.block [Codegen.str "DSeq.generator", Pretty.brk 1,
wenzelm@43307
   528
                      gen_p, Codegen.str " :->", Pretty.brk 1, rest])
wenzelm@43307
   529
                    (exists (not o is_exhaustive) out_ts''), gr3)
wenzelm@43307
   530
                end)
berghofe@11537
   531
          end;
berghofe@11537
   532
haftmann@28537
   533
    val (prem_p, gr') = compile_prems in_ts' arg_vs all_vs' ps gr ;
berghofe@11537
   534
  in
wenzelm@41704
   535
    (Pretty.block [Codegen.str "DSeq.single", Pretty.brk 1, inp,
wenzelm@41704
   536
       Codegen.str " :->", Pretty.brk 1, prem_p], gr')
berghofe@11537
   537
  end;
berghofe@11537
   538
wenzelm@43292
   539
fun compile_pred thy codegen_mode defs dep module prfx all_vs arg_vs modes s cls mode gr =
berghofe@22271
   540
  let
wenzelm@41704
   541
    val xs = map Codegen.str (Name.variant_list arg_vs
berghofe@22271
   542
      (map (fn i => "x" ^ string_of_int i) (snd mode)));
haftmann@28537
   543
    val ((cl_ps, mode_id), gr') = gr |>
wenzelm@43292
   544
      fold_map (fn cl => compile_clause thy codegen_mode defs
haftmann@28537
   545
        dep module all_vs arg_vs modes mode cl (mk_tuple xs)) cls ||>>
berghofe@22271
   546
      modename module s mode
berghofe@11537
   547
  in
haftmann@28537
   548
    (Pretty.block
berghofe@11537
   549
      ([Pretty.block (separate (Pretty.brk 1)
wenzelm@41704
   550
         (Codegen.str (prfx ^ mode_id) ::
wenzelm@41704
   551
           map Codegen.str arg_vs @
wenzelm@41704
   552
           (case mode of ([], []) => [Codegen.str "()"] | _ => xs)) @
wenzelm@41704
   553
         [Codegen.str " ="]),
berghofe@11537
   554
        Pretty.brk 1] @
wenzelm@41704
   555
       flat (separate [Codegen.str " ++", Pretty.brk 1] (map single cl_ps))), (gr', "and "))
berghofe@11537
   556
  end;
berghofe@11537
   557
wenzelm@43292
   558
fun compile_preds thy codegen_mode defs dep module all_vs arg_vs modes preds gr =
haftmann@28537
   559
  let val (prs, (gr', _)) = fold_map (fn (s, cls) =>
wenzelm@43292
   560
    fold_map (fn (mode, _) => fn (gr', prfx') => compile_pred thy codegen_mode defs
haftmann@28537
   561
      dep module prfx' all_vs arg_vs modes s cls mode gr')
haftmann@28537
   562
        (((the o AList.lookup (op =) modes) s))) preds (gr, "fun ")
berghofe@11537
   563
  in
wenzelm@41704
   564
    (space_implode "\n\n" (map Codegen.string_of (flat prs)) ^ ";\n\n", gr')
berghofe@11537
   565
  end;
berghofe@11537
   566
berghofe@11537
   567
(**** processing of introduction rules ****)
berghofe@11537
   568
berghofe@12557
   569
exception Modes of
berghofe@33765
   570
  (string * ((int list option list * int list) * bool) list) list *
berghofe@22271
   571
  (string * (int option list * int)) list;
berghofe@12557
   572
wenzelm@32952
   573
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip
wenzelm@41704
   574
  (map ((fn (SOME (Modes x), _, _) => x | _ => ([], [])) o Codegen.get_node gr)
berghofe@17144
   575
    (Graph.all_preds (fst gr) [dep]))));
berghofe@12557
   576
wenzelm@41704
   577
fun print_arities arities = Codegen.message ("Arities:\n" ^
wenzelm@26931
   578
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
berghofe@12557
   579
    space_implode " -> " (map
berghofe@22271
   580
      (fn NONE => "X" | SOME k' => string_of_int k')
berghofe@22271
   581
        (ks @ [SOME k]))) arities));
berghofe@11537
   582
wenzelm@41704
   583
fun prep_intrs intrs =
wenzelm@41704
   584
  map (Codegen.rename_term o #prop o rep_thm o Drule.export_without_context) intrs;
berghofe@15031
   585
berghofe@15031
   586
fun constrain cs [] = []
wenzelm@33260
   587
  | constrain cs ((s, xs) :: ys) =
wenzelm@33260
   588
      (s,
wenzelm@43307
   589
        (case AList.lookup (op =) cs (s : string) of
wenzelm@33260
   590
          NONE => xs
wenzelm@43307
   591
        | SOME xs' => inter (op = o apfst fst) xs' xs)) :: constrain cs ys;
berghofe@15031
   592
wenzelm@43292
   593
fun mk_extra_defs thy codegen_mode defs gr dep names module ts =
wenzelm@33260
   594
  fold (fn name => fn gr =>
haftmann@36677
   595
    if member (op =) names name then gr
wenzelm@33260
   596
    else
wenzelm@33260
   597
      (case get_clauses thy name of
skalberg@15531
   598
        NONE => gr
berghofe@22271
   599
      | SOME (names, thyname, nparms, intrs) =>
wenzelm@43292
   600
          mk_ind_def thy codegen_mode defs gr dep names
wenzelm@43292
   601
            (Codegen.if_library codegen_mode thyname module)
berghofe@22271
   602
            [] (prep_intrs intrs) nparms))
wenzelm@33260
   603
    (fold Term.add_const_names ts []) gr
berghofe@12557
   604
wenzelm@43292
   605
and mk_ind_def thy codegen_mode defs gr dep names module modecs intrs nparms =
wenzelm@41704
   606
  Codegen.add_edge_acyclic (hd names, dep) gr handle
berghofe@24129
   607
    Graph.CYCLES (xs :: _) =>
haftmann@37365
   608
      error ("Inductive_Codegen: illegal cyclic dependencies:\n" ^ commas xs)
berghofe@24129
   609
  | Graph.UNDEF _ =>
berghofe@11537
   610
    let
berghofe@22271
   611
      val _ $ u = Logic.strip_imp_concl (hd intrs);
berghofe@22271
   612
      val args = List.take (snd (strip_comb u), nparms);
wenzelm@32952
   613
      val arg_vs = maps term_vs args;
berghofe@15031
   614
haftmann@36677
   615
      fun get_nparms s = if member (op =) names s then SOME nparms else
berghofe@22271
   616
        Option.map #3 (get_clauses thy s);
berghofe@11537
   617
haftmann@37677
   618
      fun dest_prem (_ $ (Const (@{const_name Set.member}, _) $ t $ u)) =
wenzelm@35364
   619
            Prem ([t], Envir.beta_eta_contract u, true)
haftmann@39093
   620
        | dest_prem (_ $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u)) =
wenzelm@35364
   621
            Prem ([t, u], eq, false)
berghofe@22271
   622
        | dest_prem (_ $ t) =
berghofe@22271
   623
            (case strip_comb t of
wenzelm@43307
   624
              (v as Var _, ts) =>
wenzelm@43307
   625
                if member (op =) args v then Prem (ts, v, false) else Sidecond t
wenzelm@35364
   626
            | (c as Const (s, _), ts) =>
wenzelm@35364
   627
                (case get_nparms s of
wenzelm@35364
   628
                  NONE => Sidecond t
wenzelm@35364
   629
                | SOME k =>
wenzelm@35364
   630
                    let val (ts1, ts2) = chop k ts
wenzelm@35364
   631
                    in Prem (ts2, list_comb (c, ts1), false) end)
wenzelm@35364
   632
            | _ => Sidecond t);
berghofe@22271
   633
berghofe@22271
   634
      fun add_clause intr (clauses, arities) =
berghofe@11537
   635
        let
berghofe@22271
   636
          val _ $ t = Logic.strip_imp_concl intr;
berghofe@22271
   637
          val (Const (name, T), ts) = strip_comb t;
berghofe@22271
   638
          val (ts1, ts2) = chop nparms ts;
berghofe@22271
   639
          val prems = map dest_prem (Logic.strip_imp_prems intr);
berghofe@22271
   640
          val (Ts, Us) = chop nparms (binder_types T)
berghofe@11537
   641
        in
berghofe@22271
   642
          (AList.update op = (name, these (AList.lookup op = clauses name) @
berghofe@22271
   643
             [(ts2, prems)]) clauses,
wenzelm@43307
   644
           AList.update op = (name, (map (fn U =>
wenzelm@43307
   645
              (case strip_type U of
wenzelm@43307
   646
                (Rs as _ :: _, @{typ bool}) => SOME (length Rs)
wenzelm@43307
   647
              | _ => NONE)) Ts,
berghofe@22271
   648
             length Us)) arities)
berghofe@11537
   649
        end;
berghofe@11537
   650
wenzelm@43292
   651
      val gr' = mk_extra_defs thy codegen_mode defs
wenzelm@41704
   652
        (Codegen.add_edge (hd names, dep)
wenzelm@41704
   653
          (Codegen.new_node (hd names, (NONE, "", "")) gr)) (hd names) names module intrs;
berghofe@22271
   654
      val (extra_modes, extra_arities) = lookup_modes gr' (hd names);
berghofe@22271
   655
      val (clauses, arities) = fold add_clause intrs ([], []);
berghofe@15031
   656
      val modes = constrain modecs
wenzelm@43292
   657
        (infer_modes thy codegen_mode extra_modes arities arg_vs clauses);
berghofe@22271
   658
      val _ = print_arities arities;
berghofe@11537
   659
      val _ = print_modes modes;
wenzelm@43307
   660
      val (s, gr'') =
wenzelm@43307
   661
        compile_preds thy codegen_mode defs (hd names) module (terms_vs intrs)
wenzelm@43307
   662
          arg_vs (modes @ extra_modes) clauses gr';
berghofe@11537
   663
    in
wenzelm@41704
   664
      (Codegen.map_node (hd names)
berghofe@22271
   665
        (K (SOME (Modes (modes, arities)), module, s)) gr'')
berghofe@16645
   666
    end;
berghofe@11537
   667
wenzelm@43307
   668
fun find_mode gr dep s u modes is =
wenzelm@43307
   669
  (case find_first (fn Mode (_, js, _) => is = js) (modes_of modes u handle Option => []) of
wenzelm@43307
   670
    NONE =>
wenzelm@43307
   671
      Codegen.codegen_error gr dep
wenzelm@43307
   672
        ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
wenzelm@43307
   673
  | mode => mode);
berghofe@15031
   674
wenzelm@43292
   675
fun mk_ind_call thy codegen_mode defs dep module is_query s T ts names thyname k intrs gr =
berghofe@22271
   676
  let
berghofe@22271
   677
    val (ts1, ts2) = chop k ts;
berghofe@22271
   678
    val u = list_comb (Const (s, T), ts1);
berghofe@11537
   679
wenzelm@43307
   680
    fun mk_mode (Const (@{const_name dummy_pattern}, _)) ((ts, mode), i) =
wenzelm@43307
   681
          ((ts, mode), i + 1)
wenzelm@33260
   682
      | mk_mode t ((ts, mode), i) = ((ts @ [t], mode @ [i]), i + 1);
berghofe@11537
   683
wenzelm@43292
   684
    val module' = Codegen.if_library codegen_mode thyname module;
wenzelm@43307
   685
    val gr1 =
wenzelm@43307
   686
      mk_extra_defs thy codegen_mode defs
wenzelm@43307
   687
        (mk_ind_def thy codegen_mode defs gr dep names module'
wenzelm@43307
   688
        [] (prep_intrs intrs) k) dep names module' [u];
berghofe@22271
   689
    val (modes, _) = lookup_modes gr1 dep;
wenzelm@33260
   690
    val (ts', is) =
wenzelm@33260
   691
      if is_query then fst (fold mk_mode ts2 (([], []), 1))
berghofe@22271
   692
      else (ts2, 1 upto length (binder_types T) - k);
berghofe@22271
   693
    val mode = find_mode gr1 dep s u modes is;
berghofe@33765
   694
    val _ = if is_query orelse not (needs_random (the mode)) then ()
berghofe@33765
   695
      else warning ("Illegal use of random data generators in " ^ s);
wenzelm@43292
   696
    val (in_ps, gr2) =
wenzelm@43307
   697
      fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
wenzelm@43307
   698
        ts' gr1;
wenzelm@43307
   699
    val (ps, gr3) =
wenzelm@43307
   700
      compile_expr thy codegen_mode defs dep module false modes (mode, u) gr2;
berghofe@22271
   701
  in
haftmann@28537
   702
    (Pretty.block (ps @ (if null in_ps then [] else [Pretty.brk 1]) @
haftmann@28537
   703
       separate (Pretty.brk 1) in_ps), gr3)
berghofe@22271
   704
  end;
berghofe@15031
   705
berghofe@15031
   706
fun clause_of_eqn eqn =
berghofe@15031
   707
  let
berghofe@15031
   708
    val (t, u) = HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of eqn));
berghofe@15031
   709
    val (Const (s, T), ts) = strip_comb t;
berghofe@15031
   710
    val (Ts, U) = strip_type T
berghofe@15031
   711
  in
wenzelm@41704
   712
    Codegen.rename_term (Logic.list_implies (prems_of eqn, HOLogic.mk_Trueprop
berghofe@22271
   713
      (list_comb (Const (s ^ "_aux", Ts @ [U] ---> HOLogic.boolT), ts @ [u]))))
berghofe@15031
   714
  end;
berghofe@15031
   715
wenzelm@43292
   716
fun mk_fun thy codegen_mode defs name eqns dep module module' gr =
wenzelm@43307
   717
  (case try (Codegen.get_node gr) name of
berghofe@17144
   718
    NONE =>
wenzelm@43307
   719
      let
wenzelm@43307
   720
        val clauses = map clause_of_eqn eqns;
wenzelm@43307
   721
        val pname = name ^ "_aux";
wenzelm@43307
   722
        val arity =
wenzelm@43307
   723
          length (snd (strip_comb (fst (HOLogic.dest_eq
wenzelm@43307
   724
            (HOLogic.dest_Trueprop (concl_of (hd eqns)))))));
wenzelm@43307
   725
        val mode = 1 upto arity;
wenzelm@43307
   726
        val ((fun_id, mode_id), gr') = gr |>
wenzelm@43307
   727
          Codegen.mk_const_id module' name ||>>
wenzelm@43307
   728
          modename module' pname ([], mode);
wenzelm@43307
   729
        val vars = map (fn i => Codegen.str ("x" ^ string_of_int i)) mode;
wenzelm@43307
   730
        val s = Codegen.string_of (Pretty.block
wenzelm@43307
   731
          [Codegen.mk_app false (Codegen.str ("fun " ^ snd fun_id)) vars, Codegen.str " =",
wenzelm@43307
   732
           Pretty.brk 1, Codegen.str "DSeq.hd", Pretty.brk 1,
wenzelm@43307
   733
           Codegen.parens (Pretty.block (separate (Pretty.brk 1) (Codegen.str mode_id ::
wenzelm@43307
   734
             vars)))]) ^ ";\n\n";
wenzelm@43307
   735
        val gr'' = mk_ind_def thy codegen_mode defs (Codegen.add_edge (name, dep)
wenzelm@43307
   736
          (Codegen.new_node (name, (NONE, module', s)) gr')) name [pname] module'
wenzelm@43307
   737
          [(pname, [([], mode)])] clauses 0;
wenzelm@43307
   738
        val (modes, _) = lookup_modes gr'' dep;
wenzelm@43307
   739
        val _ = find_mode gr'' dep pname (head_of (HOLogic.dest_Trueprop
wenzelm@43307
   740
          (Logic.strip_imp_concl (hd clauses)))) modes mode
wenzelm@43307
   741
      in (Codegen.mk_qual_id module fun_id, gr'') end
berghofe@17144
   742
  | SOME _ =>
wenzelm@43307
   743
      (Codegen.mk_qual_id module (Codegen.get_const_id gr name),
wenzelm@43307
   744
        Codegen.add_edge (name, dep) gr));
berghofe@15031
   745
berghofe@23761
   746
(* convert n-tuple to nested pairs *)
berghofe@23761
   747
berghofe@23761
   748
fun conv_ntuple fs ts p =
berghofe@23761
   749
  let
berghofe@23761
   750
    val k = length fs;
wenzelm@41704
   751
    val xs = map_range (fn i => Codegen.str ("x" ^ string_of_int i)) (k + 1);
berghofe@23761
   752
    val xs' = map (fn Bound i => nth xs (k - i)) ts;
berghofe@23761
   753
    fun conv xs js =
haftmann@36677
   754
      if member (op =) fs js then
berghofe@23761
   755
        let
berghofe@23761
   756
          val (p, xs') = conv xs (1::js);
berghofe@23761
   757
          val (q, xs'') = conv xs' (2::js)
berghofe@23761
   758
        in (mk_tuple [p, q], xs'') end
berghofe@23761
   759
      else (hd xs, tl xs)
berghofe@23761
   760
  in
berghofe@23761
   761
    if k > 0 then
berghofe@23761
   762
      Pretty.block
wenzelm@41704
   763
        [Codegen.str "DSeq.map (fn", Pretty.brk 1,
wenzelm@41704
   764
         mk_tuple xs', Codegen.str " =>", Pretty.brk 1, fst (conv xs []),
wenzelm@41704
   765
         Codegen.str ")", Pretty.brk 1, Codegen.parens p]
berghofe@23761
   766
    else p
berghofe@23761
   767
  end;
berghofe@23761
   768
wenzelm@43307
   769
fun inductive_codegen thy codegen_mode defs dep module brack t gr =
wenzelm@43307
   770
  (case strip_comb t of
wenzelm@35364
   771
    (Const (@{const_name Collect}, _), [u]) =>
wenzelm@43307
   772
      let val (r, Ts, fs) = HOLogic.strip_psplits u in
wenzelm@43307
   773
        (case strip_comb r of
berghofe@23761
   774
          (Const (s, T), ts) =>
wenzelm@41704
   775
            (case (get_clauses thy s, Codegen.get_assoc_code thy (s, T)) of
berghofe@23761
   776
              (SOME (names, thyname, k, intrs), NONE) =>
berghofe@23761
   777
                let
berghofe@23761
   778
                  val (ts1, ts2) = chop k ts;
berghofe@23761
   779
                  val ts2' = map
haftmann@34949
   780
                    (fn Bound i => Term.dummy_pattern (nth Ts (length Ts - i - 1)) | t => t) ts2;
berghofe@23761
   781
                  val (ots, its) = List.partition is_Bound ts2;
wenzelm@43307
   782
                  val closed = forall (not o Term.is_open);
berghofe@23761
   783
                in
berghofe@23761
   784
                  if null (duplicates op = ots) andalso
wenzelm@42989
   785
                    closed ts1 andalso closed its
berghofe@23761
   786
                  then
wenzelm@43307
   787
                    let
wenzelm@43307
   788
                      val (call_p, gr') =
wenzelm@43307
   789
                        mk_ind_call thy codegen_mode defs dep module true
wenzelm@43307
   790
                          s T (ts1 @ ts2') names thyname k intrs gr;
wenzelm@43307
   791
                    in
wenzelm@43307
   792
                      SOME ((if brack then Codegen.parens else I) (Pretty.block
wenzelm@43307
   793
                        [Codegen.str "Set", Pretty.brk 1, Codegen.str "(DSeq.list_of", Pretty.brk 1,
wenzelm@43307
   794
                         Codegen.str "(", conv_ntuple fs ots call_p, Codegen.str "))"]),
wenzelm@43307
   795
                         (*this is a very strong assumption about the generated code!*)
wenzelm@43307
   796
                         gr')
berghofe@23761
   797
                    end
berghofe@23761
   798
                  else NONE
berghofe@23761
   799
                end
berghofe@23761
   800
            | _ => NONE)
wenzelm@43307
   801
        | _ => NONE)
berghofe@23761
   802
      end
wenzelm@41704
   803
  | (Const (s, T), ts) =>
wenzelm@43307
   804
      (case Symtab.lookup (#eqns (CodegenData.get thy)) s of
wenzelm@43307
   805
        NONE =>
wenzelm@43307
   806
          (case (get_clauses thy s, Codegen.get_assoc_code thy (s, T)) of
wenzelm@43307
   807
            (SOME (names, thyname, k, intrs), NONE) =>
wenzelm@43307
   808
              if length ts < k then NONE else
wenzelm@43307
   809
                SOME
wenzelm@43307
   810
                  (let
wenzelm@43307
   811
                    val (call_p, gr') = mk_ind_call thy codegen_mode defs dep module false
wenzelm@43307
   812
                      s T (map Term.no_dummy_patterns ts) names thyname k intrs gr
wenzelm@43307
   813
                   in
wenzelm@43307
   814
                    (mk_funcomp brack "?!"
wenzelm@43307
   815
                      (length (binder_types T) - length ts) (Codegen.parens call_p), gr')
wenzelm@43307
   816
                   end
wenzelm@43307
   817
                   handle TERM _ =>
wenzelm@43307
   818
                    mk_ind_call thy codegen_mode defs dep module true
wenzelm@43307
   819
                      s T ts names thyname k intrs gr)
wenzelm@43307
   820
          | _ => NONE)
wenzelm@43307
   821
      | SOME eqns =>
wenzelm@43307
   822
          let
wenzelm@43307
   823
            val (_, thyname) :: _ = eqns;
wenzelm@43307
   824
            val (id, gr') =
wenzelm@43307
   825
              mk_fun thy codegen_mode defs s (Codegen.preprocess thy (map fst (rev eqns)))
wenzelm@43307
   826
                dep module (Codegen.if_library codegen_mode thyname module) gr;
wenzelm@43307
   827
            val (ps, gr'') =
wenzelm@43307
   828
              fold_map (Codegen.invoke_codegen thy codegen_mode defs dep module true)
wenzelm@43307
   829
                ts gr';
wenzelm@43307
   830
          in SOME (Codegen.mk_app brack (Codegen.str id) ps, gr'') end)
berghofe@22271
   831
  | _ => NONE);
berghofe@11537
   832
berghofe@12557
   833
val setup =
wenzelm@41704
   834
  Codegen.add_codegen "inductive" inductive_codegen #>
wenzelm@33260
   835
  Attrib.setup @{binding code_ind}
wenzelm@33260
   836
    (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
wenzelm@36970
   837
      Scan.option (Args.$$$ "params" |-- Args.colon |-- Parse.nat) >> uncurry add))
haftmann@31998
   838
    "introduction rules for executable predicates";
berghofe@11537
   839
wenzelm@43306
   840
berghofe@33765
   841
(**** Quickcheck involving inductive predicates ****)
berghofe@33765
   842
wenzelm@43306
   843
structure Result = Proof_Data
wenzelm@43306
   844
(
wenzelm@43306
   845
  type T = int * int * int -> term list option;
wenzelm@43306
   846
  fun init _ = (fn _ => NONE);
wenzelm@43306
   847
);
wenzelm@43306
   848
wenzelm@43306
   849
val get_test_fn = Result.get;
wenzelm@43306
   850
fun poke_test_fn f = Context.>> (Context.map_proof (Result.put f));
wenzelm@43306
   851
berghofe@33765
   852
berghofe@33765
   853
fun strip_imp p =
berghofe@33765
   854
  let val (q, r) = HOLogic.dest_imp p
berghofe@33765
   855
  in strip_imp r |>> cons q end
berghofe@33765
   856
  handle TERM _ => ([], p);
berghofe@33765
   857
berghofe@33765
   858
fun deepen bound f i =
berghofe@33765
   859
  if i > bound then NONE
wenzelm@43307
   860
  else
wenzelm@43307
   861
    (case f i of
berghofe@33765
   862
      NONE => deepen bound f (i + 1)
berghofe@33765
   863
    | SOME x => SOME x);
berghofe@33765
   864
wenzelm@43487
   865
val depth_bound = Attrib.setup_config_int @{binding ind_quickcheck_depth} (K 10);
wenzelm@43487
   866
val depth_start = Attrib.setup_config_int @{binding ind_quickcheck_depth_start} (K 1);
wenzelm@43487
   867
val random_values = Attrib.setup_config_int @{binding ind_quickcheck_random} (K 5);
wenzelm@43487
   868
val size_offset = Attrib.setup_config_int @{binding ind_quickcheck_size_offset} (K 0);
berghofe@33765
   869
bulwahn@43030
   870
fun test_term ctxt [(t, [])] =
wenzelm@43306
   871
      let
wenzelm@43306
   872
        val t' = list_abs_free (Term.add_frees t [], t)
wenzelm@43306
   873
        val thy = Proof_Context.theory_of ctxt;
wenzelm@43306
   874
        val (xs, p) = strip_abs t';
wenzelm@43306
   875
        val args' = map_index (fn (i, (_, T)) => ("arg" ^ string_of_int i, T)) xs;
wenzelm@43306
   876
        val args = map Free args';
wenzelm@43306
   877
        val (ps, q) = strip_imp p;
wenzelm@43306
   878
        val Ts = map snd xs;
wenzelm@43306
   879
        val T = Ts ---> HOLogic.boolT;
wenzelm@43306
   880
        val rl = Logic.list_implies
wenzelm@43306
   881
          (map (HOLogic.mk_Trueprop o curry subst_bounds (rev args)) ps @
wenzelm@43306
   882
           [HOLogic.mk_Trueprop (HOLogic.mk_not (subst_bounds (rev args, q)))],
wenzelm@43306
   883
           HOLogic.mk_Trueprop (list_comb (Free ("quickcheckp", T), args)));
wenzelm@43306
   884
        val (_, thy') = Inductive.add_inductive_global
wenzelm@43306
   885
          {quiet_mode=true, verbose=false, alt_name=Binding.empty, coind=false,
wenzelm@43306
   886
           no_elim=true, no_ind=false, skip_mono=false, fork_mono=false}
wenzelm@44500
   887
          [((@{binding quickcheckp}, T), NoSyn)] []
wenzelm@43306
   888
          [(Attrib.empty_binding, rl)] [] (Theory.copy thy);
wenzelm@43306
   889
        val pred = HOLogic.mk_Trueprop (list_comb
wenzelm@43306
   890
          (Const (Sign.intern_const thy' "quickcheckp", T),
wenzelm@43306
   891
           map Term.dummy_pattern Ts));
wenzelm@43306
   892
        val (code, gr) =
wenzelm@43306
   893
          Codegen.generate_code_i thy' ["term_of", "test", "random_ind"] [] "Generated"
wenzelm@43306
   894
            [("testf", pred)];
wenzelm@43306
   895
        val s = "structure Test_Term =\nstruct\n\n" ^
wenzelm@43306
   896
          cat_lines (map snd code) ^
wenzelm@43306
   897
          "\nopen Generated;\n\n" ^ Codegen.string_of
wenzelm@43306
   898
            (Pretty.block [Codegen.str "val () = Inductive_Codegen.poke_test_fn",
wenzelm@43306
   899
              Pretty.brk 1, Codegen.str "(fn p =>", Pretty.brk 1,
wenzelm@43306
   900
              Codegen.str "case Seq.pull (testf p) of", Pretty.brk 1,
wenzelm@43306
   901
              Codegen.str "SOME ",
wenzelm@43306
   902
              mk_tuple [mk_tuple (map (Codegen.str o fst) args'), Codegen.str "_"],
wenzelm@43306
   903
              Codegen.str " =>", Pretty.brk 1, Codegen.str "SOME ",
wenzelm@43306
   904
              Pretty.enum "," "[" "]"
wenzelm@43306
   905
                (map (fn (s, T) => Pretty.block
wenzelm@43306
   906
                  [Codegen.mk_term_of gr "Generated" false T, Pretty.brk 1, Codegen.str s]) args'),
wenzelm@43306
   907
              Pretty.brk 1,
wenzelm@43306
   908
              Codegen.str "| NONE => NONE);"]) ^
wenzelm@43306
   909
          "\n\nend;\n";
wenzelm@43306
   910
        val test_fn =
wenzelm@43306
   911
          ctxt
wenzelm@43306
   912
          |> Context.proof_map
wenzelm@43306
   913
              (ML_Context.exec (fn () => ML_Context.eval_text false Position.none s))
wenzelm@43306
   914
          |> get_test_fn;
wenzelm@43306
   915
        val values = Config.get ctxt random_values;
wenzelm@43306
   916
        val bound = Config.get ctxt depth_bound;
wenzelm@43306
   917
        val start = Config.get ctxt depth_start;
wenzelm@43306
   918
        val offset = Config.get ctxt size_offset;
wenzelm@43306
   919
        fun test [k] = (deepen bound (fn i =>
wenzelm@43306
   920
          (Output.urgent_message ("Search depth: " ^ string_of_int i);
wenzelm@43306
   921
           test_fn (i, values, k+offset))) start, NONE);
wenzelm@43306
   922
      in test end
wenzelm@43306
   923
  | test_term ctxt [_] = error "Option eval is not supported by tester SML_inductive"
wenzelm@43306
   924
  | test_term ctxt _ =
wenzelm@43306
   925
      error "Compilation of multiple instances is not supported by tester SML_inductive";
berghofe@33765
   926
bulwahn@44748
   927
val test_goal = Quickcheck.generator_test_goal_terms test_term;
bulwahn@44748
   928
berghofe@33765
   929
val quickcheck_setup =
bulwahn@44748
   930
  Context.theory_map (Quickcheck.add_tester ("SML_inductive", test_goal));
berghofe@33765
   931
berghofe@11537
   932
end;