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