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