author | wenzelm |
Thu, 02 Oct 2008 19:59:00 +0200 | |
changeset 28464 | dcc030b52583 |
parent 27809 | a1e409db516b |
child 29269 | 5c25a2012975 |
permissions | -rw-r--r-- |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
1 |
(* Title: Provers/eqsubst.ML |
16434 | 2 |
ID: $Id$ |
18598 | 3 |
Author: Lucas Dixon, University of Edinburgh, lucas.dixon@ed.ac.uk |
15481 | 4 |
|
18598 | 5 |
A proof method to perform a substiution using an equation. |
6 |
*) |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
7 |
|
18591 | 8 |
signature EQSUBST = |
15481 | 9 |
sig |
19871 | 10 |
(* a type abriviation for match information *) |
11 |
type match = |
|
12 |
((indexname * (sort * typ)) list (* type instantiations *) |
|
13 |
* (indexname * (typ * term)) list) (* term instantiations *) |
|
14 |
* (string * typ) list (* fake named type abs env *) |
|
15 |
* (string * typ) list (* type abs env *) |
|
16 |
* term (* outer term *) |
|
17 |
||
18 |
type searchinfo = |
|
19 |
theory |
|
20 |
* int (* maxidx *) |
|
21 |
* Zipper.T (* focusterm to search under *) |
|
22 |
||
23 |
exception eqsubst_occL_exp of |
|
24 |
string * int list * Thm.thm list * int * Thm.thm |
|
25 |
||
26 |
(* low level substitution functions *) |
|
27 |
val apply_subst_in_asm : |
|
28 |
int -> |
|
29 |
Thm.thm -> |
|
30 |
Thm.thm -> |
|
31 |
(Thm.cterm list * int * 'a * Thm.thm) * match -> Thm.thm Seq.seq |
|
32 |
val apply_subst_in_concl : |
|
33 |
int -> |
|
34 |
Thm.thm -> |
|
35 |
Thm.cterm list * Thm.thm -> |
|
36 |
Thm.thm -> match -> Thm.thm Seq.seq |
|
37 |
||
38 |
(* matching/unification within zippers *) |
|
39 |
val clean_match_z : |
|
40 |
Context.theory -> Term.term -> Zipper.T -> match option |
|
41 |
val clean_unify_z : |
|
42 |
Context.theory -> int -> Term.term -> Zipper.T -> match Seq.seq |
|
43 |
||
44 |
(* skipping things in seq seq's *) |
|
45 |
||
46 |
(* skipping non-empty sub-sequences but when we reach the end |
|
47 |
of the seq, remembering how much we have left to skip. *) |
|
48 |
datatype 'a skipseq = SkipMore of int |
|
49 |
| SkipSeq of 'a Seq.seq Seq.seq; |
|
50 |
||
51 |
val skip_first_asm_occs_search : |
|
52 |
('a -> 'b -> 'c Seq.seq Seq.seq) -> |
|
53 |
'a -> int -> 'b -> 'c skipseq |
|
54 |
val skip_first_occs_search : |
|
55 |
int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq |
|
56 |
val skipto_skipseq : int -> 'a Seq.seq Seq.seq -> 'a skipseq |
|
57 |
||
58 |
(* tactics *) |
|
59 |
val eqsubst_asm_tac : |
|
60 |
Proof.context -> |
|
61 |
int list -> Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq |
|
62 |
val eqsubst_asm_tac' : |
|
63 |
Proof.context -> |
|
64 |
(searchinfo -> int -> Term.term -> match skipseq) -> |
|
65 |
int -> Thm.thm -> int -> Thm.thm -> Thm.thm Seq.seq |
|
66 |
val eqsubst_tac : |
|
67 |
Proof.context -> |
|
22727 | 68 |
int list -> (* list of occurences to rewrite, use [0] for any *) |
69 |
Thm.thm list -> int -> Thm.thm -> Thm.thm Seq.seq |
|
19871 | 70 |
val eqsubst_tac' : |
22727 | 71 |
Proof.context -> (* proof context *) |
72 |
(searchinfo -> Term.term -> match Seq.seq) (* search function *) |
|
73 |
-> Thm.thm (* equation theorem to rewrite with *) |
|
74 |
-> int (* subgoal number in goal theorem *) |
|
75 |
-> Thm.thm (* goal theorem *) |
|
76 |
-> Thm.thm Seq.seq (* rewritten goal theorem *) |
|
19871 | 77 |
|
78 |
||
79 |
val fakefree_badbounds : |
|
80 |
(string * Term.typ) list -> |
|
81 |
Term.term -> |
|
82 |
(string * Term.typ) list * (string * Term.typ) list * Term.term |
|
83 |
||
84 |
val mk_foo_match : |
|
85 |
(Term.term -> Term.term) -> |
|
86 |
('a * Term.typ) list -> Term.term -> Term.term |
|
87 |
||
88 |
(* preparing substitution *) |
|
89 |
val prep_meta_eq : Proof.context -> Thm.thm -> Thm.thm list |
|
90 |
val prep_concl_subst : |
|
91 |
int -> Thm.thm -> (Thm.cterm list * Thm.thm) * searchinfo |
|
92 |
val prep_subst_in_asm : |
|
93 |
int -> Thm.thm -> int -> |
|
94 |
(Thm.cterm list * int * int * Thm.thm) * searchinfo |
|
95 |
val prep_subst_in_asms : |
|
96 |
int -> Thm.thm -> |
|
97 |
((Thm.cterm list * int * int * Thm.thm) * searchinfo) list |
|
98 |
val prep_zipper_match : |
|
99 |
Zipper.T -> Term.term * ((string * Term.typ) list * (string * Term.typ) list * Term.term) |
|
100 |
||
101 |
(* search for substitutions *) |
|
102 |
val valid_match_start : Zipper.T -> bool |
|
103 |
val search_lr_all : Zipper.T -> Zipper.T Seq.seq |
|
104 |
val search_lr_valid : (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq |
|
105 |
val searchf_lr_unify_all : |
|
106 |
searchinfo -> Term.term -> match Seq.seq Seq.seq |
|
107 |
val searchf_lr_unify_valid : |
|
108 |
searchinfo -> Term.term -> match Seq.seq Seq.seq |
|
23064 | 109 |
val searchf_bt_unify_valid : |
110 |
searchinfo -> Term.term -> match Seq.seq Seq.seq |
|
19871 | 111 |
|
112 |
(* syntax tools *) |
|
113 |
val ith_syntax : Args.T list -> int list * Args.T list |
|
114 |
val options_syntax : Args.T list -> bool * Args.T list |
|
115 |
||
116 |
(* Isar level hooks *) |
|
20289 | 117 |
val eqsubst_asm_meth : Proof.context -> int list -> Thm.thm list -> Proof.method |
118 |
val eqsubst_meth : Proof.context -> int list -> Thm.thm list -> Proof.method |
|
119 |
val subst_meth : Method.src -> Proof.context -> Proof.method |
|
19871 | 120 |
val setup : theory -> theory |
121 |
||
15481 | 122 |
end; |
123 |
||
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
124 |
structure EqSubst |
19871 | 125 |
: EQSUBST |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
126 |
= struct |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
127 |
|
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
128 |
structure Z = Zipper; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
129 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
130 |
(* changes object "=" to meta "==" which prepares a given rewrite rule *) |
18598 | 131 |
fun prep_meta_eq ctxt = |
132 |
let val (_, {mk_rews = {mk, ...}, ...}) = Simplifier.rep_ss (Simplifier.local_simpset_of ctxt) |
|
133 |
in mk #> map Drule.zero_var_indexes end; |
|
134 |
||
15481 | 135 |
|
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
136 |
(* a type abriviation for match information *) |
16978 | 137 |
type match = |
138 |
((indexname * (sort * typ)) list (* type instantiations *) |
|
139 |
* (indexname * (typ * term)) list) (* term instantiations *) |
|
140 |
* (string * typ) list (* fake named type abs env *) |
|
141 |
* (string * typ) list (* type abs env *) |
|
142 |
* term (* outer term *) |
|
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
143 |
|
16978 | 144 |
type searchinfo = |
18598 | 145 |
theory |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
146 |
* int (* maxidx *) |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
147 |
* Zipper.T (* focusterm to search under *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
148 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
149 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
150 |
(* skipping non-empty sub-sequences but when we reach the end |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
151 |
of the seq, remembering how much we have left to skip. *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
152 |
datatype 'a skipseq = SkipMore of int |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
153 |
| SkipSeq of 'a Seq.seq Seq.seq; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
154 |
(* given a seqseq, skip the first m non-empty seq's, note deficit *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
155 |
fun skipto_skipseq m s = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
156 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
157 |
fun skip_occs n sq = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
158 |
case Seq.pull sq of |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
159 |
NONE => SkipMore n |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
160 |
| SOME (h,t) => |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
161 |
(case Seq.pull h of NONE => skip_occs n t |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
162 |
| SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
163 |
else skip_occs (n - 1) t) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
164 |
in (skip_occs m s) end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
165 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
166 |
(* note: outerterm is the taget with the match replaced by a bound |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
167 |
variable : ie: "P lhs" beocmes "%x. P x" |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
168 |
insts is the types of instantiations of vars in lhs |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
169 |
and typinsts is the type instantiations of types in the lhs |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
170 |
Note: Final rule is the rule lifted into the ontext of the |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
171 |
taget thm. *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
172 |
fun mk_foo_match mkuptermfunc Ts t = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
173 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
174 |
val ty = Term.type_of t |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
175 |
val bigtype = (rev (map snd Ts)) ---> ty |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
176 |
fun mk_foo 0 t = t |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
177 |
| mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
178 |
val num_of_bnds = (length Ts) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
179 |
(* foo_term = "fooabs y0 ... yn" where y's are local bounds *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
180 |
val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
181 |
in Abs("fooabs", bigtype, mkuptermfunc foo_term) end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
182 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
183 |
(* T is outer bound vars, n is number of locally bound vars *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
184 |
(* THINK: is order of Ts correct...? or reversed? *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
185 |
fun fakefree_badbounds Ts t = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
186 |
let val (FakeTs,Ts,newnames) = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
187 |
List.foldr (fn ((n,ty),(FakeTs,Ts,usednames)) => |
20071
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
wenzelm
parents:
19975
diff
changeset
|
188 |
let val newname = Name.variant usednames n |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
189 |
in ((RWTools.mk_fake_bound_name newname,ty)::FakeTs, |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
190 |
(newname,ty)::Ts, |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
191 |
newname::usednames) end) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
192 |
([],[],[]) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
193 |
Ts |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
194 |
in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
195 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
196 |
(* before matching we need to fake the bound vars that are missing an |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
197 |
abstraction. In this function we additionally construct the |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
198 |
abstraction environment, and an outer context term (with the focus |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
199 |
abstracted out) for use in rewriting with RWInst.rw *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
200 |
fun prep_zipper_match z = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
201 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
202 |
val t = Z.trm z |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
203 |
val c = Z.ctxt z |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
204 |
val Ts = Z.C.nty_ctxt c |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
205 |
val (FakeTs', Ts', t') = fakefree_badbounds Ts t |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
206 |
val absterm = mk_foo_match (Z.C.apply c) Ts' t' |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
207 |
in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
208 |
(t', (FakeTs', Ts', absterm)) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
209 |
end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
210 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
211 |
(* Matching and Unification with exception handled *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
212 |
fun clean_match thy (a as (pat, t)) = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
213 |
let val (tyenv, tenv) = Pattern.match thy a (Vartab.empty, Vartab.empty) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
214 |
in SOME (Vartab.dest tyenv, Vartab.dest tenv) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
215 |
end handle Pattern.MATCH => NONE; |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
216 |
|
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
217 |
(* given theory, max var index, pat, tgt; returns Seq of instantiations *) |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
218 |
fun clean_unify thry ix (a as (pat, tgt)) = |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
219 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
220 |
(* type info will be re-derived, maybe this can be cached |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
221 |
for efficiency? *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
222 |
val pat_ty = Term.type_of pat; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
223 |
val tgt_ty = Term.type_of tgt; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
224 |
(* is it OK to ignore the type instantiation info? |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
225 |
or should I be using it? *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
226 |
val typs_unify = |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
227 |
SOME (Sign.typ_unify thry (pat_ty, tgt_ty) (Term.Vartab.empty, ix)) |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
228 |
handle Type.TUNIFY => NONE; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
229 |
in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
230 |
case typs_unify of |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
231 |
SOME (typinsttab, ix2) => |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
232 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
233 |
(* is it right to throw away the flexes? |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
234 |
or should I be using them somehow? *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
235 |
fun mk_insts env = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
236 |
(Vartab.dest (Envir.type_env env), |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
237 |
Envir.alist_of env); |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
238 |
val initenv = Envir.Envir {asol = Vartab.empty, |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
239 |
iTs = typinsttab, maxidx = ix2}; |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
240 |
val useq = Unify.smash_unifiers thry [a] initenv |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
241 |
handle UnequalLengths => Seq.empty |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
242 |
| Term.TERM _ => Seq.empty; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
243 |
fun clean_unify' useq () = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
244 |
(case (Seq.pull useq) of |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
245 |
NONE => NONE |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
246 |
| SOME (h,t) => SOME (mk_insts h, Seq.make (clean_unify' t))) |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
247 |
handle UnequalLengths => NONE |
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
248 |
| Term.TERM _ => NONE |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
249 |
in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
250 |
(Seq.make (clean_unify' useq)) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
251 |
end |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
252 |
| NONE => Seq.empty |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
253 |
end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
254 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
255 |
(* Matching and Unification for zippers *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
256 |
(* Note: Ts is a modified version of the original names of the outer |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
257 |
bound variables. New names have been introduced to make sure they are |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
258 |
unique w.r.t all names in the term and each other. usednames' is |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
259 |
oldnames + new names. *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
260 |
fun clean_match_z thy pat z = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
261 |
let val (t, (FakeTs,Ts,absterm)) = prep_zipper_match z in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
262 |
case clean_match thy (pat, t) of |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
263 |
NONE => NONE |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
264 |
| SOME insts => SOME (insts, FakeTs, Ts, absterm) end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
265 |
(* ix = max var index *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
266 |
fun clean_unify_z sgn ix pat z = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
267 |
let val (t, (FakeTs, Ts,absterm)) = prep_zipper_match z in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
268 |
Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
269 |
(clean_unify sgn ix (t, pat)) end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
270 |
|
15550
806214035275
lucas - added more comments and an extra type to clarify the code.
dixon
parents:
15538
diff
changeset
|
271 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
272 |
(* FOR DEBUGGING... |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
273 |
type trace_subst_errT = int (* subgoal *) |
16978 | 274 |
* thm (* thm with all goals *) |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
275 |
* (Thm.cterm list (* certified free var placeholders for vars *) |
16978 | 276 |
* thm) (* trivial thm of goal concl *) |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
277 |
(* possible matches/unifiers *) |
16978 | 278 |
* thm (* rule *) |
279 |
* (((indexname * typ) list (* type instantiations *) |
|
280 |
* (indexname * term) list ) (* term instantiations *) |
|
281 |
* (string * typ) list (* Type abs env *) |
|
282 |
* term) (* outer term *); |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
283 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
284 |
val trace_subst_err = (ref NONE : trace_subst_errT option ref); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
285 |
val trace_subst_search = ref false; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
286 |
exception trace_subst_exp of trace_subst_errT; |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
287 |
*) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
288 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
289 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
290 |
fun bot_left_leaf_of (l $ r) = bot_left_leaf_of l |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
291 |
| bot_left_leaf_of (Abs(s,ty,t)) = bot_left_leaf_of t |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
292 |
| bot_left_leaf_of x = x; |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
293 |
|
19975
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents:
19871
diff
changeset
|
294 |
(* Avoid considering replacing terms which have a var at the head as |
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents:
19871
diff
changeset
|
295 |
they always succeed trivially, and uninterestingly. *) |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
296 |
fun valid_match_start z = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
297 |
(case bot_left_leaf_of (Z.trm z) of |
19975
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents:
19871
diff
changeset
|
298 |
Var _ => false |
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents:
19871
diff
changeset
|
299 |
| _ => true); |
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
dixon
parents:
19871
diff
changeset
|
300 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
301 |
(* search from top, left to right, then down *) |
19871 | 302 |
val search_lr_all = ZipperSearch.all_bl_ur; |
15481 | 303 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
304 |
(* search from top, left to right, then down *) |
19871 | 305 |
fun search_lr_valid validf = |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
306 |
let |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
307 |
fun sf_valid_td_lr z = |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
308 |
let val here = if validf z then [Z.Here z] else [] in |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
309 |
case Z.trm z |
19871 | 310 |
of _ $ _ => [Z.LookIn (Z.move_down_left z)] |
311 |
@ here |
|
312 |
@ [Z.LookIn (Z.move_down_right z)] |
|
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
313 |
| Abs _ => here @ [Z.LookIn (Z.move_down_abs z)] |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
314 |
| _ => here |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
315 |
end; |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
316 |
in Z.lzy_search sf_valid_td_lr end; |
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
317 |
|
23064 | 318 |
(* search from bottom to top, left to right *) |
319 |
||
320 |
fun search_bt_valid validf = |
|
321 |
let |
|
322 |
fun sf_valid_td_lr z = |
|
323 |
let val here = if validf z then [Z.Here z] else [] in |
|
324 |
case Z.trm z |
|
325 |
of _ $ _ => [Z.LookIn (Z.move_down_left z), |
|
326 |
Z.LookIn (Z.move_down_right z)] @ here |
|
327 |
| Abs _ => [Z.LookIn (Z.move_down_abs z)] @ here |
|
328 |
| _ => here |
|
329 |
end; |
|
330 |
in Z.lzy_search sf_valid_td_lr end; |
|
331 |
||
332 |
fun searchf_unify_gen f (sgn, maxidx, z) lhs = |
|
333 |
Seq.map (clean_unify_z sgn maxidx lhs) |
|
334 |
(Z.limit_apply f z); |
|
335 |
||
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
336 |
(* search all unifications *) |
23064 | 337 |
val searchf_lr_unify_all = |
338 |
searchf_unify_gen search_lr_all; |
|
15481 | 339 |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
340 |
(* search only for 'valid' unifiers (non abs subterms and non vars) *) |
23064 | 341 |
val searchf_lr_unify_valid = |
342 |
searchf_unify_gen (search_lr_valid valid_match_start); |
|
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
343 |
|
23064 | 344 |
val searchf_bt_unify_valid = |
345 |
searchf_unify_gen (search_bt_valid valid_match_start); |
|
15814
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
dixon
parents:
15550
diff
changeset
|
346 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
347 |
(* apply a substitution in the conclusion of the theorem th *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
348 |
(* cfvs are certified free var placeholders for goal params *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
349 |
(* conclthm is a theorem of for just the conclusion *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
350 |
(* m is instantiation/match information *) |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
351 |
(* rule is the equation for substitution *) |
16978 | 352 |
fun apply_subst_in_concl i th (cfvs, conclthm) rule m = |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
353 |
(RWInst.rw m rule conclthm) |
15855 | 354 |
|> IsaND.unfix_frees cfvs |
15915
b0e8b37642a4
lucas - improved interface to isand.ML and cleaned up clean-unification code, and added some better comments.
dixon
parents:
15855
diff
changeset
|
355 |
|> RWInst.beta_eta_contract |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
356 |
|> (fn r => Tactic.rtac r i th); |
15481 | 357 |
|
358 |
(* substitute within the conclusion of goal i of gth, using a meta |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
359 |
equation rule. Note that we assume rule has var indicies zero'd *) |
16978 | 360 |
fun prep_concl_subst i gth = |
361 |
let |
|
15481 | 362 |
val th = Thm.incr_indexes 1 gth; |
363 |
val tgt_term = Thm.prop_of th; |
|
364 |
||
22578 | 365 |
val sgn = Thm.theory_of_thm th; |
15481 | 366 |
val ctermify = Thm.cterm_of sgn; |
367 |
val trivify = Thm.trivial o ctermify; |
|
368 |
||
369 |
val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; |
|
370 |
val cfvs = rev (map ctermify fvs); |
|
371 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
372 |
val conclterm = Logic.strip_imp_concl fixedbody; |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
373 |
val conclthm = trivify conclterm; |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
374 |
val maxidx = Thm.maxidx_of th; |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
375 |
val ft = ((Z.move_down_right (* ==> *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
376 |
o Z.move_down_left (* Trueprop *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
377 |
o Z.mktop |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
378 |
o Thm.prop_of) conclthm) |
15481 | 379 |
in |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
380 |
((cfvs, conclthm), (sgn, maxidx, ft)) |
15481 | 381 |
end; |
382 |
||
383 |
(* substitute using an object or meta level equality *) |
|
18598 | 384 |
fun eqsubst_tac' ctxt searchf instepthm i th = |
16978 | 385 |
let |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
386 |
val (cvfsconclthm, searchinfo) = prep_concl_subst i th; |
18598 | 387 |
val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
388 |
fun rewrite_with_thm r = |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
389 |
let val (lhs,_) = Logic.dest_equals (Thm.concl_of r); |
18598 | 390 |
in searchf searchinfo lhs |
391 |
|> Seq.maps (apply_subst_in_concl i th cvfsconclthm r) end; |
|
392 |
in stepthms |> Seq.maps rewrite_with_thm end; |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
393 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
394 |
|
19047 | 395 |
(* distinct subgoals *) |
15959
366d39e95d3c
lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents:
15936
diff
changeset
|
396 |
fun distinct_subgoals th = |
19047 | 397 |
the_default th (SINGLE distinct_subgoals_tac th); |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
398 |
|
19047 | 399 |
(* General substitution of multiple occurances using one of |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
400 |
the given theorems*) |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
401 |
|
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
402 |
|
16978 | 403 |
exception eqsubst_occL_exp of |
404 |
string * (int list) * (thm list) * int * thm; |
|
405 |
fun skip_first_occs_search occ srchf sinfo lhs = |
|
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
406 |
case (skipto_skipseq occ (srchf sinfo lhs)) of |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
407 |
SkipMore _ => Seq.empty |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
408 |
| SkipSeq ss => Seq.flat ss; |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
409 |
|
22727 | 410 |
(* The occL is a list of integers indicating which occurence |
411 |
w.r.t. the search order, to rewrite. Backtracking will also find later |
|
412 |
occurences, but all earlier ones are skipped. Thus you can use [0] to |
|
413 |
just find all rewrites. *) |
|
414 |
||
18598 | 415 |
fun eqsubst_tac ctxt occL thms i th = |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
416 |
let val nprems = Thm.nprems_of th in |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
417 |
if nprems < i then Seq.empty else |
16978 | 418 |
let val thmseq = (Seq.of_list thms) |
419 |
fun apply_occ occ th = |
|
18598 | 420 |
thmseq |> Seq.maps |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
421 |
(fn r => eqsubst_tac' |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
422 |
ctxt |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
423 |
(skip_first_occs_search |
19871 | 424 |
occ searchf_lr_unify_valid) r |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
425 |
(i + ((Thm.nprems_of th) - nprems)) |
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
426 |
th); |
16978 | 427 |
val sortedoccL = |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
428 |
Library.sort (Library.rev_order o Library.int_ord) occL; |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
429 |
in |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
430 |
Seq.map distinct_subgoals (Seq.EVERY (map apply_occ sortedoccL) th) |
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
431 |
end |
15959
366d39e95d3c
lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents:
15936
diff
changeset
|
432 |
end |
366d39e95d3c
lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents:
15936
diff
changeset
|
433 |
handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); |
366d39e95d3c
lucas - fixed bug with uninstantiated type contexts in eqsubst and added the automatic removal of duplicate subgoals (when there are no flex-flex constraints)
dixon
parents:
15936
diff
changeset
|
434 |
|
15481 | 435 |
|
436 |
(* inthms are the given arguments in Isar, and treated as eqstep with |
|
437 |
the first one, then the second etc *) |
|
18598 | 438 |
fun eqsubst_meth ctxt occL inthms = |
21588 | 439 |
Method.SIMPLE_METHOD' (eqsubst_tac ctxt occL inthms); |
15481 | 440 |
|
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
441 |
(* apply a substitution inside assumption j, keeps asm in the same place *) |
16978 | 442 |
fun apply_subst_in_asm i th rule ((cfvs, j, ngoalprems, pth),m) = |
443 |
let |
|
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
444 |
val th2 = Thm.rotate_rule (j - 1) i th; (* put premice first *) |
16978 | 445 |
val preelimrule = |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
446 |
(RWInst.rw m rule pth) |
21708 | 447 |
|> (Seq.hd o prune_params_tac) |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
448 |
|> Thm.permute_prems 0 ~1 (* put old asm first *) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
449 |
|> IsaND.unfix_frees cfvs (* unfix any global params *) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
450 |
|> RWInst.beta_eta_contract; (* normal form *) |
16978 | 451 |
(* val elimrule = |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
452 |
preelimrule |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
453 |
|> Tactic.make_elim (* make into elim rule *) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
454 |
|> Thm.lift_rule (th2, i); (* lift into context *) |
16007
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
455 |
*) |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
456 |
in |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
457 |
(* ~j because new asm starts at back, thus we subtract 1 *) |
16007
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
458 |
Seq.map (Thm.rotate_rule (~j) ((Thm.nprems_of rule) + i)) |
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
459 |
(Tactic.dtac preelimrule i th2) |
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
460 |
|
16978 | 461 |
(* (Thm.bicompose |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
462 |
false (* use unification *) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
463 |
(true, (* elim resolution *) |
16007
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
464 |
elimrule, (2 + (Thm.nprems_of rule)) - ngoalprems) |
4dcccaa11a13
lucas - bugfix to subst in assumptions: fixed index error for conditional rules.
dixon
parents:
16004
diff
changeset
|
465 |
i th2) *) |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
466 |
end; |
15481 | 467 |
|
468 |
||
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
469 |
(* prepare to substitute within the j'th premise of subgoal i of gth, |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
470 |
using a meta-level equation. Note that we assume rule has var indicies |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
471 |
zero'd. Note that we also assume that premt is the j'th premice of |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
472 |
subgoal i of gth. Note the repetition of work done for each |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
473 |
assumption, i.e. this can be made more efficient for search over |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
474 |
multiple assumptions. *) |
16978 | 475 |
fun prep_subst_in_asm i gth j = |
476 |
let |
|
15481 | 477 |
val th = Thm.incr_indexes 1 gth; |
478 |
val tgt_term = Thm.prop_of th; |
|
479 |
||
22578 | 480 |
val sgn = Thm.theory_of_thm th; |
15481 | 481 |
val ctermify = Thm.cterm_of sgn; |
482 |
val trivify = Thm.trivial o ctermify; |
|
483 |
||
484 |
val (fixedbody, fvs) = IsaND.fix_alls_term i tgt_term; |
|
485 |
val cfvs = rev (map ctermify fvs); |
|
486 |
||
18011
685d95c793ff
cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents:
17795
diff
changeset
|
487 |
val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
488 |
val asm_nprems = length (Logic.strip_imp_prems asmt); |
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
489 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
490 |
val pth = trivify asmt; |
27033
6ef5134fc631
fixed bug: maxidx was wrongly calculuated from term, now calculated
dixon
parents:
23064
diff
changeset
|
491 |
val maxidx = Thm.maxidx_of th; |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
492 |
|
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
493 |
val ft = ((Z.move_down_right (* trueprop *) |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
494 |
o Z.mktop |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
495 |
o Thm.prop_of) pth) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
496 |
in ((cfvs, j, asm_nprems, pth), (sgn, maxidx, ft)) end; |
15481 | 497 |
|
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
498 |
(* prepare subst in every possible assumption *) |
16978 | 499 |
fun prep_subst_in_asms i gth = |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
500 |
map (prep_subst_in_asm i gth) |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
501 |
((fn l => Library.upto (1, length l)) |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
502 |
(Logic.prems_of_goal (Thm.prop_of gth) i)); |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
503 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
504 |
|
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
505 |
(* substitute in an assumption using an object or meta level equality *) |
18598 | 506 |
fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i th = |
16978 | 507 |
let |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
508 |
val asmpreps = prep_subst_in_asms i th; |
18598 | 509 |
val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
510 |
fun rewrite_with_thm r = |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
511 |
let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
512 |
fun occ_search occ [] = Seq.empty |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
513 |
| occ_search occ ((asminfo, searchinfo)::moreasms) = |
16978 | 514 |
(case searchf searchinfo occ lhs of |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
515 |
SkipMore i => occ_search i moreasms |
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
516 |
| SkipSeq ss => |
19861 | 517 |
Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) |
518 |
(occ_search 1 moreasms)) |
|
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
519 |
(* find later substs also *) |
16978 | 520 |
in |
18598 | 521 |
occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm i th r) |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
522 |
end; |
18598 | 523 |
in stepthms |> Seq.maps rewrite_with_thm end; |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
524 |
|
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
525 |
|
16978 | 526 |
fun skip_first_asm_occs_search searchf sinfo occ lhs = |
19835
81d6dc597559
added updated version of IsaPlanner and substitution.
dixon
parents:
19473
diff
changeset
|
527 |
skipto_skipseq occ (searchf sinfo lhs); |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
528 |
|
18598 | 529 |
fun eqsubst_asm_tac ctxt occL thms i th = |
16978 | 530 |
let val nprems = Thm.nprems_of th |
15538
d8edf54cc28c
lucas - re-arranged code and added comments. Also added check to make sure the subgoal that we are being applied to exists. If it does not, empty seq is returned.
dixon
parents:
15486
diff
changeset
|
531 |
in |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
532 |
if nprems < i then Seq.empty else |
16978 | 533 |
let val thmseq = (Seq.of_list thms) |
534 |
fun apply_occ occK th = |
|
18598 | 535 |
thmseq |> Seq.maps |
16978 | 536 |
(fn r => |
18598 | 537 |
eqsubst_asm_tac' ctxt (skip_first_asm_occs_search |
19871 | 538 |
searchf_lr_unify_valid) occK r |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
539 |
(i + ((Thm.nprems_of th) - nprems)) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
540 |
th); |
16978 | 541 |
val sortedoccs = |
16004
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
542 |
Library.sort (Library.rev_order o Library.int_ord) occL |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
543 |
in |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
544 |
Seq.map distinct_subgoals |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
545 |
(Seq.EVERY (map apply_occ sortedoccs) th) |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
546 |
end |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
547 |
end |
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
dixon
parents:
15959
diff
changeset
|
548 |
handle THM _ => raise eqsubst_occL_exp ("THM",occL,thms,i,th); |
15481 | 549 |
|
550 |
(* inthms are the given arguments in Isar, and treated as eqstep with |
|
551 |
the first one, then the second etc *) |
|
18598 | 552 |
fun eqsubst_asm_meth ctxt occL inthms = |
21588 | 553 |
Method.SIMPLE_METHOD' (eqsubst_asm_tac ctxt occL inthms); |
15481 | 554 |
|
555 |
(* syntax for options, given "(asm)" will give back true, without |
|
556 |
gives back false *) |
|
557 |
val options_syntax = |
|
558 |
(Args.parens (Args.$$$ "asm") >> (K true)) || |
|
559 |
(Scan.succeed false); |
|
15936
817ac93ee786
lucas - added ability to provide multiple replacements for subst: syntax is now: subst (1 3) myrule
dixon
parents:
15929
diff
changeset
|
560 |
|
15929
68bd1e16552a
lucas - added option to select occurance to rewrite e.g. (occ 4)
dixon
parents:
15915
diff
changeset
|
561 |
val ith_syntax = |
27809
a1e409db516b
unified Args.T with OuterLex.token, renamed some operations;
wenzelm
parents:
27033
diff
changeset
|
562 |
Scan.optional (Args.parens (Scan.repeat OuterParse.nat)) [0]; |
15481 | 563 |
|
18598 | 564 |
(* combination method that takes a flag (true indicates that subst |
565 |
should be done to an assumption, false = apply to the conclusion of |
|
566 |
the goal) as well as the theorems to use *) |
|
567 |
fun subst_meth src = |
|
18988 | 568 |
Method.syntax ((Scan.lift options_syntax) -- (Scan.lift ith_syntax) -- Attrib.thms) src |
21879 | 569 |
#> (fn (((asmflag, occL), inthms), ctxt) => |
18598 | 570 |
(if asmflag then eqsubst_asm_meth else eqsubst_meth) ctxt occL inthms); |
15481 | 571 |
|
18598 | 572 |
|
16978 | 573 |
val setup = |
18833 | 574 |
Method.add_method ("subst", subst_meth, "single-step substitution"); |
15481 | 575 |
|
16978 | 576 |
end; |