src/HOL/Tools/inductive_codegen.ML
author haftmann
Tue, 20 Oct 2009 16:13:01 +0200
changeset 33037 b22e44496dc2
parent 32959 675c0c7e6a37
child 33038 8f9594c31de4
permissions -rw-r--r--
replaced old_style infixes eq_set, subset, union, inter and variants by generic versions
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
berghofe@22271
     9
  val add : string option -> int option -> attribute
wenzelm@18708
    10
  val setup : theory -> theory
berghofe@11537
    11
end;
berghofe@11537
    12
berghofe@11537
    13
structure InductiveCodegen : INDUCTIVE_CODEGEN =
berghofe@11537
    14
struct
berghofe@11537
    15
berghofe@11537
    16
open Codegen;
berghofe@11537
    17
berghofe@12557
    18
(**** theory data ****)
berghofe@11537
    19
wenzelm@18930
    20
fun merge_rules tabs =
haftmann@22642
    21
  Symtab.join (fn _ => AList.merge (Thm.eq_thm_prop) (K true)) tabs;
wenzelm@18930
    22
wenzelm@16424
    23
structure CodegenData = TheoryDataFun
wenzelm@22846
    24
(
berghofe@15031
    25
  type T =
berghofe@22271
    26
    {intros : (thm * (string * int)) list Symtab.table,
berghofe@15031
    27
     graph : unit Graph.T,
berghofe@16645
    28
     eqns : (thm * string) list Symtab.table};
berghofe@15031
    29
  val empty =
berghofe@15031
    30
    {intros = Symtab.empty, graph = Graph.empty, eqns = Symtab.empty};
berghofe@12557
    31
  val copy = I;
wenzelm@16424
    32
  val extend = I;
wenzelm@16424
    33
  fun merge _ ({intros=intros1, graph=graph1, eqns=eqns1},
berghofe@15031
    34
    {intros=intros2, graph=graph2, eqns=eqns2}) =
wenzelm@18930
    35
    {intros = merge_rules (intros1, intros2),
berghofe@15031
    36
     graph = Graph.merge (K true) (graph1, graph2),
wenzelm@18930
    37
     eqns = merge_rules (eqns1, eqns2)};
wenzelm@22846
    38
);
berghofe@12557
    39
berghofe@12557
    40
berghofe@12557
    41
fun warn thm = warning ("InductiveCodegen: Not a proper clause:\n" ^
wenzelm@32111
    42
  Display.string_of_thm_without_context thm);
berghofe@12557
    43
berghofe@14162
    44
fun add_node (g, x) = Graph.new_node (x, ()) g handle Graph.DUP _ => g;
berghofe@14162
    45
berghofe@22271
    46
fun add optmod optnparms = Thm.declaration_attribute (fn thm => Context.mapping (fn thy =>
berghofe@16645
    47
  let
berghofe@16645
    48
    val {intros, graph, eqns} = CodegenData.get thy;
berghofe@16645
    49
    fun thyname_of s = (case optmod of
haftmann@27398
    50
      NONE => Codegen.thyname_of_const thy s | SOME s => s);
berghofe@22271
    51
  in (case Option.map strip_comb (try HOLogic.dest_Trueprop (concl_of thm)) of
berghofe@22271
    52
      SOME (Const ("op =", _), [t, _]) => (case head_of t of
berghofe@15031
    53
        Const (s, _) =>
wenzelm@18728
    54
          CodegenData.put {intros = intros, graph = graph,
haftmann@22642
    55
             eqns = eqns |> Symtab.map_default (s, [])
haftmann@22642
    56
               (AList.update Thm.eq_thm_prop (thm, thyname_of s))} thy
wenzelm@18728
    57
      | _ => (warn thm; thy))
berghofe@22271
    58
    | SOME (Const (s, _), _) =>
berghofe@22271
    59
        let
wenzelm@29287
    60
          val cs = fold Term.add_const_names (Thm.prems_of thm) [];
berghofe@22271
    61
          val rules = Symtab.lookup_list intros s;
berghofe@22271
    62
          val nparms = (case optnparms of
berghofe@22271
    63
            SOME k => k
berghofe@22271
    64
          | NONE => (case rules of
haftmann@31723
    65
             [] => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
berghofe@22791
    66
                 SOME (_, {raw_induct, ...}) =>
haftmann@31723
    67
                   length (Inductive.params_of raw_induct)
berghofe@22271
    68
               | NONE => 0)
berghofe@22271
    69
            | xs => snd (snd (snd (split_last xs)))))
berghofe@22271
    70
        in CodegenData.put
berghofe@22271
    71
          {intros = intros |>
haftmann@22642
    72
           Symtab.update (s, (AList.update Thm.eq_thm_prop
haftmann@22642
    73
             (thm, (thyname_of s, nparms)) rules)),
wenzelm@30193
    74
           graph = List.foldr (uncurry (Graph.add_edge o pair s))
berghofe@22271
    75
             (Library.foldl add_node (graph, s :: cs)) cs,
berghofe@22271
    76
           eqns = eqns} thy
berghofe@22271
    77
        end
wenzelm@18728
    78
    | _ => (warn thm; thy))
wenzelm@20897
    79
  end) I);
berghofe@12557
    80
berghofe@12557
    81
fun get_clauses thy s =
berghofe@15031
    82
  let val {intros, graph, ...} = CodegenData.get thy
wenzelm@17412
    83
  in case Symtab.lookup intros s of
haftmann@31723
    84
      NONE => (case try (Inductive.the_inductive (ProofContext.init thy)) s of
skalberg@15531
    85
        NONE => NONE
berghofe@22271
    86
      | SOME ({names, ...}, {intrs, raw_induct, ...}) =>
haftmann@27398
    87
          SOME (names, Codegen.thyname_of_const thy s,
haftmann@31723
    88
            length (Inductive.params_of raw_induct),
berghofe@22661
    89
            preprocess thy intrs))
skalberg@15531
    90
    | SOME _ =>
berghofe@16645
    91
        let
berghofe@16645
    92
          val SOME names = find_first
haftmann@22642
    93
            (fn xs => member (op =) xs s) (Graph.strong_conn graph);
haftmann@22642
    94
          val intrs as (_, (thyname, nparms)) :: _ =
haftmann@22642
    95
            maps (the o Symtab.lookup intros) names;
haftmann@22642
    96
        in SOME (names, thyname, nparms, preprocess thy (map fst (rev intrs))) end
berghofe@14162
    97
  end;
berghofe@12557
    98
berghofe@12557
    99
berghofe@11537
   100
(**** check if a term contains only constructor functions ****)
berghofe@11537
   101
berghofe@11537
   102
fun is_constrt thy =
berghofe@11537
   103
  let
haftmann@31784
   104
    val cnstrs = flat (maps
berghofe@11537
   105
      (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd)
haftmann@31784
   106
      (Symtab.dest (Datatype.get_all thy)));
berghofe@11537
   107
    fun check t = (case strip_comb t of
berghofe@11537
   108
        (Var _, []) => true
haftmann@17521
   109
      | (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of
skalberg@15531
   110
            NONE => false
skalberg@15531
   111
          | SOME i => length ts = i andalso forall check ts)
berghofe@11537
   112
      | _ => false)
berghofe@11537
   113
  in check end;
berghofe@11537
   114
berghofe@11537
   115
(**** check if a type is an equality type (i.e. doesn't contain fun) ****)
berghofe@11537
   116
berghofe@11537
   117
fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
berghofe@11537
   118
  | is_eqT _ = true;
berghofe@11537
   119
berghofe@11537
   120
(**** mode inference ****)
berghofe@11537
   121
berghofe@15204
   122
fun string_of_mode (iss, is) = space_implode " -> " (map
skalberg@15531
   123
  (fn NONE => "X"
skalberg@15531
   124
    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
skalberg@15531
   125
       (iss @ [SOME is]));
berghofe@15204
   126
berghofe@15204
   127
fun print_modes modes = message ("Inferred modes:\n" ^
wenzelm@26931
   128
  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
berghofe@15204
   129
    string_of_mode ms)) modes));
berghofe@15204
   130
wenzelm@29265
   131
val term_vs = map (fst o fst o dest_Var) o OldTerm.term_vars;
wenzelm@32952
   132
val terms_vs = distinct (op =) o maps term_vs;
berghofe@11537
   133
berghofe@11537
   134
(** collect all Vars in a term (with duplicates!) **)
wenzelm@16861
   135
fun term_vTs tm =
wenzelm@16861
   136
  fold_aterms (fn Var ((x, _), T) => cons (x, T) | _ => I) tm [];
berghofe@11537
   137
berghofe@11537
   138
fun get_args _ _ [] = ([], [])
berghofe@11537
   139
  | get_args is i (x::xs) = (if i mem is then apfst else apsnd) (cons x)
berghofe@11537
   140
      (get_args is (i+1) xs);
berghofe@11537
   141
berghofe@11537
   142
fun merge xs [] = xs
berghofe@11537
   143
  | merge [] ys = ys
berghofe@11537
   144
  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
berghofe@11537
   145
      else y::merge (x::xs) ys;
berghofe@11537
   146
berghofe@11537
   147
fun subsets i j = if i <= j then
berghofe@11537
   148
       let val is = subsets (i+1) j
berghofe@11537
   149
       in merge (map (fn ks => i::ks) is) is end
berghofe@11537
   150
     else [[]];
berghofe@11537
   151
berghofe@12557
   152
fun cprod ([], ys) = []
berghofe@12557
   153
  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
berghofe@12557
   154
wenzelm@30193
   155
fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
berghofe@12557
   156
berghofe@22271
   157
datatype mode = Mode of (int list option list * int list) * int list * mode option list;
berghofe@12557
   158
berghofe@12557
   159
fun modes_of modes t =
berghofe@12557
   160
  let
berghofe@22271
   161
    val ks = 1 upto length (binder_types (fastype_of t));
berghofe@22271
   162
    val default = [Mode (([], ks), ks, [])];
wenzelm@32952
   163
    fun mk_modes name args = Option.map
wenzelm@32952
   164
     (maps (fn (m as (iss, is)) =>
berghofe@22271
   165
        let
berghofe@22271
   166
          val (args1, args2) =
berghofe@22271
   167
            if length args < length iss then
berghofe@22271
   168
              error ("Too few arguments for inductive predicate " ^ name)
berghofe@22271
   169
            else chop (length iss) args;
berghofe@22271
   170
          val k = length args2;
berghofe@22271
   171
          val prfx = 1 upto k
berghofe@22271
   172
        in
berghofe@22271
   173
          if not (is_prefix op = prfx is) then [] else
berghofe@22271
   174
          let val is' = map (fn i => i - k) (List.drop (is, k))
berghofe@22271
   175
          in map (fn x => Mode (m, is', x)) (cprods (map
berghofe@22271
   176
            (fn (NONE, _) => [NONE]
berghofe@22271
   177
              | (SOME js, arg) => map SOME (List.filter
berghofe@22271
   178
                  (fn Mode (_, js', _) => js=js') (modes_of modes arg)))
berghofe@22271
   179
                    (iss ~~ args1)))
berghofe@22271
   180
          end
berghofe@22271
   181
        end)) (AList.lookup op = modes name)
berghofe@12557
   182
berghofe@12557
   183
  in (case strip_comb t of
berghofe@14163
   184
      (Const ("op =", Type (_, [T, _])), _) =>
berghofe@22271
   185
        [Mode (([], [1]), [1], []), Mode (([], [2]), [2], [])] @
berghofe@22271
   186
        (if is_eqT T then [Mode (([], [1, 2]), [1, 2], [])] else [])
berghofe@22271
   187
    | (Const (name, _), args) => the_default default (mk_modes name args)
berghofe@22271
   188
    | (Var ((name, _), _), args) => the (mk_modes name args)
berghofe@22271
   189
    | (Free (name, _), args) => the (mk_modes name args)
berghofe@22271
   190
    | _ => default)
berghofe@12557
   191
  end;
berghofe@12557
   192
berghofe@26806
   193
datatype indprem = Prem of term list * term * bool | Sidecond of term;
berghofe@12557
   194
berghofe@11537
   195
fun select_mode_prem thy modes vs ps =
wenzelm@19861
   196
  find_first (is_some o snd) (ps ~~ map
berghofe@26806
   197
    (fn Prem (us, t, is_set) => find_first (fn Mode (_, is, _) =>
berghofe@11537
   198
          let
berghofe@15031
   199
            val (in_ts, out_ts) = get_args is 1 us;
skalberg@15570
   200
            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
wenzelm@32952
   201
            val vTs = maps term_vTs out_ts';
wenzelm@18964
   202
            val dupTs = map snd (duplicates (op =) vTs) @
wenzelm@32952
   203
              map_filter (AList.lookup (op =) vTs) vs;
berghofe@11537
   204
          in
haftmann@33037
   205
            gen_subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
berghofe@15031
   206
            forall (is_eqT o fastype_of) in_ts' andalso
haftmann@33037
   207
            gen_subset (op =) (term_vs t, vs) andalso
berghofe@11537
   208
            forall is_eqT dupTs
berghofe@11537
   209
          end)
berghofe@26806
   210
            (if is_set then [Mode (([], []), [], [])]
berghofe@26806
   211
             else modes_of modes t handle Option =>
wenzelm@26939
   212
               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
haftmann@33037
   213
      | Sidecond t => if gen_subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
skalberg@15531
   214
          else NONE) ps);
berghofe@11537
   215
berghofe@12557
   216
fun check_mode_clause thy arg_vs modes (iss, is) (ts, ps) =
berghofe@11537
   217
  let
wenzelm@32952
   218
    val modes' = modes @ map_filter
skalberg@15531
   219
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
berghofe@12557
   220
        (arg_vs ~~ iss);
skalberg@15531
   221
    fun check_mode_prems vs [] = SOME vs
berghofe@12557
   222
      | check_mode_prems vs ps = (case select_mode_prem thy modes' vs ps of
skalberg@15531
   223
          NONE => NONE
skalberg@15531
   224
        | SOME (x, _) => check_mode_prems
haftmann@33037
   225
            (case x of Prem (us, _, _) => gen_union (op =) (vs, terms_vs us) | _ => vs)
berghofe@11537
   226
            (filter_out (equal x) ps));
skalberg@15570
   227
    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (get_args is 1 ts));
berghofe@11537
   228
    val in_vs = terms_vs in_ts;
berghofe@11537
   229
    val concl_vs = terms_vs ts
berghofe@11537
   230
  in
wenzelm@32952
   231
    forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
berghofe@15204
   232
    forall (is_eqT o fastype_of) in_ts' andalso
haftmann@33037
   233
    (case check_mode_prems (gen_union (op =) (arg_vs, in_vs)) ps of
skalberg@15531
   234
       NONE => false
haftmann@33037
   235
     | SOME vs => gen_subset (op =) (concl_vs, vs))
berghofe@11537
   236
  end;
berghofe@11537
   237
berghofe@11537
   238
fun check_modes_pred thy arg_vs preds modes (p, ms) =
haftmann@17521
   239
  let val SOME rs = AList.lookup (op =) preds p
skalberg@15570
   240
  in (p, List.filter (fn m => case find_index
berghofe@15204
   241
    (not o check_mode_clause thy arg_vs modes m) rs of
berghofe@15204
   242
      ~1 => true
berghofe@15204
   243
    | i => (message ("Clause " ^ string_of_int (i+1) ^ " of " ^
berghofe@15204
   244
      p ^ " violates mode " ^ string_of_mode m); false)) ms)
berghofe@15204
   245
  end;
berghofe@11537
   246
haftmann@22642
   247
fun fixp f (x : (string * (int list option list * int list) list) list) =
berghofe@11537
   248
  let val y = f x
berghofe@11537
   249
  in if x = y then x else fixp f y end;
berghofe@11537
   250
berghofe@22271
   251
fun infer_modes thy extra_modes arities arg_vs preds = fixp (fn modes =>
berghofe@11537
   252
  map (check_modes_pred thy arg_vs preds (modes @ extra_modes)) modes)
berghofe@22271
   253
    (map (fn (s, (ks, k)) => (s, cprod (cprods (map
skalberg@15531
   254
      (fn NONE => [NONE]
berghofe@22271
   255
        | SOME k' => map SOME (subsets 1 k')) ks),
berghofe@22271
   256
      subsets 1 k))) arities);
berghofe@11537
   257
berghofe@11537
   258
(**** code generation ****)
berghofe@11537
   259
berghofe@11537
   260
fun mk_eq (x::xs) =
berghofe@11537
   261
  let fun mk_eqs _ [] = []
berghofe@26975
   262
        | mk_eqs a (b::cs) = str (a ^ " = " ^ b) :: mk_eqs b cs
berghofe@11537
   263
  in mk_eqs x xs end;
berghofe@11537
   264
berghofe@26975
   265
fun mk_tuple xs = Pretty.block (str "(" ::
wenzelm@32952
   266
  flat (separate [str ",", Pretty.brk 1] (map single xs)) @
berghofe@26975
   267
  [str ")"]);
berghofe@11537
   268
haftmann@17521
   269
fun mk_v ((names, vs), s) = (case AList.lookup (op =) vs s of
skalberg@15531
   270
      NONE => ((names, (s, [s])::vs), s)
wenzelm@20071
   271
    | SOME xs => let val s' = Name.variant names s in
haftmann@17521
   272
        ((s'::names, AList.update (op =) (s, s'::xs) vs), s') end);
berghofe@11537
   273
berghofe@11537
   274
fun distinct_v (nvs, Var ((s, 0), T)) =
berghofe@11537
   275
      apsnd (Var o rpair T o rpair 0) (mk_v (nvs, s))
berghofe@11537
   276
  | distinct_v (nvs, t $ u) =
berghofe@11537
   277
      let
berghofe@11537
   278
        val (nvs', t') = distinct_v (nvs, t);
berghofe@11537
   279
        val (nvs'', u') = distinct_v (nvs', u);
berghofe@11537
   280
      in (nvs'', t' $ u') end
berghofe@11537
   281
  | distinct_v x = x;
berghofe@11537
   282
berghofe@15061
   283
fun is_exhaustive (Var _) = true
berghofe@15061
   284
  | is_exhaustive (Const ("Pair", _) $ t $ u) =
berghofe@15061
   285
      is_exhaustive t andalso is_exhaustive u
berghofe@15061
   286
  | is_exhaustive _ = false;
berghofe@15061
   287
berghofe@15061
   288
fun compile_match nvs eq_ps out_ps success_p can_fail =
wenzelm@32952
   289
  let val eqs = flat (separate [str " andalso", Pretty.brk 1]
wenzelm@32952
   290
    (map single (maps (mk_eq o snd) nvs @ eq_ps)));
berghofe@11537
   291
  in
berghofe@11537
   292
    Pretty.block
berghofe@26975
   293
     ([str "(fn ", mk_tuple out_ps, str " =>", Pretty.brk 1] @
berghofe@26975
   294
      (Pretty.block ((if eqs=[] then [] else str "if " ::
berghofe@26975
   295
         [Pretty.block eqs, Pretty.brk 1, str "then "]) @
berghofe@11537
   296
         (success_p ::
berghofe@26975
   297
          (if eqs=[] then [] else [Pretty.brk 1, str "else DSeq.empty"]))) ::
berghofe@15061
   298
       (if can_fail then
berghofe@26975
   299
          [Pretty.brk 1, str "| _ => DSeq.empty)"]
berghofe@26975
   300
        else [str ")"])))
berghofe@11537
   301
  end;
berghofe@11537
   302
berghofe@17144
   303
fun modename module s (iss, is) gr =
haftmann@28537
   304
  let val (id, gr') = if s = @{const_name "op ="} then (("", "equal"), gr)
berghofe@17144
   305
    else mk_const_id module s gr
haftmann@28537
   306
  in (space_implode "__"
berghofe@17144
   307
    (mk_qual_id module id ::
wenzelm@32952
   308
      map (space_implode "_" o map string_of_int) (map_filter I iss @ [is])), gr')
berghofe@17144
   309
  end;
berghofe@16645
   310
berghofe@22271
   311
fun mk_funcomp brack s k p = (if brack then parens else I)
berghofe@26975
   312
  (Pretty.block [Pretty.block ((if k = 0 then [] else [str "("]) @
berghofe@26975
   313
    separate (Pretty.brk 1) (str s :: replicate k (str "|> ???")) @
berghofe@26975
   314
    (if k = 0 then [] else [str ")"])), Pretty.brk 1, p]);
berghofe@22271
   315
haftmann@28537
   316
fun compile_expr thy defs dep module brack modes (NONE, t) gr =
haftmann@28537
   317
      apfst single (invoke_codegen thy defs dep module brack t gr)
haftmann@28537
   318
  | compile_expr _ _ _ _ _ _ (SOME _, Var ((name, _), _)) gr =
haftmann@28537
   319
      ([str name], gr)
haftmann@28537
   320
  | compile_expr thy defs dep module brack modes (SOME (Mode (mode, _, ms)), t) gr =
berghofe@22271
   321
      (case strip_comb t of
berghofe@22271
   322
         (Const (name, _), args) =>
haftmann@28537
   323
           if name = @{const_name "op ="} orelse AList.defined op = modes name then
berghofe@22271
   324
             let
berghofe@22271
   325
               val (args1, args2) = chop (length ms) args;
haftmann@28537
   326
               val ((ps, mode_id), gr') = gr |> fold_map
haftmann@28537
   327
                   (compile_expr thy defs dep module true modes) (ms ~~ args1)
haftmann@28537
   328
                   ||>> modename module name mode;
haftmann@28537
   329
               val (ps', gr'') = (case mode of
haftmann@28537
   330
                   ([], []) => ([str "()"], gr')
haftmann@28537
   331
                 | _ => fold_map
haftmann@28537
   332
                     (invoke_codegen thy defs dep module true) args2 gr')
haftmann@28537
   333
             in ((if brack andalso not (null ps andalso null ps') then
berghofe@22271
   334
               single o parens o Pretty.block else I)
wenzelm@32952
   335
                 (flat (separate [Pretty.brk 1]
haftmann@28537
   336
                   ([str mode_id] :: ps @ map single ps'))), gr')
berghofe@22271
   337
             end
haftmann@28537
   338
           else apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
haftmann@28537
   339
             (invoke_codegen thy defs dep module true t gr)
haftmann@28537
   340
       | _ => apfst (single o mk_funcomp brack "??" (length (binder_types (fastype_of t))))
haftmann@28537
   341
           (invoke_codegen thy defs dep module true t gr));
berghofe@12557
   342
haftmann@28537
   343
fun compile_clause thy defs dep module all_vs arg_vs modes (iss, is) (ts, ps) inp gr =
berghofe@11537
   344
  let
wenzelm@32952
   345
    val modes' = modes @ map_filter
skalberg@15531
   346
      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
berghofe@12557
   347
        (arg_vs ~~ iss);
berghofe@12557
   348
berghofe@11537
   349
    fun check_constrt ((names, eqs), t) =
berghofe@11537
   350
      if is_constrt thy t then ((names, eqs), t) else
wenzelm@20071
   351
        let val s = Name.variant names "x";
berghofe@11537
   352
        in ((s::names, (s, t)::eqs), Var ((s, 0), fastype_of t)) end;
berghofe@11537
   353
haftmann@28537
   354
    fun compile_eq (s, t) gr =
haftmann@28537
   355
      apfst (Pretty.block o cons (str (s ^ " = ")) o single)
haftmann@28537
   356
        (invoke_codegen thy defs dep module false t gr);
berghofe@15031
   357
berghofe@12557
   358
    val (in_ts, out_ts) = get_args is 1 ts;
berghofe@11537
   359
    val ((all_vs', eqs), in_ts') =
wenzelm@30193
   360
      Library.foldl_map check_constrt ((all_vs, []), in_ts);
berghofe@11537
   361
haftmann@28537
   362
    fun compile_prems out_ts' vs names [] gr =
berghofe@11537
   363
          let
haftmann@28537
   364
            val (out_ps, gr2) = fold_map
haftmann@28537
   365
              (invoke_codegen thy defs dep module false) out_ts gr;
haftmann@28537
   366
            val (eq_ps, gr3) = fold_map compile_eq eqs gr2;
berghofe@15031
   367
            val ((names', eqs'), out_ts'') =
wenzelm@30193
   368
              Library.foldl_map check_constrt ((names, []), out_ts');
wenzelm@30193
   369
            val (nvs, out_ts''') = Library.foldl_map distinct_v
berghofe@15031
   370
              ((names', map (fn x => (x, [x])) vs), out_ts'');
haftmann@28537
   371
            val (out_ps', gr4) = fold_map
haftmann@28537
   372
              (invoke_codegen thy defs dep module false) (out_ts''') gr3;
haftmann@28537
   373
            val (eq_ps', gr5) = fold_map compile_eq eqs' gr4;
berghofe@11537
   374
          in
haftmann@28537
   375
            (compile_match (snd nvs) (eq_ps @ eq_ps') out_ps'
berghofe@26975
   376
              (Pretty.block [str "DSeq.single", Pretty.brk 1, mk_tuple out_ps])
haftmann@28537
   377
              (exists (not o is_exhaustive) out_ts'''), gr5)
berghofe@11537
   378
          end
haftmann@28537
   379
      | compile_prems out_ts vs names ps gr =
berghofe@11537
   380
          let
haftmann@31852
   381
            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
berghofe@22271
   382
            val SOME (p, mode as SOME (Mode (_, js, _))) =
berghofe@15126
   383
              select_mode_prem thy modes' vs' ps;
berghofe@11537
   384
            val ps' = filter_out (equal p) ps;
berghofe@15031
   385
            val ((names', eqs), out_ts') =
wenzelm@30193
   386
              Library.foldl_map check_constrt ((names, []), out_ts);
wenzelm@30193
   387
            val (nvs, out_ts'') = Library.foldl_map distinct_v
berghofe@15031
   388
              ((names', map (fn x => (x, [x])) vs), out_ts');
haftmann@28537
   389
            val (out_ps, gr0) = fold_map
haftmann@28537
   390
              (invoke_codegen thy defs dep module false) out_ts'' gr;
haftmann@28537
   391
            val (eq_ps, gr1) = fold_map compile_eq eqs gr0;
berghofe@11537
   392
          in
berghofe@11537
   393
            (case p of
berghofe@26806
   394
               Prem (us, t, is_set) =>
berghofe@11537
   395
                 let
berghofe@15031
   396
                   val (in_ts, out_ts''') = get_args js 1 us;
haftmann@28537
   397
                   val (in_ps, gr2) = fold_map
haftmann@28537
   398
                     (invoke_codegen thy defs dep module true) in_ts gr1;
haftmann@28537
   399
                   val (ps, gr3) =
berghofe@26806
   400
                     if not is_set then
haftmann@28537
   401
                       apfst (fn ps => ps @
berghofe@24129
   402
                           (if null in_ps then [] else [Pretty.brk 1]) @
berghofe@22271
   403
                           separate (Pretty.brk 1) in_ps)
berghofe@22271
   404
                         (compile_expr thy defs dep module false modes
haftmann@28537
   405
                           (mode, t) gr2)
berghofe@26806
   406
                     else
haftmann@31852
   407
                       apfst (fn p => Pretty.breaks [str "DSeq.of_list", str "(case", p,
haftmann@31852
   408
                         str "of", str "Set", str "xs", str "=>", str "xs)"])
haftmann@31852
   409
                         (*this is a very strong assumption about the generated code!*)
haftmann@28537
   410
                           (invoke_codegen thy defs dep module true t gr2);
haftmann@28537
   411
                   val (rest, gr4) = compile_prems out_ts''' vs' (fst nvs) ps' gr3;
berghofe@11537
   412
                 in
haftmann@28537
   413
                   (compile_match (snd nvs) eq_ps out_ps
berghofe@12557
   414
                      (Pretty.block (ps @
berghofe@26975
   415
                         [str " :->", Pretty.brk 1, rest]))
haftmann@28537
   416
                      (exists (not o is_exhaustive) out_ts''), gr4)
berghofe@11537
   417
                 end
berghofe@11537
   418
             | Sidecond t =>
berghofe@11537
   419
                 let
haftmann@28537
   420
                   val (side_p, gr2) = invoke_codegen thy defs dep module true t gr1;
haftmann@28537
   421
                   val (rest, gr3) = compile_prems [] vs' (fst nvs) ps' gr2;
berghofe@11537
   422
                 in
haftmann@28537
   423
                   (compile_match (snd nvs) eq_ps out_ps
berghofe@26975
   424
                      (Pretty.block [str "?? ", side_p,
berghofe@26975
   425
                        str " :->", Pretty.brk 1, rest])
haftmann@28537
   426
                      (exists (not o is_exhaustive) out_ts''), gr3)
berghofe@11537
   427
                 end)
berghofe@11537
   428
          end;
berghofe@11537
   429
haftmann@28537
   430
    val (prem_p, gr') = compile_prems in_ts' arg_vs all_vs' ps gr ;
berghofe@11537
   431
  in
haftmann@28537
   432
    (Pretty.block [str "DSeq.single", Pretty.brk 1, inp,
haftmann@28537
   433
       str " :->", Pretty.brk 1, prem_p], gr')
berghofe@11537
   434
  end;
berghofe@11537
   435
haftmann@28537
   436
fun compile_pred thy defs dep module prfx all_vs arg_vs modes s cls mode gr =
berghofe@22271
   437
  let
berghofe@26975
   438
    val xs = map str (Name.variant_list arg_vs
berghofe@22271
   439
      (map (fn i => "x" ^ string_of_int i) (snd mode)));
haftmann@28537
   440
    val ((cl_ps, mode_id), gr') = gr |>
haftmann@28537
   441
      fold_map (fn cl => compile_clause thy defs
haftmann@28537
   442
        dep module all_vs arg_vs modes mode cl (mk_tuple xs)) cls ||>>
berghofe@22271
   443
      modename module s mode
berghofe@11537
   444
  in
haftmann@28537
   445
    (Pretty.block
berghofe@11537
   446
      ([Pretty.block (separate (Pretty.brk 1)
berghofe@26975
   447
         (str (prfx ^ mode_id) ::
berghofe@26975
   448
           map str arg_vs @
berghofe@26975
   449
           (case mode of ([], []) => [str "()"] | _ => xs)) @
berghofe@26975
   450
         [str " ="]),
berghofe@11537
   451
        Pretty.brk 1] @
wenzelm@32952
   452
       flat (separate [str " ++", Pretty.brk 1] (map single cl_ps))), (gr', "and "))
berghofe@11537
   453
  end;
berghofe@11537
   454
haftmann@28537
   455
fun compile_preds thy defs dep module all_vs arg_vs modes preds gr =
haftmann@28537
   456
  let val (prs, (gr', _)) = fold_map (fn (s, cls) =>
haftmann@28537
   457
    fold_map (fn mode => fn (gr', prfx') => compile_pred thy defs
haftmann@28537
   458
      dep module prfx' all_vs arg_vs modes s cls mode gr')
haftmann@28537
   459
        (((the o AList.lookup (op =) modes) s))) preds (gr, "fun ")
berghofe@11537
   460
  in
wenzelm@32952
   461
    (space_implode "\n\n" (map string_of (flat prs)) ^ ";\n\n", gr')
berghofe@11537
   462
  end;
berghofe@11537
   463
berghofe@11537
   464
(**** processing of introduction rules ****)
berghofe@11537
   465
berghofe@12557
   466
exception Modes of
berghofe@12557
   467
  (string * (int list option list * int list) list) list *
berghofe@22271
   468
  (string * (int option list * int)) list;
berghofe@12557
   469
wenzelm@32952
   470
fun lookup_modes gr dep = apfst flat (apsnd flat (ListPair.unzip
berghofe@17144
   471
  (map ((fn (SOME (Modes x), _, _) => x | _ => ([], [])) o get_node gr)
berghofe@17144
   472
    (Graph.all_preds (fst gr) [dep]))));
berghofe@12557
   473
berghofe@22271
   474
fun print_arities arities = message ("Arities:\n" ^
wenzelm@26931
   475
  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
berghofe@12557
   476
    space_implode " -> " (map
berghofe@22271
   477
      (fn NONE => "X" | SOME k' => string_of_int k')
berghofe@22271
   478
        (ks @ [SOME k]))) arities));
berghofe@11537
   479
wenzelm@32959
   480
fun prep_intrs intrs = map (rename_term o #prop o rep_thm o Drule.standard) intrs;
berghofe@15031
   481
berghofe@15031
   482
fun constrain cs [] = []
haftmann@22642
   483
  | constrain cs ((s, xs) :: ys) = (s, case AList.lookup (op =) cs (s : string) of
skalberg@15531
   484
      NONE => xs
haftmann@33037
   485
    | SOME xs' => gen_inter (op =) (xs, xs')) :: constrain cs ys;
berghofe@15031
   486
berghofe@17144
   487
fun mk_extra_defs thy defs gr dep names module ts =
skalberg@15570
   488
  Library.foldl (fn (gr, name) =>
berghofe@12557
   489
    if name mem names then gr
berghofe@12557
   490
    else (case get_clauses thy name of
skalberg@15531
   491
        NONE => gr
berghofe@22271
   492
      | SOME (names, thyname, nparms, intrs) =>
berghofe@17144
   493
          mk_ind_def thy defs gr dep names (if_library thyname module)
berghofe@22271
   494
            [] (prep_intrs intrs) nparms))
wenzelm@29287
   495
            (gr, fold Term.add_const_names ts [])
berghofe@12557
   496
berghofe@22271
   497
and mk_ind_def thy defs gr dep names module modecs intrs nparms =
berghofe@24129
   498
  add_edge_acyclic (hd names, dep) gr handle
berghofe@24129
   499
    Graph.CYCLES (xs :: _) =>
berghofe@24129
   500
      error ("InductiveCodegen: illegal cyclic dependencies:\n" ^ commas xs)
berghofe@24129
   501
  | Graph.UNDEF _ =>
berghofe@11537
   502
    let
berghofe@22271
   503
      val _ $ u = Logic.strip_imp_concl (hd intrs);
berghofe@22271
   504
      val args = List.take (snd (strip_comb u), nparms);
wenzelm@32952
   505
      val arg_vs = maps term_vs args;
berghofe@15031
   506
berghofe@22271
   507
      fun get_nparms s = if s mem names then SOME nparms else
berghofe@22271
   508
        Option.map #3 (get_clauses thy s);
berghofe@11537
   509
berghofe@26806
   510
      fun dest_prem (_ $ (Const ("op :", _) $ t $ u)) = Prem ([t], Envir.beta_eta_contract u, true)
berghofe@26806
   511
        | dest_prem (_ $ ((eq as Const ("op =", _)) $ t $ u)) = Prem ([t, u], eq, false)
berghofe@22271
   512
        | dest_prem (_ $ t) =
berghofe@22271
   513
            (case strip_comb t of
berghofe@26806
   514
               (v as Var _, ts) => if v mem args then Prem (ts, v, false) else Sidecond t
berghofe@22271
   515
             | (c as Const (s, _), ts) => (case get_nparms s of
berghofe@22271
   516
                 NONE => Sidecond t
berghofe@22271
   517
               | SOME k =>
berghofe@22271
   518
                   let val (ts1, ts2) = chop k ts
berghofe@26806
   519
                   in Prem (ts2, list_comb (c, ts1), false) end)
berghofe@22271
   520
             | _ => Sidecond t);
berghofe@22271
   521
berghofe@22271
   522
      fun add_clause intr (clauses, arities) =
berghofe@11537
   523
        let
berghofe@22271
   524
          val _ $ t = Logic.strip_imp_concl intr;
berghofe@22271
   525
          val (Const (name, T), ts) = strip_comb t;
berghofe@22271
   526
          val (ts1, ts2) = chop nparms ts;
berghofe@22271
   527
          val prems = map dest_prem (Logic.strip_imp_prems intr);
berghofe@22271
   528
          val (Ts, Us) = chop nparms (binder_types T)
berghofe@11537
   529
        in
berghofe@22271
   530
          (AList.update op = (name, these (AList.lookup op = clauses name) @
berghofe@22271
   531
             [(ts2, prems)]) clauses,
berghofe@22271
   532
           AList.update op = (name, (map (fn U => (case strip_type U of
berghofe@22271
   533
                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
berghofe@22271
   534
               | _ => NONE)) Ts,
berghofe@22271
   535
             length Us)) arities)
berghofe@11537
   536
        end;
berghofe@11537
   537
berghofe@16645
   538
      val gr' = mk_extra_defs thy defs
berghofe@17144
   539
        (add_edge (hd names, dep)
berghofe@17144
   540
          (new_node (hd names, (NONE, "", "")) gr)) (hd names) names module intrs;
berghofe@22271
   541
      val (extra_modes, extra_arities) = lookup_modes gr' (hd names);
berghofe@22271
   542
      val (clauses, arities) = fold add_clause intrs ([], []);
berghofe@15031
   543
      val modes = constrain modecs
berghofe@22271
   544
        (infer_modes thy extra_modes arities arg_vs clauses);
berghofe@22271
   545
      val _ = print_arities arities;
berghofe@11537
   546
      val _ = print_modes modes;
haftmann@28537
   547
      val (s, gr'') = compile_preds thy defs (hd names) module (terms_vs intrs)
haftmann@28537
   548
        arg_vs (modes @ extra_modes) clauses gr';
berghofe@11537
   549
    in
berghofe@17144
   550
      (map_node (hd names)
berghofe@22271
   551
        (K (SOME (Modes (modes, arities)), module, s)) gr'')
berghofe@16645
   552
    end;
berghofe@11537
   553
berghofe@22271
   554
fun find_mode gr dep s u modes is = (case find_first (fn Mode (_, js, _) => is=js)
wenzelm@15660
   555
  (modes_of modes u handle Option => []) of
berghofe@17144
   556
     NONE => codegen_error gr dep
berghofe@17144
   557
       ("No such mode for " ^ s ^ ": " ^ string_of_mode ([], is))
berghofe@15031
   558
   | mode => mode);
berghofe@15031
   559
haftmann@28537
   560
fun mk_ind_call thy defs dep module is_query s T ts names thyname k intrs gr =
berghofe@22271
   561
  let
berghofe@22271
   562
    val (ts1, ts2) = chop k ts;
berghofe@22271
   563
    val u = list_comb (Const (s, T), ts1);
berghofe@11537
   564
berghofe@22271
   565
    fun mk_mode (((ts, mode), i), Const ("dummy_pattern", _)) =
berghofe@22271
   566
          ((ts, mode), i+1)
berghofe@22271
   567
      | mk_mode (((ts, mode), i), t) = ((ts @ [t], mode @ [i]), i+1);
berghofe@11537
   568
berghofe@22271
   569
    val module' = if_library thyname module;
berghofe@22271
   570
    val gr1 = mk_extra_defs thy defs
berghofe@22271
   571
      (mk_ind_def thy defs gr dep names module'
berghofe@22271
   572
      [] (prep_intrs intrs) k) dep names module' [u];
berghofe@22271
   573
    val (modes, _) = lookup_modes gr1 dep;
berghofe@22271
   574
    val (ts', is) = if is_query then
berghofe@22271
   575
        fst (Library.foldl mk_mode ((([], []), 1), ts2))
berghofe@22271
   576
      else (ts2, 1 upto length (binder_types T) - k);
berghofe@22271
   577
    val mode = find_mode gr1 dep s u modes is;
haftmann@28537
   578
    val (in_ps, gr2) = fold_map (invoke_codegen thy defs dep module true) ts' gr1;
haftmann@28537
   579
    val (ps, gr3) = compile_expr thy defs dep module false modes (mode, u) gr2;
berghofe@22271
   580
  in
haftmann@28537
   581
    (Pretty.block (ps @ (if null in_ps then [] else [Pretty.brk 1]) @
haftmann@28537
   582
       separate (Pretty.brk 1) in_ps), gr3)
berghofe@22271
   583
  end;
berghofe@15031
   584
berghofe@15031
   585
fun clause_of_eqn eqn =
berghofe@15031
   586
  let
berghofe@15031
   587
    val (t, u) = HOLogic.dest_eq (HOLogic.dest_Trueprop (concl_of eqn));
berghofe@15031
   588
    val (Const (s, T), ts) = strip_comb t;
berghofe@15031
   589
    val (Ts, U) = strip_type T
berghofe@15031
   590
  in
berghofe@22271
   591
    rename_term (Logic.list_implies (prems_of eqn, HOLogic.mk_Trueprop
berghofe@22271
   592
      (list_comb (Const (s ^ "_aux", Ts @ [U] ---> HOLogic.boolT), ts @ [u]))))
berghofe@15031
   593
  end;
berghofe@15031
   594
berghofe@17144
   595
fun mk_fun thy defs name eqns dep module module' gr =
berghofe@17144
   596
  case try (get_node gr) name of
berghofe@17144
   597
    NONE =>
berghofe@15031
   598
    let
berghofe@15031
   599
      val clauses = map clause_of_eqn eqns;
berghofe@17144
   600
      val pname = name ^ "_aux";
berghofe@15031
   601
      val arity = length (snd (strip_comb (fst (HOLogic.dest_eq
berghofe@15031
   602
        (HOLogic.dest_Trueprop (concl_of (hd eqns)))))));
berghofe@15031
   603
      val mode = 1 upto arity;
haftmann@28537
   604
      val ((fun_id, mode_id), gr') = gr |>
haftmann@28537
   605
        mk_const_id module' name ||>>
berghofe@17144
   606
        modename module' pname ([], mode);
berghofe@26975
   607
      val vars = map (fn i => str ("x" ^ string_of_int i)) mode;
berghofe@26975
   608
      val s = string_of (Pretty.block
berghofe@26975
   609
        [mk_app false (str ("fun " ^ snd fun_id)) vars, str " =",
berghofe@26975
   610
         Pretty.brk 1, str "DSeq.hd", Pretty.brk 1,
berghofe@26975
   611
         parens (Pretty.block (separate (Pretty.brk 1) (str mode_id ::
berghofe@22271
   612
           vars)))]) ^ ";\n\n";
berghofe@17144
   613
      val gr'' = mk_ind_def thy defs (add_edge (name, dep)
berghofe@17144
   614
        (new_node (name, (NONE, module', s)) gr')) name [pname] module'
berghofe@22271
   615
        [(pname, [([], mode)])] clauses 0;
berghofe@17144
   616
      val (modes, _) = lookup_modes gr'' dep;
berghofe@22271
   617
      val _ = find_mode gr'' dep pname (head_of (HOLogic.dest_Trueprop
berghofe@22271
   618
        (Logic.strip_imp_concl (hd clauses)))) modes mode
haftmann@28537
   619
    in (mk_qual_id module fun_id, gr'') end
berghofe@17144
   620
  | SOME _ =>
haftmann@28537
   621
      (mk_qual_id module (get_const_id gr name), add_edge (name, dep) gr);
berghofe@15031
   622
berghofe@23761
   623
(* convert n-tuple to nested pairs *)
berghofe@23761
   624
berghofe@23761
   625
fun conv_ntuple fs ts p =
berghofe@23761
   626
  let
berghofe@23761
   627
    val k = length fs;
berghofe@26975
   628
    val xs = map (fn i => str ("x" ^ string_of_int i)) (0 upto k);
berghofe@23761
   629
    val xs' = map (fn Bound i => nth xs (k - i)) ts;
berghofe@23761
   630
    fun conv xs js =
berghofe@23761
   631
      if js mem fs then
berghofe@23761
   632
        let
berghofe@23761
   633
          val (p, xs') = conv xs (1::js);
berghofe@23761
   634
          val (q, xs'') = conv xs' (2::js)
berghofe@23761
   635
        in (mk_tuple [p, q], xs'') end
berghofe@23761
   636
      else (hd xs, tl xs)
berghofe@23761
   637
  in
berghofe@23761
   638
    if k > 0 then
berghofe@23761
   639
      Pretty.block
berghofe@26975
   640
        [str "DSeq.map (fn", Pretty.brk 1,
berghofe@26975
   641
         mk_tuple xs', str " =>", Pretty.brk 1, fst (conv xs []),
berghofe@26975
   642
         str ")", Pretty.brk 1, parens p]
berghofe@23761
   643
    else p
berghofe@23761
   644
  end;
berghofe@23761
   645
haftmann@28537
   646
fun inductive_codegen thy defs dep module brack t gr  = (case strip_comb t of
berghofe@23761
   647
    (Const ("Collect", _), [u]) =>
haftmann@32342
   648
      let val (r, Ts, fs) = HOLogic.strip_psplits u
berghofe@23761
   649
      in case strip_comb r of
berghofe@23761
   650
          (Const (s, T), ts) =>
berghofe@23761
   651
            (case (get_clauses thy s, get_assoc_code thy (s, T)) of
berghofe@23761
   652
              (SOME (names, thyname, k, intrs), NONE) =>
berghofe@23761
   653
                let
berghofe@23761
   654
                  val (ts1, ts2) = chop k ts;
berghofe@23761
   655
                  val ts2' = map
berghofe@23761
   656
                    (fn Bound i => Term.dummy_pattern (nth Ts i) | t => t) ts2;
berghofe@23761
   657
                  val (ots, its) = List.partition is_Bound ts2;
berghofe@23761
   658
                  val no_loose = forall (fn t => not (loose_bvar (t, 0)))
berghofe@23761
   659
                in
berghofe@23761
   660
                  if null (duplicates op = ots) andalso
berghofe@23761
   661
                    no_loose ts1 andalso no_loose its
berghofe@23761
   662
                  then
haftmann@28537
   663
                    let val (call_p, gr') = mk_ind_call thy defs dep module true
haftmann@28537
   664
                      s T (ts1 @ ts2') names thyname k intrs gr 
haftmann@28537
   665
                    in SOME ((if brack then parens else I) (Pretty.block
haftmann@31852
   666
                      [str "Set", Pretty.brk 1, str "(DSeq.list_of", Pretty.brk 1, str "(",
haftmann@31852
   667
                       conv_ntuple fs ots call_p, str "))"]),
haftmann@31852
   668
                       (*this is a very strong assumption about the generated code!*)
haftmann@31852
   669
                       gr')
berghofe@23761
   670
                    end
berghofe@23761
   671
                  else NONE
berghofe@23761
   672
                end
berghofe@23761
   673
            | _ => NONE)
berghofe@23761
   674
        | _ => NONE
berghofe@23761
   675
      end
berghofe@22271
   676
  | (Const (s, T), ts) => (case Symtab.lookup (#eqns (CodegenData.get thy)) s of
haftmann@22921
   677
      NONE => (case (get_clauses thy s, get_assoc_code thy (s, T)) of
berghofe@22271
   678
        (SOME (names, thyname, k, intrs), NONE) =>
berghofe@22271
   679
          if length ts < k then NONE else SOME
haftmann@28537
   680
            (let val (call_p, gr') = mk_ind_call thy defs dep module false
haftmann@28537
   681
               s T (map Term.no_dummy_patterns ts) names thyname k intrs gr
haftmann@28537
   682
             in (mk_funcomp brack "?!"
haftmann@28537
   683
               (length (binder_types T) - length ts) (parens call_p), gr')
haftmann@28537
   684
             end handle TERM _ => mk_ind_call thy defs dep module true
haftmann@28537
   685
               s T ts names thyname k intrs gr )
berghofe@22271
   686
      | _ => NONE)
berghofe@22271
   687
    | SOME eqns =>
berghofe@22271
   688
        let
haftmann@22642
   689
          val (_, thyname) :: _ = eqns;
haftmann@28537
   690
          val (id, gr') = mk_fun thy defs s (preprocess thy (map fst (rev eqns)))
berghofe@22271
   691
            dep module (if_library thyname module) gr;
haftmann@28537
   692
          val (ps, gr'') = fold_map
haftmann@28537
   693
            (invoke_codegen thy defs dep module true) ts gr';
haftmann@28537
   694
        in SOME (mk_app brack (str id) ps, gr'')
berghofe@22271
   695
        end)
berghofe@22271
   696
  | _ => NONE);
berghofe@11537
   697
berghofe@12557
   698
val setup =
wenzelm@18708
   699
  add_codegen "inductive" inductive_codegen #>
haftmann@31998
   700
  Attrib.setup @{binding code_ind} (Scan.lift (Scan.option (Args.$$$ "target" |-- Args.colon |-- Args.name) --
haftmann@31998
   701
    Scan.option (Args.$$$ "params" |-- Args.colon |-- OuterParse.nat) >> uncurry add))
haftmann@31998
   702
    "introduction rules for executable predicates";
berghofe@11537
   703
berghofe@11537
   704
end;