author | wenzelm |
Wed, 11 Aug 2010 18:41:06 +0200 | |
changeset 38353 | d98baa2cf589 |
parent 36945 | 9bec62c10714 |
child 43324 | 2b47822868e4 |
permissions | -rw-r--r-- |
23175 | 1 |
(* Title: Tools/IsaPlanner/rw_inst.ML |
23171 | 2 |
Author: Lucas Dixon, University of Edinburgh |
3 |
||
23175 | 4 |
Rewriting using a conditional meta-equality theorem which supports |
5 |
schematic variable instantiation. |
|
6 |
*) |
|
23171 | 7 |
|
8 |
signature RW_INST = |
|
9 |
sig |
|
10 |
||
11 |
(* Rewrite: give it instantiation infromation, a rule, and the |
|
12 |
target thm, and it will return the rewritten target thm *) |
|
13 |
val rw : |
|
14 |
((Term.indexname * (Term.sort * Term.typ)) list * (* type var instantiations *) |
|
15 |
(Term.indexname * (Term.typ * Term.term)) list) (* schematic var instantiations *) |
|
16 |
* (string * Term.typ) list (* Fake named bounds + types *) |
|
17 |
* (string * Term.typ) list (* names of bound + types *) |
|
18 |
* Term.term -> (* outer term for instantiation *) |
|
19 |
Thm.thm -> (* rule with indexies lifted *) |
|
20 |
Thm.thm -> (* target thm *) |
|
21 |
Thm.thm (* rewritten theorem possibly |
|
22 |
with additional premises for |
|
23 |
rule conditions *) |
|
24 |
||
25 |
(* used tools *) |
|
26 |
val mk_abstractedrule : |
|
27 |
(string * Term.typ) list (* faked outer bound *) |
|
28 |
-> (string * Term.typ) list (* hopeful name of outer bounds *) |
|
29 |
-> Thm.thm -> Thm.cterm list * Thm.thm |
|
30 |
val mk_fixtvar_tyinsts : |
|
31 |
(Term.indexname * (Term.sort * Term.typ)) list -> |
|
32 |
Term.term list -> ((string * int) * (Term.sort * Term.typ)) list |
|
33 |
* (string * Term.sort) list |
|
34 |
val mk_renamings : |
|
35 |
Term.term -> Thm.thm -> (((string * int) * Term.typ) * Term.term) list |
|
36 |
val new_tfree : |
|
37 |
((string * int) * Term.sort) * |
|
38 |
(((string * int) * (Term.sort * Term.typ)) list * string list) -> |
|
39 |
((string * int) * (Term.sort * Term.typ)) list * string list |
|
40 |
val cross_inst : (Term.indexname * (Term.typ * Term.term)) list |
|
41 |
-> (Term.indexname *(Term.typ * Term.term)) list |
|
42 |
val cross_inst_typs : (Term.indexname * (Term.sort * Term.typ)) list |
|
43 |
-> (Term.indexname * (Term.sort * Term.typ)) list |
|
44 |
||
45 |
val beta_contract : Thm.thm -> Thm.thm |
|
46 |
val beta_eta_contract : Thm.thm -> Thm.thm |
|
47 |
||
48 |
end; |
|
49 |
||
50 |
structure RWInst |
|
51 |
: RW_INST |
|
52 |
= struct |
|
53 |
||
54 |
||
55 |
(* beta contract the theorem *) |
|
56 |
fun beta_contract thm = |
|
36945 | 57 |
Thm.equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm; |
23171 | 58 |
|
59 |
(* beta-eta contract the theorem *) |
|
60 |
fun beta_eta_contract thm = |
|
61 |
let |
|
36945 | 62 |
val thm2 = Thm.equal_elim (Thm.beta_conversion true (Thm.cprop_of thm)) thm |
63 |
val thm3 = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of thm2)) thm2 |
|
23171 | 64 |
in thm3 end; |
65 |
||
66 |
||
67 |
(* to get the free names of a theorem (including hyps and flexes) *) |
|
68 |
fun usednames_of_thm th = |
|
69 |
let val rep = Thm.rep_thm th |
|
70 |
val hyps = #hyps rep |
|
71 |
val (tpairl,tpairr) = Library.split_list (#tpairs rep) |
|
72 |
val prop = #prop rep |
|
73 |
in |
|
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
74 |
List.foldr OldTerm.add_term_names [] (prop :: (tpairl @ (tpairr @ hyps))) |
23171 | 75 |
end; |
76 |
||
77 |
(* Given a list of variables that were bound, and a that has been |
|
78 |
instantiated with free variable placeholders for the bound vars, it |
|
79 |
creates an abstracted version of the theorem, with local bound vars as |
|
80 |
lambda-params: |
|
81 |
||
82 |
Ts: |
|
83 |
("x", ty) |
|
84 |
||
85 |
rule:: |
|
86 |
C :x ==> P :x = Q :x |
|
87 |
||
88 |
results in: |
|
89 |
("!! x. C x", (%x. p x = %y. p y) [!! x. C x]) |
|
90 |
||
91 |
note: assumes rule is instantiated |
|
92 |
*) |
|
93 |
(* Note, we take abstraction in the order of last abstraction first *) |
|
94 |
fun mk_abstractedrule TsFake Ts rule = |
|
95 |
let |
|
96 |
val ctermify = Thm.cterm_of (Thm.theory_of_thm rule); |
|
97 |
||
98 |
(* now we change the names of temporary free vars that represent |
|
99 |
bound vars with binders outside the redex *) |
|
100 |
val prop = Thm.prop_of rule; |
|
101 |
val names = usednames_of_thm rule; |
|
102 |
val (fromnames,tonames,names2,Ts') = |
|
103 |
Library.foldl (fn ((rnf,rnt,names, Ts''),((faken,_),(n,ty))) => |
|
104 |
let val n2 = Name.variant names n in |
|
105 |
(ctermify (Free(faken,ty)) :: rnf, |
|
106 |
ctermify (Free(n2,ty)) :: rnt, |
|
107 |
n2 :: names, |
|
108 |
(n2,ty) :: Ts'') |
|
109 |
end) |
|
110 |
(([],[],names, []), TsFake~~Ts); |
|
111 |
||
112 |
(* rename conflicting free's in the rule to avoid cconflicts |
|
113 |
with introduced vars from bounds outside in redex *) |
|
114 |
val rule' = rule |> Drule.forall_intr_list fromnames |
|
115 |
|> Drule.forall_elim_list tonames; |
|
116 |
||
117 |
(* make unconditional rule and prems *) |
|
118 |
val (uncond_rule, cprems) = IsaND.allify_conditions ctermify (rev Ts') |
|
119 |
rule'; |
|
120 |
||
121 |
(* using these names create lambda-abstracted version of the rule *) |
|
122 |
val abstractions = rev (Ts' ~~ tonames); |
|
123 |
val abstract_rule = Library.foldl (fn (th,((n,ty),ct)) => |
|
124 |
Thm.abstract_rule n ct th) |
|
125 |
(uncond_rule, abstractions); |
|
126 |
in (cprems, abstract_rule) end; |
|
127 |
||
128 |
||
129 |
(* given names to avoid, and vars that need to be fixed, it gives |
|
130 |
unique new names to the vars so that they can be fixed as free |
|
131 |
variables *) |
|
132 |
(* make fixed unique free variable instantiations for non-ground vars *) |
|
133 |
(* Create a table of vars to be renamed after instantiation - ie |
|
134 |
other uninstantiated vars in the hyps of the rule |
|
135 |
ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) |
|
136 |
fun mk_renamings tgt rule_inst = |
|
137 |
let |
|
138 |
val rule_conds = Thm.prems_of rule_inst |
|
30190 | 139 |
val names = List.foldr OldTerm.add_term_names [] (tgt :: rule_conds); |
23171 | 140 |
val (conds_tyvs,cond_vs) = |
141 |
Library.foldl (fn ((tyvs, vs), t) => |
|
33042 | 142 |
(union (op =) (OldTerm.term_tvars t) tyvs, |
143 |
union (op =) (map Term.dest_Var (OldTerm.term_vars t)) vs)) |
|
23171 | 144 |
(([],[]), rule_conds); |
29265
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
wenzelm
parents:
23175
diff
changeset
|
145 |
val termvars = map Term.dest_Var (OldTerm.term_vars tgt); |
33042 | 146 |
val vars_to_fix = union (op =) termvars cond_vs; |
23171 | 147 |
val (renamings, names2) = |
30190 | 148 |
List.foldr (fn (((n,i),ty), (vs, names')) => |
23171 | 149 |
let val n' = Name.variant names' n in |
150 |
((((n,i),ty), Free (n', ty)) :: vs, n'::names') |
|
151 |
end) |
|
152 |
([], names) vars_to_fix; |
|
153 |
in renamings end; |
|
154 |
||
155 |
(* make a new fresh typefree instantiation for the given tvar *) |
|
156 |
fun new_tfree (tv as (ix,sort), (pairs,used)) = |
|
157 |
let val v = Name.variant used (string_of_indexname ix) |
|
158 |
in ((ix,(sort,TFree(v,sort)))::pairs, v::used) end; |
|
159 |
||
160 |
||
161 |
(* make instantiations to fix type variables that are not |
|
162 |
already instantiated (in ignore_ixs) from the list of terms. *) |
|
163 |
fun mk_fixtvar_tyinsts ignore_insts ts = |
|
164 |
let |
|
165 |
val ignore_ixs = map fst ignore_insts; |
|
166 |
val (tvars, tfrees) = |
|
30190 | 167 |
List.foldr (fn (t, (varixs, tfrees)) => |
29270
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
168 |
(OldTerm.add_term_tvars (t,varixs), |
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
wenzelm
parents:
29265
diff
changeset
|
169 |
OldTerm.add_term_tfrees (t,tfrees))) |
23171 | 170 |
([],[]) ts; |
171 |
val unfixed_tvars = |
|
33317 | 172 |
filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars; |
30190 | 173 |
val (fixtyinsts, _) = List.foldr new_tfree ([], map fst tfrees) unfixed_tvars |
23171 | 174 |
in (fixtyinsts, tfrees) end; |
175 |
||
176 |
||
177 |
(* cross-instantiate the instantiations - ie for each instantiation |
|
178 |
replace all occurances in other instantiations - no loops are possible |
|
179 |
and thus only one-parsing of the instantiations is necessary. *) |
|
180 |
fun cross_inst insts = |
|
181 |
let |
|
182 |
fun instL (ix, (ty,t)) = |
|
183 |
map (fn (ix2,(ty2,t2)) => |
|
184 |
(ix2, (ty2,Term.subst_vars ([], [(ix, t)]) t2))); |
|
185 |
||
186 |
fun cross_instL ([], l) = rev l |
|
187 |
| cross_instL ((ix, t) :: insts, l) = |
|
188 |
cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); |
|
189 |
||
190 |
in cross_instL (insts, []) end; |
|
191 |
||
192 |
(* as above but for types -- I don't know if this is needed, will we ever incur mixed up types? *) |
|
193 |
fun cross_inst_typs insts = |
|
194 |
let |
|
195 |
fun instL (ix, (srt,ty)) = |
|
196 |
map (fn (ix2,(srt2,ty2)) => |
|
197 |
(ix2, (srt2,Term.typ_subst_TVars [(ix, ty)] ty2))); |
|
198 |
||
199 |
fun cross_instL ([], l) = rev l |
|
200 |
| cross_instL ((ix, t) :: insts, l) = |
|
201 |
cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); |
|
202 |
||
203 |
in cross_instL (insts, []) end; |
|
204 |
||
205 |
||
206 |
(* assume that rule and target_thm have distinct var names. THINK: |
|
207 |
efficient version with tables for vars for: target vars, introduced |
|
208 |
vars, and rule vars, for quicker instantiation? The outerterm defines |
|
209 |
which part of the target_thm was modified. Note: we take Ts in the |
|
210 |
upterm order, ie last abstraction first., and with an outeterm where |
|
211 |
the abstracted subterm has the arguments in the revered order, ie |
|
212 |
first abstraction first. FakeTs has abstractions using the fake name |
|
213 |
- ie the name distinct from all other abstractions. *) |
|
214 |
||
215 |
fun rw ((nonfixed_typinsts, unprepinsts), FakeTs, Ts, outerterm) rule target_thm = |
|
216 |
let |
|
217 |
(* general signature info *) |
|
218 |
val target_sign = (Thm.theory_of_thm target_thm); |
|
219 |
val ctermify = Thm.cterm_of target_sign; |
|
220 |
val ctypeify = Thm.ctyp_of target_sign; |
|
221 |
||
222 |
(* fix all non-instantiated tvars *) |
|
223 |
val (fixtyinsts, othertfrees) = |
|
224 |
mk_fixtvar_tyinsts nonfixed_typinsts |
|
225 |
[Thm.prop_of rule, Thm.prop_of target_thm]; |
|
226 |
val new_fixed_typs = map (fn ((s,i),(srt,ty)) => (Term.dest_TFree ty)) |
|
227 |
fixtyinsts; |
|
228 |
val typinsts = cross_inst_typs (nonfixed_typinsts @ fixtyinsts); |
|
229 |
||
230 |
(* certified instantiations for types *) |
|
231 |
val ctyp_insts = |
|
232 |
map (fn (ix,(s,ty)) => (ctypeify (TVar (ix,s)), ctypeify ty)) |
|
233 |
typinsts; |
|
234 |
||
235 |
(* type instantiated versions *) |
|
236 |
val tgt_th_tyinst = Thm.instantiate (ctyp_insts,[]) target_thm; |
|
237 |
val rule_tyinst = Thm.instantiate (ctyp_insts,[]) rule; |
|
238 |
||
239 |
val term_typ_inst = map (fn (ix,(srt,ty)) => (ix,ty)) typinsts; |
|
240 |
(* type instanitated outer term *) |
|
241 |
val outerterm_tyinst = Term.subst_TVars term_typ_inst outerterm; |
|
242 |
||
243 |
val FakeTs_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) |
|
244 |
FakeTs; |
|
245 |
val Ts_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) |
|
246 |
Ts; |
|
247 |
||
248 |
(* type-instantiate the var instantiations *) |
|
30190 | 249 |
val insts_tyinst = List.foldr (fn ((ix,(ty,t)),insts_tyinst) => |
23171 | 250 |
(ix, (Term.typ_subst_TVars term_typ_inst ty, |
251 |
Term.subst_TVars term_typ_inst t)) |
|
252 |
:: insts_tyinst) |
|
253 |
[] unprepinsts; |
|
254 |
||
255 |
(* cross-instantiate *) |
|
256 |
val insts_tyinst_inst = cross_inst insts_tyinst; |
|
257 |
||
258 |
(* create certms of instantiations *) |
|
259 |
val cinsts_tyinst = |
|
260 |
map (fn (ix,(ty,t)) => (ctermify (Var (ix, ty)), |
|
261 |
ctermify t)) insts_tyinst_inst; |
|
262 |
||
263 |
(* The instantiated rule *) |
|
264 |
val rule_inst = rule_tyinst |> Thm.instantiate ([], cinsts_tyinst); |
|
265 |
||
266 |
(* Create a table of vars to be renamed after instantiation - ie |
|
267 |
other uninstantiated vars in the hyps the *instantiated* rule |
|
268 |
ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) |
|
269 |
val renamings = mk_renamings (Thm.prop_of tgt_th_tyinst) |
|
270 |
rule_inst; |
|
271 |
val cterm_renamings = |
|
272 |
map (fn (x,y) => (ctermify (Var x), ctermify y)) renamings; |
|
273 |
||
274 |
(* Create the specific version of the rule for this target application *) |
|
275 |
val outerterm_inst = |
|
276 |
outerterm_tyinst |
|
277 |
|> Term.subst_Vars (map (fn (ix,(ty,t)) => (ix,t)) insts_tyinst_inst) |
|
278 |
|> Term.subst_Vars (map (fn ((ix,ty),t) => (ix,t)) renamings); |
|
279 |
val couter_inst = Thm.reflexive (ctermify outerterm_inst); |
|
280 |
val (cprems, abstract_rule_inst) = |
|
281 |
rule_inst |> Thm.instantiate ([], cterm_renamings) |
|
282 |
|> mk_abstractedrule FakeTs_tyinst Ts_tyinst; |
|
283 |
val specific_tgt_rule = |
|
284 |
beta_eta_contract |
|
285 |
(Thm.combination couter_inst abstract_rule_inst); |
|
286 |
||
287 |
(* create an instantiated version of the target thm *) |
|
288 |
val tgt_th_inst = |
|
289 |
tgt_th_tyinst |> Thm.instantiate ([], cinsts_tyinst) |
|
290 |
|> Thm.instantiate ([], cterm_renamings); |
|
291 |
||
292 |
val (vars,frees_of_fixed_vars) = Library.split_list cterm_renamings; |
|
293 |
||
294 |
in |
|
295 |
(beta_eta_contract tgt_th_inst) |
|
296 |
|> Thm.equal_elim specific_tgt_rule |
|
297 |
|> Drule.implies_intr_list cprems |
|
298 |
|> Drule.forall_intr_list frees_of_fixed_vars |
|
299 |
|> Drule.forall_elim_list vars |
|
35845
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
wenzelm
parents:
33317
diff
changeset
|
300 |
|> Thm.varifyT_global' othertfrees |
23171 | 301 |
|-> K Drule.zero_var_indexes |
302 |
end; |
|
303 |
||
304 |
||
305 |
end; (* struct *) |