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