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