| author | wenzelm | 
| Thu, 01 Jan 2009 22:20:29 +0100 | |
| changeset 29299 | df4300a1acd3 | 
| parent 26830 | 7b7139f961bd | 
| child 32449 | 696d64ed85da | 
| permissions | -rw-r--r-- | 
| 1459 | 1  | 
(* Title: FOLP/hypsubst  | 
| 1022 | 2  | 
ID: $Id$  | 
| 1459 | 3  | 
Author: Martin D Coen, Cambridge University Computer Laboratory  | 
| 1022 | 4  | 
Copyright 1995 University of Cambridge  | 
5  | 
||
6  | 
Original version of Provers/hypsubst. Cannot use new version because it  | 
|
7  | 
relies on the new simplifier!  | 
|
8  | 
||
9  | 
Martin Coen's tactic for substitution in the hypotheses  | 
|
10  | 
*)  | 
|
11  | 
||
12  | 
signature HYPSUBST_DATA =  | 
|
13  | 
sig  | 
|
| 1459 | 14  | 
val dest_eq : term -> term*term  | 
15  | 
val imp_intr : thm (* (P ==> Q) ==> P-->Q *)  | 
|
16  | 
val rev_mp : thm (* [| P; P-->Q |] ==> Q *)  | 
|
17  | 
val subst : thm (* [| a=b; P(a) |] ==> P(b) *)  | 
|
18  | 
val sym : thm (* a=b ==> b=a *)  | 
|
| 1022 | 19  | 
end;  | 
20  | 
||
21  | 
signature HYPSUBST =  | 
|
22  | 
sig  | 
|
23  | 
val bound_hyp_subst_tac : int -> tactic  | 
|
24  | 
val hyp_subst_tac : int -> tactic  | 
|
25  | 
(*exported purely for debugging purposes*)  | 
|
26  | 
val eq_var : bool -> term -> int * thm  | 
|
27  | 
val inspect_pair : bool -> term * term -> thm  | 
|
28  | 
end;  | 
|
29  | 
||
30  | 
functor HypsubstFun(Data: HYPSUBST_DATA): HYPSUBST =  | 
|
31  | 
struct  | 
|
32  | 
||
33  | 
local open Data in  | 
|
34  | 
||
35  | 
exception EQ_VAR;  | 
|
36  | 
||
37  | 
fun loose (i,t) = 0 mem add_loose_bnos(t,i,[]);  | 
|
38  | 
||
39  | 
(*It's not safe to substitute for a constant; consider 0=1.  | 
|
40  | 
It's not safe to substitute for x=t[x] since x is not eliminated.  | 
|
41  | 
It's not safe to substitute for a Var; what if it appears in other goals?  | 
|
42  | 
It's not safe to substitute for a variable free in the premises,  | 
|
43  | 
but how could we check for this?*)  | 
|
44  | 
fun inspect_pair bnd (t,u) =  | 
|
| 
26830
 
7b7139f961bd
Replaced Pattern.eta_contract_atom by Envir.eta_contract.
 
berghofe 
parents: 
3537 
diff
changeset
 | 
45  | 
case (Envir.eta_contract t, Envir.eta_contract u) of  | 
| 1022 | 46  | 
(Bound i, _) => if loose(i,u) then raise Match  | 
| 1459 | 47  | 
else sym (*eliminates t*)  | 
| 1022 | 48  | 
| (_, Bound i) => if loose(i,t) then raise Match  | 
| 1459 | 49  | 
else asm_rl (*eliminates u*)  | 
| 1022 | 50  | 
| (Free _, _) => if bnd orelse Logic.occs(t,u) then raise Match  | 
| 1459 | 51  | 
else sym (*eliminates t*)  | 
| 1022 | 52  | 
| (_, Free _) => if bnd orelse Logic.occs(u,t) then raise Match  | 
| 1459 | 53  | 
else asm_rl (*eliminates u*)  | 
| 1022 | 54  | 
| _ => raise Match;  | 
55  | 
||
56  | 
(*Locates a substitutable variable on the left (resp. right) of an equality  | 
|
57  | 
assumption. Returns the number of intervening assumptions, paried with  | 
|
58  | 
the rule asm_rl (resp. sym). *)  | 
|
59  | 
fun eq_var bnd =  | 
|
60  | 
  let fun eq_var_aux k (Const("all",_) $ Abs(_,_,t)) = eq_var_aux k t
 | 
|
| 1459 | 61  | 
        | eq_var_aux k (Const("==>",_) $ A $ B) = 
 | 
62  | 
((k, inspect_pair bnd (dest_eq A))  | 
|
63  | 
(*Exception Match comes from inspect_pair or dest_eq*)  | 
|
64  | 
handle Match => eq_var_aux (k+1) B)  | 
|
65  | 
| eq_var_aux k _ = raise EQ_VAR  | 
|
| 1022 | 66  | 
in eq_var_aux 0 end;  | 
67  | 
||
68  | 
(*Select a suitable equality assumption and substitute throughout the subgoal  | 
|
69  | 
Replaces only Bound variables if bnd is true*)  | 
|
| 3537 | 70  | 
fun gen_hyp_subst_tac bnd = SUBGOAL(fn (Bi,i) =>  | 
71  | 
let val n = length(Logic.strip_assums_hyp Bi) - 1  | 
|
| 1459 | 72  | 
val (k,symopt) = eq_var bnd Bi  | 
| 1022 | 73  | 
in  | 
| 3537 | 74  | 
DETERM  | 
75  | 
(EVERY [REPEAT_DETERM_N k (etac rev_mp i),  | 
|
76  | 
etac revcut_rl i,  | 
|
77  | 
REPEAT_DETERM_N (n-k) (etac rev_mp i),  | 
|
78  | 
etac (symopt RS subst) i,  | 
|
79  | 
REPEAT_DETERM_N n (rtac imp_intr i)])  | 
|
| 1022 | 80  | 
end  | 
| 3537 | 81  | 
handle THM _ => no_tac | EQ_VAR => no_tac);  | 
| 1022 | 82  | 
|
83  | 
(*Substitutes for Free or Bound variables*)  | 
|
84  | 
val hyp_subst_tac = gen_hyp_subst_tac false;  | 
|
85  | 
||
86  | 
(*Substitutes for Bound variables only -- this is always safe*)  | 
|
87  | 
val bound_hyp_subst_tac = gen_hyp_subst_tac true;  | 
|
88  | 
||
89  | 
end  | 
|
90  | 
end;  | 
|
91  |