| author | chaieb | 
| Fri, 27 Mar 2009 17:35:21 +0000 | |
| changeset 30748 | fe67d729a61c | 
| parent 30528 | 7173bf123335 | 
| child 31241 | b3c7044d47b6 | 
| permissions | -rw-r--r-- | 
| 
14620
 
1be590fd2422
Minor cleanup of headers and some speedup of the HOL4 import.
 
skalberg 
parents: 
14518 
diff
changeset
 | 
1  | 
(* Title: HOL/Import/shuffler.ML  | 
| 14516 | 2  | 
Author: Sebastian Skalberg, TU Muenchen  | 
3  | 
||
4  | 
Package for proving two terms equal by normalizing (hence the  | 
|
5  | 
"shuffler" name). Uses the simplifier for the normalization.  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature Shuffler =  | 
|
9  | 
sig  | 
|
10  | 
val debug : bool ref  | 
|
11  | 
||
12  | 
val norm_term : theory -> term -> thm  | 
|
13  | 
val make_equal : theory -> term -> term -> thm option  | 
|
14  | 
val set_prop : theory -> term -> (string * thm) list -> (string * thm) option  | 
|
15  | 
||
16  | 
val find_potential: theory -> term -> (string * thm) list  | 
|
17  | 
||
18  | 
val gen_shuffle_tac: theory -> bool -> (string * thm) list -> int -> tactic  | 
|
19  | 
||
20  | 
val shuffle_tac: (string * thm) list -> int -> tactic  | 
|
21  | 
val search_tac : (string * thm) list -> int -> tactic  | 
|
22  | 
||
23  | 
val print_shuffles: theory -> unit  | 
|
24  | 
||
25  | 
val add_shuffle_rule: thm -> theory -> theory  | 
|
| 18728 | 26  | 
val shuffle_attr: attribute  | 
| 14516 | 27  | 
|
| 18708 | 28  | 
val setup : theory -> theory  | 
| 14516 | 29  | 
end  | 
30  | 
||
31  | 
structure Shuffler :> Shuffler =  | 
|
32  | 
struct  | 
|
33  | 
||
34  | 
val debug = ref false  | 
|
35  | 
||
36  | 
fun if_debug f x = if !debug then f x else ()  | 
|
37  | 
val message = if_debug writeln  | 
|
38  | 
||
39  | 
(*Prints exceptions readably to users*)  | 
|
| 21588 | 40  | 
fun print_sign_exn_unit sign e =  | 
| 14516 | 41  | 
case e of  | 
42  | 
THM (msg,i,thms) =>  | 
|
| 21588 | 43  | 
         (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
 | 
| 26928 | 44  | 
List.app Display.print_thm thms)  | 
| 14516 | 45  | 
| THEORY (msg,thys) =>  | 
| 21588 | 46  | 
         (writeln ("Exception THEORY raised:\n" ^ msg);
 | 
47  | 
List.app (writeln o Context.str_of_thy) thys)  | 
|
| 14516 | 48  | 
| TERM (msg,ts) =>  | 
| 21588 | 49  | 
         (writeln ("Exception TERM raised:\n" ^ msg);
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
50  | 
List.app (writeln o Syntax.string_of_term_global sign) ts)  | 
| 14516 | 51  | 
| TYPE (msg,Ts,ts) =>  | 
| 21588 | 52  | 
         (writeln ("Exception TYPE raised:\n" ^ msg);
 | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
53  | 
List.app (writeln o Syntax.string_of_typ_global sign) Ts;  | 
| 
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26928 
diff
changeset
 | 
54  | 
List.app (writeln o Syntax.string_of_term_global sign) ts)  | 
| 14516 | 55  | 
| e => raise e  | 
56  | 
||
57  | 
(*Prints an exception, then fails*)  | 
|
58  | 
fun print_sign_exn sign e = (print_sign_exn_unit sign e; raise e)  | 
|
59  | 
||
| 26928 | 60  | 
val string_of_thm = PrintMode.setmp [] Display.string_of_thm;  | 
61  | 
val string_of_cterm = PrintMode.setmp [] Display.string_of_cterm;  | 
|
| 14516 | 62  | 
|
63  | 
fun mk_meta_eq th =  | 
|
64  | 
(case concl_of th of  | 
|
| 21588 | 65  | 
         Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th RS eq_reflection
 | 
| 14516 | 66  | 
       | Const("==",_) $ _ $ _ => th
 | 
67  | 
       | _ => raise THM("Not an equality",0,[th]))
 | 
|
| 
28397
 
389c5e494605
handle _ should be avoided (spurious Interrupt will spoil the game);
 
wenzelm 
parents: 
27865 
diff
changeset
 | 
68  | 
    handle _ => raise THM("Couldn't make meta equality",0,[th])  (* FIXME avoid handle _ *)
 | 
| 21588 | 69  | 
|
| 14516 | 70  | 
fun mk_obj_eq th =  | 
71  | 
(case concl_of th of  | 
|
| 21588 | 72  | 
         Const("Trueprop",_) $ (Const("op =",_) $ _ $ _) => th
 | 
| 14516 | 73  | 
       | Const("==",_) $ _ $ _ => th RS meta_eq_to_obj_eq
 | 
74  | 
       | _ => raise THM("Not an equality",0,[th]))
 | 
|
| 
28397
 
389c5e494605
handle _ should be avoided (spurious Interrupt will spoil the game);
 
wenzelm 
parents: 
27865 
diff
changeset
 | 
75  | 
    handle _ => raise THM("Couldn't make object equality",0,[th])  (* FIXME avoid handle _ *)
 | 
| 14516 | 76  | 
|
| 22846 | 77  | 
structure ShuffleData = TheoryDataFun  | 
78  | 
(  | 
|
79  | 
type T = thm list  | 
|
80  | 
val empty = []  | 
|
81  | 
val copy = I  | 
|
82  | 
val extend = I  | 
|
83  | 
fun merge _ = Library.gen_union Thm.eq_thm  | 
|
84  | 
)  | 
|
| 14516 | 85  | 
|
| 22846 | 86  | 
fun print_shuffles thy =  | 
87  | 
Pretty.writeln (Pretty.big_list "Shuffle theorems:"  | 
|
88  | 
(map Display.pretty_thm (ShuffleData.get thy)))  | 
|
| 14516 | 89  | 
|
90  | 
val weaken =  | 
|
91  | 
let  | 
|
| 26424 | 92  | 
val cert = cterm_of Pure.thy  | 
| 21588 | 93  | 
        val P = Free("P",propT)
 | 
94  | 
        val Q = Free("Q",propT)
 | 
|
95  | 
val PQ = Logic.mk_implies(P,Q)  | 
|
96  | 
val PPQ = Logic.mk_implies(P,PQ)  | 
|
97  | 
val cP = cert P  | 
|
98  | 
val cQ = cert Q  | 
|
99  | 
val cPQ = cert PQ  | 
|
100  | 
val cPPQ = cert PPQ  | 
|
101  | 
val th1 = assume cPQ |> implies_intr_list [cPQ,cP]  | 
|
102  | 
val th3 = assume cP  | 
|
103  | 
val th4 = implies_elim_list (assume cPPQ) [th3,th3]  | 
|
104  | 
|> implies_intr_list [cPPQ,cP]  | 
|
| 14516 | 105  | 
in  | 
| 21588 | 106  | 
equal_intr th4 th1 |> standard  | 
| 14516 | 107  | 
end  | 
108  | 
||
109  | 
val imp_comm =  | 
|
110  | 
let  | 
|
| 26424 | 111  | 
val cert = cterm_of Pure.thy  | 
| 21588 | 112  | 
        val P = Free("P",propT)
 | 
113  | 
        val Q = Free("Q",propT)
 | 
|
114  | 
        val R = Free("R",propT)
 | 
|
115  | 
val PQR = Logic.mk_implies(P,Logic.mk_implies(Q,R))  | 
|
116  | 
val QPR = Logic.mk_implies(Q,Logic.mk_implies(P,R))  | 
|
117  | 
val cP = cert P  | 
|
118  | 
val cQ = cert Q  | 
|
119  | 
val cPQR = cert PQR  | 
|
120  | 
val cQPR = cert QPR  | 
|
121  | 
val th1 = implies_elim_list (assume cPQR) [assume cP,assume cQ]  | 
|
122  | 
|> implies_intr_list [cPQR,cQ,cP]  | 
|
123  | 
val th2 = implies_elim_list (assume cQPR) [assume cQ,assume cP]  | 
|
124  | 
|> implies_intr_list [cQPR,cP,cQ]  | 
|
| 14516 | 125  | 
in  | 
| 21588 | 126  | 
equal_intr th1 th2 |> standard  | 
| 14516 | 127  | 
end  | 
128  | 
||
129  | 
val def_norm =  | 
|
130  | 
let  | 
|
| 26424 | 131  | 
val cert = cterm_of Pure.thy  | 
| 21588 | 132  | 
        val aT = TFree("'a",[])
 | 
133  | 
        val bT = TFree("'b",[])
 | 
|
134  | 
        val v = Free("v",aT)
 | 
|
135  | 
        val P = Free("P",aT-->bT)
 | 
|
136  | 
        val Q = Free("Q",aT-->bT)
 | 
|
137  | 
        val cvPQ = cert (list_all ([("v",aT)],Logic.mk_equals(P $ Bound 0,Q $ Bound 0)))
 | 
|
138  | 
val cPQ = cert (Logic.mk_equals(P,Q))  | 
|
139  | 
val cv = cert v  | 
|
140  | 
val rew = assume cvPQ  | 
|
141  | 
|> forall_elim cv  | 
|
142  | 
|> abstract_rule "v" cv  | 
|
143  | 
val (lhs,rhs) = Logic.dest_equals(concl_of rew)  | 
|
144  | 
val th1 = transitive (transitive  | 
|
145  | 
(eta_conversion (cert lhs) |> symmetric)  | 
|
146  | 
rew)  | 
|
147  | 
(eta_conversion (cert rhs))  | 
|
148  | 
|> implies_intr cvPQ  | 
|
149  | 
val th2 = combination (assume cPQ) (reflexive cv)  | 
|
150  | 
|> forall_intr cv  | 
|
151  | 
|> implies_intr cPQ  | 
|
| 14516 | 152  | 
in  | 
| 21588 | 153  | 
equal_intr th1 th2 |> standard  | 
| 14516 | 154  | 
end  | 
155  | 
||
156  | 
val all_comm =  | 
|
157  | 
let  | 
|
| 26424 | 158  | 
val cert = cterm_of Pure.thy  | 
| 21588 | 159  | 
        val xT = TFree("'a",[])
 | 
160  | 
        val yT = TFree("'b",[])
 | 
|
| 27330 | 161  | 
        val x = Free("x",xT)
 | 
162  | 
        val y = Free("y",yT)
 | 
|
| 21588 | 163  | 
        val P = Free("P",xT-->yT-->propT)
 | 
| 27330 | 164  | 
val lhs = Logic.all x (Logic.all y (P $ x $ y))  | 
165  | 
val rhs = Logic.all y (Logic.all x (P $ x $ y))  | 
|
| 21588 | 166  | 
val cl = cert lhs  | 
167  | 
val cr = cert rhs  | 
|
| 27330 | 168  | 
val cx = cert x  | 
169  | 
val cy = cert y  | 
|
| 21588 | 170  | 
val th1 = assume cr  | 
171  | 
|> forall_elim_list [cy,cx]  | 
|
172  | 
|> forall_intr_list [cx,cy]  | 
|
173  | 
|> implies_intr cr  | 
|
174  | 
val th2 = assume cl  | 
|
175  | 
|> forall_elim_list [cx,cy]  | 
|
176  | 
|> forall_intr_list [cy,cx]  | 
|
177  | 
|> implies_intr cl  | 
|
| 14516 | 178  | 
in  | 
| 21588 | 179  | 
equal_intr th1 th2 |> standard  | 
| 14516 | 180  | 
end  | 
181  | 
||
182  | 
val equiv_comm =  | 
|
183  | 
let  | 
|
| 26424 | 184  | 
val cert = cterm_of Pure.thy  | 
| 21588 | 185  | 
        val T    = TFree("'a",[])
 | 
186  | 
        val t    = Free("t",T)
 | 
|
187  | 
        val u    = Free("u",T)
 | 
|
188  | 
val ctu = cert (Logic.mk_equals(t,u))  | 
|
189  | 
val cut = cert (Logic.mk_equals(u,t))  | 
|
190  | 
val th1 = assume ctu |> symmetric |> implies_intr ctu  | 
|
191  | 
val th2 = assume cut |> symmetric |> implies_intr cut  | 
|
| 14516 | 192  | 
in  | 
| 21588 | 193  | 
equal_intr th1 th2 |> standard  | 
| 14516 | 194  | 
end  | 
195  | 
||
196  | 
(* This simplification procedure rewrites !!x y. P x y  | 
|
197  | 
deterministicly, in order for the normalization function, defined  | 
|
198  | 
below, to handle nested quantifiers robustly *)  | 
|
199  | 
||
200  | 
local  | 
|
201  | 
||
202  | 
exception RESULT of int  | 
|
203  | 
||
204  | 
fun find_bound n (Bound i) = if i = n then raise RESULT 0  | 
|
| 21588 | 205  | 
else if i = n+1 then raise RESULT 1  | 
206  | 
else ()  | 
|
| 14516 | 207  | 
| find_bound n (t $ u) = (find_bound n t; find_bound n u)  | 
208  | 
| find_bound n (Abs(_,_,t)) = find_bound (n+1) t  | 
|
209  | 
| find_bound _ _ = ()  | 
|
210  | 
||
211  | 
fun swap_bound n (Bound i) = if i = n then Bound (n+1)  | 
|
| 21588 | 212  | 
else if i = n+1 then Bound n  | 
213  | 
else Bound i  | 
|
| 14516 | 214  | 
| swap_bound n (t $ u) = (swap_bound n t $ swap_bound n u)  | 
215  | 
| swap_bound n (Abs(x,xT,t)) = Abs(x,xT,swap_bound (n+1) t)  | 
|
216  | 
| swap_bound n t = t  | 
|
217  | 
||
| 21078 | 218  | 
fun rew_th thy (xv as (x,xT)) (yv as (y,yT)) t =  | 
| 14516 | 219  | 
let  | 
| 21588 | 220  | 
val lhs = list_all ([xv,yv],t)  | 
221  | 
val rhs = list_all ([yv,xv],swap_bound 0 t)  | 
|
222  | 
val rew = Logic.mk_equals (lhs,rhs)  | 
|
223  | 
val init = trivial (cterm_of thy rew)  | 
|
| 14516 | 224  | 
in  | 
| 21588 | 225  | 
(all_comm RS init handle e => (message "rew_th"; OldGoals.print_exn e))  | 
| 14516 | 226  | 
end  | 
227  | 
||
| 21078 | 228  | 
fun quant_rewrite thy assumes (t as Const("all",T1) $ (Abs(x,xT,Const("all",T2) $ Abs(y,yT,body)))) =
 | 
| 14516 | 229  | 
let  | 
| 21588 | 230  | 
val res = (find_bound 0 body;2) handle RESULT i => i  | 
| 14516 | 231  | 
in  | 
| 21588 | 232  | 
case res of  | 
233  | 
0 => SOME (rew_th thy (x,xT) (y,yT) body)  | 
|
234  | 
| 1 => if string_ord(y,x) = LESS  | 
|
235  | 
then  | 
|
236  | 
let  | 
|
237  | 
                         val newt = Const("all",T1) $ (Abs(y,xT,Const("all",T2) $ Abs(x,yT,body)))
 | 
|
238  | 
val t_th = reflexive (cterm_of thy t)  | 
|
239  | 
val newt_th = reflexive (cterm_of thy newt)  | 
|
240  | 
in  | 
|
241  | 
SOME (transitive t_th newt_th)  | 
|
242  | 
end  | 
|
243  | 
else NONE  | 
|
244  | 
| _ => error "norm_term (quant_rewrite) internal error"  | 
|
| 14516 | 245  | 
end  | 
| 15531 | 246  | 
| quant_rewrite _ _ _ = (warning "quant_rewrite: Unknown lhs"; NONE)  | 
| 14516 | 247  | 
|
248  | 
fun freeze_thaw_term t =  | 
|
249  | 
let  | 
|
| 
29270
 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
250  | 
val tvars = OldTerm.term_tvars t  | 
| 
 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
251  | 
val tfree_names = OldTerm.add_term_tfree_names(t,[])  | 
| 21588 | 252  | 
val (type_inst,_) =  | 
253  | 
Library.foldl (fn ((inst,used),(w as (v,_),S)) =>  | 
|
254  | 
let  | 
|
255  | 
val v' = Name.variant used v  | 
|
256  | 
in  | 
|
257  | 
((w,TFree(v',S))::inst,v'::used)  | 
|
258  | 
end)  | 
|
259  | 
(([],tfree_names),tvars)  | 
|
260  | 
val t' = subst_TVars type_inst t  | 
|
| 14516 | 261  | 
in  | 
| 21588 | 262  | 
(t',map (fn (w,TFree(v,S)) => (v,TVar(w,S))  | 
263  | 
| _ => error "Internal error in Shuffler.freeze_thaw") type_inst)  | 
|
| 14516 | 264  | 
end  | 
265  | 
||
| 21078 | 266  | 
fun inst_tfrees thy [] thm = thm  | 
| 21588 | 267  | 
| inst_tfrees thy ((name,U)::rest) thm =  | 
| 14516 | 268  | 
let  | 
| 21588 | 269  | 
val cU = ctyp_of thy U  | 
| 
29270
 
0eade173f77e
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
270  | 
val tfrees = OldTerm.add_term_tfrees (prop_of thm,[])  | 
| 21588 | 271  | 
val (rens, thm') = Thm.varifyT'  | 
| 
20951
 
868120282837
gen_rem(s) abandoned in favour of remove / subtract
 
haftmann 
parents: 
20897 
diff
changeset
 | 
272  | 
(remove (op = o apsnd fst) name tfrees) thm  | 
| 21588 | 273  | 
val mid =  | 
274  | 
case rens of  | 
|
275  | 
[] => thm'  | 
|
276  | 
| [((_, S), idx)] => instantiate  | 
|
| 21078 | 277  | 
([(ctyp_of thy (TVar (idx, S)), cU)], []) thm'  | 
| 21588 | 278  | 
| _ => error "Shuffler.inst_tfrees internal error"  | 
| 14516 | 279  | 
in  | 
| 21588 | 280  | 
inst_tfrees thy rest mid  | 
| 14516 | 281  | 
end  | 
282  | 
||
283  | 
fun is_Abs (Abs _) = true  | 
|
284  | 
| is_Abs _ = false  | 
|
285  | 
||
286  | 
fun eta_redex (t $ Bound 0) =  | 
|
287  | 
let  | 
|
| 21588 | 288  | 
fun free n (Bound i) = i = n  | 
289  | 
| free n (t $ u) = free n t orelse free n u  | 
|
290  | 
| free n (Abs(_,_,t)) = free (n+1) t  | 
|
291  | 
| free n _ = false  | 
|
| 14516 | 292  | 
in  | 
| 21588 | 293  | 
not (free 0 t)  | 
| 14516 | 294  | 
end  | 
295  | 
| eta_redex _ = false  | 
|
296  | 
||
| 21078 | 297  | 
fun eta_contract thy assumes origt =  | 
| 14516 | 298  | 
let  | 
| 21588 | 299  | 
val (typet,Tinst) = freeze_thaw_term origt  | 
300  | 
val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet))  | 
|
301  | 
val final = inst_tfrees thy Tinst o thaw  | 
|
302  | 
val t = #1 (Logic.dest_equals (prop_of init))  | 
|
303  | 
val _ =  | 
|
304  | 
let  | 
|
305  | 
val lhs = #1 (Logic.dest_equals (prop_of (final init)))  | 
|
306  | 
in  | 
|
307  | 
if not (lhs aconv origt)  | 
|
308  | 
then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";  | 
|
| 26928 | 309  | 
writeln (Display.string_of_cterm (cterm_of thy origt));  | 
310  | 
writeln (Display.string_of_cterm (cterm_of thy lhs));  | 
|
311  | 
writeln (Display.string_of_cterm (cterm_of thy typet));  | 
|
312  | 
writeln (Display.string_of_cterm (cterm_of thy t));  | 
|
313  | 
app (fn (n,T) => writeln (n ^ ": " ^ (Display.string_of_ctyp (ctyp_of thy T)))) Tinst;  | 
|
| 21588 | 314  | 
writeln "done")  | 
315  | 
else ()  | 
|
316  | 
end  | 
|
| 14516 | 317  | 
in  | 
| 21588 | 318  | 
case t of  | 
319  | 
            Const("all",_) $ (Abs(x,xT,Const("==",eqT) $ P $ Q)) =>
 | 
|
320  | 
((if eta_redex P andalso eta_redex Q  | 
|
321  | 
then  | 
|
322  | 
let  | 
|
323  | 
val cert = cterm_of thy  | 
|
| 29281 | 324  | 
val v = Free (Name.variant (Term.add_free_names t []) "v", xT)  | 
| 21588 | 325  | 
val cv = cert v  | 
326  | 
val ct = cert t  | 
|
327  | 
val th = (assume ct)  | 
|
328  | 
|> forall_elim cv  | 
|
329  | 
|> abstract_rule x cv  | 
|
330  | 
val ext_th = eta_conversion (cert (Abs(x,xT,P)))  | 
|
331  | 
val th' = transitive (symmetric ext_th) th  | 
|
332  | 
val cu = cert (prop_of th')  | 
|
333  | 
val uth = combination (assume cu) (reflexive cv)  | 
|
334  | 
val uth' = (beta_conversion false (cert (Abs(x,xT,Q) $ v)))  | 
|
335  | 
|> transitive uth  | 
|
336  | 
|> forall_intr cv  | 
|
337  | 
|> implies_intr cu  | 
|
338  | 
val rew_th = equal_intr (th' |> implies_intr ct) uth'  | 
|
339  | 
val res = final rew_th  | 
|
340  | 
val lhs = (#1 (Logic.dest_equals (prop_of res)))  | 
|
341  | 
in  | 
|
342  | 
SOME res  | 
|
343  | 
end  | 
|
344  | 
else NONE)  | 
|
345  | 
handle e => OldGoals.print_exn e)  | 
|
346  | 
| _ => NONE  | 
|
| 
17440
 
df77edc4f5d0
fixed HOL-light/Isabelle syntax incompatability via more protect_xxx functions
 
obua 
parents: 
17188 
diff
changeset
 | 
347  | 
end  | 
| 14516 | 348  | 
|
| 21078 | 349  | 
fun beta_fun thy assume t =  | 
350  | 
SOME (beta_conversion true (cterm_of thy t))  | 
|
| 14516 | 351  | 
|
| 17188 | 352  | 
val meta_sym_rew = thm "refl"  | 
353  | 
||
| 21078 | 354  | 
fun equals_fun thy assume t =  | 
| 17188 | 355  | 
case t of  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29265 
diff
changeset
 | 
356  | 
        Const("op ==",_) $ u $ v => if TermOrd.term_ord (u,v) = LESS then SOME (meta_sym_rew) else NONE
 | 
| 17188 | 357  | 
| _ => NONE  | 
358  | 
||
| 21078 | 359  | 
fun eta_expand thy assumes origt =  | 
| 14516 | 360  | 
let  | 
| 21588 | 361  | 
val (typet,Tinst) = freeze_thaw_term origt  | 
362  | 
val (init,thaw) = freeze_thaw (reflexive (cterm_of thy typet))  | 
|
363  | 
val final = inst_tfrees thy Tinst o thaw  | 
|
364  | 
val t = #1 (Logic.dest_equals (prop_of init))  | 
|
365  | 
val _ =  | 
|
366  | 
let  | 
|
367  | 
val lhs = #1 (Logic.dest_equals (prop_of (final init)))  | 
|
368  | 
in  | 
|
369  | 
if not (lhs aconv origt)  | 
|
370  | 
then (writeln "Something is utterly wrong: (orig,lhs,frozen type,t,tinst)";  | 
|
| 26928 | 371  | 
writeln (Display.string_of_cterm (cterm_of thy origt));  | 
372  | 
writeln (Display.string_of_cterm (cterm_of thy lhs));  | 
|
373  | 
writeln (Display.string_of_cterm (cterm_of thy typet));  | 
|
374  | 
writeln (Display.string_of_cterm (cterm_of thy t));  | 
|
375  | 
app (fn (n,T) => writeln (n ^ ": " ^ (Display.string_of_ctyp (ctyp_of thy T)))) Tinst;  | 
|
| 21588 | 376  | 
writeln "done")  | 
377  | 
else ()  | 
|
378  | 
end  | 
|
| 14516 | 379  | 
in  | 
| 21588 | 380  | 
case t of  | 
381  | 
            Const("==",T) $ P $ Q =>
 | 
|
382  | 
if is_Abs P orelse is_Abs Q  | 
|
383  | 
then (case domain_type T of  | 
|
384  | 
                      Type("fun",[aT,bT]) =>
 | 
|
385  | 
let  | 
|
386  | 
val cert = cterm_of thy  | 
|
| 29281 | 387  | 
val vname = Name.variant (Term.add_free_names t []) "v"  | 
| 21588 | 388  | 
val v = Free(vname,aT)  | 
389  | 
val cv = cert v  | 
|
390  | 
val ct = cert t  | 
|
391  | 
val th1 = (combination (assume ct) (reflexive cv))  | 
|
392  | 
|> forall_intr cv  | 
|
393  | 
|> implies_intr ct  | 
|
394  | 
val concl = cert (concl_of th1)  | 
|
395  | 
val th2 = (assume concl)  | 
|
396  | 
|> forall_elim cv  | 
|
397  | 
|> abstract_rule vname cv  | 
|
398  | 
val (lhs,rhs) = Logic.dest_equals (prop_of th2)  | 
|
399  | 
val elhs = eta_conversion (cert lhs)  | 
|
400  | 
val erhs = eta_conversion (cert rhs)  | 
|
401  | 
val th2' = transitive  | 
|
402  | 
(transitive (symmetric elhs) th2)  | 
|
403  | 
erhs  | 
|
404  | 
val res = equal_intr th1 (th2' |> implies_intr concl)  | 
|
405  | 
val res' = final res  | 
|
406  | 
in  | 
|
407  | 
SOME res'  | 
|
408  | 
end  | 
|
409  | 
| _ => NONE)  | 
|
410  | 
else NONE  | 
|
411  | 
          | _ => (error ("Bad eta_expand argument" ^ (string_of_cterm (cterm_of thy t))); NONE)
 | 
|
| 14516 | 412  | 
end  | 
| 17959 | 413  | 
handle e => (writeln "eta_expand internal error"; OldGoals.print_exn e)  | 
| 14516 | 414  | 
|
| 14854 | 415  | 
fun mk_tfree s = TFree("'"^s,[])
 | 
| 20326 | 416  | 
fun mk_free s t = Free (s,t)  | 
| 14516 | 417  | 
val xT = mk_tfree "a"  | 
418  | 
val yT = mk_tfree "b"  | 
|
| 27330 | 419  | 
val x = Free ("x", xT)
 | 
420  | 
val y = Free ("y", yT)
 | 
|
| 20326 | 421  | 
val P = mk_free "P" (xT-->yT-->propT)  | 
422  | 
val Q = mk_free "Q" (xT-->yT)  | 
|
423  | 
val R = mk_free "R" (xT-->yT)  | 
|
424  | 
val S = mk_free "S" xT  | 
|
425  | 
val S' = mk_free "S'" xT  | 
|
| 14516 | 426  | 
in  | 
| 21078 | 427  | 
fun beta_simproc thy = Simplifier.simproc_i  | 
| 21588 | 428  | 
thy  | 
429  | 
"Beta-contraction"  | 
|
430  | 
                      [Abs("x",xT,Q) $ S]
 | 
|
431  | 
beta_fun  | 
|
| 14516 | 432  | 
|
| 21078 | 433  | 
fun equals_simproc thy = Simplifier.simproc_i  | 
| 21588 | 434  | 
thy  | 
435  | 
"Ordered rewriting of meta equalities"  | 
|
436  | 
                      [Const("op ==",xT) $ S $ S']
 | 
|
437  | 
equals_fun  | 
|
| 17188 | 438  | 
|
| 21078 | 439  | 
fun quant_simproc thy = Simplifier.simproc_i  | 
| 21588 | 440  | 
thy  | 
441  | 
"Ordered rewriting of nested quantifiers"  | 
|
| 27330 | 442  | 
[Logic.all x (Logic.all y (P $ x $ y))]  | 
| 21588 | 443  | 
quant_rewrite  | 
| 21078 | 444  | 
fun eta_expand_simproc thy = Simplifier.simproc_i  | 
| 21588 | 445  | 
thy  | 
446  | 
"Smart eta-expansion by equivalences"  | 
|
447  | 
[Logic.mk_equals(Q,R)]  | 
|
448  | 
eta_expand  | 
|
| 21078 | 449  | 
fun eta_contract_simproc thy = Simplifier.simproc_i  | 
| 21588 | 450  | 
thy  | 
451  | 
"Smart handling of eta-contractions"  | 
|
| 27330 | 452  | 
[Logic.all x (Logic.mk_equals (Q $ x, R $ x))]  | 
| 21588 | 453  | 
eta_contract  | 
| 14516 | 454  | 
end  | 
455  | 
||
456  | 
(* Disambiguates the names of bound variables in a term, returning t  | 
|
457  | 
== t' where all the names of bound variables in t' are unique *)  | 
|
458  | 
||
| 21078 | 459  | 
fun disamb_bound thy t =  | 
| 14516 | 460  | 
let  | 
| 21588 | 461  | 
|
462  | 
fun F (t $ u,idx) =  | 
|
463  | 
let  | 
|
464  | 
val (t',idx') = F (t,idx)  | 
|
465  | 
val (u',idx'') = F (u,idx')  | 
|
466  | 
in  | 
|
467  | 
(t' $ u',idx'')  | 
|
468  | 
end  | 
|
469  | 
| F (Abs(x,xT,t),idx) =  | 
|
470  | 
let  | 
|
471  | 
val x' = "x" ^ (LargeInt.toString idx) (* amazing *)  | 
|
472  | 
val (t',idx') = F (t,idx+1)  | 
|
473  | 
in  | 
|
474  | 
(Abs(x',xT,t'),idx')  | 
|
475  | 
end  | 
|
476  | 
| F arg = arg  | 
|
477  | 
val (t',_) = F (t,0)  | 
|
478  | 
val ct = cterm_of thy t  | 
|
479  | 
val ct' = cterm_of thy t'  | 
|
480  | 
val res = transitive (reflexive ct) (reflexive ct')  | 
|
481  | 
        val _ = message ("disamb_term: " ^ (string_of_thm res))
 | 
|
| 14516 | 482  | 
in  | 
| 21588 | 483  | 
res  | 
| 14516 | 484  | 
end  | 
485  | 
||
486  | 
(* Transforms a term t to some normal form t', returning the theorem t  | 
|
487  | 
== t'. This is originally a help function for make_equal, but might  | 
|
488  | 
be handy in its own right, for example for indexing terms. *)  | 
|
489  | 
||
490  | 
fun norm_term thy t =  | 
|
491  | 
let  | 
|
| 21588 | 492  | 
val norms = ShuffleData.get thy  | 
493  | 
val ss = Simplifier.theory_context thy empty_ss  | 
|
| 17892 | 494  | 
setmksimps single  | 
| 21588 | 495  | 
addsimps (map (Thm.transfer thy) norms)  | 
| 21078 | 496  | 
addsimprocs [quant_simproc thy, eta_expand_simproc thy,eta_contract_simproc thy]  | 
| 21588 | 497  | 
fun chain f th =  | 
498  | 
let  | 
|
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22846 
diff
changeset
 | 
499  | 
val rhs = Thm.rhs_of th  | 
| 21588 | 500  | 
in  | 
501  | 
transitive th (f rhs)  | 
|
502  | 
end  | 
|
503  | 
val th =  | 
|
| 21078 | 504  | 
t |> disamb_bound thy  | 
| 21588 | 505  | 
|> chain (Simplifier.full_rewrite ss)  | 
| 20326 | 506  | 
|> chain eta_conversion  | 
| 21588 | 507  | 
|> strip_shyps  | 
508  | 
        val _ = message ("norm_term: " ^ (string_of_thm th))
 | 
|
| 14516 | 509  | 
in  | 
| 21588 | 510  | 
th  | 
| 17463 | 511  | 
end  | 
| 21078 | 512  | 
handle e => (writeln "norm_term internal error"; print_sign_exn thy e)  | 
| 14516 | 513  | 
|
514  | 
||
515  | 
(* Closes a theorem with respect to free and schematic variables (does  | 
|
516  | 
not touch type variables, though). *)  | 
|
517  | 
||
518  | 
fun close_thm th =  | 
|
519  | 
let  | 
|
| 22578 | 520  | 
val thy = Thm.theory_of_thm th  | 
| 21588 | 521  | 
val c = prop_of th  | 
| 
29265
 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 
wenzelm 
parents: 
28397 
diff
changeset
 | 
522  | 
val vars = OldTerm.add_term_frees (c, OldTerm.add_term_vars(c,[]))  | 
| 14516 | 523  | 
in  | 
| 21588 | 524  | 
Drule.forall_intr_list (map (cterm_of thy) vars) th  | 
| 14516 | 525  | 
end  | 
| 17959 | 526  | 
handle e => (writeln "close_thm internal error"; OldGoals.print_exn e)  | 
| 14516 | 527  | 
|
528  | 
(* Normalizes a theorem's conclusion using norm_term. *)  | 
|
529  | 
||
530  | 
fun norm_thm thy th =  | 
|
531  | 
let  | 
|
| 21588 | 532  | 
val c = prop_of th  | 
| 14516 | 533  | 
in  | 
| 21588 | 534  | 
equal_elim (norm_term thy c) th  | 
| 14516 | 535  | 
end  | 
536  | 
||
| 21078 | 537  | 
(* make_equal thy t u tries to construct the theorem t == u under the  | 
538  | 
signature thy. If it succeeds, SOME (t == u) is returned, otherwise  | 
|
| 15531 | 539  | 
NONE is returned. *)  | 
| 14516 | 540  | 
|
| 21078 | 541  | 
fun make_equal thy t u =  | 
| 14516 | 542  | 
let  | 
| 21588 | 543  | 
val t_is_t' = norm_term thy t  | 
544  | 
val u_is_u' = norm_term thy u  | 
|
545  | 
val th = transitive t_is_t' (symmetric u_is_u')  | 
|
546  | 
        val _ = message ("make_equal: SOME " ^ (string_of_thm th))
 | 
|
| 14516 | 547  | 
in  | 
| 21588 | 548  | 
SOME th  | 
| 14516 | 549  | 
end  | 
| 15531 | 550  | 
handle e as THM _ => (message "make_equal: NONE";NONE)  | 
| 21588 | 551  | 
|
| 14516 | 552  | 
fun match_consts ignore t (* th *) =  | 
553  | 
let  | 
|
| 21588 | 554  | 
fun add_consts (Const (c, _), cs) =  | 
555  | 
if c mem_string ignore  | 
|
556  | 
then cs  | 
|
557  | 
else insert (op =) c cs  | 
|
558  | 
| add_consts (t $ u, cs) = add_consts (t, add_consts (u, cs))  | 
|
559  | 
| add_consts (Abs (_, _, t), cs) = add_consts (t, cs)  | 
|
560  | 
| add_consts (_, cs) = cs  | 
|
561  | 
val t_consts = add_consts(t,[])  | 
|
| 14516 | 562  | 
in  | 
563  | 
fn (name,th) =>  | 
|
| 21588 | 564  | 
let  | 
565  | 
val th_consts = add_consts(prop_of th,[])  | 
|
566  | 
in  | 
|
567  | 
eq_set(t_consts,th_consts)  | 
|
568  | 
end  | 
|
| 14516 | 569  | 
end  | 
| 21588 | 570  | 
|
| 14516 | 571  | 
val collect_ignored =  | 
| 21078 | 572  | 
fold_rev (fn thm => fn cs =>  | 
| 21588 | 573  | 
let  | 
574  | 
val (lhs,rhs) = Logic.dest_equals (prop_of thm)  | 
|
| 29287 | 575  | 
val ignore_lhs = Term.add_const_names lhs [] \\ Term.add_const_names rhs []  | 
576  | 
val ignore_rhs = Term.add_const_names rhs [] \\ Term.add_const_names lhs []  | 
|
| 21588 | 577  | 
in  | 
578  | 
fold_rev (insert (op =)) cs (ignore_lhs @ ignore_rhs)  | 
|
579  | 
end)  | 
|
| 14516 | 580  | 
|
581  | 
(* set_prop t thms tries to make a theorem with the proposition t from  | 
|
582  | 
one of the theorems thms, by shuffling the propositions around. If it  | 
|
| 15531 | 583  | 
succeeds, SOME theorem is returned, otherwise NONE. *)  | 
| 14516 | 584  | 
|
585  | 
fun set_prop thy t =  | 
|
586  | 
let  | 
|
| 
29265
 
5b4247055bd7
moved old add_term_vars, add_term_frees etc. to structure OldTerm;
 
wenzelm 
parents: 
28397 
diff
changeset
 | 
587  | 
val vars = OldTerm.add_term_frees (t, OldTerm.add_term_vars (t,[]))  | 
| 27330 | 588  | 
val closed_t = fold_rev Logic.all vars t  | 
| 21588 | 589  | 
val rew_th = norm_term thy closed_t  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22846 
diff
changeset
 | 
590  | 
val rhs = Thm.rhs_of rew_th  | 
| 14516 | 591  | 
|
| 21588 | 592  | 
val shuffles = ShuffleData.get thy  | 
593  | 
fun process [] = NONE  | 
|
594  | 
| process ((name,th)::thms) =  | 
|
595  | 
let  | 
|
596  | 
val norm_th = Thm.varifyT (norm_thm thy (close_thm (Thm.transfer thy th)))  | 
|
597  | 
val triv_th = trivial rhs  | 
|
598  | 
                val _ = message ("Shuffler.set_prop: Gluing together " ^ (string_of_thm norm_th) ^ " and " ^ (string_of_thm triv_th))
 | 
|
599  | 
val mod_th = case Seq.pull (bicompose false (*true*) (false,norm_th,0) 1 triv_th) of  | 
|
600  | 
SOME(th,_) => SOME th  | 
|
601  | 
| NONE => NONE  | 
|
602  | 
in  | 
|
603  | 
case mod_th of  | 
|
604  | 
SOME mod_th =>  | 
|
605  | 
let  | 
|
606  | 
val closed_th = equal_elim (symmetric rew_th) mod_th  | 
|
607  | 
in  | 
|
608  | 
                        message ("Shuffler.set_prop succeeded by " ^ name);
 | 
|
609  | 
SOME (name,forall_elim_list (map (cterm_of thy) vars) closed_th)  | 
|
610  | 
end  | 
|
611  | 
| NONE => process thms  | 
|
612  | 
end  | 
|
613  | 
handle e as THM _ => process thms  | 
|
| 14516 | 614  | 
in  | 
| 21588 | 615  | 
fn thms =>  | 
616  | 
case process thms of  | 
|
617  | 
res as SOME (name,th) => if (prop_of th) aconv t  | 
|
618  | 
then res  | 
|
619  | 
else error "Internal error in set_prop"  | 
|
620  | 
| NONE => NONE  | 
|
| 14516 | 621  | 
end  | 
| 17959 | 622  | 
handle e => (writeln "set_prop internal error"; OldGoals.print_exn e)  | 
| 14516 | 623  | 
|
624  | 
fun find_potential thy t =  | 
|
625  | 
let  | 
|
| 21588 | 626  | 
val shuffles = ShuffleData.get thy  | 
627  | 
val ignored = collect_ignored shuffles []  | 
|
| 26662 | 628  | 
val all_thms =  | 
| 
27865
 
27a8ad9612a3
moved basic thm operations from structure PureThy to Thm (cf. more_thm.ML);
 
wenzelm 
parents: 
27330 
diff
changeset
 | 
629  | 
map (`Thm.get_name_hint) (maps #2 (Facts.dest_static [] (PureThy.facts_of thy)))  | 
| 14516 | 630  | 
in  | 
| 26277 | 631  | 
List.filter (match_consts ignored t) all_thms  | 
| 14516 | 632  | 
end  | 
633  | 
||
634  | 
fun gen_shuffle_tac thy search thms i st =  | 
|
635  | 
let  | 
|
| 21588 | 636  | 
        val _ = message ("Shuffling " ^ (string_of_thm st))
 | 
637  | 
val t = List.nth(prems_of st,i-1)  | 
|
638  | 
val set = set_prop thy t  | 
|
639  | 
fun process_tac thms st =  | 
|
640  | 
case set thms of  | 
|
641  | 
SOME (_,th) => Seq.of_list (compose (th,i,st))  | 
|
642  | 
| NONE => Seq.empty  | 
|
| 14516 | 643  | 
in  | 
| 21588 | 644  | 
(process_tac thms APPEND (if search  | 
645  | 
then process_tac (find_potential thy t)  | 
|
646  | 
else no_tac)) st  | 
|
| 14516 | 647  | 
end  | 
648  | 
||
649  | 
fun shuffle_tac thms i st =  | 
|
650  | 
gen_shuffle_tac (the_context()) false thms i st  | 
|
651  | 
||
652  | 
fun search_tac thms i st =  | 
|
653  | 
gen_shuffle_tac (the_context()) true thms i st  | 
|
654  | 
||
655  | 
fun shuffle_meth (thms:thm list) ctxt =  | 
|
656  | 
let  | 
|
| 21588 | 657  | 
val thy = ProofContext.theory_of ctxt  | 
| 14516 | 658  | 
in  | 
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
659  | 
SIMPLE_METHOD' (gen_shuffle_tac thy false (map (pair "") thms))  | 
| 14516 | 660  | 
end  | 
661  | 
||
662  | 
fun search_meth ctxt =  | 
|
663  | 
let  | 
|
| 21588 | 664  | 
val thy = ProofContext.theory_of ctxt  | 
| 
30473
 
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
 
wenzelm 
parents: 
29287 
diff
changeset
 | 
665  | 
val prems = Assumption.all_prems_of ctxt  | 
| 14516 | 666  | 
in  | 
| 
30510
 
4120fc59dd85
unified type Proof.method and pervasive METHOD combinators;
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
667  | 
SIMPLE_METHOD' (gen_shuffle_tac thy true (map (pair "premise") prems))  | 
| 14516 | 668  | 
end  | 
669  | 
||
670  | 
fun add_shuffle_rule thm thy =  | 
|
671  | 
let  | 
|
| 21588 | 672  | 
val shuffles = ShuffleData.get thy  | 
| 14516 | 673  | 
in  | 
| 21588 | 674  | 
if exists (curry Thm.eq_thm thm) shuffles  | 
675  | 
then (warning ((string_of_thm thm) ^ " already known to the shuffler");  | 
|
676  | 
thy)  | 
|
677  | 
else ShuffleData.put (thm::shuffles) thy  | 
|
| 14516 | 678  | 
end  | 
679  | 
||
| 20897 | 680  | 
val shuffle_attr = Thm.declaration_attribute (fn th => Context.mapping (add_shuffle_rule th) I);  | 
| 14516 | 681  | 
|
| 18708 | 682  | 
val setup =  | 
| 22846 | 683  | 
  Method.add_method ("shuffle_tac",
 | 
684  | 
Method.thms_ctxt_args shuffle_meth,"solve goal by shuffling terms around") #>  | 
|
685  | 
  Method.add_method ("search_tac",
 | 
|
686  | 
Method.ctxt_args search_meth,"search for suitable theorems") #>  | 
|
| 18708 | 687  | 
add_shuffle_rule weaken #>  | 
688  | 
add_shuffle_rule equiv_comm #>  | 
|
689  | 
add_shuffle_rule imp_comm #>  | 
|
690  | 
add_shuffle_rule Drule.norm_hhf_eq #>  | 
|
691  | 
add_shuffle_rule Drule.triv_forall_equality #>  | 
|
| 30528 | 692  | 
  Attrib.setup @{binding shuffle_rule} (Scan.succeed shuffle_attr) "declare rule for shuffler";
 | 
| 18708 | 693  | 
|
| 14516 | 694  | 
end  |