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