| author | wenzelm | 
| Thu, 30 May 2013 16:31:53 +0200 | |
| changeset 52240 | 066c2ff17f7c | 
| parent 51717 | 9e7d1c139569 | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 17456 | 1  | 
(* Title: CCL/CCL.thy  | 
| 1474 | 2  | 
Author: Martin Coen  | 
| 0 | 3  | 
Copyright 1993 University of Cambridge  | 
4  | 
*)  | 
|
5  | 
||
| 17456 | 6  | 
header {* Classical Computational Logic for Untyped Lambda Calculus
 | 
7  | 
with reduction to weak head-normal form *}  | 
|
| 0 | 8  | 
|
| 17456 | 9  | 
theory CCL  | 
10  | 
imports Gfp  | 
|
11  | 
begin  | 
|
| 0 | 12  | 
|
| 17456 | 13  | 
text {*
 | 
14  | 
Based on FOL extended with set collection, a primitive higher-order  | 
|
15  | 
logic. HOL is too strong - descriptions prevent a type of programs  | 
|
16  | 
being defined which contains only executable terms.  | 
|
17  | 
*}  | 
|
| 0 | 18  | 
|
| 17456 | 19  | 
classes prog < "term"  | 
| 36452 | 20  | 
default_sort prog  | 
| 17456 | 21  | 
|
| 24825 | 22  | 
arities "fun" :: (prog, prog) prog  | 
| 17456 | 23  | 
|
24  | 
typedecl i  | 
|
25  | 
arities i :: prog  | 
|
26  | 
||
| 0 | 27  | 
|
28  | 
consts  | 
|
29  | 
(*** Evaluation Judgement ***)  | 
|
| 24825 | 30  | 
Eval :: "[i,i]=>prop" (infixl "--->" 20)  | 
| 0 | 31  | 
|
32  | 
(*** Bisimulations for pre-order and equality ***)  | 
|
| 24825 | 33  | 
po :: "['a,'a]=>o" (infixl "[=" 50)  | 
| 0 | 34  | 
SIM :: "[i,i,i set]=>o"  | 
| 17456 | 35  | 
POgen :: "i set => i set"  | 
36  | 
EQgen :: "i set => i set"  | 
|
37  | 
PO :: "i set"  | 
|
38  | 
EQ :: "i set"  | 
|
| 0 | 39  | 
|
40  | 
(*** Term Formers ***)  | 
|
| 17456 | 41  | 
true :: "i"  | 
42  | 
false :: "i"  | 
|
| 0 | 43  | 
  pair        ::       "[i,i]=>i"             ("(1<_,/_>)")
 | 
44  | 
lambda :: "(i=>i)=>i" (binder "lam " 55)  | 
|
| 17456 | 45  | 
"case" :: "[i,i,i,[i,i]=>i,(i=>i)=>i]=>i"  | 
| 24825 | 46  | 
"apply" :: "[i,i]=>i" (infixl "`" 56)  | 
| 0 | 47  | 
bot :: "i"  | 
| 17456 | 48  | 
"fix" :: "(i=>i)=>i"  | 
| 0 | 49  | 
|
50  | 
(*** Defined Predicates ***)  | 
|
| 17456 | 51  | 
Trm :: "i => o"  | 
52  | 
Dvg :: "i => o"  | 
|
| 0 | 53  | 
|
54  | 
(******* EVALUATION SEMANTICS *******)  | 
|
55  | 
||
56  | 
(** This is the evaluation semantics from which the axioms below were derived. **)  | 
|
57  | 
(** It is included here just as an evaluator for FUN and has no influence on **)  | 
|
58  | 
(** inference in the theory CCL. **)  | 
|
59  | 
||
| 51307 | 60  | 
axiomatization where  | 
61  | 
trueV: "true ---> true" and  | 
|
62  | 
falseV: "false ---> false" and  | 
|
63  | 
pairV: "<a,b> ---> <a,b>" and  | 
|
64  | 
lamV: "\<And>b. lam x. b(x) ---> lam x. b(x)" and  | 
|
65  | 
||
66  | 
caseVtrue: "[| t ---> true; d ---> c |] ==> case(t,d,e,f,g) ---> c" and  | 
|
67  | 
caseVfalse: "[| t ---> false; e ---> c |] ==> case(t,d,e,f,g) ---> c" and  | 
|
68  | 
caseVpair: "[| t ---> <a,b>; f(a,b) ---> c |] ==> case(t,d,e,f,g) ---> c" and  | 
|
69  | 
caseVlam: "\<And>b. [| t ---> lam x. b(x); g(b) ---> c |] ==> case(t,d,e,f,g) ---> c"  | 
|
| 0 | 70  | 
|
71  | 
(*** Properties of evaluation: note that "t ---> c" impies that c is canonical ***)  | 
|
72  | 
||
| 51307 | 73  | 
axiomatization where  | 
| 17456 | 74  | 
canonical: "[| t ---> c; c==true ==> u--->v;  | 
75  | 
c==false ==> u--->v;  | 
|
76  | 
!!a b. c==<a,b> ==> u--->v;  | 
|
77  | 
!!f. c==lam x. f(x) ==> u--->v |] ==>  | 
|
| 1149 | 78  | 
u--->v"  | 
| 0 | 79  | 
|
80  | 
(* Should be derivable - but probably a bitch! *)  | 
|
| 51307 | 81  | 
axiomatization where  | 
| 17456 | 82  | 
substitute: "[| a==a'; t(a)--->c(a) |] ==> t(a')--->c(a')"  | 
| 0 | 83  | 
|
84  | 
(************** LOGIC ***************)  | 
|
85  | 
||
86  | 
(*** Definitions used in the following rules ***)  | 
|
87  | 
||
| 51307 | 88  | 
axiomatization where  | 
89  | 
bot_def: "bot == (lam x. x`x)`(lam x. x`x)" and  | 
|
| 17456 | 90  | 
apply_def: "f ` t == case(f,bot,bot,%x y. bot,%u. u(t))"  | 
| 42156 | 91  | 
|
92  | 
defs  | 
|
| 17456 | 93  | 
fix_def: "fix(f) == (lam x. f(x`x))`(lam x. f(x`x))"  | 
| 0 | 94  | 
|
95  | 
(* The pre-order ([=) is defined as a simulation, and behavioural equivalence (=) *)  | 
|
96  | 
(* as a bisimulation. They can both be expressed as (bi)simulations up to *)  | 
|
97  | 
(* behavioural equivalence (ie the relations PO and EQ defined below). *)  | 
|
98  | 
||
| 17456 | 99  | 
SIM_def:  | 
100  | 
"SIM(t,t',R) == (t=true & t'=true) | (t=false & t'=false) |  | 
|
101  | 
(EX a a' b b'. t=<a,b> & t'=<a',b'> & <a,a'> : R & <b,b'> : R) |  | 
|
| 3837 | 102  | 
(EX f f'. t=lam x. f(x) & t'=lam x. f'(x) & (ALL x.<f(x),f'(x)> : R))"  | 
| 0 | 103  | 
|
| 17456 | 104  | 
  POgen_def:  "POgen(R) == {p. EX t t'. p=<t,t'> & (t = bot | SIM(t,t',R))}"
 | 
105  | 
  EQgen_def:  "EQgen(R) == {p. EX t t'. p=<t,t'> & (t = bot & t' = bot | SIM(t,t',R))}"
 | 
|
| 0 | 106  | 
|
| 17456 | 107  | 
PO_def: "PO == gfp(POgen)"  | 
108  | 
EQ_def: "EQ == gfp(EQgen)"  | 
|
| 0 | 109  | 
|
110  | 
(*** Rules ***)  | 
|
111  | 
||
112  | 
(** Partial Order **)  | 
|
113  | 
||
| 51307 | 114  | 
axiomatization where  | 
115  | 
po_refl: "a [= a" and  | 
|
116  | 
po_trans: "[| a [= b; b [= c |] ==> a [= c" and  | 
|
117  | 
po_cong: "a [= b ==> f(a) [= f(b)" and  | 
|
| 0 | 118  | 
|
119  | 
(* Extend definition of [= to program fragments of higher type *)  | 
|
| 17456 | 120  | 
po_abstractn: "(!!x. f(x) [= g(x)) ==> (%x. f(x)) [= (%x. g(x))"  | 
| 0 | 121  | 
|
122  | 
(** Equality - equivalence axioms inherited from FOL.thy **)  | 
|
123  | 
(** - congruence of "=" is axiomatised implicitly **)  | 
|
124  | 
||
| 51307 | 125  | 
axiomatization where  | 
| 17456 | 126  | 
eq_iff: "t = t' <-> t [= t' & t' [= t"  | 
| 0 | 127  | 
|
128  | 
(** Properties of canonical values given by greatest fixed point definitions **)  | 
|
| 17456 | 129  | 
|
| 51307 | 130  | 
axiomatization where  | 
131  | 
PO_iff: "t [= t' <-> <t,t'> : PO" and  | 
|
| 17456 | 132  | 
EQ_iff: "t = t' <-> <t,t'> : EQ"  | 
| 0 | 133  | 
|
134  | 
(** Behaviour of non-canonical terms (ie case) given by the following beta-rules **)  | 
|
135  | 
||
| 51307 | 136  | 
axiomatization where  | 
137  | 
caseBtrue: "case(true,d,e,f,g) = d" and  | 
|
138  | 
caseBfalse: "case(false,d,e,f,g) = e" and  | 
|
139  | 
caseBpair: "case(<a,b>,d,e,f,g) = f(a,b)" and  | 
|
140  | 
caseBlam: "\<And>b. case(lam x. b(x),d,e,f,g) = g(b)" and  | 
|
141  | 
caseBbot: "case(bot,d,e,f,g) = bot" (* strictness *)  | 
|
| 0 | 142  | 
|
143  | 
(** The theory is non-trivial **)  | 
|
| 51307 | 144  | 
axiomatization where  | 
| 17456 | 145  | 
distinctness: "~ lam x. b(x) = bot"  | 
| 0 | 146  | 
|
147  | 
(*** Definitions of Termination and Divergence ***)  | 
|
148  | 
||
| 42156 | 149  | 
defs  | 
| 17456 | 150  | 
Dvg_def: "Dvg(t) == t = bot"  | 
151  | 
Trm_def: "Trm(t) == ~ Dvg(t)"  | 
|
| 0 | 152  | 
|
| 17456 | 153  | 
text {*
 | 
| 0 | 154  | 
Would be interesting to build a similar theory for a typed programming language:  | 
155  | 
    ie.     true :: bool,      fix :: ('a=>'a)=>'a  etc......
 | 
|
156  | 
||
157  | 
This is starting to look like LCF.  | 
|
| 17456 | 158  | 
What are the advantages of this approach?  | 
159  | 
- less axiomatic  | 
|
| 0 | 160  | 
- wfd induction / coinduction and fixed point induction available  | 
| 17456 | 161  | 
*}  | 
162  | 
||
| 20140 | 163  | 
|
164  | 
lemmas ccl_data_defs = apply_def fix_def  | 
|
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
165  | 
|
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
166  | 
declare po_refl [simp]  | 
| 20140 | 167  | 
|
168  | 
||
169  | 
subsection {* Congruence Rules *}
 | 
|
170  | 
||
171  | 
(*similar to AP_THM in Gordon's HOL*)  | 
|
172  | 
lemma fun_cong: "(f::'a=>'b) = g ==> f(x)=g(x)"  | 
|
173  | 
by simp  | 
|
174  | 
||
175  | 
(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*)  | 
|
176  | 
lemma arg_cong: "x=y ==> f(x)=f(y)"  | 
|
177  | 
by simp  | 
|
178  | 
||
179  | 
lemma abstractn: "(!!x. f(x) = g(x)) ==> f = g"  | 
|
180  | 
apply (simp add: eq_iff)  | 
|
181  | 
apply (blast intro: po_abstractn)  | 
|
182  | 
done  | 
|
183  | 
||
184  | 
lemmas caseBs = caseBtrue caseBfalse caseBpair caseBlam caseBbot  | 
|
185  | 
||
186  | 
||
187  | 
subsection {* Termination and Divergence *}
 | 
|
188  | 
||
189  | 
lemma Trm_iff: "Trm(t) <-> ~ t = bot"  | 
|
190  | 
by (simp add: Trm_def Dvg_def)  | 
|
191  | 
||
192  | 
lemma Dvg_iff: "Dvg(t) <-> t = bot"  | 
|
193  | 
by (simp add: Trm_def Dvg_def)  | 
|
194  | 
||
195  | 
||
196  | 
subsection {* Constructors are injective *}
 | 
|
197  | 
||
198  | 
lemma eq_lemma: "[| x=a; y=b; x=y |] ==> a=b"  | 
|
199  | 
by simp  | 
|
200  | 
||
201  | 
ML {*
 | 
|
| 32154 | 202  | 
fun inj_rl_tac ctxt rews i =  | 
| 24825 | 203  | 
let  | 
204  | 
      fun mk_inj_lemmas r = [@{thm arg_cong}] RL [r RS (r RS @{thm eq_lemma})]
 | 
|
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
205  | 
val inj_lemmas = maps mk_inj_lemmas rews  | 
| 32154 | 206  | 
in  | 
| 35409 | 207  | 
      CHANGED (REPEAT (ares_tac [@{thm iffI}, @{thm allI}, @{thm conjI}] i ORELSE
 | 
| 32154 | 208  | 
eresolve_tac inj_lemmas i ORELSE  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51670 
diff
changeset
 | 
209  | 
asm_simp_tac (ctxt addsimps rews) i))  | 
| 32154 | 210  | 
end;  | 
| 20140 | 211  | 
*}  | 
212  | 
||
| 32154 | 213  | 
method_setup inj_rl = {*
 | 
214  | 
Attrib.thms >> (fn rews => fn ctxt => SIMPLE_METHOD' (inj_rl_tac ctxt rews))  | 
|
| 42814 | 215  | 
*}  | 
| 32154 | 216  | 
|
217  | 
lemma ccl_injs:  | 
|
218  | 
"<a,b> = <a',b'> <-> (a=a' & b=b')"  | 
|
219  | 
"!!b b'. (lam x. b(x) = lam x. b'(x)) <-> ((ALL z. b(z)=b'(z)))"  | 
|
220  | 
by (inj_rl caseBs)  | 
|
| 20140 | 221  | 
|
222  | 
||
223  | 
lemma pair_inject: "<a,b> = <a',b'> \<Longrightarrow> (a = a' \<Longrightarrow> b = b' \<Longrightarrow> R) \<Longrightarrow> R"  | 
|
224  | 
by (simp add: ccl_injs)  | 
|
225  | 
||
226  | 
||
227  | 
subsection {* Constructors are distinct *}
 | 
|
228  | 
||
229  | 
lemma lem: "t=t' ==> case(t,b,c,d,e) = case(t',b,c,d,e)"  | 
|
230  | 
by simp  | 
|
231  | 
||
232  | 
ML {*
 | 
|
233  | 
||
234  | 
local  | 
|
235  | 
fun pairs_of f x [] = []  | 
|
236  | 
| pairs_of f x (y::ys) = (f x y) :: (f y x) :: (pairs_of f x ys)  | 
|
237  | 
||
238  | 
fun mk_combs ff [] = []  | 
|
239  | 
| mk_combs ff (x::xs) = (pairs_of ff x xs) @ mk_combs ff xs  | 
|
240  | 
||
241  | 
(* Doesn't handle binder types correctly *)  | 
|
242  | 
fun saturate thy sy name =  | 
|
243  | 
let fun arg_str 0 a s = s  | 
|
244  | 
         | arg_str 1 a s = "(" ^ a ^ "a" ^ s ^ ")"
 | 
|
245  | 
         | arg_str n a s = arg_str (n-1) a ("," ^ a ^ (chr((ord "a")+n-1)) ^ s)
 | 
|
246  | 
val T = Sign.the_const_type thy (Sign.intern_const thy sy);  | 
|
| 40844 | 247  | 
val arity = length (binder_types T)  | 
| 20140 | 248  | 
in sy ^ (arg_str arity name "") end  | 
249  | 
||
250  | 
fun mk_thm_str thy a b = "~ " ^ (saturate thy a "a") ^ " = " ^ (saturate thy b "b")  | 
|
251  | 
||
| 39159 | 252  | 
  val lemma = @{thm lem};
 | 
253  | 
  val eq_lemma = @{thm eq_lemma};
 | 
|
254  | 
  val distinctness = @{thm distinctness};
 | 
|
| 42480 | 255  | 
fun mk_lemma (ra,rb) =  | 
256  | 
[lemma] RL [ra RS (rb RS eq_lemma)] RL  | 
|
257  | 
    [distinctness RS @{thm notE}, @{thm sym} RS (distinctness RS @{thm notE})]
 | 
|
| 20140 | 258  | 
in  | 
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
259  | 
fun mk_lemmas rls = maps mk_lemma (mk_combs pair rls)  | 
| 20140 | 260  | 
fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs  | 
261  | 
end  | 
|
262  | 
||
263  | 
*}  | 
|
264  | 
||
265  | 
ML {*
 | 
|
266  | 
||
| 32010 | 267  | 
val caseB_lemmas = mk_lemmas @{thms caseBs}
 | 
| 20140 | 268  | 
|
269  | 
val ccl_dstncts =  | 
|
| 32175 | 270  | 
let  | 
271  | 
fun mk_raw_dstnct_thm rls s =  | 
|
272  | 
      Goal.prove_global @{theory} [] [] (Syntax.read_prop_global @{theory} s)
 | 
|
| 35409 | 273  | 
        (fn _=> rtac @{thm notI} 1 THEN eresolve_tac rls 1)
 | 
| 32175 | 274  | 
in map (mk_raw_dstnct_thm caseB_lemmas)  | 
275  | 
    (mk_dstnct_rls @{theory} ["bot","true","false","pair","lambda"]) end
 | 
|
| 20140 | 276  | 
|
| 51670 | 277  | 
fun mk_dstnct_thms ctxt defs inj_rls xs =  | 
| 32175 | 278  | 
let  | 
| 51670 | 279  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 32175 | 280  | 
fun mk_dstnct_thm rls s =  | 
| 51670 | 281  | 
Goal.prove_global thy [] [] (Syntax.read_prop ctxt s)  | 
| 32175 | 282  | 
(fn _ =>  | 
283  | 
rewrite_goals_tac defs THEN  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51670 
diff
changeset
 | 
284  | 
simp_tac (ctxt addsimps (rls @ inj_rls)) 1)  | 
| 
32149
 
ef59550a55d3
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32010 
diff
changeset
 | 
285  | 
in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end  | 
| 20140 | 286  | 
|
| 51670 | 287  | 
fun mkall_dstnct_thms ctxt defs i_rls xss = maps (mk_dstnct_thms ctxt defs i_rls) xss  | 
| 20140 | 288  | 
|
289  | 
(*** Rewriting and Proving ***)  | 
|
290  | 
||
| 42480 | 291  | 
fun XH_to_I rl = rl RS @{thm iffD2}
 | 
292  | 
fun XH_to_D rl = rl RS @{thm iffD1}
 | 
|
| 20140 | 293  | 
val XH_to_E = make_elim o XH_to_D  | 
294  | 
val XH_to_Is = map XH_to_I  | 
|
295  | 
val XH_to_Ds = map XH_to_D  | 
|
296  | 
val XH_to_Es = map XH_to_E;  | 
|
297  | 
||
| 32154 | 298  | 
bind_thms ("ccl_rews", @{thms caseBs} @ @{thms ccl_injs} @ ccl_dstncts);
 | 
| 42480 | 299  | 
bind_thms ("ccl_dstnctsEs", ccl_dstncts RL [@{thm notE}]);
 | 
| 32010 | 300  | 
bind_thms ("ccl_injDs", XH_to_Ds @{thms ccl_injs});
 | 
| 20140 | 301  | 
*}  | 
302  | 
||
303  | 
lemmas [simp] = ccl_rews  | 
|
304  | 
and [elim!] = pair_inject ccl_dstnctsEs  | 
|
305  | 
and [dest!] = ccl_injDs  | 
|
306  | 
||
307  | 
||
308  | 
subsection {* Facts from gfp Definition of @{text "[="} and @{text "="} *}
 | 
|
309  | 
||
310  | 
lemma XHlemma1: "[| A=B; a:B <-> P |] ==> a:A <-> P"  | 
|
311  | 
by simp  | 
|
312  | 
||
313  | 
lemma XHlemma2: "(P(t,t') <-> Q) ==> (<t,t'> : {p. EX t t'. p=<t,t'> &  P(t,t')} <-> Q)"
 | 
|
314  | 
by blast  | 
|
315  | 
||
316  | 
||
317  | 
subsection {* Pre-Order *}
 | 
|
318  | 
||
319  | 
lemma POgen_mono: "mono(%X. POgen(X))"  | 
|
320  | 
apply (unfold POgen_def SIM_def)  | 
|
321  | 
apply (rule monoI)  | 
|
322  | 
apply blast  | 
|
323  | 
done  | 
|
324  | 
||
325  | 
lemma POgenXH:  | 
|
326  | 
"<t,t'> : POgen(R) <-> t= bot | (t=true & t'=true) | (t=false & t'=false) |  | 
|
327  | 
(EX a a' b b'. t=<a,b> & t'=<a',b'> & <a,a'> : R & <b,b'> : R) |  | 
|
328  | 
(EX f f'. t=lam x. f(x) & t'=lam x. f'(x) & (ALL x. <f(x),f'(x)> : R))"  | 
|
329  | 
apply (unfold POgen_def SIM_def)  | 
|
330  | 
apply (rule iff_refl [THEN XHlemma2])  | 
|
331  | 
done  | 
|
332  | 
||
333  | 
lemma poXH:  | 
|
334  | 
"t [= t' <-> t=bot | (t=true & t'=true) | (t=false & t'=false) |  | 
|
335  | 
(EX a a' b b'. t=<a,b> & t'=<a',b'> & a [= a' & b [= b') |  | 
|
336  | 
(EX f f'. t=lam x. f(x) & t'=lam x. f'(x) & (ALL x. f(x) [= f'(x)))"  | 
|
337  | 
apply (simp add: PO_iff del: ex_simps)  | 
|
338  | 
apply (rule POgen_mono  | 
|
339  | 
[THEN PO_def [THEN def_gfp_Tarski], THEN XHlemma1, unfolded POgen_def SIM_def])  | 
|
340  | 
apply (rule iff_refl [THEN XHlemma2])  | 
|
341  | 
done  | 
|
342  | 
||
343  | 
lemma po_bot: "bot [= b"  | 
|
344  | 
apply (rule poXH [THEN iffD2])  | 
|
345  | 
apply simp  | 
|
346  | 
done  | 
|
347  | 
||
348  | 
lemma bot_poleast: "a [= bot ==> a=bot"  | 
|
349  | 
apply (drule poXH [THEN iffD1])  | 
|
350  | 
apply simp  | 
|
351  | 
done  | 
|
352  | 
||
353  | 
lemma po_pair: "<a,b> [= <a',b'> <-> a [= a' & b [= b'"  | 
|
354  | 
apply (rule poXH [THEN iff_trans])  | 
|
355  | 
apply simp  | 
|
356  | 
done  | 
|
357  | 
||
358  | 
lemma po_lam: "lam x. f(x) [= lam x. f'(x) <-> (ALL x. f(x) [= f'(x))"  | 
|
359  | 
apply (rule poXH [THEN iff_trans])  | 
|
| 47966 | 360  | 
apply fastforce  | 
| 20140 | 361  | 
done  | 
362  | 
||
363  | 
lemmas ccl_porews = po_bot po_pair po_lam  | 
|
364  | 
||
365  | 
lemma case_pocong:  | 
|
366  | 
assumes 1: "t [= t'"  | 
|
367  | 
and 2: "a [= a'"  | 
|
368  | 
and 3: "b [= b'"  | 
|
369  | 
and 4: "!!x y. c(x,y) [= c'(x,y)"  | 
|
370  | 
and 5: "!!u. d(u) [= d'(u)"  | 
|
371  | 
shows "case(t,a,b,c,d) [= case(t',a',b',c',d')"  | 
|
372  | 
apply (rule 1 [THEN po_cong, THEN po_trans])  | 
|
373  | 
apply (rule 2 [THEN po_cong, THEN po_trans])  | 
|
374  | 
apply (rule 3 [THEN po_cong, THEN po_trans])  | 
|
375  | 
apply (rule 4 [THEN po_abstractn, THEN po_abstractn, THEN po_cong, THEN po_trans])  | 
|
376  | 
apply (rule_tac f1 = "%d. case (t',a',b',c',d)" in  | 
|
377  | 
5 [THEN po_abstractn, THEN po_cong, THEN po_trans])  | 
|
378  | 
apply (rule po_refl)  | 
|
379  | 
done  | 
|
380  | 
||
381  | 
lemma apply_pocong: "[| f [= f'; a [= a' |] ==> f ` a [= f' ` a'"  | 
|
382  | 
unfolding ccl_data_defs  | 
|
383  | 
apply (rule case_pocong, (rule po_refl | assumption)+)  | 
|
384  | 
apply (erule po_cong)  | 
|
385  | 
done  | 
|
386  | 
||
387  | 
lemma npo_lam_bot: "~ lam x. b(x) [= bot"  | 
|
388  | 
apply (rule notI)  | 
|
389  | 
apply (drule bot_poleast)  | 
|
390  | 
apply (erule distinctness [THEN notE])  | 
|
391  | 
done  | 
|
392  | 
||
393  | 
lemma po_lemma: "[| x=a; y=b; x[=y |] ==> a[=b"  | 
|
394  | 
by simp  | 
|
395  | 
||
396  | 
lemma npo_pair_lam: "~ <a,b> [= lam x. f(x)"  | 
|
397  | 
apply (rule notI)  | 
|
398  | 
apply (rule npo_lam_bot [THEN notE])  | 
|
399  | 
apply (erule case_pocong [THEN caseBlam [THEN caseBpair [THEN po_lemma]]])  | 
|
400  | 
apply (rule po_refl npo_lam_bot)+  | 
|
401  | 
done  | 
|
402  | 
||
403  | 
lemma npo_lam_pair: "~ lam x. f(x) [= <a,b>"  | 
|
404  | 
apply (rule notI)  | 
|
405  | 
apply (rule npo_lam_bot [THEN notE])  | 
|
406  | 
apply (erule case_pocong [THEN caseBpair [THEN caseBlam [THEN po_lemma]]])  | 
|
407  | 
apply (rule po_refl npo_lam_bot)+  | 
|
408  | 
done  | 
|
409  | 
||
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
410  | 
lemma npo_rls1:  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
411  | 
"~ true [= false"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
412  | 
"~ false [= true"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
413  | 
"~ true [= <a,b>"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
414  | 
"~ <a,b> [= true"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
415  | 
"~ true [= lam x. f(x)"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
416  | 
"~ lam x. f(x) [= true"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
417  | 
"~ false [= <a,b>"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
418  | 
"~ <a,b> [= false"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
419  | 
"~ false [= lam x. f(x)"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
420  | 
"~ lam x. f(x) [= false"  | 
| 
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
421  | 
  by (tactic {*
 | 
| 35409 | 422  | 
    REPEAT (rtac @{thm notI} 1 THEN
 | 
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
423  | 
      dtac @{thm case_pocong} 1 THEN
 | 
| 35409 | 424  | 
      etac @{thm rev_mp} 5 THEN
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51670 
diff
changeset
 | 
425  | 
      ALLGOALS (simp_tac @{context}) THEN
 | 
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
426  | 
      REPEAT (resolve_tac [@{thm po_refl}, @{thm npo_lam_bot}] 1)) *})
 | 
| 20140 | 427  | 
|
| 
32153
 
a0e57fb1b930
misc modernization: proper method setup instead of adhoc ML proofs;
 
wenzelm 
parents: 
32149 
diff
changeset
 | 
428  | 
lemmas npo_rls = npo_pair_lam npo_lam_pair npo_rls1  | 
| 20140 | 429  | 
|
430  | 
||
431  | 
subsection {* Coinduction for @{text "[="} *}
 | 
|
432  | 
||
433  | 
lemma po_coinduct: "[| <t,u> : R; R <= POgen(R) |] ==> t [= u"  | 
|
434  | 
apply (rule PO_def [THEN def_coinduct, THEN PO_iff [THEN iffD2]])  | 
|
435  | 
apply assumption+  | 
|
436  | 
done  | 
|
437  | 
||
438  | 
||
439  | 
subsection {* Equality *}
 | 
|
440  | 
||
441  | 
lemma EQgen_mono: "mono(%X. EQgen(X))"  | 
|
442  | 
apply (unfold EQgen_def SIM_def)  | 
|
443  | 
apply (rule monoI)  | 
|
444  | 
apply blast  | 
|
445  | 
done  | 
|
446  | 
||
447  | 
lemma EQgenXH:  | 
|
448  | 
"<t,t'> : EQgen(R) <-> (t=bot & t'=bot) | (t=true & t'=true) |  | 
|
449  | 
(t=false & t'=false) |  | 
|
450  | 
(EX a a' b b'. t=<a,b> & t'=<a',b'> & <a,a'> : R & <b,b'> : R) |  | 
|
451  | 
(EX f f'. t=lam x. f(x) & t'=lam x. f'(x) & (ALL x.<f(x),f'(x)> : R))"  | 
|
452  | 
apply (unfold EQgen_def SIM_def)  | 
|
453  | 
apply (rule iff_refl [THEN XHlemma2])  | 
|
454  | 
done  | 
|
455  | 
||
456  | 
lemma eqXH:  | 
|
457  | 
"t=t' <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) |  | 
|
458  | 
(EX a a' b b'. t=<a,b> & t'=<a',b'> & a=a' & b=b') |  | 
|
459  | 
(EX f f'. t=lam x. f(x) & t'=lam x. f'(x) & (ALL x. f(x)=f'(x)))"  | 
|
460  | 
apply (subgoal_tac "<t,t'> : EQ <-> (t=bot & t'=bot) | (t=true & t'=true) | (t=false & t'=false) | (EX a a' b b'. t=<a,b> & t'=<a',b'> & <a,a'> : EQ & <b,b'> : EQ) | (EX f f'. t=lam x. f (x) & t'=lam x. f' (x) & (ALL x. <f (x) ,f' (x) > : EQ))")  | 
|
461  | 
apply (erule rev_mp)  | 
|
462  | 
apply (simp add: EQ_iff [THEN iff_sym])  | 
|
463  | 
apply (rule EQgen_mono [THEN EQ_def [THEN def_gfp_Tarski], THEN XHlemma1,  | 
|
464  | 
unfolded EQgen_def SIM_def])  | 
|
465  | 
apply (rule iff_refl [THEN XHlemma2])  | 
|
466  | 
done  | 
|
467  | 
||
468  | 
lemma eq_coinduct: "[| <t,u> : R; R <= EQgen(R) |] ==> t = u"  | 
|
469  | 
apply (rule EQ_def [THEN def_coinduct, THEN EQ_iff [THEN iffD2]])  | 
|
470  | 
apply assumption+  | 
|
471  | 
done  | 
|
472  | 
||
473  | 
lemma eq_coinduct3:  | 
|
474  | 
"[| <t,u> : R; R <= EQgen(lfp(%x. EQgen(x) Un R Un EQ)) |] ==> t = u"  | 
|
475  | 
apply (rule EQ_def [THEN def_coinduct3, THEN EQ_iff [THEN iffD2]])  | 
|
476  | 
apply (rule EQgen_mono | assumption)+  | 
|
477  | 
done  | 
|
478  | 
||
479  | 
ML {*
 | 
|
| 27239 | 480  | 
  fun eq_coinduct_tac ctxt s i = res_inst_tac ctxt [(("R", 0), s)] @{thm eq_coinduct} i
 | 
481  | 
  fun eq_coinduct3_tac ctxt s i = res_inst_tac ctxt [(("R", 0), s)] @{thm eq_coinduct3} i
 | 
|
| 20140 | 482  | 
*}  | 
483  | 
||
484  | 
||
485  | 
subsection {* Untyped Case Analysis and Other Facts *}
 | 
|
486  | 
||
487  | 
lemma cond_eta: "(EX f. t=lam x. f(x)) ==> t = lam x.(t ` x)"  | 
|
488  | 
by (auto simp: apply_def)  | 
|
489  | 
||
490  | 
lemma exhaustion: "(t=bot) | (t=true) | (t=false) | (EX a b. t=<a,b>) | (EX f. t=lam x. f(x))"  | 
|
491  | 
apply (cut_tac refl [THEN eqXH [THEN iffD1]])  | 
|
492  | 
apply blast  | 
|
493  | 
done  | 
|
494  | 
||
495  | 
lemma term_case:  | 
|
496  | 
"[| P(bot); P(true); P(false); !!x y. P(<x,y>); !!b. P(lam x. b(x)) |] ==> P(t)"  | 
|
497  | 
using exhaustion [of t] by blast  | 
|
| 17456 | 498  | 
|
499  | 
end  |