author | blanchet |
Wed, 15 Dec 2010 18:10:32 +0100 | |
changeset 41171 | 043f8dc3b51f |
parent 41063 | 0828bfa70b20 |
child 41174 | 10eb369f8c01 |
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 |
|
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
|
31 |
val monomorph: (int * thm) list -> Proof.context -> |
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
|
32 |
(int * thm) list * Proof.context |
36898 | 33 |
end |
34 |
||
35 |
structure SMT_Monomorph: SMT_MONOMORPH = |
|
36 |
struct |
|
37 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
38 |
(* utility functions *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
39 |
|
36898 | 40 |
val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false) |
41 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
42 |
val ignored = member (op =) [@{const_name All}, @{const_name Ex}, |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
43 |
@{const_name Let}, @{const_name If}, @{const_name HOL.eq}] |
36898 | 44 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
45 |
fun is_const pred (n, T) = not (ignored n) andalso pred T |
36898 | 46 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
47 |
fun collect_consts_if pred f = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
48 |
Thm.prop_of #> |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
49 |
Term.fold_aterms (fn Const c => if is_const pred c then f c else I | _ => I) |
36898 | 50 |
|
39687 | 51 |
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
|
52 |
|
36898 | 53 |
fun tvar_consts_of thm = collect_consts_if typ_has_tvars insert_const thm [] |
54 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
55 |
fun add_const_types pred = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
56 |
collect_consts_if pred (fn (n, T) => Symtab.map_entry n (insert (op =) T)) |
36898 | 57 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
58 |
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
|
59 |
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
|
60 |
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
|
61 |
((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
|
62 |
in fst (fold_map inc ithms 0) end |
36898 | 63 |
|
64 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
65 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
66 |
(* search for necessary substitutions *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
67 |
|
36898 | 68 |
fun new_substitutions thy grounds (n, T) subst = |
69 |
if not (typ_has_tvars T) then [subst] |
|
70 |
else |
|
71 |
Symtab.lookup_list grounds n |
|
72 |
|> map_filter (try (fn U => Sign.typ_match thy (T, U) subst)) |
|
73 |
|> cons subst |
|
74 |
||
75 |
fun apply_subst grounds consts subst = |
|
76 |
let |
|
77 |
fun is_new_ground (n, T) = not (typ_has_tvars T) andalso |
|
78 |
not (member (op =) (Symtab.lookup_list grounds n) T) |
|
79 |
||
80 |
fun apply_const (n, T) new_grounds = |
|
81 |
let val c = (n, Envir.subst_type subst T) |
|
82 |
in |
|
83 |
new_grounds |
|
84 |
|> is_new_ground c ? Symtab.insert_list (op =) c |
|
85 |
|> pair c |
|
86 |
end |
|
87 |
in fold_map apply_const consts #>> pair subst end |
|
88 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
89 |
fun specialize thy all_grounds new_grounds scs = |
36898 | 90 |
let |
91 |
fun spec (subst, consts) next_grounds = |
|
92 |
[subst] |
|
93 |
|> fold (maps o new_substitutions thy new_grounds) consts |
|
94 |
|> rpair next_grounds |
|
95 |
|-> fold_map (apply_subst all_grounds consts) |
|
96 |
in |
|
97 |
fold_map spec scs #>> (fn scss => |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
98 |
fold (fold (insert (eq_snd (op =)))) scss []) |
36898 | 99 |
end |
100 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
101 |
val limit_reached_warning = "Warning: Monomorphization limit reached" |
36898 | 102 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
103 |
fun search_substitutions ctxt limit all_grounds new_grounds scss = |
36898 | 104 |
let |
40424
7550b2cba1cb
better modularization: moved SMT configuration options and diagnostics as well as SMT failure and exception into separate structures (both of which are loaded first and consequently are available to other SMT structures)
boehmes
parents:
40161
diff
changeset
|
105 |
val thy = ProofContext.theory_of ctxt |
36898 | 106 |
val all_grounds' = Symtab.merge_list (op =) (all_grounds, new_grounds) |
107 |
val spec = specialize thy all_grounds' new_grounds |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
108 |
val (scss', new_grounds') = fold_map spec scss Symtab.empty |
36898 | 109 |
in |
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
110 |
if Symtab.is_empty new_grounds' then scss' |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
111 |
else if limit > 0 then |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
112 |
search_substitutions ctxt (limit-1) all_grounds' new_grounds' scss' |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
113 |
else (SMT_Config.verbose_msg ctxt (K limit_reached_warning) (); scss') |
36898 | 114 |
end |
115 |
||
116 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
117 |
|
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
118 |
(* instantiation *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
119 |
|
36898 | 120 |
fun filter_most_specific thy = |
121 |
let |
|
122 |
fun typ_match (_, T) (_, U) = Sign.typ_match thy (T, U) |
|
123 |
||
124 |
fun is_trivial subst = Vartab.is_empty subst orelse |
|
125 |
forall (fn (v, (S, T)) => TVar (v, S) = T) (Vartab.dest subst) |
|
126 |
||
127 |
fun match general specific = |
|
128 |
(case try (fold2 typ_match general specific) Vartab.empty of |
|
129 |
NONE => false |
|
130 |
| SOME subst => not (is_trivial subst)) |
|
131 |
||
132 |
fun most_specific _ [] = [] |
|
133 |
| most_specific css ((ss, cs) :: scs) = |
|
134 |
let val substs = most_specific (cs :: css) scs |
|
135 |
in |
|
136 |
if exists (match cs) css orelse exists (match cs o snd) scs |
|
137 |
then substs else ss :: substs |
|
138 |
end |
|
139 |
||
140 |
in most_specific [] end |
|
141 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
142 |
fun instantiate (i, thm) substs (ithms, ctxt) = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
143 |
let |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
144 |
val (vs, Ss) = split_list (Term.add_tvars (Thm.prop_of thm) []) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
145 |
val (Tenv, ctxt') = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
146 |
ctxt |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
147 |
|> Variable.invent_types Ss |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
148 |
|>> map2 (fn v => fn (n, S) => (v, (S, TFree (n, S)))) vs |
36898 | 149 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
150 |
val thy = ProofContext.theory_of ctxt' |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
151 |
|
36898 | 152 |
fun replace (v, (_, T)) (U as TVar (u, _)) = if u = v then T else U |
153 |
| replace _ T = T |
|
154 |
||
155 |
fun complete (vT as (v, _)) subst = |
|
156 |
subst |
|
157 |
|> not (Vartab.defined subst v) ? Vartab.update vT |
|
39020 | 158 |
|> Vartab.map (K (apsnd (Term.map_atyps (replace vT)))) |
36898 | 159 |
|
160 |
fun cert (ix, (S, T)) = pairself (Thm.ctyp_of thy) (TVar (ix, S), T) |
|
161 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
162 |
fun inst subst = |
36898 | 163 |
let val cTs = Vartab.fold (cons o cert) (fold complete Tenv subst) [] |
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
|
164 |
in (i, Thm.instantiate (cTs, []) thm) end |
36898 | 165 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
166 |
in (map inst substs @ ithms, ctxt') end |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
167 |
|
36898 | 168 |
|
169 |
||
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
170 |
(* overall procedure *) |
36898 | 171 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
172 |
fun mono_all ctxt polys monos = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
173 |
let |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
174 |
val scss = map (single o pair Vartab.empty o tvar_consts_of o snd) polys |
36898 | 175 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
176 |
(* all known non-schematic instances of polymorphic constants: find all |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
177 |
names of polymorphic constants, then add all known ground types *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
178 |
val grounds = |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
179 |
Symtab.empty |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
180 |
|> fold (fold (fold (Symtab.update o rpair [] o fst) o snd)) scss |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
181 |
|> fold (add_const_types (K true) o snd) monos |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
182 |
|> fold (add_const_types (not o typ_has_tvars) o snd) polys |
36898 | 183 |
|
41063
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
184 |
val limit = Config.get ctxt SMT_Config.monomorph_limit |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
185 |
in |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
186 |
scss |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
187 |
|> search_substitutions ctxt limit Symtab.empty grounds |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
188 |
|> map (filter_most_specific (ProofContext.theory_of ctxt)) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
189 |
|> rpair (monos, ctxt) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
190 |
|-> fold2 instantiate polys |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
191 |
end |
36898 | 192 |
|
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
|
193 |
fun monomorph irules ctxt = |
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
|
194 |
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
|
195 |
|> 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
|
196 |
|>> incr_indexes (* avoid clashes of schematic type variables *) |
0828bfa70b20
reduced unnecessary complexity; improved documentation; tuned
boehmes
parents:
40512
diff
changeset
|
197 |
|-> (fn [] => rpair ctxt | polys => mono_all ctxt polys) |
36898 | 198 |
|
199 |
end |