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