1(*  Title:      CCL/CCL.thy
2    Author:     Martin Coen
3    Copyright   1993  University of Cambridge
4*)
5
6section \<open>Classical Computational Logic for Untyped Lambda Calculus
7  with reduction to weak head-normal form\<close>
8
9theory CCL
10imports Gfp
11begin
12
13text \<open>
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\<close>
18
19class prog = "term"
20default_sort prog
21
22instance "fun" :: (prog, prog) prog ..
23
24typedecl i
25instance i :: prog ..
26
27
28consts
29  (*** Evaluation Judgement ***)
30  Eval      ::       "[i,i]\<Rightarrow>prop"          (infixl "\<longlongrightarrow>" 20)
31
32  (*** Bisimulations for pre-order and equality ***)
33  po          ::       "['a,'a]\<Rightarrow>o"           (infixl "[=" 50)
34
35  (*** Term Formers ***)
36  true        ::       "i"
37  false       ::       "i"
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)
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
50axiomatization where
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
55
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"
60
61  (*** Properties of evaluation: note that "t \<longlongrightarrow> c" impies that c is canonical ***)
62
63axiomatization where
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"
69
70  (* Should be derivable - but probably a bitch! *)
71axiomatization where
72  substitute: "\<lbrakk>a==a'; t(a)\<longlongrightarrow>c(a)\<rbrakk> \<Longrightarrow> t(a')\<longlongrightarrow>c(a')"
73
74  (************** LOGIC ***************)
75
76  (*** Definitions used in the following rules ***)
77
78axiomatization where
79  bot_def:         "bot == (lam x. x`x)`(lam x. x`x)" and
80  apply_def:     "f ` t == case(f, bot, bot, \<lambda>x y. bot, \<lambda>u. u(t))"
81
82definition "fix" :: "(i\<Rightarrow>i)\<Rightarrow>i"
83  where "fix(f) == (lam x. f(x`x))`(lam x. f(x`x))"
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
89definition SIM :: "[i,i,i set]\<Rightarrow>o"
90  where
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))"
94
95definition POgen :: "i set \<Rightarrow> i set"
96  where "POgen(R) == {p. \<exists>t t'. p=<t,t'> \<and> (t = bot | SIM(t,t',R))}"
97
98definition 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))}"
100
101definition PO :: "i set"
102  where "PO == gfp(POgen)"
103
104definition EQ :: "i set"
105  where "EQ == gfp(EQgen)"
106
107
108  (*** Rules ***)
109
110  (** Partial Order **)
111
112axiomatization where
113  po_refl:        "a [= a" and
114  po_trans:       "\<lbrakk>a [= b;  b [= c\<rbrakk> \<Longrightarrow> a [= c" and
115  po_cong:        "a [= b \<Longrightarrow> f(a) [= f(b)" and
116
117  (* Extend definition of [= to program fragments of higher type *)
118  po_abstractn:   "(\<And>x. f(x) [= g(x)) \<Longrightarrow> (\<lambda>x. f(x)) [= (\<lambda>x. g(x))"
119
120  (** Equality - equivalence axioms inherited from FOL.thy   **)
121  (**          - congruence of "=" is axiomatised implicitly **)
122
123axiomatization where
124  eq_iff:         "t = t' \<longleftrightarrow> t [= t' \<and> t' [= t"
125
126  (** Properties of canonical values given by greatest fixed point definitions **)
127
128axiomatization where
129  PO_iff:         "t [= t' \<longleftrightarrow> <t,t'> : PO" and
130  EQ_iff:         "t =  t' \<longleftrightarrow> <t,t'> : EQ"
131
132  (** Behaviour of non-canonical terms (ie case) given by the following beta-rules **)
133
134axiomatization 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 *)
140
141  (** The theory is non-trivial **)
142axiomatization where
143  distinctness:   "\<not> lam x. b(x) = bot"
144
145  (*** Definitions of Termination and Divergence ***)
146
147definition Dvg :: "i \<Rightarrow> o"
148  where "Dvg(t) == t = bot"
149
150definition Trm :: "i \<Rightarrow> o"
151  where "Trm(t) == \<not> Dvg(t)"
152
153text \<open>
154Would be interesting to build a similar theory for a typed programming language:
155    ie.     true :: bool,      fix :: ('a\<Rightarrow>'a)\<Rightarrow>'a  etc......
156
157This is starting to look like LCF.
158What are the advantages of this approach?
159        - less axiomatic
160        - wfd induction / coinduction and fixed point induction available
161\<close>
162
163
164lemmas ccl_data_defs = apply_def fix_def
165
166declare po_refl [simp]
167
168
169subsection \<open>Congruence Rules\<close>
170
171(*similar to AP_THM in Gordon's HOL*)
172lemma fun_cong: "(f::'a\<Rightarrow>'b) = g \<Longrightarrow> f(x)=g(x)"
173  by simp
174
175(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*)
176lemma arg_cong: "x=y \<Longrightarrow> f(x)=f(y)"
177  by simp
178
179lemma abstractn: "(\<And>x. f(x) = g(x)) \<Longrightarrow> f = g"
180  apply (simp add: eq_iff)
181  apply (blast intro: po_abstractn)
182  done
183
184lemmas caseBs = caseBtrue caseBfalse caseBpair caseBlam caseBbot
185
186
187subsection \<open>Termination and Divergence\<close>
188
189lemma Trm_iff: "Trm(t) \<longleftrightarrow> \<not> t = bot"
190  by (simp add: Trm_def Dvg_def)
191
192lemma Dvg_iff: "Dvg(t) \<longleftrightarrow> t = bot"
193  by (simp add: Trm_def Dvg_def)
194
195
196subsection \<open>Constructors are injective\<close>
197
198lemma eq_lemma: "\<lbrakk>x=a; y=b; x=y\<rbrakk> \<Longrightarrow> a=b"
199  by simp
200
201ML \<open>
202  fun inj_rl_tac ctxt rews i =
203    let
204      fun mk_inj_lemmas r = [@{thm arg_cong}] RL [r RS (r RS @{thm eq_lemma})]
205      val inj_lemmas = maps mk_inj_lemmas rews
206    in
207      CHANGED (REPEAT (assume_tac ctxt i ORELSE
208        resolve_tac ctxt @{thms iffI allI conjI} i ORELSE
209        eresolve_tac ctxt inj_lemmas i ORELSE
210        asm_simp_tac (ctxt addsimps rews) i))
211    end;
212\<close>
213
214method_setup inj_rl = \<open>
215  Attrib.thms >> (fn rews => fn ctxt => SIMPLE_METHOD' (inj_rl_tac ctxt rews))
216\<close>
217
218lemma ccl_injs:
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)))"
221  by (inj_rl caseBs)
222
223
224lemma pair_inject: "<a,b> = <a',b'> \<Longrightarrow> (a = a' \<Longrightarrow> b = b' \<Longrightarrow> R) \<Longrightarrow> R"
225  by (simp add: ccl_injs)
226
227
228subsection \<open>Constructors are distinct\<close>
229
230lemma lem: "t=t' \<Longrightarrow> case(t,b,c,d,e) = case(t',b,c,d,e)"
231  by simp
232
233ML \<open>
234local
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);
247           val arity = length (binder_types T)
248       in sy ^ (arg_str arity name "") end
249
250  fun mk_thm_str thy a b = "\<not> " ^ (saturate thy a "a") ^ " = " ^ (saturate thy b "b")
251
252  val lemma = @{thm lem};
253  val eq_lemma = @{thm eq_lemma};
254  val distinctness = @{thm distinctness};
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})]
258in
259  fun mk_lemmas rls = maps mk_lemma (mk_combs pair rls)
260  fun mk_dstnct_rls thy xs = mk_combs (mk_thm_str thy) xs
261end
262\<close>
263
264ML \<open>
265val caseB_lemmas = mk_lemmas @{thms caseBs}
266
267val ccl_dstncts =
268  let
269    fun mk_raw_dstnct_thm rls s =
270      Goal.prove_global \<^theory> [] [] (Syntax.read_prop_global \<^theory> s)
271        (fn {context = ctxt, ...} => resolve_tac ctxt @{thms notI} 1 THEN eresolve_tac ctxt rls 1)
272  in map (mk_raw_dstnct_thm caseB_lemmas)
273    (mk_dstnct_rls \<^theory> ["bot","true","false","pair","lambda"]) end
274
275fun mk_dstnct_thms ctxt defs inj_rls xs =
276  let
277    val thy = Proof_Context.theory_of ctxt;
278    fun mk_dstnct_thm rls s =
279      Goal.prove_global thy [] [] (Syntax.read_prop ctxt s)
280        (fn _ =>
281          rewrite_goals_tac ctxt defs THEN
282          simp_tac (ctxt addsimps (rls @ inj_rls)) 1)
283  in map (mk_dstnct_thm ccl_dstncts) (mk_dstnct_rls thy xs) end
284
285fun mkall_dstnct_thms ctxt defs i_rls xss = maps (mk_dstnct_thms ctxt defs i_rls) xss
286
287(*** Rewriting and Proving ***)
288
289fun XH_to_I rl = rl RS @{thm iffD2}
290fun XH_to_D rl = rl RS @{thm iffD1}
291val XH_to_E = make_elim o XH_to_D
292val XH_to_Is = map XH_to_I
293val XH_to_Ds = map XH_to_D
294val XH_to_Es = map XH_to_E;
295
296ML_Thms.bind_thms ("ccl_rews", @{thms caseBs} @ @{thms ccl_injs} @ ccl_dstncts);
297ML_Thms.bind_thms ("ccl_dstnctsEs", ccl_dstncts RL [@{thm notE}]);
298ML_Thms.bind_thms ("ccl_injDs", XH_to_Ds @{thms ccl_injs});
299\<close>
300
301lemmas [simp] = ccl_rews
302  and [elim!] = pair_inject ccl_dstnctsEs
303  and [dest!] = ccl_injDs
304
305
306subsection \<open>Facts from gfp Definition of \<open>[=\<close> and \<open>=\<close>\<close>
307
308lemma XHlemma1: "\<lbrakk>A=B; a:B \<longleftrightarrow> P\<rbrakk> \<Longrightarrow> a:A \<longleftrightarrow> P"
309  by simp
310
311lemma XHlemma2: "(P(t,t') \<longleftrightarrow> Q) \<Longrightarrow> (<t,t'> : {p. \<exists>t t'. p=<t,t'> \<and>  P(t,t')} \<longleftrightarrow> Q)"
312  by blast
313
314
315subsection \<open>Pre-Order\<close>
316
317lemma POgen_mono: "mono(\<lambda>X. POgen(X))"
318  apply (unfold POgen_def SIM_def)
319  apply (rule monoI)
320  apply blast
321  done
322
323lemma 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) |
326           (EX f f'. t=lam x. f(x) \<and> t'=lam x. f'(x) \<and> (ALL x. <f(x),f'(x)> : R))"
327  apply (unfold POgen_def SIM_def)
328  apply (rule iff_refl [THEN XHlemma2])
329  done
330
331lemma 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') |
334                 (EX f f'. t=lam x. f(x) \<and> t'=lam x. f'(x) \<and> (ALL x. f(x) [= f'(x)))"
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
341lemma po_bot: "bot [= b"
342  apply (rule poXH [THEN iffD2])
343  apply simp
344  done
345
346lemma bot_poleast: "a [= bot \<Longrightarrow> a=bot"
347  apply (drule poXH [THEN iffD1])
348  apply simp
349  done
350
351lemma po_pair: "<a,b> [= <a',b'> \<longleftrightarrow>  a [= a' \<and> b [= b'"
352  apply (rule poXH [THEN iff_trans])
353  apply simp
354  done
355
356lemma po_lam: "lam x. f(x) [= lam x. f'(x) \<longleftrightarrow> (ALL x. f(x) [= f'(x))"
357  apply (rule poXH [THEN iff_trans])
358  apply fastforce
359  done
360
361lemmas ccl_porews = po_bot po_pair po_lam
362
363lemma case_pocong:
364  assumes 1: "t [= t'"
365    and 2: "a [= a'"
366    and 3: "b [= b'"
367    and 4: "\<And>x y. c(x,y) [= c'(x,y)"
368    and 5: "\<And>u. d(u) [= d'(u)"
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])
374  apply (rule_tac f1 = "\<lambda>d. case (t',a',b',c',d)" in
375    5 [THEN po_abstractn, THEN po_cong, THEN po_trans])
376  apply (rule po_refl)
377  done
378
379lemma apply_pocong: "\<lbrakk>f [= f'; a [= a'\<rbrakk> \<Longrightarrow> f ` a [= f' ` a'"
380  unfolding ccl_data_defs
381  apply (rule case_pocong, (rule po_refl | assumption)+)
382  apply (erule po_cong)
383  done
384
385lemma npo_lam_bot: "\<not> lam x. b(x) [= bot"
386  apply (rule notI)
387  apply (drule bot_poleast)
388  apply (erule distinctness [THEN notE])
389  done
390
391lemma po_lemma: "\<lbrakk>x=a; y=b; x[=y\<rbrakk> \<Longrightarrow> a[=b"
392  by simp
393
394lemma npo_pair_lam: "\<not> <a,b> [= lam x. f(x)"
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
401lemma npo_lam_pair: "\<not> lam x. f(x) [= <a,b>"
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
408lemma npo_rls1:
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"
419  by (rule notI, drule case_pocong, erule_tac [5] rev_mp, simp_all,
420    (rule po_refl npo_lam_bot)+)+
421
422lemmas npo_rls = npo_pair_lam npo_lam_pair npo_rls1
423
424
425subsection \<open>Coinduction for \<open>[=\<close>\<close>
426
427lemma po_coinduct: "\<lbrakk><t,u> : R; R <= POgen(R)\<rbrakk> \<Longrightarrow> t [= u"
428  apply (rule PO_def [THEN def_coinduct, THEN PO_iff [THEN iffD2]])
429   apply assumption+
430  done
431
432
433subsection \<open>Equality\<close>
434
435lemma EQgen_mono: "mono(\<lambda>X. EQgen(X))"
436  apply (unfold EQgen_def SIM_def)
437  apply (rule monoI)
438  apply blast
439  done
440
441lemma 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) |
445                 (EX f f'. t=lam x. f(x) \<and> t'=lam x. f'(x) \<and> (ALL x.<f(x),f'(x)> : R))"
446  apply (unfold EQgen_def SIM_def)
447  apply (rule iff_refl [THEN XHlemma2])
448  done
449
450lemma 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') |
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))")
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
465lemma eq_coinduct: "\<lbrakk><t,u> : R; R <= EQgen(R)\<rbrakk> \<Longrightarrow> t = u"
466  apply (rule EQ_def [THEN def_coinduct, THEN EQ_iff [THEN iffD2]])
467   apply assumption+
468  done
469
470lemma eq_coinduct3:
471  "\<lbrakk><t,u> : R;  R <= EQgen(lfp(\<lambda>x. EQgen(x) Un R Un EQ))\<rbrakk> \<Longrightarrow> t = u"
472  apply (rule EQ_def [THEN def_coinduct3, THEN EQ_iff [THEN iffD2]])
473  apply (rule EQgen_mono | assumption)+
474  done
475
476method_setup eq_coinduct3 = \<open>
477  Scan.lift Args.embedded_inner_syntax >> (fn s => fn ctxt =>
478    SIMPLE_METHOD'
479      (Rule_Insts.res_inst_tac ctxt [((("R", 0), Position.none), s)] [] @{thm eq_coinduct3}))
480\<close>
481
482
483subsection \<open>Untyped Case Analysis and Other Facts\<close>
484
485lemma cond_eta: "(EX f. t=lam x. f(x)) \<Longrightarrow> t = lam x.(t ` x)"
486  by (auto simp: apply_def)
487
488lemma 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
493lemma term_case:
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)"
495  using exhaustion [of t] by blast
496
497end
498