| author | wenzelm | 
| Wed, 25 Nov 2020 16:14:16 +0100 | |
| changeset 72709 | cb9d5af781b4 | 
| parent 61394 | 6142b282b164 | 
| child 76213 | e44d86131648 | 
| permissions | -rw-r--r-- | 
| 1478 | 1  | 
(* Title: ZF/AC/WO6_WO1.thy  | 
2  | 
Author: Krzysztof Grabczewski  | 
|
| 1019 | 3  | 
|
| 12776 | 4  | 
Proofs needed to state that formulations WO1,...,WO6 are all equivalent.  | 
5  | 
The only hard one is WO6 ==> WO1.  | 
|
| 1019 | 6  | 
|
| 12776 | 7  | 
Every proof (except WO6 ==> WO1 and WO1 ==> WO2) are described as "clear"  | 
8  | 
by Rubin & Rubin (page 2).  | 
|
| 40945 | 9  | 
They refer reader to a book by Gödel to see the proof WO1 ==> WO2.  | 
| 12776 | 10  | 
Fortunately order types made this proof also very easy.  | 
| 1019 | 11  | 
*)  | 
12  | 
||
| 27678 | 13  | 
theory WO6_WO1  | 
14  | 
imports Cardinal_aux  | 
|
15  | 
begin  | 
|
| 1019 | 16  | 
|
17  | 
(* Auxiliary definitions used in proof *)  | 
|
| 24893 | 18  | 
definition  | 
19  | 
NN :: "i => i" where  | 
|
| 12776 | 20  | 
    "NN(y) == {m \<in> nat. \<exists>a. \<exists>f. Ord(a) & domain(f)=a  &  
 | 
21  | 
(\<Union>b<a. f`b) = y & (\<forall>b<a. f`b \<lesssim> m)}"  | 
|
22  | 
||
| 24893 | 23  | 
definition  | 
24  | 
uu :: "[i, i, i, i] => i" where  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
25  | 
"uu(f, beta, gamma, delta) == (f`beta * f`gamma) \<inter> f`delta"  | 
| 12776 | 26  | 
|
| 24893 | 27  | 
|
28  | 
(** Definitions for case 1 **)  | 
|
29  | 
definition  | 
|
30  | 
vv1 :: "[i, i, i] => i" where  | 
|
| 12776 | 31  | 
"vv1(f,m,b) ==  | 
| 61394 | 32  | 
let g = \<mu> g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 &  | 
| 12776 | 33  | 
domain(uu(f,b,g,d)) \<lesssim> m));  | 
| 61394 | 34  | 
d = \<mu> d. domain(uu(f,b,g,d)) \<noteq> 0 &  | 
| 12776 | 35  | 
domain(uu(f,b,g,d)) \<lesssim> m  | 
36  | 
in if f`b \<noteq> 0 then domain(uu(f,b,g,d)) else 0"  | 
|
37  | 
||
| 24893 | 38  | 
definition  | 
39  | 
ww1 :: "[i, i, i] => i" where  | 
|
| 12776 | 40  | 
"ww1(f,m,b) == f`b - vv1(f,m,b)"  | 
41  | 
||
| 24893 | 42  | 
definition  | 
43  | 
gg1 :: "[i, i, i] => i" where  | 
|
| 12776 | 44  | 
"gg1(f,a,m) == \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"  | 
45  | 
||
| 24893 | 46  | 
|
47  | 
(** Definitions for case 2 **)  | 
|
48  | 
definition  | 
|
49  | 
vv2 :: "[i, i, i, i] => i" where  | 
|
| 12776 | 50  | 
"vv2(f,b,g,s) ==  | 
| 61394 | 51  | 
           if f`g \<noteq> 0 then {uu(f, b, g, \<mu> d. uu(f,b,g,d) \<noteq> 0)`s} else 0"
 | 
| 12776 | 52  | 
|
| 24893 | 53  | 
definition  | 
54  | 
ww2 :: "[i, i, i, i] => i" where  | 
|
| 12776 | 55  | 
"ww2(f,b,g,s) == f`g - vv2(f,b,g,s)"  | 
56  | 
||
| 24893 | 57  | 
definition  | 
58  | 
gg2 :: "[i, i, i, i] => i" where  | 
|
| 12776 | 59  | 
"gg2(f,a,b,s) ==  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
27678 
diff
changeset
 | 
60  | 
\<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"  | 
| 12776 | 61  | 
|
62  | 
||
63  | 
lemma WO2_WO3: "WO2 ==> WO3"  | 
|
64  | 
by (unfold WO2_def WO3_def, fast)  | 
|
65  | 
||
66  | 
(* ********************************************************************** *)  | 
|
67  | 
||
68  | 
lemma WO3_WO1: "WO3 ==> WO1"  | 
|
69  | 
apply (unfold eqpoll_def WO1_def WO3_def)  | 
|
70  | 
apply (intro allI)  | 
|
71  | 
apply (drule_tac x=A in spec)  | 
|
72  | 
apply (blast intro: bij_is_inj well_ord_rvimage  | 
|
73  | 
well_ord_Memrel [THEN well_ord_subset])  | 
|
74  | 
done  | 
|
75  | 
||
76  | 
(* ********************************************************************** *)  | 
|
77  | 
||
78  | 
lemma WO1_WO2: "WO1 ==> WO2"  | 
|
79  | 
apply (unfold eqpoll_def WO1_def WO2_def)  | 
|
80  | 
apply (blast intro!: Ord_ordertype ordermap_bij)  | 
|
81  | 
done  | 
|
82  | 
||
83  | 
(* ********************************************************************** *)  | 
|
84  | 
||
85  | 
lemma lam_sets: "f \<in> A->B ==> (\<lambda>x \<in> A. {f`x}): A -> {{b}. b \<in> B}"
 | 
|
86  | 
by (fast intro!: lam_type apply_type)  | 
|
87  | 
||
| 24783 | 88  | 
lemma surj_imp_eq': "f \<in> surj(A,B) ==> (\<Union>a \<in> A. {f`a}) = B"
 | 
| 12776 | 89  | 
apply (unfold surj_def)  | 
90  | 
apply (fast elim!: apply_type)  | 
|
91  | 
done  | 
|
92  | 
||
93  | 
lemma surj_imp_eq: "[| f \<in> surj(A,B); Ord(A) |] ==> (\<Union>a<A. {f`a}) = B"
 | 
|
| 24783 | 94  | 
by (fast dest!: surj_imp_eq' intro!: ltI elim!: ltE)  | 
| 12776 | 95  | 
|
96  | 
lemma WO1_WO4: "WO1 ==> WO4(1)"  | 
|
97  | 
apply (unfold WO1_def WO4_def)  | 
|
98  | 
apply (rule allI)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
13175 
diff
changeset
 | 
99  | 
apply (erule_tac x = A in allE)  | 
| 12776 | 100  | 
apply (erule exE)  | 
101  | 
apply (intro exI conjI)  | 
|
102  | 
apply (erule Ord_ordertype)  | 
|
103  | 
apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_is_fun, THEN lam_sets, THEN domain_of_fun])  | 
|
104  | 
apply (simp_all add: singleton_eqpoll_1 eqpoll_imp_lepoll Ord_ordertype  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
105  | 
ordermap_bij [THEN bij_converse_bij, THEN bij_is_surj, THEN surj_imp_eq]  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
106  | 
ltD)  | 
| 12776 | 107  | 
done  | 
108  | 
||
109  | 
(* ********************************************************************** *)  | 
|
110  | 
||
111  | 
lemma WO4_mono: "[| m\<le>n; WO4(m) |] ==> WO4(n)"  | 
|
112  | 
apply (unfold WO4_def)  | 
|
113  | 
apply (blast dest!: spec intro: lepoll_trans [OF _ le_imp_lepoll])  | 
|
114  | 
done  | 
|
115  | 
||
116  | 
(* ********************************************************************** *)  | 
|
117  | 
||
118  | 
lemma WO4_WO5: "[| m \<in> nat; 1\<le>m; WO4(m) |] ==> WO5"  | 
|
119  | 
by (unfold WO4_def WO5_def, blast)  | 
|
120  | 
||
121  | 
(* ********************************************************************** *)  | 
|
122  | 
||
123  | 
lemma WO5_WO6: "WO5 ==> WO6"  | 
|
124  | 
by (unfold WO4_def WO5_def WO6_def, blast)  | 
|
125  | 
||
126  | 
||
127  | 
(* **********************************************************************  | 
|
128  | 
The proof of "WO6 ==> WO1". Simplified by L C Paulson.  | 
|
129  | 
||
130  | 
From the book "Equivalents of the Axiom of Choice" by Rubin & Rubin,  | 
|
131  | 
pages 2-5  | 
|
132  | 
************************************************************************* *)  | 
|
133  | 
||
134  | 
lemma lt_oadd_odiff_disj:  | 
|
135  | 
"[| k < i++j; Ord(i); Ord(j) |]  | 
|
136  | 
==> k < i | (~ k<i & k = i ++ (k--i) & (k--i)<j)"  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
13175 
diff
changeset
 | 
137  | 
apply (rule_tac i = k and j = i in Ord_linear2)  | 
| 12776 | 138  | 
prefer 4  | 
139  | 
apply (drule odiff_lt_mono2, assumption)  | 
|
140  | 
apply (simp add: oadd_odiff_inverse odiff_oadd_inverse)  | 
|
141  | 
apply (auto elim!: lt_Ord)  | 
|
142  | 
done  | 
|
143  | 
||
144  | 
||
145  | 
(* ********************************************************************** *)  | 
|
146  | 
(* The most complicated part of the proof - lemma ii - p. 2-4 *)  | 
|
147  | 
(* ********************************************************************** *)  | 
|
148  | 
||
149  | 
(* ********************************************************************** *)  | 
|
150  | 
(* some properties of relation uu(beta, gamma, delta) - p. 2 *)  | 
|
151  | 
(* ********************************************************************** *)  | 
|
| 1019 | 152  | 
|
| 12776 | 153  | 
lemma domain_uu_subset: "domain(uu(f,b,g,d)) \<subseteq> f`b"  | 
154  | 
by (unfold uu_def, blast)  | 
|
155  | 
||
156  | 
lemma quant_domain_uu_lepoll_m:  | 
|
157  | 
"\<forall>b<a. f`b \<lesssim> m ==> \<forall>b<a. \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<lesssim> m"  | 
|
158  | 
by (blast intro: domain_uu_subset [THEN subset_imp_lepoll] lepoll_trans)  | 
|
159  | 
||
160  | 
lemma uu_subset1: "uu(f,b,g,d) \<subseteq> f`b * f`g"  | 
|
161  | 
by (unfold uu_def, blast)  | 
|
162  | 
||
163  | 
lemma uu_subset2: "uu(f,b,g,d) \<subseteq> f`d"  | 
|
164  | 
by (unfold uu_def, blast)  | 
|
165  | 
||
166  | 
lemma uu_lepoll_m: "[| \<forall>b<a. f`b \<lesssim> m; d<a |] ==> uu(f,b,g,d) \<lesssim> m"  | 
|
167  | 
by (blast intro: uu_subset2 [THEN subset_imp_lepoll] lepoll_trans)  | 
|
168  | 
||
169  | 
(* ********************************************************************** *)  | 
|
170  | 
(* Two cases for lemma ii *)  | 
|
171  | 
(* ********************************************************************** *)  | 
|
172  | 
lemma cases:  | 
|
173  | 
"\<forall>b<a. \<forall>g<a. \<forall>d<a. u(f,b,g,d) \<lesssim> m  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
174  | 
==> (\<forall>b<a. f`b \<noteq> 0 \<longrightarrow>  | 
| 12776 | 175  | 
(\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 & u(f,b,g,d) \<prec> m))  | 
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
176  | 
| (\<exists>b<a. f`b \<noteq> 0 & (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 \<longrightarrow>  | 
| 12776 | 177  | 
u(f,b,g,d) \<approx> m))"  | 
178  | 
apply (unfold lesspoll_def)  | 
|
179  | 
apply (blast del: equalityI)  | 
|
180  | 
done  | 
|
181  | 
||
182  | 
(* ********************************************************************** *)  | 
|
183  | 
(* Lemmas used in both cases *)  | 
|
184  | 
(* ********************************************************************** *)  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
185  | 
lemma UN_oadd: "Ord(a) ==> (\<Union>b<a++a. C(b)) = (\<Union>b<a. C(b) \<union> C(a++b))"  | 
| 12776 | 186  | 
by (blast intro: ltI lt_oadd1 oadd_lt_mono2 dest!: lt_oadd_disj)  | 
187  | 
||
188  | 
||
189  | 
(* ********************************************************************** *)  | 
|
190  | 
(* Case 1: lemmas *)  | 
|
191  | 
(* ********************************************************************** *)  | 
|
192  | 
||
193  | 
lemma vv1_subset: "vv1(f,m,b) \<subseteq> f`b"  | 
|
194  | 
by (simp add: vv1_def Let_def domain_uu_subset)  | 
|
195  | 
||
196  | 
(* ********************************************************************** *)  | 
|
197  | 
(* Case 1: Union of images is the whole "y" *)  | 
|
198  | 
(* ********************************************************************** *)  | 
|
199  | 
lemma UN_gg1_eq:  | 
|
200  | 
"[| Ord(a); m \<in> nat |] ==> (\<Union>b<a++a. gg1(f,a,m)`b) = (\<Union>b<a. f`b)"  | 
|
201  | 
by (simp add: gg1_def UN_oadd lt_oadd1 oadd_le_self [THEN le_imp_not_lt]  | 
|
202  | 
lt_Ord odiff_oadd_inverse ltD vv1_subset [THEN Diff_partition]  | 
|
203  | 
ww1_def)  | 
|
204  | 
||
205  | 
lemma domain_gg1: "domain(gg1(f,a,m)) = a++a"  | 
|
206  | 
by (simp add: lam_funtype [THEN domain_of_fun] gg1_def)  | 
|
| 992 | 207  | 
|
| 12776 | 208  | 
(* ********************************************************************** *)  | 
209  | 
(* every value of defined function is less than or equipollent to m *)  | 
|
210  | 
(* ********************************************************************** *)  | 
|
211  | 
lemma nested_LeastI:  | 
|
212  | 
"[| P(a, b); Ord(a); Ord(b);  | 
|
| 61394 | 213  | 
Least_a = (\<mu> a. \<exists>x. Ord(x) & P(a, x)) |]  | 
214  | 
==> P(Least_a, \<mu> b. P(Least_a, b))"  | 
|
| 12776 | 215  | 
apply (erule ssubst)  | 
| 61394 | 216  | 
apply (rule_tac Q = "%z. P (z, \<mu> b. P (z, b))" in LeastI2)  | 
| 12776 | 217  | 
apply (fast elim!: LeastI)+  | 
218  | 
done  | 
|
219  | 
||
| 45602 | 220  | 
lemmas nested_Least_instance =  | 
| 12776 | 221  | 
nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \<noteq> 0 &  | 
| 45602 | 222  | 
domain(uu(f,b,g,d)) \<lesssim> m"] for f b m  | 
| 1019 | 223  | 
|
| 12776 | 224  | 
lemma gg1_lepoll_m:  | 
225  | 
"[| Ord(a); m \<in> nat;  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
226  | 
\<forall>b<a. f`b \<noteq>0 \<longrightarrow>  | 
| 12776 | 227  | 
(\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0 &  | 
228  | 
domain(uu(f,b,g,d)) \<lesssim> m);  | 
|
229  | 
\<forall>b<a. f`b \<lesssim> succ(m); b<a++a |]  | 
|
230  | 
==> gg1(f,a,m)`b \<lesssim> m"  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
231  | 
apply (simp add: gg1_def empty_lepollI)  | 
| 12776 | 232  | 
apply (safe dest!: lt_oadd_odiff_disj)  | 
233  | 
(*Case b<a \<in> show vv1(f,m,b) \<lesssim> m *)  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
234  | 
apply (simp add: vv1_def Let_def empty_lepollI)  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
235  | 
apply (fast intro: nested_Least_instance [THEN conjunct2]  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
236  | 
elim!: lt_Ord)  | 
| 12776 | 237  | 
(*Case a\<le>b \<in> show ww1(f,m,b--a) \<lesssim> m *)  | 
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
238  | 
apply (simp add: ww1_def empty_lepollI)  | 
| 12776 | 239  | 
apply (case_tac "f` (b--a) = 0", simp add: empty_lepollI)  | 
240  | 
apply (rule Diff_lepoll, blast)  | 
|
241  | 
apply (rule vv1_subset)  | 
|
242  | 
apply (drule ospec [THEN mp], assumption+)  | 
|
243  | 
apply (elim oexE conjE)  | 
|
244  | 
apply (simp add: vv1_def Let_def lt_Ord nested_Least_instance [THEN conjunct1])  | 
|
245  | 
done  | 
|
246  | 
||
247  | 
(* ********************************************************************** *)  | 
|
248  | 
(* Case 2: lemmas *)  | 
|
249  | 
(* ********************************************************************** *)  | 
|
| 1042 | 250  | 
|
| 12776 | 251  | 
(* ********************************************************************** *)  | 
252  | 
(* Case 2: vv2_subset *)  | 
|
253  | 
(* ********************************************************************** *)  | 
|
254  | 
||
255  | 
lemma ex_d_uu_not_empty:  | 
|
256  | 
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0;  | 
|
257  | 
y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]  | 
|
258  | 
==> \<exists>d<a. uu(f,b,g,d) \<noteq> 0"  | 
|
259  | 
by (unfold uu_def, blast)  | 
|
260  | 
||
261  | 
lemma uu_not_empty:  | 
|
262  | 
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]  | 
|
| 61394 | 263  | 
==> uu(f,b,g,\<mu> d. (uu(f,b,g,d) \<noteq> 0)) \<noteq> 0"  | 
| 12776 | 264  | 
apply (drule ex_d_uu_not_empty, assumption+)  | 
265  | 
apply (fast elim!: LeastI lt_Ord)  | 
|
266  | 
done  | 
|
267  | 
||
268  | 
lemma not_empty_rel_imp_domain: "[| r \<subseteq> A*B; r\<noteq>0 |] ==> domain(r)\<noteq>0"  | 
|
269  | 
by blast  | 
|
270  | 
||
271  | 
lemma Least_uu_not_empty_lt_a:  | 
|
272  | 
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |]  | 
|
| 61394 | 273  | 
==> (\<mu> d. uu(f,b,g,d) \<noteq> 0) < a"  | 
| 12776 | 274  | 
apply (erule ex_d_uu_not_empty [THEN oexE], assumption+)  | 
275  | 
apply (blast intro: Least_le [THEN lt_trans1] lt_Ord)  | 
|
276  | 
done  | 
|
277  | 
||
278  | 
lemma subset_Diff_sing: "[| B \<subseteq> A; a\<notin>B |] ==> B \<subseteq> A-{a}"
 | 
|
279  | 
by blast  | 
|
| 1042 | 280  | 
|
| 12776 | 281  | 
(*Could this be proved more directly?*)  | 
282  | 
lemma supset_lepoll_imp_eq:  | 
|
283  | 
"[| A \<lesssim> m; m \<lesssim> B; B \<subseteq> A; m \<in> nat |] ==> A=B"  | 
|
284  | 
apply (erule natE)  | 
|
285  | 
apply (fast dest!: lepoll_0_is_0 intro!: equalityI)  | 
|
286  | 
apply (safe intro!: equalityI)  | 
|
287  | 
apply (rule ccontr)  | 
|
288  | 
apply (rule succ_lepoll_natE)  | 
|
289  | 
apply (erule lepoll_trans)  | 
|
290  | 
apply (rule lepoll_trans)  | 
|
291  | 
apply (erule subset_Diff_sing [THEN subset_imp_lepoll], assumption)  | 
|
292  | 
apply (rule Diff_sing_lepoll, assumption+)  | 
|
293  | 
done  | 
|
294  | 
||
295  | 
lemma uu_Least_is_fun:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
296  | 
"[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>  | 
| 12776 | 297  | 
domain(uu(f, b, g, d)) \<approx> succ(m);  | 
298  | 
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;  | 
|
299  | 
(\<Union>b<a. f`b)=y; b<a; g<a; d<a;  | 
|
300  | 
f`b\<noteq>0; f`g\<noteq>0; m \<in> nat; s \<in> f`b |]  | 
|
| 61394 | 301  | 
==> uu(f, b, g, \<mu> d. uu(f,b,g,d)\<noteq>0) \<in> f`b -> f`g"  | 
| 12776 | 302  | 
apply (drule_tac x2=g in ospec [THEN ospec, THEN mp])  | 
303  | 
apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty])  | 
|
304  | 
apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+)  | 
|
305  | 
apply (rule rel_is_fun)  | 
|
306  | 
apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll])  | 
|
307  | 
apply (erule uu_lepoll_m)  | 
|
308  | 
apply (rule Least_uu_not_empty_lt_a, assumption+)  | 
|
309  | 
apply (rule uu_subset1)  | 
|
310  | 
apply (rule supset_lepoll_imp_eq [OF _ eqpoll_sym [THEN eqpoll_imp_lepoll]])  | 
|
311  | 
apply (fast intro!: domain_uu_subset)+  | 
|
312  | 
done  | 
|
313  | 
||
314  | 
lemma vv2_subset:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
315  | 
"[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 \<longrightarrow>  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
27678 
diff
changeset
 | 
316  | 
domain(uu(f, b, g, d)) \<approx> succ(m);  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
27678 
diff
changeset
 | 
317  | 
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
27678 
diff
changeset
 | 
318  | 
(\<Union>b<a. f`b)=y; b<a; g<a; m \<in> nat; s \<in> f`b |]  | 
| 12776 | 319  | 
==> vv2(f,b,g,s) \<subseteq> f`g"  | 
320  | 
apply (simp add: vv2_def)  | 
|
321  | 
apply (blast intro: uu_Least_is_fun [THEN apply_type])  | 
|
322  | 
done  | 
|
323  | 
||
324  | 
(* ********************************************************************** *)  | 
|
325  | 
(* Case 2: Union of images is the whole "y" *)  | 
|
326  | 
(* ********************************************************************** *)  | 
|
327  | 
lemma UN_gg2_eq:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
328  | 
"[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>  | 
| 12776 | 329  | 
domain(uu(f,b,g,d)) \<approx> succ(m);  | 
330  | 
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;  | 
|
331  | 
(\<Union>b<a. f`b)=y; Ord(a); m \<in> nat; s \<in> f`b; b<a |]  | 
|
332  | 
==> (\<Union>g<a++a. gg2(f,a,b,s) ` g) = y"  | 
|
333  | 
apply (unfold gg2_def)  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
334  | 
apply (drule sym)  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
335  | 
apply (simp add: ltD UN_oadd oadd_le_self [THEN le_imp_not_lt]  | 
| 12776 | 336  | 
lt_Ord odiff_oadd_inverse ww2_def  | 
337  | 
vv2_subset [THEN Diff_partition])  | 
|
338  | 
done  | 
|
339  | 
||
340  | 
lemma domain_gg2: "domain(gg2(f,a,b,s)) = a++a"  | 
|
341  | 
by (simp add: lam_funtype [THEN domain_of_fun] gg2_def)  | 
|
342  | 
||
| 1042 | 343  | 
|
| 12776 | 344  | 
(* ********************************************************************** *)  | 
345  | 
(* every value of defined function is less than or equipollent to m *)  | 
|
346  | 
(* ********************************************************************** *)  | 
|
347  | 
||
348  | 
lemma vv2_lepoll: "[| m \<in> nat; m\<noteq>0 |] ==> vv2(f,b,g,s) \<lesssim> m"  | 
|
349  | 
apply (unfold vv2_def)  | 
|
350  | 
apply (simp add: empty_lepollI)  | 
|
351  | 
apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0]  | 
|
352  | 
intro!: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans]  | 
|
353  | 
not_lt_imp_le [THEN le_imp_subset, THEN subset_imp_lepoll]  | 
|
354  | 
nat_into_Ord nat_1I)  | 
|
355  | 
done  | 
|
356  | 
||
357  | 
lemma ww2_lepoll:  | 
|
358  | 
"[| \<forall>b<a. f`b \<lesssim> succ(m); g<a; m \<in> nat; vv2(f,b,g,d) \<subseteq> f`g |]  | 
|
359  | 
==> ww2(f,b,g,d) \<lesssim> m"  | 
|
360  | 
apply (unfold ww2_def)  | 
|
361  | 
apply (case_tac "f`g = 0")  | 
|
362  | 
apply (simp add: empty_lepollI)  | 
|
363  | 
apply (drule ospec, assumption)  | 
|
364  | 
apply (rule Diff_lepoll, assumption+)  | 
|
365  | 
apply (simp add: vv2_def not_emptyI)  | 
|
366  | 
done  | 
|
367  | 
||
368  | 
lemma gg2_lepoll_m:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
369  | 
"[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 \<longrightarrow>  | 
| 12776 | 370  | 
domain(uu(f,b,g,d)) \<approx> succ(m);  | 
371  | 
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y;  | 
|
372  | 
(\<Union>b<a. f`b)=y; b<a; s \<in> f`b; m \<in> nat; m\<noteq> 0; g<a++a |]  | 
|
373  | 
==> gg2(f,a,b,s) ` g \<lesssim> m"  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
374  | 
apply (simp add: gg2_def empty_lepollI)  | 
| 12776 | 375  | 
apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj)  | 
376  | 
apply (simp add: vv2_lepoll)  | 
|
377  | 
apply (simp add: ww2_lepoll vv2_subset)  | 
|
378  | 
done  | 
|
379  | 
||
380  | 
(* ********************************************************************** *)  | 
|
381  | 
(* lemma ii *)  | 
|
382  | 
(* ********************************************************************** *)  | 
|
383  | 
lemma lemma_ii: "[| succ(m) \<in> NN(y); y*y \<subseteq> y; m \<in> nat; m\<noteq>0 |] ==> m \<in> NN(y)"  | 
|
384  | 
apply (unfold NN_def)  | 
|
385  | 
apply (elim CollectE exE conjE)  | 
|
386  | 
apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption)  | 
|
387  | 
(* case 1 *)  | 
|
388  | 
apply (simp add: lesspoll_succ_iff)  | 
|
389  | 
apply (rule_tac x = "a++a" in exI)  | 
|
390  | 
apply (fast intro!: Ord_oadd domain_gg1 UN_gg1_eq gg1_lepoll_m)  | 
|
391  | 
(* case 2 *)  | 
|
392  | 
apply (elim oexE conjE)  | 
|
| 59788 | 393  | 
apply (rule_tac A = "f`B" for B in not_emptyE, assumption)  | 
| 12776 | 394  | 
apply (rule CollectI)  | 
395  | 
apply (erule succ_natD)  | 
|
396  | 
apply (rule_tac x = "a++a" in exI)  | 
|
397  | 
apply (rule_tac x = "gg2 (f,a,b,x) " in exI)  | 
|
398  | 
apply (simp add: Ord_oadd domain_gg2 UN_gg2_eq gg2_lepoll_m)  | 
|
399  | 
done  | 
|
400  | 
||
401  | 
||
402  | 
(* ********************************************************************** *)  | 
|
403  | 
(* lemma iv - p. 4: *)  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
404  | 
(* For every set x there is a set y such that x \<union> (y * y) \<subseteq> y *)  | 
| 12776 | 405  | 
(* ********************************************************************** *)  | 
406  | 
||
407  | 
||
408  | 
(* The leading \<forall>-quantifier looks odd but makes the proofs shorter  | 
|
409  | 
(used only in the following two lemmas) *)  | 
|
| 1042 | 410  | 
|
| 12776 | 411  | 
lemma z_n_subset_z_succ_n:  | 
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
412  | 
"\<forall>n \<in> nat. rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(succ(n), x, %k r. r \<union> r*r)"  | 
| 12776 | 413  | 
by (fast intro: rec_succ [THEN ssubst])  | 
414  | 
||
415  | 
lemma le_subsets:  | 
|
416  | 
"[| \<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat |]  | 
|
417  | 
==> f(n)<=f(m)"  | 
|
418  | 
apply (erule_tac P = "n\<le>m" in rev_mp)  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
419  | 
apply (rule_tac P = "%z. n\<le>z \<longrightarrow> f (n) \<subseteq> f (z) " in nat_induct)  | 
| 12776 | 420  | 
apply (auto simp add: le_iff)  | 
421  | 
done  | 
|
422  | 
||
423  | 
lemma le_imp_rec_subset:  | 
|
424  | 
"[| n\<le>m; m \<in> nat |]  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
425  | 
==> rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(m, x, %k r. r \<union> r*r)"  | 
| 12776 | 426  | 
apply (rule z_n_subset_z_succ_n [THEN le_subsets])  | 
427  | 
apply (blast intro: lt_nat_in_nat)+  | 
|
428  | 
done  | 
|
429  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
430  | 
lemma lemma_iv: "\<exists>y. x \<union> y*y \<subseteq> y"  | 
| 
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
431  | 
apply (rule_tac x = "\<Union>n \<in> nat. rec (n, x, %k r. r \<union> r*r) " in exI)  | 
| 12776 | 432  | 
apply safe  | 
433  | 
apply (rule nat_0I [THEN UN_I], simp)  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
434  | 
apply (rule_tac a = "succ (n \<union> na) " in UN_I)  | 
| 12776 | 435  | 
apply (erule Un_nat_type [THEN nat_succI], assumption)  | 
436  | 
apply (auto intro: le_imp_rec_subset [THEN subsetD]  | 
|
437  | 
intro!: Un_upper1_le Un_upper2_le Un_nat_type  | 
|
438  | 
elim!: nat_into_Ord)  | 
|
439  | 
done  | 
|
| 1019 | 440  | 
|
| 12776 | 441  | 
(* ********************************************************************** *)  | 
442  | 
(* Rubin & Rubin wrote, *)  | 
|
443  | 
(* "It follows from (ii) and mathematical induction that if y*y \<subseteq> y then *)  | 
|
444  | 
(* y can be well-ordered" *)  | 
|
445  | 
||
446  | 
(* In fact we have to prove *)  | 
|
447  | 
(* * WO6 ==> NN(y) \<noteq> 0 *)  | 
|
448  | 
(* * reverse induction which lets us infer that 1 \<in> NN(y) *)  | 
|
449  | 
(* * 1 \<in> NN(y) ==> y can be well-ordered *)  | 
|
450  | 
(* ********************************************************************** *)  | 
|
451  | 
||
452  | 
(* ********************************************************************** *)  | 
|
453  | 
(* WO6 ==> NN(y) \<noteq> 0 *)  | 
|
454  | 
(* ********************************************************************** *)  | 
|
455  | 
||
456  | 
lemma WO6_imp_NN_not_empty: "WO6 ==> NN(y) \<noteq> 0"  | 
|
| 12820 | 457  | 
by (unfold WO6_def NN_def, clarify, blast)  | 
| 12776 | 458  | 
|
459  | 
(* ********************************************************************** *)  | 
|
460  | 
(* 1 \<in> NN(y) ==> y can be well-ordered *)  | 
|
461  | 
(* ********************************************************************** *)  | 
|
462  | 
||
463  | 
lemma lemma1:  | 
|
464  | 
     "[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |] ==> \<exists>c<a. f`c = {x}"
 | 
|
465  | 
by (fast elim!: lepoll_1_is_sing)  | 
|
466  | 
||
467  | 
lemma lemma2:  | 
|
468  | 
"[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |]  | 
|
| 61394 | 469  | 
      ==> f` (\<mu> i. f`i = {x}) = {x}"
 | 
| 12776 | 470  | 
apply (drule lemma1, assumption+)  | 
471  | 
apply (fast elim!: lt_Ord intro: LeastI)  | 
|
472  | 
done  | 
|
| 1019 | 473  | 
|
| 12776 | 474  | 
lemma NN_imp_ex_inj: "1 \<in> NN(y) ==> \<exists>a f. Ord(a) & f \<in> inj(y, a)"  | 
475  | 
apply (unfold NN_def)  | 
|
476  | 
apply (elim CollectE exE conjE)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
13175 
diff
changeset
 | 
477  | 
apply (rule_tac x = a in exI)  | 
| 61394 | 478  | 
apply (rule_tac x = "\<lambda>x \<in> y. \<mu> i. f`i = {x}" in exI)
 | 
| 12776 | 479  | 
apply (rule conjI, assumption)  | 
480  | 
apply (rule_tac d = "%i. THE x. x \<in> f`i" in lam_injective)  | 
|
481  | 
apply (drule lemma1, assumption+)  | 
|
482  | 
apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord)  | 
|
483  | 
apply (rule lemma2 [THEN ssubst], assumption+, blast)  | 
|
484  | 
done  | 
|
485  | 
||
486  | 
lemma y_well_ord: "[| y*y \<subseteq> y; 1 \<in> NN(y) |] ==> \<exists>r. well_ord(y, r)"  | 
|
487  | 
apply (drule NN_imp_ex_inj)  | 
|
488  | 
apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel])  | 
|
489  | 
done  | 
|
490  | 
||
491  | 
(* ********************************************************************** *)  | 
|
492  | 
(* reverse induction which lets us infer that 1 \<in> NN(y) *)  | 
|
493  | 
(* ********************************************************************** *)  | 
|
494  | 
||
495  | 
lemma rev_induct_lemma [rule_format]:  | 
|
496  | 
"[| n \<in> nat; !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
46471 
diff
changeset
 | 
497  | 
==> n\<noteq>0 \<longrightarrow> P(n) \<longrightarrow> P(1)"  | 
| 12776 | 498  | 
by (erule nat_induct, blast+)  | 
499  | 
||
500  | 
lemma rev_induct:  | 
|
501  | 
"[| n \<in> nat; P(n); n\<noteq>0;  | 
|
502  | 
!!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |]  | 
|
503  | 
==> P(1)"  | 
|
504  | 
by (rule rev_induct_lemma, blast+)  | 
|
| 1019 | 505  | 
|
| 12776 | 506  | 
lemma NN_into_nat: "n \<in> NN(y) ==> n \<in> nat"  | 
507  | 
by (simp add: NN_def)  | 
|
508  | 
||
509  | 
lemma lemma3: "[| n \<in> NN(y); y*y \<subseteq> y; n\<noteq>0 |] ==> 1 \<in> NN(y)"  | 
|
510  | 
apply (rule rev_induct [OF NN_into_nat], assumption+)  | 
|
511  | 
apply (rule lemma_ii, assumption+)  | 
|
512  | 
done  | 
|
513  | 
||
514  | 
(* ********************************************************************** *)  | 
|
515  | 
(* Main theorem "WO6 ==> WO1" *)  | 
|
516  | 
(* ********************************************************************** *)  | 
|
517  | 
||
518  | 
(* another helpful lemma *)  | 
|
519  | 
lemma NN_y_0: "0 \<in> NN(y) ==> y=0"  | 
|
520  | 
apply (unfold NN_def)  | 
|
521  | 
apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst)  | 
|
522  | 
done  | 
|
523  | 
||
524  | 
lemma WO6_imp_WO1: "WO6 ==> WO1"  | 
|
525  | 
apply (unfold WO1_def)  | 
|
526  | 
apply (rule allI)  | 
|
527  | 
apply (case_tac "A=0")  | 
|
528  | 
apply (fast intro!: well_ord_Memrel nat_0I [THEN nat_into_Ord])  | 
|
| 46471 | 529  | 
apply (rule_tac x = A in lemma_iv [elim_format])  | 
| 12776 | 530  | 
apply (erule exE)  | 
531  | 
apply (drule WO6_imp_NN_not_empty)  | 
|
532  | 
apply (erule Un_subset_iff [THEN iffD1, THEN conjE])  | 
|
533  | 
apply (erule_tac A = "NN (y) " in not_emptyE)  | 
|
534  | 
apply (frule y_well_ord)  | 
|
535  | 
apply (fast intro!: lemma3 dest!: NN_y_0 elim!: not_emptyE)  | 
|
536  | 
apply (fast elim: well_ord_subset)  | 
|
537  | 
done  | 
|
538  | 
||
| 1019 | 539  | 
end  |