author | bulwahn |
Wed, 20 Apr 2011 16:00:45 +0200 | |
changeset 42434 | 1914fd5d7c0e |
parent 42361 | 23f352990944 |
child 42744 | 59753d5448e0 |
permissions | -rw-r--r-- |
36898 | 1 |
(* Title: HOL/Tools/SMT/smt_monomorph.ML |
2 |
Author: Sascha Boehme, TU Muenchen |
|
3 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
4 |
Monomorphization of theorems, i.e., computation of all (necessary) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
5 |
instances. This procedure is incomplete in general, but works well for |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
6 |
most practical problems. |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
7 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
8 |
For a list of universally closed theorems (without schematic term |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
9 |
variables), monomorphization computes a list of theorems with schematic |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
10 |
term variables: all polymorphic constants (i.e., constants occurring both |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
11 |
with ground types and schematic type variables) are instantiated with all |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
12 |
(necessary) ground types; thereby theorems containing these constants are |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
13 |
copied. To prevent non-termination, there is an upper limit for the number |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
14 |
of iterations involved in the fixpoint construction. |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
15 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
16 |
The search for instances is performed on the constants with schematic |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
17 |
types, which are extracted from the initial set of theorems. The search |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
18 |
constructs, for each theorem with those constants, a set of substitutions, |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
19 |
which, in the end, is applied to all corresponding theorems. Remaining |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
20 |
schematic type variables are substituted with fresh types. |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
21 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
22 |
Searching for necessary substitutions is an iterative fixpoint |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
23 |
construction: each iteration computes all required instances required by |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
24 |
the ground instances computed in the previous step and which haven't been |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
25 |
found before. Computed substitutions are always nontrivial: schematic type |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
26 |
variables are never mapped to schematic type variables. |
36898 | 27 |
*) |
28 |
||
29 |
signature SMT_MONOMORPH = |
|
30 |
sig |
|
42181
8f25605e646c
temporary workaround: filter out spurious monomorphic instances
blanchet
parents:
42180
diff
changeset
|
31 |
val typ_has_tvars: typ -> bool |
42190
b6b5846504cd
make configuration of (SMT) full monomorphization more flexible: turn boolean argument into a configuration option
boehmes
parents:
42183
diff
changeset
|
32 |
val monomorph: ('a * thm) list -> Proof.context -> |
42180
a6c141925a8a
added monomorphization option to Sledgehammer ATPs -- this looks promising but is still off by default
blanchet
parents:
41762
diff
changeset
|
33 |
('a * thm) list * Proof.context |
36898 | 34 |
end |
35 |
||
36 |
structure SMT_Monomorph: SMT_MONOMORPH = |
|
37 |
struct |
|
38 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
39 |
(* utility functions *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
40 |
|
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
41 |
fun fold_maps f = fold (fn x => uncurry (fold_map (f x)) #>> flat) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
42 |
|
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
43 |
fun pair_trans ((x, y), z) = (x, (y, z)) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
44 |
|
36898 | 45 |
val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false) |
46 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
47 |
val ignored = member (op =) [@{const_name All}, @{const_name Ex}, |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
48 |
@{const_name Let}, @{const_name If}, @{const_name HOL.eq}] |
36898 | 49 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
50 |
fun is_const pred (n, T) = not (ignored n) andalso pred T |
36898 | 51 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
52 |
fun collect_consts_if pred f = |
41174
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
53 |
let |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
54 |
fun collect (@{const trigger} $ p $ t) = collect_trigger p #> collect t |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
55 |
| collect (t $ u) = collect t #> collect u |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
56 |
| collect (Abs (_, _, t)) = collect t |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
57 |
| collect (Const c) = if is_const pred c then f c else I |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
58 |
| collect _ = I |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
59 |
and collect_trigger t = |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
60 |
let val dest = these o try HOLogic.dest_list |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
61 |
in fold (fold collect_pat o dest) (dest t) end |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
62 |
and collect_pat (Const (@{const_name SMT.pat}, _) $ t) = collect t |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
63 |
| collect_pat (Const (@{const_name SMT.nopat}, _) $ t) = collect t |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
64 |
| collect_pat _ = I |
10eb369f8c01
fixed trigger inference: testing if a theorem already has a trigger was too strict;
boehmes
parents:
41063
diff
changeset
|
65 |
in collect o Thm.prop_of end |
36898 | 66 |
|
39687 | 67 |
val insert_const = Ord_List.insert (prod_ord fast_string_ord Term_Ord.typ_ord) |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
68 |
|
36898 | 69 |
fun tvar_consts_of thm = collect_consts_if typ_has_tvars insert_const thm [] |
70 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
71 |
fun add_const_types pred = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
72 |
collect_consts_if pred (fn (n, T) => Symtab.map_entry n (insert (op =) T)) |
36898 | 73 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
74 |
fun incr_indexes ithms = |
40161
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39687
diff
changeset
|
75 |
let |
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39687
diff
changeset
|
76 |
fun inc (i, thm) idx = |
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39687
diff
changeset
|
77 |
((i, Thm.incr_indexes idx thm), Thm.maxidx_of thm + idx + 1) |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
78 |
in fst (fold_map inc ithms 0) end |
36898 | 79 |
|
80 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
81 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
82 |
(* search for necessary substitutions *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
83 |
|
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
84 |
fun new_substitutions thy limit grounds (n, T) subst instances = |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
85 |
if not (typ_has_tvars T) then ([subst], instances) |
36898 | 86 |
else |
87 |
Symtab.lookup_list grounds n |
|
88 |
|> map_filter (try (fn U => Sign.typ_match thy (T, U) subst)) |
|
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
89 |
|> (fn substs => (substs, instances - length substs)) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
90 |
|>> take limit (* limit the breadth of the search as well as the width *) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
91 |
|>> cons subst |
36898 | 92 |
|
93 |
fun apply_subst grounds consts subst = |
|
94 |
let |
|
95 |
fun is_new_ground (n, T) = not (typ_has_tvars T) andalso |
|
96 |
not (member (op =) (Symtab.lookup_list grounds n) T) |
|
97 |
||
98 |
fun apply_const (n, T) new_grounds = |
|
99 |
let val c = (n, Envir.subst_type subst T) |
|
100 |
in |
|
101 |
new_grounds |
|
102 |
|> is_new_ground c ? Symtab.insert_list (op =) c |
|
103 |
|> pair c |
|
104 |
end |
|
105 |
in fold_map apply_const consts #>> pair subst end |
|
106 |
||
41212
2781e8c76165
impose a limit on the breadth of monomorphization (in addition to on the depth) to prevent an explosion of the number of monomorphic instances
blanchet
parents:
41174
diff
changeset
|
107 |
fun specialize thy limit all_grounds new_grounds scs = |
36898 | 108 |
let |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
109 |
fun spec (subst, consts) (next_grounds, instances) = |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
110 |
([subst], instances) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
111 |
|> fold_maps (new_substitutions thy limit new_grounds) consts |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
112 |
|>> rpair next_grounds |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
113 |
|>> uncurry (fold_map (apply_subst all_grounds consts)) |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
114 |
|> pair_trans |
36898 | 115 |
in |
116 |
fold_map spec scs #>> (fn scss => |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
117 |
fold (fold (insert (eq_snd (op =)))) scss []) |
36898 | 118 |
end |
119 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
120 |
val limit_reached_warning = "Warning: Monomorphization limit reached" |
36898 | 121 |
|
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
122 |
fun search_substitutions ctxt limit instances all_grounds new_grounds scss = |
36898 | 123 |
let |
42361 | 124 |
val thy = Proof_Context.theory_of ctxt |
36898 | 125 |
val all_grounds' = Symtab.merge_list (op =) (all_grounds, new_grounds) |
41212
2781e8c76165
impose a limit on the breadth of monomorphization (in addition to on the depth) to prevent an explosion of the number of monomorphic instances
blanchet
parents:
41174
diff
changeset
|
126 |
val spec = specialize thy limit all_grounds' new_grounds |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
127 |
val (scss', (new_grounds', instances')) = |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
128 |
fold_map spec scss (Symtab.empty, instances) |
36898 | 129 |
in |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
130 |
if Symtab.is_empty new_grounds' then scss' |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
131 |
else if limit > 0 andalso instances' > 0 then |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
132 |
search_substitutions ctxt (limit-1) instances' all_grounds' new_grounds' |
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
133 |
scss' |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
134 |
else (SMT_Config.verbose_msg ctxt (K limit_reached_warning) (); scss') |
36898 | 135 |
end |
136 |
||
137 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
138 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
139 |
(* instantiation *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
140 |
|
36898 | 141 |
fun filter_most_specific thy = |
142 |
let |
|
143 |
fun typ_match (_, T) (_, U) = Sign.typ_match thy (T, U) |
|
144 |
||
145 |
fun is_trivial subst = Vartab.is_empty subst orelse |
|
146 |
forall (fn (v, (S, T)) => TVar (v, S) = T) (Vartab.dest subst) |
|
147 |
||
148 |
fun match general specific = |
|
149 |
(case try (fold2 typ_match general specific) Vartab.empty of |
|
150 |
NONE => false |
|
151 |
| SOME subst => not (is_trivial subst)) |
|
152 |
||
153 |
fun most_specific _ [] = [] |
|
154 |
| most_specific css ((ss, cs) :: scs) = |
|
155 |
let val substs = most_specific (cs :: css) scs |
|
156 |
in |
|
157 |
if exists (match cs) css orelse exists (match cs o snd) scs |
|
158 |
then substs else ss :: substs |
|
159 |
end |
|
160 |
||
161 |
in most_specific [] end |
|
162 |
||
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
163 |
fun instantiate full (i, thm) substs (ithms, ctxt) = |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
164 |
let |
42361 | 165 |
val thy = Proof_Context.theory_of ctxt |
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
166 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
167 |
val (vs, Ss) = split_list (Term.add_tvars (Thm.prop_of thm) []) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
168 |
val (Tenv, ctxt') = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
169 |
ctxt |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
170 |
|> Variable.invent_types Ss |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
171 |
|>> map2 (fn v => fn (n, S) => (v, (S, TFree (n, S)))) vs |
36898 | 172 |
|
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
173 |
exception PARTIAL_INST of unit |
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
174 |
|
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
175 |
fun update_subst vT subst = |
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
176 |
if full then Vartab.update vT subst |
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
177 |
else raise PARTIAL_INST () |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
178 |
|
36898 | 179 |
fun replace (v, (_, T)) (U as TVar (u, _)) = if u = v then T else U |
180 |
| replace _ T = T |
|
181 |
||
182 |
fun complete (vT as (v, _)) subst = |
|
183 |
subst |
|
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
184 |
|> not (Vartab.defined subst v) ? update_subst vT |
39020 | 185 |
|> Vartab.map (K (apsnd (Term.map_atyps (replace vT)))) |
36898 | 186 |
|
187 |
fun cert (ix, (S, T)) = pairself (Thm.ctyp_of thy) (TVar (ix, S), T) |
|
188 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
189 |
fun inst subst = |
36898 | 190 |
let val cTs = Vartab.fold (cons o cert) (fold complete Tenv subst) [] |
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
191 |
in SOME (i, Thm.instantiate (cTs, []) thm) end |
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
192 |
handle PARTIAL_INST () => NONE |
36898 | 193 |
|
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
194 |
in (map_filter inst substs @ ithms, if full then ctxt' else ctxt) end |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
195 |
|
36898 | 196 |
|
197 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
198 |
(* overall procedure *) |
36898 | 199 |
|
42190
b6b5846504cd
make configuration of (SMT) full monomorphization more flexible: turn boolean argument into a configuration option
boehmes
parents:
42183
diff
changeset
|
200 |
fun mono_all ctxt polys monos = |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
201 |
let |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
202 |
val scss = map (single o pair Vartab.empty o tvar_consts_of o snd) polys |
36898 | 203 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
204 |
(* all known non-schematic instances of polymorphic constants: find all |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
205 |
names of polymorphic constants, then add all known ground types *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
206 |
val grounds = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
207 |
Symtab.empty |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
208 |
|> fold (fold (fold (Symtab.update o rpair [] o fst) o snd)) scss |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
209 |
|> fold (add_const_types (K true) o snd) monos |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
210 |
|> fold (add_const_types (not o typ_has_tvars) o snd) polys |
36898 | 211 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
212 |
val limit = Config.get ctxt SMT_Config.monomorph_limit |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
213 |
val instances = Config.get ctxt SMT_Config.monomorph_instances |
42190
b6b5846504cd
make configuration of (SMT) full monomorphization more flexible: turn boolean argument into a configuration option
boehmes
parents:
42183
diff
changeset
|
214 |
val full = Config.get ctxt SMT_Config.monomorph_full |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
215 |
in |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
216 |
scss |
41762
00060198de12
limit the number of generated SMT monomorphization instances (in addition to the existing monomorphization limit which does not prevent seemingly infinite loops);
boehmes
parents:
41212
diff
changeset
|
217 |
|> search_substitutions ctxt limit instances Symtab.empty grounds |
42361 | 218 |
|> map (filter_most_specific (Proof_Context.theory_of ctxt)) |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
219 |
|> rpair (monos, ctxt) |
42183
173b0f488428
provide a flag controlling whether all provided facts should be instantiated, possibly inventing new types (which does not work well with Sledgehammer)
boehmes
parents:
42181
diff
changeset
|
220 |
|-> fold2 (instantiate full) polys |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
221 |
end |
36898 | 222 |
|
42190
b6b5846504cd
make configuration of (SMT) full monomorphization more flexible: turn boolean argument into a configuration option
boehmes
parents:
42183
diff
changeset
|
223 |
fun monomorph irules ctxt = |
40161
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39687
diff
changeset
|
224 |
irules |
539d07b00e5f
keep track of theorems initially given to SMT (even if they are rewritten); provide interface to filter theorems necessary for SMT proofs
boehmes
parents:
39687
diff
changeset
|
225 |
|> List.partition (Term.exists_type typ_has_tvars o Thm.prop_of o snd) |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
226 |
|>> incr_indexes (* avoid clashes of schematic type variables *) |
42190
b6b5846504cd
make configuration of (SMT) full monomorphization more flexible: turn boolean argument into a configuration option
boehmes
parents:
42183
diff
changeset
|
227 |
|-> (fn [] => rpair ctxt | polys => mono_all ctxt polys) |
36898 | 228 |
|
229 |
end |