| author | blanchet | 
| Thu, 25 Jun 2015 12:41:43 +0200 | |
| changeset 60568 | a9b71c82647b | 
| parent 60358 | aebfbcab1eb8 | 
| child 67149 | e61557884799 | 
| permissions | -rw-r--r-- | 
| 30160 
5f7b17941730
moved some generic tools to src/Tools/ -- src/Provers is essentially obsolete;
 wenzelm parents: 
29269diff
changeset | 1 | (* Title: Tools/eqsubst.ML | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
27809diff
changeset | 2 | Author: Lucas Dixon, University of Edinburgh | 
| 15481 | 3 | |
| 52235 | 4 | Perform a substitution using an equation. | 
| 18598 | 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: 
15486diff
changeset | 6 | |
| 18591 | 7 | signature EQSUBST = | 
| 15481 | 8 | sig | 
| 19871 | 9 | type match = | 
| 52234 | 10 | ((indexname * (sort * typ)) list (* type instantiations *) | 
| 11 | * (indexname * (typ * term)) list) (* term instantiations *) | |
| 12 | * (string * typ) list (* fake named type abs env *) | |
| 13 | * (string * typ) list (* type abs env *) | |
| 14 | * term (* outer term *) | |
| 19871 | 15 | |
| 16 | type searchinfo = | |
| 60358 | 17 | Proof.context | 
| 52234 | 18 | * int (* maxidx *) | 
| 19 | * Zipper.T (* focusterm to search under *) | |
| 19871 | 20 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 21 | datatype 'a skipseq = SkipMore of int | SkipSeq of 'a Seq.seq Seq.seq | 
| 19871 | 22 | |
| 52234 | 23 |   val skip_first_asm_occs_search: ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> int -> 'b -> 'c skipseq
 | 
| 24 |   val skip_first_occs_search: int -> ('a -> 'b -> 'c Seq.seq Seq.seq) -> 'a -> 'b -> 'c Seq.seq
 | |
| 25 | val skipto_skipseq: int -> 'a Seq.seq Seq.seq -> 'a skipseq | |
| 19871 | 26 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 27 | (* tactics *) | 
| 52234 | 28 | val eqsubst_asm_tac: Proof.context -> int list -> thm list -> int -> tactic | 
| 29 | val eqsubst_asm_tac': Proof.context -> | |
| 30 | (searchinfo -> int -> term -> match skipseq) -> int -> thm -> int -> tactic | |
| 31 | val eqsubst_tac: Proof.context -> | |
| 58318 | 32 | int list -> (* list of occurrences to rewrite, use [0] for any *) | 
| 52234 | 33 | thm list -> int -> tactic | 
| 34 | val eqsubst_tac': Proof.context -> | |
| 35 | (searchinfo -> term -> match Seq.seq) (* search function *) | |
| 36 | -> thm (* equation theorem to rewrite with *) | |
| 37 | -> int (* subgoal number in goal theorem *) | |
| 38 | -> thm (* goal theorem *) | |
| 39 | -> thm Seq.seq (* rewritten goal theorem *) | |
| 19871 | 40 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 41 | (* search for substitutions *) | 
| 52234 | 42 | val valid_match_start: Zipper.T -> bool | 
| 43 | val search_lr_all: Zipper.T -> Zipper.T Seq.seq | |
| 44 | val search_lr_valid: (Zipper.T -> bool) -> Zipper.T -> Zipper.T Seq.seq | |
| 45 | val searchf_lr_unify_all: searchinfo -> term -> match Seq.seq Seq.seq | |
| 46 | val searchf_lr_unify_valid: searchinfo -> term -> match Seq.seq Seq.seq | |
| 47 | val searchf_bt_unify_valid: searchinfo -> term -> match Seq.seq Seq.seq | |
| 15481 | 48 | end; | 
| 49 | ||
| 41164 
6854e9a40edc
avoid ML structure aliases (especially single-letter abbreviations);
 wenzelm parents: 
40722diff
changeset | 50 | structure EqSubst: EQSUBST = | 
| 
6854e9a40edc
avoid ML structure aliases (especially single-letter abbreviations);
 wenzelm parents: 
40722diff
changeset | 51 | struct | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 52 | |
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 53 | (* changes object "=" to meta "==" which prepares a given rewrite rule *) | 
| 18598 | 54 | fun prep_meta_eq ctxt = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49340diff
changeset | 55 | Simplifier.mksimps ctxt #> map Drule.zero_var_indexes; | 
| 18598 | 56 | |
| 52246 | 57 | (* make free vars into schematic vars with index zero *) | 
| 58 | fun unfix_frees frees = | |
| 59 | fold (K (Thm.forall_elim_var 0)) frees o Drule.forall_intr_list frees; | |
| 60 | ||
| 15481 | 61 | |
| 52234 | 62 | type match = | 
| 52235 | 63 | ((indexname * (sort * typ)) list (* type instantiations *) | 
| 64 | * (indexname * (typ * term)) list) (* term instantiations *) | |
| 65 | * (string * typ) list (* fake named type abs env *) | |
| 66 | * (string * typ) list (* type abs env *) | |
| 67 | * term; (* outer term *) | |
| 15550 
806214035275
lucas - added more comments and an extra type to clarify the code.
 dixon parents: 
15538diff
changeset | 68 | |
| 52234 | 69 | type searchinfo = | 
| 60358 | 70 | Proof.context | 
| 52235 | 71 | * int (* maxidx *) | 
| 72 | * Zipper.T; (* focusterm to search under *) | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 73 | |
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 74 | |
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 75 | (* skipping non-empty sub-sequences but when we reach the end | 
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 76 | of the seq, remembering how much we have left to skip. *) | 
| 52234 | 77 | datatype 'a skipseq = | 
| 78 | SkipMore of int | | |
| 79 | SkipSeq of 'a Seq.seq Seq.seq; | |
| 80 | ||
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 81 | (* given a seqseq, skip the first m non-empty seq's, note deficit *) | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 82 | fun skipto_skipseq m s = | 
| 52234 | 83 | let | 
| 84 | fun skip_occs n sq = | |
| 85 | (case Seq.pull sq of | |
| 86 | NONE => SkipMore n | |
| 52237 | 87 | | SOME (h, t) => | 
| 52234 | 88 | (case Seq.pull h of | 
| 89 | NONE => skip_occs n t | |
| 90 | | SOME _ => if n <= 1 then SkipSeq (Seq.cons h t) else skip_occs (n - 1) t)) | |
| 91 | in skip_occs m s end; | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 92 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 93 | (* note: outerterm is the taget with the match replaced by a bound | 
| 52234 | 94 | variable : ie: "P lhs" beocmes "%x. P x" | 
| 95 | insts is the types of instantiations of vars in lhs | |
| 96 | and typinsts is the type instantiations of types in the lhs | |
| 97 | Note: Final rule is the rule lifted into the ontext of the | |
| 98 | taget thm. *) | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 99 | fun mk_foo_match mkuptermfunc Ts t = | 
| 52234 | 100 | let | 
| 101 | val ty = Term.type_of t | |
| 52235 | 102 | val bigtype = rev (map snd Ts) ---> ty | 
| 52234 | 103 | fun mk_foo 0 t = t | 
| 104 | | mk_foo i t = mk_foo (i - 1) (t $ (Bound (i - 1))) | |
| 52235 | 105 | val num_of_bnds = length Ts | 
| 52234 | 106 | (* foo_term = "fooabs y0 ... yn" where y's are local bounds *) | 
| 107 | val foo_term = mk_foo num_of_bnds (Bound num_of_bnds) | |
| 108 |   in Abs ("fooabs", bigtype, mkuptermfunc foo_term) end;
 | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 109 | |
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 110 | (* T is outer bound vars, n is number of locally bound vars *) | 
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 111 | (* THINK: is order of Ts correct...? or reversed? *) | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 112 | fun mk_fake_bound_name n = ":b_" ^ n; | 
| 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 113 | fun fakefree_badbounds Ts t = | 
| 52234 | 114 | let val (FakeTs, Ts, newnames) = | 
| 52242 | 115 | fold_rev (fn (n, ty) => fn (FakeTs, Ts, usednames) => | 
| 52234 | 116 | let | 
| 117 | val newname = singleton (Name.variant_list usednames) n | |
| 118 | in | |
| 119 | ((mk_fake_bound_name newname, ty) :: FakeTs, | |
| 120 | (newname, ty) :: Ts, | |
| 121 | newname :: usednames) | |
| 52242 | 122 | end) Ts ([], [], []) | 
| 52234 | 123 | in (FakeTs, Ts, Term.subst_bounds (map Free FakeTs, t)) end; | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 124 | |
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 125 | (* before matching we need to fake the bound vars that are missing an | 
| 52235 | 126 | abstraction. In this function we additionally construct the | 
| 127 | abstraction environment, and an outer context term (with the focus | |
| 52240 | 128 | abstracted out) for use in rewriting with RW_Inst.rw *) | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 129 | fun prep_zipper_match z = | 
| 52234 | 130 | let | 
| 131 | val t = Zipper.trm z | |
| 132 | val c = Zipper.ctxt z | |
| 133 | val Ts = Zipper.C.nty_ctxt c | |
| 134 | val (FakeTs', Ts', t') = fakefree_badbounds Ts t | |
| 135 | val absterm = mk_foo_match (Zipper.C.apply c) Ts' t' | |
| 136 | in | |
| 137 | (t', (FakeTs', Ts', absterm)) | |
| 138 | end; | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 139 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 140 | (* Unification with exception handled *) | 
| 60358 | 141 | (* given context, max var index, pat, tgt; returns Seq of instantiations *) | 
| 142 | fun clean_unify ctxt ix (a as (pat, tgt)) = | |
| 52234 | 143 | let | 
| 144 | (* type info will be re-derived, maybe this can be cached | |
| 145 | for efficiency? *) | |
| 146 | val pat_ty = Term.type_of pat; | |
| 147 | val tgt_ty = Term.type_of tgt; | |
| 52235 | 148 | (* FIXME is it OK to ignore the type instantiation info? | 
| 52234 | 149 | or should I be using it? *) | 
| 150 | val typs_unify = | |
| 60358 | 151 | SOME (Sign.typ_unify (Proof_Context.theory_of ctxt) (pat_ty, tgt_ty) (Vartab.empty, ix)) | 
| 52234 | 152 | handle Type.TUNIFY => NONE; | 
| 153 | in | |
| 154 | (case typs_unify of | |
| 155 | SOME (typinsttab, ix2) => | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 156 | let | 
| 52234 | 157 | (* FIXME is it right to throw away the flexes? | 
| 158 | or should I be using them somehow? *) | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 159 | fun mk_insts env = | 
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 160 | (Vartab.dest (Envir.type_env env), | 
| 32032 | 161 | Vartab.dest (Envir.term_env env)); | 
| 162 | val initenv = | |
| 163 |             Envir.Envir {maxidx = ix2, tenv = Vartab.empty, tyenv = typinsttab};
 | |
| 60358 | 164 | val useq = Unify.smash_unifiers (Context.Proof ctxt) [a] initenv | 
| 52234 | 165 | handle ListPair.UnequalLengths => Seq.empty | 
| 166 | | Term.TERM _ => Seq.empty; | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 167 | fun clean_unify' useq () = | 
| 52234 | 168 | (case (Seq.pull useq) of | 
| 169 | NONE => NONE | |
| 170 | | SOME (h, t) => SOME (mk_insts h, Seq.make (clean_unify' t))) | |
| 171 | handle ListPair.UnequalLengths => NONE | |
| 172 | | Term.TERM _ => NONE; | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 173 | in | 
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 174 | (Seq.make (clean_unify' useq)) | 
| 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 175 | end | 
| 52234 | 176 | | NONE => Seq.empty) | 
| 177 | end; | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 178 | |
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 179 | (* Unification for zippers *) | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 180 | (* Note: Ts is a modified version of the original names of the outer | 
| 52235 | 181 | bound variables. New names have been introduced to make sure they are | 
| 182 | unique w.r.t all names in the term and each other. usednames' is | |
| 183 | oldnames + new names. *) | |
| 60358 | 184 | fun clean_unify_z ctxt maxidx pat z = | 
| 52235 | 185 | let val (t, (FakeTs, Ts, absterm)) = prep_zipper_match z in | 
| 49339 
d1fcb4de8349
eliminated some old material that is unused in the visible universe;
 wenzelm parents: 
44095diff
changeset | 186 | Seq.map (fn insts => (insts, FakeTs, Ts, absterm)) | 
| 60358 | 187 | (clean_unify ctxt maxidx (t, pat)) | 
| 52234 | 188 | end; | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 189 | |
| 15550 
806214035275
lucas - added more comments and an extra type to clarify the code.
 dixon parents: 
15538diff
changeset | 190 | |
| 52234 | 191 | fun bot_left_leaf_of (l $ _) = bot_left_leaf_of l | 
| 192 | | bot_left_leaf_of (Abs (_, _, t)) = bot_left_leaf_of t | |
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 193 | | 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: 
15486diff
changeset | 194 | |
| 19975 
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
 dixon parents: 
19871diff
changeset | 195 | (* 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: 
19871diff
changeset | 196 | they always succeed trivially, and uninterestingly. *) | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 197 | fun valid_match_start z = | 
| 52234 | 198 | (case bot_left_leaf_of (Zipper.trm z) of | 
| 199 | Var _ => false | |
| 200 | | _ => true); | |
| 19975 
ecd684d62808
fix to subst in order to allow subst when head of a term is a bound variable.
 dixon parents: 
19871diff
changeset | 201 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15550diff
changeset | 202 | (* search from top, left to right, then down *) | 
| 19871 | 203 | val search_lr_all = ZipperSearch.all_bl_ur; | 
| 15481 | 204 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15550diff
changeset | 205 | (* search from top, left to right, then down *) | 
| 19871 | 206 | fun search_lr_valid validf = | 
| 52234 | 207 | let | 
| 208 | fun sf_valid_td_lr z = | |
| 209 | let val here = if validf z then [Zipper.Here z] else [] in | |
| 210 | (case Zipper.trm z of | |
| 211 | _ $ _ => | |
| 212 | [Zipper.LookIn (Zipper.move_down_left z)] @ here @ | |
| 213 | [Zipper.LookIn (Zipper.move_down_right z)] | |
| 214 | | Abs _ => here @ [Zipper.LookIn (Zipper.move_down_abs z)] | |
| 215 | | _ => here) | |
| 216 | end; | |
| 217 | 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: 
15550diff
changeset | 218 | |
| 23064 | 219 | (* search from bottom to top, left to right *) | 
| 220 | fun search_bt_valid validf = | |
| 52234 | 221 | let | 
| 222 | fun sf_valid_td_lr z = | |
| 223 | let val here = if validf z then [Zipper.Here z] else [] in | |
| 224 | (case Zipper.trm z of | |
| 225 | _ $ _ => | |
| 226 | [Zipper.LookIn (Zipper.move_down_left z), | |
| 227 | Zipper.LookIn (Zipper.move_down_right z)] @ here | |
| 228 | | Abs _ => [Zipper.LookIn (Zipper.move_down_abs z)] @ here | |
| 229 | | _ => here) | |
| 230 | end; | |
| 231 | in Zipper.lzy_search sf_valid_td_lr end; | |
| 23064 | 232 | |
| 60358 | 233 | fun searchf_unify_gen f (ctxt, maxidx, z) lhs = | 
| 234 | Seq.map (clean_unify_z ctxt maxidx lhs) (Zipper.limit_apply f z); | |
| 23064 | 235 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15550diff
changeset | 236 | (* search all unifications *) | 
| 52234 | 237 | val searchf_lr_unify_all = searchf_unify_gen search_lr_all; | 
| 15481 | 238 | |
| 15814 
d65f461c8672
lucas - fixed a big with renaming of bound variables. Other small changes.
 dixon parents: 
15550diff
changeset | 239 | (* search only for 'valid' unifiers (non abs subterms and non vars) *) | 
| 52234 | 240 | val searchf_lr_unify_valid = 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: 
15915diff
changeset | 241 | |
| 52234 | 242 | val searchf_bt_unify_valid = 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: 
15550diff
changeset | 243 | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 244 | (* apply a substitution in the conclusion of the theorem *) | 
| 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: 
15486diff
changeset | 245 | (* 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: 
15486diff
changeset | 246 | (* 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: 
15486diff
changeset | 247 | (* 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: 
15486diff
changeset | 248 | (* rule is the equation for substitution *) | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 249 | fun apply_subst_in_concl ctxt i st (cfvs, conclthm) rule m = | 
| 52240 | 250 | RW_Inst.rw ctxt m rule conclthm | 
| 52246 | 251 | |> unfix_frees cfvs | 
| 52239 | 252 | |> Conv.fconv_rule Drule.beta_eta_conversion | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58950diff
changeset | 253 | |> (fn r => resolve_tac ctxt [r] i st); | 
| 15481 | 254 | |
| 255 | (* 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: 
15486diff
changeset | 256 | equation rule. Note that we assume rule has var indicies zero'd *) | 
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 257 | fun prep_concl_subst ctxt i gth = | 
| 52234 | 258 | let | 
| 259 | val th = Thm.incr_indexes 1 gth; | |
| 260 | val tgt_term = Thm.prop_of th; | |
| 15481 | 261 | |
| 52234 | 262 | val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; | 
| 59642 | 263 | val cfvs = rev (map (Thm.cterm_of ctxt) fvs); | 
| 15481 | 264 | |
| 52234 | 265 | val conclterm = Logic.strip_imp_concl fixedbody; | 
| 59642 | 266 | val conclthm = Thm.trivial (Thm.cterm_of ctxt conclterm); | 
| 52234 | 267 | val maxidx = Thm.maxidx_of th; | 
| 268 | val ft = | |
| 269 | (Zipper.move_down_right (* ==> *) | |
| 270 | o Zipper.move_down_left (* Trueprop *) | |
| 271 | o Zipper.mktop | |
| 272 | o Thm.prop_of) conclthm | |
| 273 | in | |
| 60358 | 274 | ((cfvs, conclthm), (ctxt, maxidx, ft)) | 
| 52234 | 275 | end; | 
| 15481 | 276 | |
| 277 | (* substitute using an object or meta level equality *) | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 278 | fun eqsubst_tac' ctxt searchf instepthm i st = | 
| 52234 | 279 | let | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 280 | val (cvfsconclthm, searchinfo) = prep_concl_subst ctxt i st; | 
| 52234 | 281 | val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); | 
| 282 | fun rewrite_with_thm r = | |
| 283 | let val (lhs,_) = Logic.dest_equals (Thm.concl_of r) in | |
| 284 | searchf searchinfo lhs | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 285 | |> Seq.maps (apply_subst_in_concl ctxt i st cvfsconclthm r) | 
| 52234 | 286 | end; | 
| 287 | 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: 
15486diff
changeset | 288 | |
| 
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: 
15486diff
changeset | 289 | |
| 58318 | 290 | (* General substitution of multiple occurrences using one of | 
| 52235 | 291 | the given theorems *) | 
| 19835 
81d6dc597559
added updated version of IsaPlanner and substitution.
 dixon parents: 
19473diff
changeset | 292 | |
| 16978 | 293 | fun skip_first_occs_search occ srchf sinfo lhs = | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 294 | (case skipto_skipseq occ (srchf sinfo lhs) of | 
| 52234 | 295 | SkipMore _ => Seq.empty | 
| 296 | | SkipSeq ss => Seq.flat ss); | |
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 297 | |
| 58318 | 298 | (* The "occs" argument is a list of integers indicating which occurrence | 
| 22727 | 299 | w.r.t. the search order, to rewrite. Backtracking will also find later | 
| 58318 | 300 | occurrences, but all earlier ones are skipped. Thus you can use [0] to | 
| 22727 | 301 | just find all rewrites. *) | 
| 302 | ||
| 52238 | 303 | fun eqsubst_tac ctxt occs thms i st = | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 304 | let val nprems = Thm.nprems_of st in | 
| 52234 | 305 | if nprems < i then Seq.empty else | 
| 306 | let | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 307 | val thmseq = Seq.of_list thms; | 
| 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 308 | fun apply_occ occ st = | 
| 52234 | 309 | thmseq |> Seq.maps (fn r => | 
| 310 | eqsubst_tac' ctxt | |
| 311 | (skip_first_occs_search occ searchf_lr_unify_valid) r | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 312 | (i + (Thm.nprems_of st - nprems)) st); | 
| 52238 | 313 | val sorted_occs = Library.sort (rev_order o int_ord) occs; | 
| 52234 | 314 | in | 
| 52238 | 315 | Seq.maps distinct_subgoals_tac (Seq.EVERY (map apply_occ sorted_occs) st) | 
| 52234 | 316 | end | 
| 317 | 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: 
15936diff
changeset | 318 | |
| 15481 | 319 | |
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 320 | (* apply a substitution inside assumption j, keeps asm in the same place *) | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 321 | fun apply_subst_in_asm ctxt i st rule ((cfvs, j, _, pth),m) = | 
| 52234 | 322 | let | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 323 | val st2 = Thm.rotate_rule (j - 1) i st; (* put premice first *) | 
| 52234 | 324 | val preelimrule = | 
| 52240 | 325 | RW_Inst.rw ctxt m rule pth | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
53168diff
changeset | 326 | |> (Seq.hd o prune_params_tac ctxt) | 
| 52234 | 327 | |> Thm.permute_prems 0 ~1 (* put old asm first *) | 
| 52246 | 328 | |> unfix_frees cfvs (* unfix any global params *) | 
| 52239 | 329 | |> Conv.fconv_rule Drule.beta_eta_conversion; (* normal form *) | 
| 52234 | 330 | in | 
| 331 | (* ~j because new asm starts at back, thus we subtract 1 *) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58950diff
changeset | 332 | Seq.map (Thm.rotate_rule (~ j) (Thm.nprems_of rule + i)) | 
| 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58950diff
changeset | 333 | (dresolve_tac ctxt [preelimrule] i st2) | 
| 52234 | 334 | end; | 
| 15481 | 335 | |
| 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: 
15486diff
changeset | 337 | (* 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: 
15486diff
changeset | 338 | 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: 
15486diff
changeset | 339 | 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: 
15486diff
changeset | 340 | 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: 
15486diff
changeset | 341 | 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: 
15486diff
changeset | 342 | multiple assumptions. *) | 
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 343 | fun prep_subst_in_asm ctxt i gth j = | 
| 52234 | 344 | let | 
| 345 | val th = Thm.incr_indexes 1 gth; | |
| 346 | val tgt_term = Thm.prop_of th; | |
| 15481 | 347 | |
| 52234 | 348 | val (fixedbody, fvs) = IsaND.fix_alls_term ctxt i tgt_term; | 
| 60358 | 349 | val cfvs = rev (map (Thm.cterm_of ctxt) fvs); | 
| 15481 | 350 | |
| 52234 | 351 | val asmt = nth (Logic.strip_imp_prems fixedbody) (j - 1); | 
| 352 | val asm_nprems = length (Logic.strip_imp_prems asmt); | |
| 353 | ||
| 60358 | 354 | val pth = Thm.trivial ((Thm.cterm_of ctxt) asmt); | 
| 52234 | 355 | 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: 
15486diff
changeset | 356 | |
| 52234 | 357 | val ft = | 
| 358 | (Zipper.move_down_right (* trueprop *) | |
| 359 | o Zipper.mktop | |
| 360 | o Thm.prop_of) pth | |
| 60358 | 361 | in ((cfvs, j, asm_nprems, pth), (ctxt, maxidx, ft)) end; | 
| 15481 | 362 | |
| 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: 
15486diff
changeset | 363 | (* prepare subst in every possible assumption *) | 
| 49340 
25fc6e0da459
observe context more carefully when producing "fresh" variables -- for increased chances that method "subst" works in local context (including that of forked proofs);
 wenzelm parents: 
49339diff
changeset | 364 | fun prep_subst_in_asms ctxt i gth = | 
| 52234 | 365 | map (prep_subst_in_asm ctxt i gth) | 
| 366 | ((fn l => Library.upto (1, length l)) | |
| 367 | (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: 
15486diff
changeset | 368 | |
| 
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: 
15486diff
changeset | 369 | |
| 
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: 
15486diff
changeset | 370 | (* substitute in an assumption using an object or meta level equality *) | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 371 | fun eqsubst_asm_tac' ctxt searchf skipocc instepthm i st = | 
| 52234 | 372 | let | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 373 | val asmpreps = prep_subst_in_asms ctxt i st; | 
| 52234 | 374 | val stepthms = Seq.of_list (prep_meta_eq ctxt instepthm); | 
| 375 | fun rewrite_with_thm r = | |
| 376 | let | |
| 377 | val (lhs,_) = Logic.dest_equals (Thm.concl_of r); | |
| 378 | fun occ_search occ [] = Seq.empty | |
| 379 | | occ_search occ ((asminfo, searchinfo)::moreasms) = | |
| 380 | (case searchf searchinfo occ lhs of | |
| 381 | SkipMore i => occ_search i moreasms | |
| 382 | | SkipSeq ss => | |
| 383 | Seq.append (Seq.map (Library.pair asminfo) (Seq.flat ss)) | |
| 384 | (occ_search 1 moreasms)) (* find later substs also *) | |
| 385 | in | |
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 386 | occ_search skipocc asmpreps |> Seq.maps (apply_subst_in_asm ctxt i st r) | 
| 52234 | 387 | end; | 
| 388 | 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: 
15486diff
changeset | 389 | |
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 390 | |
| 16978 | 391 | fun skip_first_asm_occs_search searchf sinfo occ lhs = | 
| 52234 | 392 | skipto_skipseq occ (searchf sinfo lhs); | 
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 393 | |
| 52238 | 394 | fun eqsubst_asm_tac ctxt occs thms i st = | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 395 | let val nprems = Thm.nprems_of st in | 
| 52234 | 396 | if nprems < i then Seq.empty | 
| 397 | else | |
| 398 | let | |
| 399 | val thmseq = Seq.of_list thms; | |
| 52238 | 400 | fun apply_occ occ st = | 
| 52234 | 401 | thmseq |> Seq.maps (fn r => | 
| 402 | eqsubst_asm_tac' ctxt | |
| 52238 | 403 | (skip_first_asm_occs_search searchf_lr_unify_valid) occ r | 
| 52236 
fb82b42eb498
tuned -- prefer terminology of tactic / goal state;
 wenzelm parents: 
52235diff
changeset | 404 | (i + (Thm.nprems_of st - nprems)) st); | 
| 52238 | 405 | val sorted_occs = Library.sort (rev_order o int_ord) occs; | 
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 406 | in | 
| 52238 | 407 | Seq.maps distinct_subgoals_tac (Seq.EVERY (map apply_occ sorted_occs) st) | 
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15959diff
changeset | 408 | end | 
| 52234 | 409 | end; | 
| 15481 | 410 | |
| 18598 | 411 | (* combination method that takes a flag (true indicates that subst | 
| 31301 | 412 | should be done to an assumption, false = apply to the conclusion of | 
| 413 | the goal) as well as the theorems to use *) | |
| 58826 | 414 | val _ = | 
| 415 | Theory.setup | |
| 416 |     (Method.setup @{binding subst}
 | |
| 417 | (Scan.lift (Args.mode "asm" -- Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) -- | |
| 418 | Attrib.thms >> (fn ((asm, occs), inthms) => fn ctxt => | |
| 419 | SIMPLE_METHOD' ((if asm then eqsubst_asm_tac else eqsubst_tac) ctxt occs inthms))) | |
| 420 | "single-step substitution"); | |
| 15481 | 421 | |
| 16978 | 422 | end; |