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