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