author | wenzelm |
Thu, 15 Feb 2001 17:18:54 +0100 | |
changeset 11145 | 3e47692e3a3e |
parent 4098 | 71e05eb27fb6 |
child 12036 | 49f6c49454c2 |
permissions | -rw-r--r-- |
2570 | 1 |
(* $Id$ *) |
2 |
||
3 |
open Dagstuhl; |
|
4 |
||
5 |
val YS_def2 = fix_prover2 Dagstuhl.thy YS_def "YS = y && YS"; |
|
6 |
val YYS_def2 = fix_prover2 Dagstuhl.thy YYS_def "YYS = y && y && YYS"; |
|
7 |
||
8 |
||
9 |
val prems = goal Dagstuhl.thy "YYS << y && YYS"; |
|
10 |
by (rtac (YYS_def RS def_fix_ind) 1); |
|
11 |
by (Simp_tac 1); |
|
3032
74c5f175aa8e
minor changes due to more powerful continuity check in Lift3.ML
mueller
parents:
2570
diff
changeset
|
12 |
by (Simp_tac 1); |
74c5f175aa8e
minor changes due to more powerful continuity check in Lift3.ML
mueller
parents:
2570
diff
changeset
|
13 |
by (Simp_tac 1); |
2570 | 14 |
by (rtac monofun_cfun_arg 1); |
15 |
by (rtac monofun_cfun_arg 1); |
|
16 |
by (atac 1); |
|
17 |
val lemma3 = result(); |
|
18 |
||
19 |
val prems = goal Dagstuhl.thy "y && YYS << YYS"; |
|
20 |
by (stac YYS_def2 1); |
|
21 |
back(); |
|
22 |
by (rtac monofun_cfun_arg 1); |
|
23 |
by (rtac lemma3 1); |
|
24 |
val lemma4=result(); |
|
25 |
||
26 |
(* val lemma5 = lemma3 RS (lemma4 RS antisym_less) *) |
|
27 |
||
28 |
val prems = goal Dagstuhl.thy "y && YYS = YYS"; |
|
29 |
by (rtac antisym_less 1); |
|
30 |
by (rtac lemma4 1); |
|
31 |
by (rtac lemma3 1); |
|
32 |
val lemma5=result(); |
|
33 |
||
34 |
val prems = goal Dagstuhl.thy "YS = YYS"; |
|
4044 | 35 |
by (resolve_tac stream.take_lemmas 1); |
2570 | 36 |
by (nat_ind_tac "n" 1); |
4098 | 37 |
by (simp_tac (simpset() addsimps stream.rews) 1); |
2570 | 38 |
by (stac YS_def2 1); |
39 |
by (stac YYS_def2 1); |
|
4098 | 40 |
by (asm_simp_tac (simpset() addsimps stream.rews) 1); |
2570 | 41 |
by (rtac (lemma5 RS sym RS subst) 1); |
42 |
by (rtac refl 1); |
|
43 |
val wir_moel=result(); |
|
44 |
||
45 |
(* ------------------------------------------------------------------------ *) |
|
46 |
(* Zweite L"osung: Bernhard M"oller *) |
|
47 |
(* statt Beweis von wir_moel "uber take_lemma beidseitige Inclusion *) |
|
48 |
(* verwendet lemma5 *) |
|
49 |
(* ------------------------------------------------------------------------ *) |
|
50 |
||
51 |
val prems = goal Dagstuhl.thy "YYS << YS"; |
|
52 |
by (rewtac YYS_def); |
|
53 |
by (rtac fix_least 1); |
|
54 |
by (stac beta_cfun 1); |
|
55 |
by (cont_tacR 1); |
|
4098 | 56 |
by (simp_tac (simpset() addsimps [YS_def2 RS sym]) 1); |
2570 | 57 |
val lemma6=result(); |
58 |
||
59 |
val prems = goal Dagstuhl.thy "YS << YYS"; |
|
60 |
by (rtac (YS_def RS def_fix_ind) 1); |
|
61 |
by (Simp_tac 1); |
|
3032
74c5f175aa8e
minor changes due to more powerful continuity check in Lift3.ML
mueller
parents:
2570
diff
changeset
|
62 |
by (Simp_tac 1); |
74c5f175aa8e
minor changes due to more powerful continuity check in Lift3.ML
mueller
parents:
2570
diff
changeset
|
63 |
by (Simp_tac 1); |
2570 | 64 |
by (stac (lemma5 RS sym) 1); |
65 |
by (etac monofun_cfun_arg 1); |
|
66 |
val lemma7 = result(); |
|
67 |
||
68 |
val wir_moel = lemma6 RS (lemma7 RS antisym_less); |
|
69 |