author | wenzelm |
Sat, 11 Feb 2006 17:17:47 +0100 | |
changeset 19012 | 2577ac76cdc6 |
parent 17456 | bcf7544875b2 |
permissions | -rw-r--r-- |
17456 | 1 |
(* Title: CCL/Coinduction.ML |
0 | 2 |
ID: $Id$ |
1459 | 3 |
Author: Martin Coen, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1993 University of Cambridge |
5 |
||
6 |
Lemmas and tactics for using the rule coinduct3 on [= and =. |
|
7 |
*) |
|
8 |
||
17456 | 9 |
val [mono,prem] = goal (the_context ()) "[| mono(f); a : f(lfp(f)) |] ==> a : lfp(f)"; |
2035 | 10 |
by (stac (mono RS lfp_Tarski) 1); |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
11 |
by (rtac prem 1); |
757 | 12 |
qed "lfpI"; |
0 | 13 |
|
17456 | 14 |
val prems = goal (the_context ()) "[| a=a'; a' : A |] ==> a : A"; |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
15 |
by (simp_tac (term_ss addsimps prems) 1); |
757 | 16 |
qed "ssubst_single"; |
0 | 17 |
|
17456 | 18 |
val prems = goal (the_context ()) "[| a=a'; b=b'; <a',b'> : A |] ==> <a,b> : A"; |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
19 |
by (simp_tac (term_ss addsimps prems) 1); |
757 | 20 |
qed "ssubst_pair"; |
0 | 21 |
|
22 |
(***) |
|
23 |
||
17456 | 24 |
local |
25 |
fun mk_thm s = prove_goal (the_context ()) s (fn mono::prems => |
|
0 | 26 |
[fast_tac (term_cs addIs ((mono RS coinduct3_mono_lemma RS lfpI)::prems)) 1]); |
27 |
in |
|
28 |
val ci3_RI = mk_thm "[| mono(Agen); a : R |] ==> a : lfp(%x. Agen(x) Un R Un A)"; |
|
29 |
val ci3_AgenI = mk_thm "[| mono(Agen); a : Agen(lfp(%x. Agen(x) Un R Un A)) |] ==> \ |
|
30 |
\ a : lfp(%x. Agen(x) Un R Un A)"; |
|
31 |
val ci3_AI = mk_thm "[| mono(Agen); a : A |] ==> a : lfp(%x. Agen(x) Un R Un A)"; |
|
32 |
end; |
|
33 |
||
17456 | 34 |
fun mk_genIs thy defs genXH gen_mono s = prove_goalw thy defs s |
0 | 35 |
(fn prems => [rtac (genXH RS iffD2) 1, |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
36 |
(simp_tac term_ss 1), |
17456 | 37 |
TRY (fast_tac (term_cs addIs |
0 | 38 |
([genXH RS iffD2,gen_mono RS coinduct3_mono_lemma RS lfpI] |
39 |
@ prems)) 1)]); |
|
40 |
||
41 |
(** POgen **) |
|
42 |
||
17456 | 43 |
goal (the_context ()) "<a,a> : PO"; |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
44 |
by (rtac (po_refl RS (XH_to_D PO_iff)) 1); |
757 | 45 |
qed "PO_refl"; |
0 | 46 |
|
17456 | 47 |
val POgenIs = map (mk_genIs (the_context ()) data_defs POgenXH POgen_mono) |
0 | 48 |
["<true,true> : POgen(R)", |
49 |
"<false,false> : POgen(R)", |
|
50 |
"[| <a,a'> : R; <b,b'> : R |] ==> <<a,b>,<a',b'>> : POgen(R)", |
|
3837 | 51 |
"[|!!x. <b(x),b'(x)> : R |] ==><lam x. b(x),lam x. b'(x)> : POgen(R)", |
0 | 52 |
"<one,one> : POgen(R)", |
53 |
"<a,a'> : lfp(%x. POgen(x) Un R Un PO) ==> \ |
|
54 |
\ <inl(a),inl(a')> : POgen(lfp(%x. POgen(x) Un R Un PO))", |
|
55 |
"<b,b'> : lfp(%x. POgen(x) Un R Un PO) ==> \ |
|
56 |
\ <inr(b),inr(b')> : POgen(lfp(%x. POgen(x) Un R Un PO))", |
|
57 |
"<zero,zero> : POgen(lfp(%x. POgen(x) Un R Un PO))", |
|
58 |
"<n,n'> : lfp(%x. POgen(x) Un R Un PO) ==> \ |
|
59 |
\ <succ(n),succ(n')> : POgen(lfp(%x. POgen(x) Un R Un PO))", |
|
60 |
"<[],[]> : POgen(lfp(%x. POgen(x) Un R Un PO))", |
|
61 |
"[| <h,h'> : lfp(%x. POgen(x) Un R Un PO); \ |
|
62 |
\ <t,t'> : lfp(%x. POgen(x) Un R Un PO) |] ==> \ |
|
289 | 63 |
\ <h$t,h'$t'> : POgen(lfp(%x. POgen(x) Un R Un PO))"]; |
0 | 64 |
|
65 |
fun POgen_tac (rla,rlb) i = |
|
66 |
SELECT_GOAL (safe_tac ccl_cs) i THEN |
|
67 |
rtac (rlb RS (rla RS ssubst_pair)) i THEN |
|
17456 | 68 |
(REPEAT (resolve_tac (POgenIs @ [PO_refl RS (POgen_mono RS ci3_AI)] @ |
0 | 69 |
(POgenIs RL [POgen_mono RS ci3_AgenI]) @ [POgen_mono RS ci3_RI]) i)); |
70 |
||
71 |
(** EQgen **) |
|
72 |
||
17456 | 73 |
goal (the_context ()) "<a,a> : EQ"; |
642
0db578095e6a
CCL/Gfp/coinduct2, coinduct3: modified proofs to suppress deep unification
lcp
parents:
289
diff
changeset
|
74 |
by (rtac (refl RS (EQ_iff RS iffD1)) 1); |
757 | 75 |
qed "EQ_refl"; |
0 | 76 |
|
17456 | 77 |
val EQgenIs = map (mk_genIs (the_context ()) data_defs EQgenXH EQgen_mono) |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
78 |
["<true,true> : EQgen(R)", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
79 |
"<false,false> : EQgen(R)", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
80 |
"[| <a,a'> : R; <b,b'> : R |] ==> <<a,b>,<a',b'>> : EQgen(R)", |
3837 | 81 |
"[|!!x. <b(x),b'(x)> : R |] ==> <lam x. b(x),lam x. b'(x)> : EQgen(R)", |
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
82 |
"<one,one> : EQgen(R)", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
83 |
"<a,a'> : lfp(%x. EQgen(x) Un R Un EQ) ==> \ |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
84 |
\ <inl(a),inl(a')> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
85 |
"<b,b'> : lfp(%x. EQgen(x) Un R Un EQ) ==> \ |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
86 |
\ <inr(b),inr(b')> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
87 |
"<zero,zero> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
88 |
"<n,n'> : lfp(%x. EQgen(x) Un R Un EQ) ==> \ |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
89 |
\ <succ(n),succ(n')> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
90 |
"<[],[]> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))", |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
91 |
"[| <h,h'> : lfp(%x. EQgen(x) Un R Un EQ); \ |
0 | 92 |
\ <t,t'> : lfp(%x. EQgen(x) Un R Un EQ) |] ==> \ |
289 | 93 |
\ <h$t,h'$t'> : EQgen(lfp(%x. EQgen(x) Un R Un EQ))"]; |
0 | 94 |
|
95 |
fun EQgen_raw_tac i = |
|
17456 | 96 |
(REPEAT (resolve_tac (EQgenIs @ [EQ_refl RS (EQgen_mono RS ci3_AI)] @ |
0 | 97 |
(EQgenIs RL [EQgen_mono RS ci3_AgenI]) @ [EQgen_mono RS ci3_RI]) i)); |
98 |
||
99 |
(* Goals of the form R <= EQgen(R) - rewrite elements <a,b> : EQgen(R) using rews and *) |
|
100 |
(* then reduce this to a goal <a',b'> : R (hopefully?) *) |
|
101 |
(* rews are rewrite rules that would cause looping in the simpifier *) |
|
102 |
||
17456 | 103 |
fun EQgen_tac simp_set rews i = |
104 |
SELECT_GOAL |
|
8
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
105 |
(TRY (safe_tac ccl_cs) THEN |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
106 |
resolve_tac ((rews@[refl]) RL ((rews@[refl]) RL [ssubst_pair])) i THEN |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
107 |
ALLGOALS (simp_tac simp_set) THEN |
c3d2c6dcf3f0
Installation of new simplfier. Previously appeared to set up the old
lcp
parents:
0
diff
changeset
|
108 |
ALLGOALS EQgen_raw_tac) i; |