| author | wenzelm | 
| Mon, 20 Sep 2021 13:51:32 +0200 | |
| changeset 74324 | 308e74afab83 | 
| parent 74282 | c2ee8d993d6a | 
| child 76052 | 6a20d0ebd5b3 | 
| permissions | -rw-r--r-- | 
| 43107 | 1 | (* Title: HOL/Tools/monomorph.ML | 
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 3 | |
| 51575 | 4 | Monomorphization of theorems, i.e., computation of ground instances for | 
| 5 | theorems with type variables. This procedure is incomplete in general, | |
| 6 | but works well for most practical problems. | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 7 | |
| 51575 | 8 | Monomorphization is essentially an enumeration of substitutions that map | 
| 9 | schematic types to ground types. Applying these substitutions to theorems | |
| 10 | with type variables results in monomorphized ground instances. The | |
| 11 | enumeration is driven by schematic constants (constants occurring with | |
| 12 | type variables) and all ground instances of such constants (occurrences | |
| 13 | without type variables). The enumeration is organized in rounds in which | |
| 14 | all substitutions for the schematic constants are computed that are induced | |
| 15 | by the ground instances. Any new ground instance may induce further | |
| 16 | substitutions in a subsequent round. To prevent nontermination, there is | |
| 17 | an upper limit of rounds involved and of the number of monomorphized ground | |
| 18 | instances computed. | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 19 | |
| 51575 | 20 | Theorems to be monomorphized must be tagged with a number indicating the | 
| 21 | initial round in which they participate first. The initial round is | |
| 22 | ignored for theorems without type variables. For any other theorem, the | |
| 23 | initial round must be greater than zero. Returned monomorphized theorems | |
| 24 | carry a number showing from which monomorphization round they emerged. | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 25 | *) | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 26 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 27 | signature MONOMORPH = | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 28 | sig | 
| 51575 | 29 | (* utility functions *) | 
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 30 | val typ_has_tvars: typ -> bool | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 31 | val all_schematic_consts_of: term -> typ list Symtab.table | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 32 | val add_schematic_consts_of: term -> typ list Symtab.table -> | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 33 | typ list Symtab.table | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 34 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 35 | (* configuration options *) | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 36 | val max_rounds: int Config.T | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 37 | val max_new_instances: int Config.T | 
| 53480 
247817dbb990
limit the number of instances of a single theorem
 blanchet parents: 
51575diff
changeset | 38 | val max_thm_instances: int Config.T | 
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 39 | val max_new_const_instances_per_round: int Config.T | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 40 | val max_duplicated_instances: int Config.T | 
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 41 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 42 | (* monomorphization *) | 
| 51575 | 43 | val monomorph: (term -> typ list Symtab.table) -> Proof.context -> | 
| 44 | (int * thm) list -> (int * thm) list list | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 45 | end | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 46 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 47 | structure Monomorph: MONOMORPH = | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 48 | struct | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 49 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 50 | (* utility functions *) | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 51 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 52 | val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false) | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 53 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 54 | fun add_schematic_const (c as (_, T)) = | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 55 | if typ_has_tvars T then Symtab.insert_list (op =) c else I | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 56 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 57 | fun add_schematic_consts_of t = | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 58 | Term.fold_aterms (fn Const c => add_schematic_const c | _ => I) t | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 59 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 60 | fun all_schematic_consts_of t = add_schematic_consts_of t Symtab.empty | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 61 | |
| 51575 | 62 | fun clear_grounds grounds = Symtab.map (K (K [])) grounds | 
| 63 | ||
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 64 | |
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 65 | (* configuration options *) | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 66 | |
| 67149 | 67 | val max_rounds = Attrib.setup_config_int \<^binding>\<open>monomorph_max_rounds\<close> (K 5) | 
| 51575 | 68 | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 69 | val max_new_instances = | 
| 67149 | 70 | Attrib.setup_config_int \<^binding>\<open>monomorph_max_new_instances\<close> (K 500) | 
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 71 | |
| 53480 
247817dbb990
limit the number of instances of a single theorem
 blanchet parents: 
51575diff
changeset | 72 | val max_thm_instances = | 
| 67149 | 73 | Attrib.setup_config_int \<^binding>\<open>monomorph_max_thm_instances\<close> (K 20) | 
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 74 | |
| 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 75 | val max_new_const_instances_per_round = | 
| 67149 | 76 | Attrib.setup_config_int \<^binding>\<open>monomorph_max_new_const_instances_per_round\<close> (K 5) | 
| 53480 
247817dbb990
limit the number of instances of a single theorem
 blanchet parents: 
51575diff
changeset | 77 | |
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 78 | val max_duplicated_instances = | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 79 | Attrib.setup_config_int \<^binding>\<open>monomorph_max_duplicated_instances\<close> (K 16000) | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 80 | |
| 43116 | 81 | fun limit_rounds ctxt f = | 
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 82 | let | 
| 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 83 | val max = Config.get ctxt max_rounds | 
| 43117 | 84 | fun round i x = if i > max then x else round (i + 1) (f ctxt i x) | 
| 85 | in round 1 end | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 86 | |
| 51575 | 87 | |
| 88 | (* theorem information and related functions *) | |
| 89 | ||
| 90 | datatype thm_info = | |
| 91 | Ground of thm | | |
| 92 | Ignored | | |
| 93 |   Schematic of {
 | |
| 94 | id: int, | |
| 95 | theorem: thm, | |
| 96 | tvars: (indexname * sort) list, | |
| 97 | schematics: (string * typ) list, | |
| 98 | initial_round: int} | |
| 99 | ||
| 100 | fun fold_grounds f = fold (fn Ground thm => f thm | _ => I) | |
| 101 | ||
| 102 | fun fold_schematic f thm_info = | |
| 103 | (case thm_info of | |
| 104 |     Schematic {id, theorem, tvars, schematics, initial_round} =>
 | |
| 105 | f id theorem tvars schematics initial_round | |
| 106 | | _ => I) | |
| 107 | ||
| 108 | fun fold_schematics pred f = | |
| 109 | let | |
| 110 | fun apply id thm tvars schematics initial_round x = | |
| 111 | if pred initial_round then f id thm tvars schematics x else x | |
| 112 | in fold (fold_schematic apply) end | |
| 113 | ||
| 114 | ||
| 115 | (* collecting data *) | |
| 116 | ||
| 117 | (* | |
| 118 | Theorems with type variables that cannot be instantiated should be ignored. | |
| 119 | A type variable has only a chance to be instantiated if it occurs in the | |
| 120 | type of one of the schematic constants. | |
| 121 | *) | |
| 122 | fun groundable all_tvars schematics = | |
| 123 | let val tvars' = Symtab.fold (fold Term.add_tvarsT o snd) schematics [] | |
| 124 | in forall (member (op =) tvars') all_tvars end | |
| 125 | ||
| 126 | ||
| 127 | fun prepare schematic_consts_of rthms = | |
| 43116 | 128 | let | 
| 51575 | 129 | fun prep (initial_round, thm) ((id, idx), consts) = | 
| 130 | if Term.exists_type typ_has_tvars (Thm.prop_of thm) then | |
| 131 | let | |
| 132 | (* increase indices to avoid clashes of type variables *) | |
| 133 | val thm' = Thm.incr_indexes idx thm | |
| 134 | val idx' = Thm.maxidx_of thm' + 1 | |
| 135 | ||
| 136 | val tvars = Term.add_tvars (Thm.prop_of thm') [] | |
| 137 | val schematics = schematic_consts_of (Thm.prop_of thm') | |
| 138 | val schematics' = | |
| 139 | Symtab.fold (fn (n, Ts) => fold (cons o pair n) Ts) schematics [] | |
| 140 | ||
| 141 | (* collect the names of all constants that need to be instantiated *) | |
| 142 | val consts' = | |
| 143 | consts | |
| 144 | |> Symtab.fold (fn (n, _) => Symtab.update (n, [])) schematics | |
| 145 | ||
| 146 | val thm_info = | |
| 147 | if not (groundable tvars schematics) then Ignored | |
| 148 | else | |
| 149 |               Schematic {
 | |
| 150 | id = id, | |
| 151 | theorem = thm', | |
| 152 | tvars = tvars, | |
| 153 | schematics = schematics', | |
| 154 | initial_round = initial_round} | |
| 155 | in (thm_info, ((id + 1, idx'), consts')) end | |
| 156 | else (Ground thm, ((id + 1, idx + Thm.maxidx_of thm + 1), consts)) | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 157 | in | 
| 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 158 | fold_map prep rthms ((0, 0), Symtab.empty) ||> snd | 
| 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 159 | end | 
| 51575 | 160 | |
| 161 | ||
| 162 | (* collecting instances *) | |
| 163 | ||
| 59642 | 164 | fun instantiate ctxt subst = | 
| 51575 | 165 | let | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
59642diff
changeset | 166 | fun cert (ix, (S, T)) = ((ix, S), Thm.ctyp_of ctxt T) | 
| 51575 | 167 | fun cert' subst = Vartab.fold (cons o cert) subst [] | 
| 74282 | 168 | in Thm.instantiate (TVars.make (cert' subst), Vars.empty) end | 
| 51575 | 169 | |
| 170 | fun add_new_grounds used_grounds new_grounds thm = | |
| 171 | let | |
| 172 | fun mem tab (n, T) = member (op =) (Symtab.lookup_list tab n) T | |
| 173 | fun add (Const (c as (n, _))) = | |
| 174 | if mem used_grounds c orelse mem new_grounds c then I | |
| 175 | else if not (Symtab.defined used_grounds n) then I | |
| 176 | else Symtab.insert_list (op =) c | |
| 177 | | add _ = I | |
| 178 | in Term.fold_aterms add (Thm.prop_of thm) end | |
| 179 | ||
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 180 | fun add_insts max_instances max_duplicated_instances max_thm_insts ctxt round used_grounds | 
| 53482 | 181 | new_grounds id thm tvars schematics cx = | 
| 51575 | 182 | let | 
| 183 | exception ENOUGH of | |
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 184 | typ list Symtab.table * (int * int * ((int * (sort * typ) Vartab.table) * thm) list Inttab.table) | 
| 51575 | 185 | |
| 186 | val thy = Proof_Context.theory_of ctxt | |
| 187 | ||
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 188 | fun add subst (cx as (next_grounds, (hits, misses, insts))) = | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 189 | if hits >= max_instances orelse misses >= max_duplicated_instances then | 
| 53481 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 190 | raise ENOUGH cx | 
| 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 191 | else | 
| 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 192 | let | 
| 59642 | 193 | val thm' = instantiate ctxt subst thm | 
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 194 | val rthm = ((round, subst), thm') | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 195 | val rthms = Inttab.lookup_list insts id | 
| 53481 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 196 | val n_insts' = | 
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 197 | if member (eq_snd Thm.eq_thm) rthms rthm then | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 198 | (hits, misses + 1, insts) | 
| 53481 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 199 | else | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 200 | let | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 201 | val (hits', misses') = | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 202 | if length rthms >= max_thm_insts then (hits, misses + 1) else (hits + 1, misses) | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 203 | in | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 204 | (hits', misses', Inttab.cons_list (id, rthm) insts) | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 205 | end | 
| 53481 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 206 | val next_grounds' = | 
| 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 207 | add_new_grounds used_grounds new_grounds thm' next_grounds | 
| 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 208 | in (next_grounds', n_insts') end | 
| 51575 | 209 | |
| 210 | fun with_grounds (n, T) f subst (n', Us) = | |
| 211 | let | |
| 212 | fun matching U = (* one-step refinement of the given substitution *) | |
| 213 | (case try (Sign.typ_match thy (T, U)) subst of | |
| 214 | NONE => I | |
| 215 | | SOME subst' => f subst') | |
| 216 | in if n = n' then fold matching Us else I end | |
| 217 | ||
| 218 | fun with_matching_ground c subst f = | |
| 219 | (* Try new grounds before already used grounds. Otherwise only | |
| 220 | substitutions already seen in previous rounds get enumerated. *) | |
| 221 | Symtab.fold (with_grounds c (f true) subst) new_grounds #> | |
| 222 | Symtab.fold (with_grounds c (f false) subst) used_grounds | |
| 223 | ||
| 224 | fun is_complete subst = | |
| 225 | (* Check if a substitution is defined for all TVars of the theorem, | |
| 226 | which guarantees that the instantiation with this substitution results | |
| 227 | in a ground theorem since all matchings that led to this substitution | |
| 228 | are with ground types only. *) | |
| 229 | subset (op =) (tvars, Vartab.fold (cons o apsnd fst) subst []) | |
| 230 | ||
| 231 | fun for_schematics _ [] _ = I | |
| 232 | | for_schematics used_new (c :: cs) subst = | |
| 233 | with_matching_ground c subst (fn new => fn subst' => | |
| 234 | if is_complete subst' then | |
| 235 | if used_new orelse new then add subst' | |
| 236 | else I | |
| 237 | else for_schematics (used_new orelse new) cs subst') #> | |
| 238 | for_schematics used_new cs subst | |
| 43116 | 239 | in | 
| 51575 | 240 | (* Enumerate all substitutions that lead to a ground instance of the | 
| 241 | theorem not seen before. A necessary condition for such a new ground | |
| 242 | instance is the usage of at least one ground from the new_grounds | |
| 243 | table. The approach used here is to match all schematics of the theorem | |
| 244 | with all relevant grounds. *) | |
| 245 | for_schematics false schematics Vartab.empty cx | |
| 246 | handle ENOUGH cx' => cx' | |
| 43116 | 247 | end | 
| 248 | ||
| 51575 | 249 | fun is_new round initial_round = (round = initial_round) | 
| 250 | fun is_active round initial_round = (round > initial_round) | |
| 251 | ||
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 252 | fun find_instances max_instances max_duplicated_instances max_thm_insts max_new_grounds thm_infos | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 253 | ctxt round (known_grounds, new_grounds0, insts) = | 
| 51575 | 254 | let | 
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 255 | val new_grounds = | 
| 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 256 | Symtab.map (fn _ => fn grounds => | 
| 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 257 | if length grounds <= max_new_grounds then grounds | 
| 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 258 | else take max_new_grounds (sort Term_Ord.typ_ord grounds)) new_grounds0 | 
| 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 259 | |
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 260 | val add_new = add_insts max_instances max_duplicated_instances max_thm_insts ctxt round | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 261 | fun consider_all pred f (cx as (_, (hits, misses, _))) = | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 262 | if hits >= max_instances orelse misses >= max_duplicated_instances then | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 263 | cx | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 264 | else | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 265 | fold_schematics pred f thm_infos cx | 
| 51575 | 266 | |
| 267 | val known_grounds' = Symtab.merge_list (op =) (known_grounds, new_grounds) | |
| 268 | val empty_grounds = clear_grounds known_grounds' | |
| 269 | ||
| 270 | val (new_grounds', insts') = | |
| 271 | (Symtab.empty, insts) | |
| 272 | |> consider_all (is_active round) (add_new known_grounds new_grounds) | |
| 273 | |> consider_all (is_new round) (add_new empty_grounds known_grounds') | |
| 57209 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 274 | in | 
| 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 275 | (known_grounds', new_grounds', insts') | 
| 
7ffa0f7e2775
removed '_new' sufffix in SMT2 solver names (in some cases)
 blanchet parents: 
54061diff
changeset | 276 | end | 
| 51575 | 277 | |
| 278 | fun add_ground_types thm = | |
| 279 | let fun add (n, T) = Symtab.map_entry n (insert (op =) T) | |
| 280 | in Term.fold_aterms (fn Const c => add c | _ => I) (Thm.prop_of thm) end | |
| 281 | ||
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 282 | fun collect_instances ctxt max_thm_insts max_new_grounds thm_infos consts = | 
| 51575 | 283 | let | 
| 284 | val known_grounds = fold_grounds add_ground_types thm_infos consts | |
| 285 | val empty_grounds = clear_grounds known_grounds | |
| 53482 | 286 | val max_instances = Config.get ctxt max_new_instances | 
| 53481 
0721fc9d0fe7
made semantics of "max_new_instances" be what the name suggests, the previous implementation did, and the Sledgehammer manual documents
 blanchet parents: 
53480diff
changeset | 287 | |> fold (fn Schematic _ => Integer.add 1 | _ => I) thm_infos | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 288 | val max_duplicated_instances = Config.get ctxt max_duplicated_instances | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 289 | val (_, _, (_, _, insts)) = | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 290 | limit_rounds ctxt (find_instances max_instances max_duplicated_instances max_thm_insts | 
| 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 291 | max_new_grounds thm_infos) (empty_grounds, known_grounds, (0, 0, Inttab.empty)) | 
| 51575 | 292 | in | 
| 73376 
96ef620c8b1e
added upper bound on monomorphisation duplicate instances
 desharna parents: 
67149diff
changeset | 293 | insts | 
| 51575 | 294 | end | 
| 295 | ||
| 296 | ||
| 297 | (* monomorphization *) | |
| 298 | ||
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 299 | fun size_of_subst subst = | 
| 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 300 | Vartab.fold (Integer.add o size_of_typ o snd o snd) subst 0 | 
| 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 301 | |
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57209diff
changeset | 302 | fun subst_ord subst = int_ord (apply2 size_of_subst subst) | 
| 51575 | 303 | |
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 304 | fun instantiated_thms _ _ (Ground thm) = [(0, thm)] | 
| 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 305 | | instantiated_thms _ _ Ignored = [] | 
| 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 306 |   | instantiated_thms max_thm_insts insts (Schematic {id, ...}) =
 | 
| 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 307 | Inttab.lookup_list insts id | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57209diff
changeset | 308 | |> (fn rthms => | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57209diff
changeset | 309 | if length rthms <= max_thm_insts then rthms | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
57209diff
changeset | 310 | else take max_thm_insts (sort (prod_ord int_ord subst_ord o apply2 fst) rthms)) | 
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 311 | |> map (apfst fst) | 
| 51575 | 312 | |
| 313 | fun monomorph schematic_consts_of ctxt rthms = | |
| 314 | let | |
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 315 | val max_thm_insts = Config.get ctxt max_thm_instances | 
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 316 | val max_new_grounds = Config.get ctxt max_new_const_instances_per_round | 
| 51575 | 317 | val (thm_infos, consts) = prepare schematic_consts_of rthms | 
| 318 | val insts = | |
| 319 | if Symtab.is_empty consts then Inttab.empty | |
| 54061 
6807b8e95adb
prevent explosion in monomorphizer (e.g. when the facts typerep_int_def typerep_num_def typerep_option_def typerep_node_def are selected)
 blanchet parents: 
53833diff
changeset | 320 | else collect_instances ctxt max_thm_insts max_new_grounds thm_infos consts | 
| 53823 
191ec7f873d5
when "max_thm_instances" is hit, choose more carefully which instances should be kept
 blanchet parents: 
53482diff
changeset | 321 | in map (instantiated_thms max_thm_insts insts) thm_infos end | 
| 51575 | 322 | |
| 43041 
218e3943d504
added re-implemented monomorphizer for types with better control for number of generated instances (strict as opposed to the SMT monomorphizer) and with fact annotations controlling in which round certain facts are considered for monomorphization
 boehmes parents: diff
changeset | 323 | end |