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