| author | wenzelm | 
| Mon, 15 Mar 2010 20:27:23 +0100 | |
| changeset 35799 | 7adb03f27b28 | 
| parent 35762 | af3ff2ba4c54 | 
| child 36543 | 0e7fc5bf38de | 
| permissions | -rw-r--r-- | 
| 9532 | 1 | (* Title: Provers/hypsubst.ML | 
| 2 | Authors: Martin D Coen, Tobias Nipkow and Lawrence C Paulson | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 3 | Copyright 1995 University of Cambridge | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 4 | |
| 15662 | 5 | Basic equational reasoning: hyp_subst_tac and methods "hypsubst", "subst". | 
| 9628 | 6 | |
| 7 | Tactic to substitute using (at least) the assumption x=t in the rest | |
| 8 | of the subgoal, and to delete (at least) that assumption. Original | |
| 9 | version due to Martin Coen. | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 10 | |
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 11 | This version uses the simplifier, and requires it to be already present. | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 12 | |
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 13 | Test data: | 
| 0 | 14 | |
| 9532 | 15 | Goal "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)"; | 
| 16 | Goal "!!x.[| Q(x,y,z); z=f(x); x=z |] ==> P(z)"; | |
| 17 | Goal "!!y. [| ?x=y; P(?x) |] ==> y = a"; | |
| 18 | Goal "!!z. [| ?x=y; P(?x) |] ==> y = a"; | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 19 | |
| 15415 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 20 | Goal "!!x a. [| x = f(b); g(a) = b |] ==> P(x)"; | 
| 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 21 | |
| 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 22 | by (bound_hyp_subst_tac 1); | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 23 | by (hyp_subst_tac 1); | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 24 | |
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 25 | Here hyp_subst_tac goes wrong; harder still to prove P(f(f(a))) & P(f(a)) | 
| 9532 | 26 | Goal "P(a) --> (EX y. a=y --> P(f(a)))"; | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 27 | |
| 9532 | 28 | Goal "!!x. [| Q(x,h1); P(a,h2); R(x,y,h3); R(y,z,h4); x=f(y); \ | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 29 | \ P(x,h5); P(y,h6); K(x,h7) |] ==> Q(x,c)"; | 
| 23908 | 30 | by (blast_hyp_subst_tac true 1); | 
| 0 | 31 | *) | 
| 32 | ||
| 33 | signature HYPSUBST_DATA = | |
| 21221 | 34 | sig | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 35 | val dest_Trueprop : term -> term | 
| 21221 | 36 | val dest_eq : term -> term * term | 
| 20974 | 37 | val dest_imp : term -> term * term | 
| 9532 | 38 | val eq_reflection : thm (* a=b ==> a==b *) | 
| 39 | val rev_eq_reflection: thm (* a==b ==> a=b *) | |
| 40 | val imp_intr : thm (* (P ==> Q) ==> P-->Q *) | |
| 41 | val rev_mp : thm (* [| P; P-->Q |] ==> Q *) | |
| 42 | val subst : thm (* [| a=b; P(a) |] ==> P(b) *) | |
| 43 | val sym : thm (* a=b ==> b=a *) | |
| 4223 | 44 | val thin_refl : thm (* [|x=x; PROP W|] ==> PROP W *) | 
| 27572 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 45 | val prop_subst : thm (* PROP P t ==> PROP prop (x = t ==> PROP P x) *) | 
| 21221 | 46 | end; | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 47 | |
| 0 | 48 | signature HYPSUBST = | 
| 21221 | 49 | sig | 
| 27572 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 50 | val single_hyp_subst_tac : int -> int -> tactic | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 51 | val single_hyp_meta_subst_tac : int -> int -> tactic | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 52 | val bound_hyp_subst_tac : int -> tactic | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 53 | val hyp_subst_tac : int -> tactic | 
| 23908 | 54 | val blast_hyp_subst_tac : bool -> int -> tactic | 
| 20945 | 55 | val stac : thm -> int -> tactic | 
| 18708 | 56 | val hypsubst_setup : theory -> theory | 
| 21221 | 57 | end; | 
| 2722 
3e07c20b967c
Now uses rotate_tac and eta_contract_atom for greater speed
 paulson parents: 
2174diff
changeset | 58 | |
| 9532 | 59 | functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST = | 
| 0 | 60 | struct | 
| 61 | ||
| 62 | exception EQ_VAR; | |
| 63 | ||
| 27734 
dcec1c564f05
meta_subst: xsymbols make it work with clean Pure;
 wenzelm parents: 
27572diff
changeset | 64 | val meta_subst = @{lemma "PROP P t \<Longrightarrow> PROP prop (x \<equiv> t \<Longrightarrow> PROP P x)"
 | 
| 27572 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 65 | by (unfold prop_def)} | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 66 | |
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 67 | (** Simple version: Just subtitute one hypothesis, specified by index k **) | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 68 | fun gen_single_hyp_subst_tac subst_rule k = CSUBGOAL (fn (csubg, i) => | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 69 | let | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 70 | val pat = fold_rev (Logic.all o Free) (Logic.strip_params (term_of csubg)) (Term.dummy_pattern propT) | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 71 | |> cterm_of (theory_of_cterm csubg) | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 72 | |
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 73 | val rule = | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 74 | Thm.lift_rule pat subst_rule (* lift just over parameters *) | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 75 |        |> Conv.fconv_rule (MetaSimplifier.rewrite true [@{thm prop_def}])
 | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 76 | in | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 77 | rotate_tac k i | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 78 | THEN Thm.compose_no_flatten false (rule, 1) i | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 79 | THEN rotate_tac (~k) i | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 80 | end) | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 81 | |
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 82 | val single_hyp_meta_subst_tac = gen_single_hyp_subst_tac meta_subst | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 83 | val single_hyp_subst_tac = gen_single_hyp_subst_tac Data.prop_subst | 
| 
67cd6ed76446
single_hyp(_meta)_subst_tac: Controlled substitution of a single hyp
 krauss parents: 
26992diff
changeset | 84 | |
| 17896 | 85 | fun loose (i,t) = member (op =) (add_loose_bnos (t, i, [])) 0; | 
| 0 | 86 | |
| 16979 | 87 | (*Simplifier turns Bound variables to special Free variables: | 
| 88 | change it back (any Bound variable will do)*) | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 89 | fun contract t = | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 90 | (case Envir.eta_contract t of | 
| 20074 | 91 | Free (a, T) => if Name.is_bound a then Bound 0 else Free (a, T) | 
| 16979 | 92 | | t' => t'); | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 93 | |
| 21221 | 94 | val has_vars = Term.exists_subterm Term.is_Var; | 
| 95 | val has_tvars = Term.exists_type (Term.exists_subtype Term.is_TVar); | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 96 | |
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 97 | (*If novars then we forbid Vars in the equality. | 
| 16979 | 98 | If bnd then we only look for Bound variables to eliminate. | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 99 | When can we safely delete the equality? | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 100 | Not if it equates two constants; consider 0=1. | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 101 | Not if it resembles x=t[x], since substitution does not eliminate x. | 
| 4299 | 102 | Not if it resembles ?x=0; consider ?x=0 ==> ?x=1 or even ?x=0 ==> P | 
| 9532 | 103 | Not if it involves a variable free in the premises, | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 104 | but we can't check for this -- hence bnd and bound_hyp_subst_tac | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 105 | Prefer to eliminate Bound variables if possible. | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 106 | Result: true = use as is, false = reorient first *) | 
| 21221 | 107 | fun inspect_pair bnd novars (t, u) = | 
| 108 | if novars andalso (has_tvars t orelse has_tvars u) | |
| 4179 
cc4b6791d5dc
hyp_subst_tac checks if the equality has type variables and uses a suitable
 paulson parents: 
3537diff
changeset | 109 | then raise Match (*variables in the type!*) | 
| 
cc4b6791d5dc
hyp_subst_tac checks if the equality has type variables and uses a suitable
 paulson parents: 
3537diff
changeset | 110 | else | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 111 | case (contract t, contract u) of | 
| 9532 | 112 | (Bound i, _) => if loose(i,u) orelse novars andalso has_vars u | 
| 113 | then raise Match | |
| 114 | else true (*eliminates t*) | |
| 115 | | (_, Bound i) => if loose(i,t) orelse novars andalso has_vars t | |
| 116 | then raise Match | |
| 117 | else false (*eliminates u*) | |
| 118 | | (Free _, _) => if bnd orelse Logic.occs(t,u) orelse | |
| 119 | novars andalso has_vars u | |
| 120 | then raise Match | |
| 121 | else true (*eliminates t*) | |
| 122 | | (_, Free _) => if bnd orelse Logic.occs(u,t) orelse | |
| 123 | novars andalso has_vars t | |
| 124 | then raise Match | |
| 125 | else false (*eliminates u*) | |
| 0 | 126 | | _ => raise Match; | 
| 127 | ||
| 680 
f9e24455bbd1
Provers/hypsubst: greatly simplified!  No longer simulates a
 lcp parents: 
646diff
changeset | 128 | (*Locates a substitutable variable on the left (resp. right) of an equality | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 129 | assumption. Returns the number of intervening assumptions. *) | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 130 | fun eq_var bnd novars = | 
| 680 
f9e24455bbd1
Provers/hypsubst: greatly simplified!  No longer simulates a
 lcp parents: 
646diff
changeset | 131 |   let fun eq_var_aux k (Const("all",_) $ Abs(_,_,t)) = eq_var_aux k t
 | 
| 9532 | 132 |         | eq_var_aux k (Const("==>",_) $ A $ B) =
 | 
| 133 | ((k, inspect_pair bnd novars | |
| 134 | (Data.dest_eq (Data.dest_Trueprop A))) | |
| 21227 | 135 | handle TERM _ => eq_var_aux (k+1) B | 
| 136 | | Match => eq_var_aux (k+1) B) | |
| 9532 | 137 | | eq_var_aux k _ = raise EQ_VAR | 
| 680 
f9e24455bbd1
Provers/hypsubst: greatly simplified!  No longer simulates a
 lcp parents: 
646diff
changeset | 138 | in eq_var_aux 0 end; | 
| 0 | 139 | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 140 | (*For the simpset. Adds ALL suitable equalities, even if not first! | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 141 | No vars are allowed here, as simpsets are built from meta-assumptions*) | 
| 15415 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 142 | fun mk_eqs bnd th = | 
| 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 143 | [ if inspect_pair bnd false (Data.dest_eq | 
| 9532 | 144 | (Data.dest_Trueprop (#prop (rep_thm th)))) | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 145 | then th RS Data.eq_reflection | 
| 9532 | 146 | else symmetric(th RS Data.eq_reflection) (*reorient*) ] | 
| 21227 | 147 | handle TERM _ => [] | Match => []; | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 148 | |
| 17896 | 149 | local | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 150 | in | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 151 | |
| 15415 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 152 | (*Select a suitable equality assumption; substitute throughout the subgoal | 
| 
6e437e276ef5
fix to bound_hyp_subst_tac, partially fixing a bug in inductive definitions
 paulson parents: 
13604diff
changeset | 153 | If bnd is true, then it replaces Bound variables only. *) | 
| 13604 | 154 | fun gen_hyp_subst_tac bnd = | 
| 17896 | 155 | let fun tac i st = SUBGOAL (fn (Bi, _) => | 
| 156 | let | |
| 157 | val (k, _) = eq_var bnd true Bi | |
| 35232 
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
 wenzelm parents: 
35021diff
changeset | 158 | val hyp_subst_ss = Simplifier.global_context (Thm.theory_of_thm st) empty_ss | 
| 17896 | 159 | setmksimps (mk_eqs bnd) | 
| 13604 | 160 | in EVERY [rotate_tac k i, asm_lr_simp_tac hyp_subst_ss i, | 
| 161 | etac thin_rl i, rotate_tac (~k) i] | |
| 17896 | 162 | end handle THM _ => no_tac | EQ_VAR => no_tac) i st | 
| 13604 | 163 | in REPEAT_DETERM1 o tac end; | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 164 | |
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 165 | end; | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 166 | |
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
32957diff
changeset | 167 | val ssubst = Drule.export_without_context (Data.sym RS Data.subst); | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 168 | |
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 169 | fun inst_subst_tac b rl = CSUBGOAL (fn (cBi, i) => | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 170 | case try (Logic.strip_assums_hyp #> hd #> | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 171 | Data.dest_Trueprop #> Data.dest_eq #> pairself contract) (Thm.term_of cBi) of | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 172 | SOME (t, t') => | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 173 | let | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 174 | val Bi = Thm.term_of cBi; | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 175 | val ps = Logic.strip_params Bi; | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 176 | val U = Term.fastype_of1 (rev (map snd ps), t); | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 177 | val Q = Data.dest_Trueprop (Logic.strip_assums_concl Bi); | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 178 | val rl' = Thm.lift_rule cBi rl; | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 179 | val Var (ixn, T) = Term.head_of (Data.dest_Trueprop | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 180 | (Logic.strip_assums_concl (Thm.prop_of rl'))); | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 181 | val (v1, v2) = Data.dest_eq (Data.dest_Trueprop | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 182 | (Logic.strip_assums_concl (hd (Thm.prems_of rl')))); | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 183 | val (Ts, V) = split_last (Term.binder_types T); | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 184 |         val u = list_abs (ps @ [("x", U)], case (if b then t else t') of
 | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 185 | Bound j => subst_bounds (map Bound | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 186 | ((1 upto j) @ 0 :: (j + 2 upto length ps)), Q) | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 187 | | t => Term.abstract_over (t, Term.incr_boundvars 1 Q)); | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 188 | val thy = Thm.theory_of_thm rl'; | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 189 | val (instT, _) = Thm.match (pairself (cterm_of thy o Logic.mk_type) (V, U)); | 
| 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 190 | in compose_tac (true, Drule.instantiate (instT, | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 191 | map (pairself (cterm_of thy)) | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 192 | [(Var (ixn, Ts ---> U --> body_type T), u), | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 193 | (Var (fst (dest_Var (head_of v1)), Ts ---> U), list_abs (ps, t)), | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 194 | (Var (fst (dest_Var (head_of v2)), Ts ---> U), list_abs (ps, t'))]) rl', | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 195 | nprems_of rl) i | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 196 | end | 
| 26992 
4508f20818af
inst_subst_tac: match types -- no longer assume that subst rule has exactly one type argument;
 wenzelm parents: 
26833diff
changeset | 197 | | NONE => no_tac); | 
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 198 | |
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 199 | val imp_intr_tac = rtac Data.imp_intr; | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 200 | |
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 201 | (* FIXME: "etac Data.rev_mp i" will not behave as expected if goal has *) | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 202 | (* premises containing meta-implications or quantifiers *) | 
| 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 203 | |
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 204 | (*Old version of the tactic above -- slower but the only way | 
| 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 205 | to handle equalities containing Vars.*) | 
| 3537 | 206 | fun vars_gen_hyp_subst_tac bnd = SUBGOAL(fn (Bi,i) => | 
| 207 | let val n = length(Logic.strip_assums_hyp Bi) - 1 | |
| 9532 | 208 | val (k,symopt) = eq_var bnd false Bi | 
| 209 | in | |
| 210 | DETERM | |
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 211 | (EVERY [REPEAT_DETERM_N k (etac Data.rev_mp i), | 
| 9532 | 212 | rotate_tac 1 i, | 
| 213 | REPEAT_DETERM_N (n-k) (etac Data.rev_mp i), | |
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 214 | inst_subst_tac symopt (if symopt then ssubst else Data.subst) i, | 
| 9532 | 215 | REPEAT_DETERM_N n (imp_intr_tac i THEN rotate_tac ~1 i)]) | 
| 0 | 216 | end | 
| 3537 | 217 | handle THM _ => no_tac | EQ_VAR => no_tac); | 
| 0 | 218 | |
| 219 | (*Substitutes for Free or Bound variables*) | |
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 220 | val hyp_subst_tac = FIRST' [ematch_tac [Data.thin_refl], | 
| 4223 | 221 | gen_hyp_subst_tac false, vars_gen_hyp_subst_tac false]; | 
| 0 | 222 | |
| 223 | (*Substitutes for Bound variables only -- this is always safe*) | |
| 9532 | 224 | val bound_hyp_subst_tac = | 
| 1011 
5c9654e2e3de
Recoded with help from Toby to use rewriting instead of the
 lcp parents: 
704diff
changeset | 225 | gen_hyp_subst_tac true ORELSE' vars_gen_hyp_subst_tac true; | 
| 0 | 226 | |
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 227 | |
| 9532 | 228 | (** Version for Blast_tac. Hyps that are affected by the substitution are | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 229 | moved to the front. Defect: even trivial changes are noticed, such as | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 230 | substitutions in the arguments of a function Var. **) | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 231 | |
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 232 | (*final re-reversal of the changed assumptions*) | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 233 | fun reverse_n_tac 0 i = all_tac | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 234 | | reverse_n_tac 1 i = rotate_tac ~1 i | 
| 9532 | 235 | | reverse_n_tac n i = | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 236 | REPEAT_DETERM_N n (rotate_tac ~1 i THEN etac Data.rev_mp i) THEN | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 237 | REPEAT_DETERM_N n (imp_intr_tac i THEN rotate_tac ~1 i); | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 238 | |
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 239 | (*Use imp_intr, comparing the old hyps with the new ones as they come out.*) | 
| 9532 | 240 | fun all_imp_intr_tac hyps i = | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 241 | let fun imptac (r, []) st = reverse_n_tac r i st | 
| 9532 | 242 | | imptac (r, hyp::hyps) st = | 
| 243 | let val (hyp',_) = List.nth (prems_of st, i-1) |> | |
| 244 | Logic.strip_assums_concl |> | |
| 245 | Data.dest_Trueprop |> Data.dest_imp | |
| 246 | val (r',tac) = if Pattern.aeconv (hyp,hyp') | |
| 247 | then (r, imp_intr_tac i THEN rotate_tac ~1 i) | |
| 248 | else (*leave affected hyps at end*) | |
| 249 | (r+1, imp_intr_tac i) | |
| 250 | in | |
| 251 | case Seq.pull(tac st) of | |
| 15531 | 252 | NONE => Seq.single(st) | 
| 253 | | SOME(st',_) => imptac (r',hyps) st' | |
| 21221 | 254 | end | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 255 | in imptac (0, rev hyps) end; | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 256 | |
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 257 | |
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 258 | fun blast_hyp_subst_tac trace = SUBGOAL(fn (Bi,i) => | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 259 | let val (k,symopt) = eq_var false false Bi | 
| 9532 | 260 | val hyps0 = map Data.dest_Trueprop (Logic.strip_assums_hyp Bi) | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 261 | (*omit selected equality, returning other hyps*) | 
| 9532 | 262 | val hyps = List.take(hyps0, k) @ List.drop(hyps0, k+1) | 
| 263 | val n = length hyps | |
| 264 | in | |
| 23908 | 265 | if trace then tracing "Substituting an equality" else (); | 
| 9532 | 266 | DETERM | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 267 | (EVERY [REPEAT_DETERM_N k (etac Data.rev_mp i), | 
| 9532 | 268 | rotate_tac 1 i, | 
| 269 | REPEAT_DETERM_N (n-k) (etac Data.rev_mp i), | |
| 26833 
7c3757fccf0e
Added function for computing instantiation for the subst rule, which is used
 berghofe parents: 
23908diff
changeset | 270 | inst_subst_tac symopt (if symopt then ssubst else Data.subst) i, | 
| 9532 | 271 | all_imp_intr_tac hyps i]) | 
| 4466 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 272 | end | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 273 | handle THM _ => no_tac | EQ_VAR => no_tac); | 
| 
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
 paulson parents: 
4299diff
changeset | 274 | |
| 9532 | 275 | |
| 276 | (*apply an equality or definition ONCE; | |
| 277 | fails unless the substitution has an effect*) | |
| 278 | fun stac th = | |
| 279 | let val th' = th RS Data.rev_eq_reflection handle THM _ => th | |
| 280 | in CHANGED_GOAL (rtac (th' RS ssubst)) end; | |
| 281 | ||
| 282 | ||
| 9628 | 283 | (* theory setup *) | 
| 284 | ||
| 9532 | 285 | val hypsubst_setup = | 
| 30515 | 286 |   Method.setup @{binding hypsubst}
 | 
| 287 | (Scan.succeed (K (SIMPLE_METHOD' (CHANGED_PROP o hyp_subst_tac)))) | |
| 288 | "substitution using an assumption (improper)" #> | |
| 289 |   Method.setup @{binding simplesubst} (Attrib.thm >> (fn th => K (SIMPLE_METHOD' (stac th))))
 | |
| 290 | "simple substitution"; | |
| 9532 | 291 | |
| 0 | 292 | end; |