# HG changeset patch # User lcp # Date 783776643 -3600 # Node ID f9e24455bbd17ebf26044803c3e18f5bee22b3fb # Parent a682bbf70dc631c0f86c3b769d41de15e6a32fba Provers/hypsubst: greatly simplified! No longer simulates a "eres_inst_tac" using rev_cut_eq; instead simply rotates the chosen equality to the end! diff -r a682bbf70dc6 -r f9e24455bbd1 src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML Wed Nov 02 10:45:14 1994 +0100 +++ b/src/Provers/hypsubst.ML Wed Nov 02 12:44:03 1994 +0100 @@ -8,12 +8,11 @@ signature HYPSUBST_DATA = sig - val dest_eq: term -> term*term - val imp_intr: thm (* (P ==> Q) ==> P-->Q *) - val rev_cut_eq: thm (* [| a=b; a=b ==> R |] ==> R *) - val rev_mp: thm (* [| P; P-->Q |] ==> Q *) - val subst: thm (* [| a=b; P(a) |] ==> P(b) *) - val sym: thm (* a=b ==> b=a *) + val dest_eq : term -> term*term + val imp_intr : thm (* (P ==> Q) ==> P-->Q *) + val rev_mp : thm (* [| P; P-->Q |] ==> Q *) + val subst : thm (* [| a=b; P(a) |] ==> P(b) *) + val sym : thm (* a=b ==> b=a *) end; signature HYPSUBST = @@ -21,9 +20,8 @@ val bound_hyp_subst_tac : int -> tactic val hyp_subst_tac : int -> tactic (*exported purely for debugging purposes*) - val eq_var : bool -> term -> term * thm - val inspect_pair : bool -> term * term -> term * thm - val liftvar : int -> typ list -> term + val eq_var : bool -> term -> int * thm + val inspect_pair : bool -> term * term -> thm end; functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST = @@ -35,14 +33,6 @@ | REPEATN n tac = Tactic(fn state => tapply(tac THEN REPEATN (n-1) tac, state)); -local - val T = case #1 (types_sorts rev_cut_eq) ("a",0) of - Some T => T - | None => error"No such variable in rev_cut_eq" -in - fun liftvar inc paramTs = Var(("a",inc), paramTs ---> incr_tvar inc T); -end; - exception EQ_VAR; fun loose (i,t) = 0 mem add_loose_bnos(t,i,[]); @@ -55,50 +45,39 @@ fun inspect_pair bnd (t,u) = case (Pattern.eta_contract t, Pattern.eta_contract u) of (Bound i, _) => if loose(i,u) then raise Match - else (t, asm_rl) + else sym (*eliminates t*) | (_, Bound i) => if loose(i,t) then raise Match - else (u, sym) + else asm_rl (*eliminates u*) | (Free _, _) => if bnd orelse Logic.occs(t,u) then raise Match - else (t, asm_rl) + else sym (*eliminates t*) | (_, Free _) => if bnd orelse Logic.occs(u,t) then raise Match - else (u, sym) + else asm_rl (*eliminates u*) | _ => raise Match; - (* Extracts the name of the variable on the left (resp. right) of an equality - assumption. Returns the rule asm_rl (resp. sym). *) -fun eq_var bnd (Const("all",_) $ Abs(_,_,t)) = eq_var bnd t - | eq_var bnd (Const("==>",_) $ A $ B) = - (inspect_pair bnd (dest_eq A) - (*Match comes from inspect_pair or dest_eq*) - handle Match => eq_var bnd B) - | eq_var bnd _ = raise EQ_VAR; - -(*Lift and instantiate a rule wrt the given state and subgoal number *) -fun lift_instpair (state, i, t, rule) = - let val {maxidx,sign,...} = rep_thm state - val (_, _, Bi, _) = dest_state(state,i) - val params = Logic.strip_params Bi - val var = liftvar (maxidx+1) (map #2 params) - and u = Unify.rlist_abs(rev params, t) - and cterm = cterm_of sign - in cterm_instantiate [(cterm var, cterm u)] (lift_rule (state,i) rule) - end; - -fun eres_instpair_tac t rule i = STATE (fn state => - compose_tac (true, lift_instpair (state, i, t, rule), - length(prems_of rule)) i); - -val ssubst = sym RS subst; +(*Locates a substitutable variable on the left (resp. right) of an equality + assumption. Returns the number of intervening assumptions, paried with + the rule asm_rl (resp. sym). *) +fun eq_var bnd = + let fun eq_var_aux k (Const("all",_) $ Abs(_,_,t)) = eq_var_aux k t + | eq_var_aux k (Const("==>",_) $ A $ B) = + ((k, inspect_pair bnd (dest_eq A)) + (*Exception Match comes from inspect_pair or dest_eq*) + handle Match => eq_var_aux (k+1) B) + | eq_var_aux k _ = raise EQ_VAR + in eq_var_aux 0 end; (*Select a suitable equality assumption and substitute throughout the subgoal Replaces only Bound variables if bnd is true*) fun gen_hyp_subst_tac bnd i = DETERM (STATE(fn state => let val (_,_,Bi,_) = dest_state(state,i) val n = length(Logic.strip_assums_hyp Bi) - 1 - val (t,symopt) = eq_var bnd Bi - in eres_instpair_tac t (symopt RS rev_cut_eq) i THEN - REPEATN n (etac rev_mp i) THEN - etac ssubst i THEN REPEATN n (rtac imp_intr i) + val (k,symopt) = eq_var bnd Bi + in + EVERY [REPEATN k (etac rev_mp i), + etac revcut_rl i, + REPEATN (n-k) (etac rev_mp i), + etac (symopt RS subst) i, + REPEATN n (rtac imp_intr i)] end handle THM _ => no_tac | EQ_VAR => no_tac));