author | wenzelm |
Wed, 28 Oct 2009 18:02:06 +0100 | |
changeset 33284 | b5347c65bcab |
parent 32960 | 69916a850301 |
child 40945 | b8703f63bfb2 |
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). |
|
9 |
They refer reader to a book by Gödel to see the proof WO1 ==> WO2. |
|
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 |
|
12776 | 25 |
"uu(f, beta, gamma, delta) == (f`beta * f`gamma) Int f`delta" |
26 |
||
24893 | 27 |
|
28 |
(** Definitions for case 1 **) |
|
29 |
definition |
|
30 |
vv1 :: "[i, i, i] => i" where |
|
12776 | 31 |
"vv1(f,m,b) == |
32 |
let g = LEAST g. (\<exists>d. Ord(d) & (domain(uu(f,b,g,d)) \<noteq> 0 & |
|
33 |
domain(uu(f,b,g,d)) \<lesssim> m)); |
|
34 |
d = LEAST d. domain(uu(f,b,g,d)) \<noteq> 0 & |
|
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) == |
51 |
if f`g \<noteq> 0 then {uu(f, b, g, LEAST d. uu(f,b,g,d) \<noteq> 0)`s} else 0" |
|
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 |
|
174 |
==> (\<forall>b<a. f`b \<noteq> 0 --> |
|
175 |
(\<exists>g<a. \<exists>d<a. u(f,b,g,d) \<noteq> 0 & u(f,b,g,d) \<prec> m)) |
|
176 |
| (\<exists>b<a. f`b \<noteq> 0 & (\<forall>g<a. \<forall>d<a. u(f,b,g,d) \<noteq> 0 --> |
|
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 |
(* ********************************************************************** *) |
|
185 |
lemma UN_oadd: "Ord(a) ==> (\<Union>b<a++a. C(b)) = (\<Union>b<a. C(b) Un C(a++b))" |
|
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); |
|
213 |
Least_a = (LEAST a. \<exists>x. Ord(x) & P(a, x)) |] |
|
214 |
==> P(Least_a, LEAST b. P(Least_a, b))" |
|
215 |
apply (erule ssubst) |
|
216 |
apply (rule_tac Q = "%z. P (z, LEAST b. P (z, b))" in LeastI2) |
|
217 |
apply (fast elim!: LeastI)+ |
|
218 |
done |
|
219 |
||
220 |
lemmas nested_Least_instance = |
|
221 |
nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \<noteq> 0 & |
|
222 |
domain(uu(f,b,g,d)) \<lesssim> m", |
|
223 |
standard] (*generalizes the Variables!*) |
|
1019 | 224 |
|
12776 | 225 |
lemma gg1_lepoll_m: |
226 |
"[| Ord(a); m \<in> nat; |
|
227 |
\<forall>b<a. f`b \<noteq>0 --> |
|
228 |
(\<exists>g<a. \<exists>d<a. domain(uu(f,b,g,d)) \<noteq> 0 & |
|
229 |
domain(uu(f,b,g,d)) \<lesssim> m); |
|
230 |
\<forall>b<a. f`b \<lesssim> succ(m); b<a++a |] |
|
231 |
==> gg1(f,a,m)`b \<lesssim> m" |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
232 |
apply (simp add: gg1_def empty_lepollI) |
12776 | 233 |
apply (safe dest!: lt_oadd_odiff_disj) |
234 |
(*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
|
235 |
apply (simp add: vv1_def Let_def empty_lepollI) |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
236 |
apply (fast intro: nested_Least_instance [THEN conjunct2] |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
237 |
elim!: lt_Ord) |
12776 | 238 |
(*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
|
239 |
apply (simp add: ww1_def empty_lepollI) |
12776 | 240 |
apply (case_tac "f` (b--a) = 0", simp add: empty_lepollI) |
241 |
apply (rule Diff_lepoll, blast) |
|
242 |
apply (rule vv1_subset) |
|
243 |
apply (drule ospec [THEN mp], assumption+) |
|
244 |
apply (elim oexE conjE) |
|
245 |
apply (simp add: vv1_def Let_def lt_Ord nested_Least_instance [THEN conjunct1]) |
|
246 |
done |
|
247 |
||
248 |
(* ********************************************************************** *) |
|
249 |
(* Case 2: lemmas *) |
|
250 |
(* ********************************************************************** *) |
|
1042 | 251 |
|
12776 | 252 |
(* ********************************************************************** *) |
253 |
(* Case 2: vv2_subset *) |
|
254 |
(* ********************************************************************** *) |
|
255 |
||
256 |
lemma ex_d_uu_not_empty: |
|
257 |
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; |
|
258 |
y*y \<subseteq> y; (\<Union>b<a. f`b)=y |] |
|
259 |
==> \<exists>d<a. uu(f,b,g,d) \<noteq> 0" |
|
260 |
by (unfold uu_def, blast) |
|
261 |
||
262 |
lemma uu_not_empty: |
|
263 |
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |] |
|
264 |
==> uu(f,b,g,LEAST d. (uu(f,b,g,d) \<noteq> 0)) \<noteq> 0" |
|
265 |
apply (drule ex_d_uu_not_empty, assumption+) |
|
266 |
apply (fast elim!: LeastI lt_Ord) |
|
267 |
done |
|
268 |
||
269 |
lemma not_empty_rel_imp_domain: "[| r \<subseteq> A*B; r\<noteq>0 |] ==> domain(r)\<noteq>0" |
|
270 |
by blast |
|
271 |
||
272 |
lemma Least_uu_not_empty_lt_a: |
|
273 |
"[| b<a; g<a; f`b\<noteq>0; f`g\<noteq>0; y*y \<subseteq> y; (\<Union>b<a. f`b)=y |] |
|
274 |
==> (LEAST d. uu(f,b,g,d) \<noteq> 0) < a" |
|
275 |
apply (erule ex_d_uu_not_empty [THEN oexE], assumption+) |
|
276 |
apply (blast intro: Least_le [THEN lt_trans1] lt_Ord) |
|
277 |
done |
|
278 |
||
279 |
lemma subset_Diff_sing: "[| B \<subseteq> A; a\<notin>B |] ==> B \<subseteq> A-{a}" |
|
280 |
by blast |
|
1042 | 281 |
|
12776 | 282 |
(*Could this be proved more directly?*) |
283 |
lemma supset_lepoll_imp_eq: |
|
284 |
"[| A \<lesssim> m; m \<lesssim> B; B \<subseteq> A; m \<in> nat |] ==> A=B" |
|
285 |
apply (erule natE) |
|
286 |
apply (fast dest!: lepoll_0_is_0 intro!: equalityI) |
|
287 |
apply (safe intro!: equalityI) |
|
288 |
apply (rule ccontr) |
|
289 |
apply (rule succ_lepoll_natE) |
|
290 |
apply (erule lepoll_trans) |
|
291 |
apply (rule lepoll_trans) |
|
292 |
apply (erule subset_Diff_sing [THEN subset_imp_lepoll], assumption) |
|
293 |
apply (rule Diff_sing_lepoll, assumption+) |
|
294 |
done |
|
295 |
||
296 |
lemma uu_Least_is_fun: |
|
297 |
"[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 --> |
|
298 |
domain(uu(f, b, g, d)) \<approx> succ(m); |
|
299 |
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y; |
|
300 |
(\<Union>b<a. f`b)=y; b<a; g<a; d<a; |
|
301 |
f`b\<noteq>0; f`g\<noteq>0; m \<in> nat; s \<in> f`b |] |
|
302 |
==> uu(f, b, g, LEAST d. uu(f,b,g,d)\<noteq>0) \<in> f`b -> f`g" |
|
303 |
apply (drule_tac x2=g in ospec [THEN ospec, THEN mp]) |
|
304 |
apply (rule_tac [3] not_empty_rel_imp_domain [OF uu_subset1 uu_not_empty]) |
|
305 |
apply (rule_tac [2] Least_uu_not_empty_lt_a, assumption+) |
|
306 |
apply (rule rel_is_fun) |
|
307 |
apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) |
|
308 |
apply (erule uu_lepoll_m) |
|
309 |
apply (rule Least_uu_not_empty_lt_a, assumption+) |
|
310 |
apply (rule uu_subset1) |
|
311 |
apply (rule supset_lepoll_imp_eq [OF _ eqpoll_sym [THEN eqpoll_imp_lepoll]]) |
|
312 |
apply (fast intro!: domain_uu_subset)+ |
|
313 |
done |
|
314 |
||
315 |
lemma vv2_subset: |
|
316 |
"[| \<forall>g<a. \<forall>d<a. domain(uu(f, b, g, d))\<noteq>0 --> |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
317 |
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
|
318 |
\<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
|
319 |
(\<Union>b<a. f`b)=y; b<a; g<a; m \<in> nat; s \<in> f`b |] |
12776 | 320 |
==> vv2(f,b,g,s) \<subseteq> f`g" |
321 |
apply (simp add: vv2_def) |
|
322 |
apply (blast intro: uu_Least_is_fun [THEN apply_type]) |
|
323 |
done |
|
324 |
||
325 |
(* ********************************************************************** *) |
|
326 |
(* Case 2: Union of images is the whole "y" *) |
|
327 |
(* ********************************************************************** *) |
|
328 |
lemma UN_gg2_eq: |
|
329 |
"[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 --> |
|
330 |
domain(uu(f,b,g,d)) \<approx> succ(m); |
|
331 |
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y; |
|
332 |
(\<Union>b<a. f`b)=y; Ord(a); m \<in> nat; s \<in> f`b; b<a |] |
|
333 |
==> (\<Union>g<a++a. gg2(f,a,b,s) ` g) = y" |
|
334 |
apply (unfold gg2_def) |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
335 |
apply (drule sym) |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
336 |
apply (simp add: ltD UN_oadd oadd_le_self [THEN le_imp_not_lt] |
12776 | 337 |
lt_Ord odiff_oadd_inverse ww2_def |
338 |
vv2_subset [THEN Diff_partition]) |
|
339 |
done |
|
340 |
||
341 |
lemma domain_gg2: "domain(gg2(f,a,b,s)) = a++a" |
|
342 |
by (simp add: lam_funtype [THEN domain_of_fun] gg2_def) |
|
343 |
||
1042 | 344 |
|
12776 | 345 |
(* ********************************************************************** *) |
346 |
(* every value of defined function is less than or equipollent to m *) |
|
347 |
(* ********************************************************************** *) |
|
348 |
||
349 |
lemma vv2_lepoll: "[| m \<in> nat; m\<noteq>0 |] ==> vv2(f,b,g,s) \<lesssim> m" |
|
350 |
apply (unfold vv2_def) |
|
351 |
apply (simp add: empty_lepollI) |
|
352 |
apply (fast dest!: le_imp_subset [THEN subset_imp_lepoll, THEN lepoll_0_is_0] |
|
353 |
intro!: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll, THEN lepoll_trans] |
|
354 |
not_lt_imp_le [THEN le_imp_subset, THEN subset_imp_lepoll] |
|
355 |
nat_into_Ord nat_1I) |
|
356 |
done |
|
357 |
||
358 |
lemma ww2_lepoll: |
|
359 |
"[| \<forall>b<a. f`b \<lesssim> succ(m); g<a; m \<in> nat; vv2(f,b,g,d) \<subseteq> f`g |] |
|
360 |
==> ww2(f,b,g,d) \<lesssim> m" |
|
361 |
apply (unfold ww2_def) |
|
362 |
apply (case_tac "f`g = 0") |
|
363 |
apply (simp add: empty_lepollI) |
|
364 |
apply (drule ospec, assumption) |
|
365 |
apply (rule Diff_lepoll, assumption+) |
|
366 |
apply (simp add: vv2_def not_emptyI) |
|
367 |
done |
|
368 |
||
369 |
lemma gg2_lepoll_m: |
|
370 |
"[| \<forall>g<a. \<forall>d<a. domain(uu(f,b,g,d)) \<noteq> 0 --> |
|
371 |
domain(uu(f,b,g,d)) \<approx> succ(m); |
|
372 |
\<forall>b<a. f`b \<lesssim> succ(m); y*y \<subseteq> y; |
|
373 |
(\<Union>b<a. f`b)=y; b<a; s \<in> f`b; m \<in> nat; m\<noteq> 0; g<a++a |] |
|
374 |
==> gg2(f,a,b,s) ` g \<lesssim> m" |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12820
diff
changeset
|
375 |
apply (simp add: gg2_def empty_lepollI) |
12776 | 376 |
apply (safe elim!: lt_Ord2 dest!: lt_oadd_odiff_disj) |
377 |
apply (simp add: vv2_lepoll) |
|
378 |
apply (simp add: ww2_lepoll vv2_subset) |
|
379 |
done |
|
380 |
||
381 |
(* ********************************************************************** *) |
|
382 |
(* lemma ii *) |
|
383 |
(* ********************************************************************** *) |
|
384 |
lemma lemma_ii: "[| succ(m) \<in> NN(y); y*y \<subseteq> y; m \<in> nat; m\<noteq>0 |] ==> m \<in> NN(y)" |
|
385 |
apply (unfold NN_def) |
|
386 |
apply (elim CollectE exE conjE) |
|
387 |
apply (rule quant_domain_uu_lepoll_m [THEN cases, THEN disjE], assumption) |
|
388 |
(* case 1 *) |
|
389 |
apply (simp add: lesspoll_succ_iff) |
|
390 |
apply (rule_tac x = "a++a" in exI) |
|
391 |
apply (fast intro!: Ord_oadd domain_gg1 UN_gg1_eq gg1_lepoll_m) |
|
392 |
(* case 2 *) |
|
393 |
apply (elim oexE conjE) |
|
394 |
apply (rule_tac A = "f`?B" in not_emptyE, assumption) |
|
395 |
apply (rule CollectI) |
|
396 |
apply (erule succ_natD) |
|
397 |
apply (rule_tac x = "a++a" in exI) |
|
398 |
apply (rule_tac x = "gg2 (f,a,b,x) " in exI) |
|
399 |
apply (simp add: Ord_oadd domain_gg2 UN_gg2_eq gg2_lepoll_m) |
|
400 |
done |
|
401 |
||
402 |
||
403 |
(* ********************************************************************** *) |
|
404 |
(* lemma iv - p. 4: *) |
|
405 |
(* For every set x there is a set y such that x Un (y * y) \<subseteq> y *) |
|
406 |
(* ********************************************************************** *) |
|
407 |
||
408 |
||
409 |
(* The leading \<forall>-quantifier looks odd but makes the proofs shorter |
|
410 |
(used only in the following two lemmas) *) |
|
1042 | 411 |
|
12776 | 412 |
lemma z_n_subset_z_succ_n: |
413 |
"\<forall>n \<in> nat. rec(n, x, %k r. r Un r*r) \<subseteq> rec(succ(n), x, %k r. r Un r*r)" |
|
414 |
by (fast intro: rec_succ [THEN ssubst]) |
|
415 |
||
416 |
lemma le_subsets: |
|
417 |
"[| \<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat |] |
|
418 |
==> f(n)<=f(m)" |
|
419 |
apply (erule_tac P = "n\<le>m" in rev_mp) |
|
420 |
apply (rule_tac P = "%z. n\<le>z --> f (n) \<subseteq> f (z) " in nat_induct) |
|
421 |
apply (auto simp add: le_iff) |
|
422 |
done |
|
423 |
||
424 |
lemma le_imp_rec_subset: |
|
425 |
"[| n\<le>m; m \<in> nat |] |
|
426 |
==> rec(n, x, %k r. r Un r*r) \<subseteq> rec(m, x, %k r. r Un r*r)" |
|
427 |
apply (rule z_n_subset_z_succ_n [THEN le_subsets]) |
|
428 |
apply (blast intro: lt_nat_in_nat)+ |
|
429 |
done |
|
430 |
||
431 |
lemma lemma_iv: "\<exists>y. x Un y*y \<subseteq> y" |
|
432 |
apply (rule_tac x = "\<Union>n \<in> nat. rec (n, x, %k r. r Un r*r) " in exI) |
|
433 |
apply safe |
|
434 |
apply (rule nat_0I [THEN UN_I], simp) |
|
435 |
apply (rule_tac a = "succ (n Un na) " in UN_I) |
|
436 |
apply (erule Un_nat_type [THEN nat_succI], assumption) |
|
437 |
apply (auto intro: le_imp_rec_subset [THEN subsetD] |
|
438 |
intro!: Un_upper1_le Un_upper2_le Un_nat_type |
|
439 |
elim!: nat_into_Ord) |
|
440 |
done |
|
1019 | 441 |
|
12776 | 442 |
(* ********************************************************************** *) |
443 |
(* Rubin & Rubin wrote, *) |
|
444 |
(* "It follows from (ii) and mathematical induction that if y*y \<subseteq> y then *) |
|
445 |
(* y can be well-ordered" *) |
|
446 |
||
447 |
(* In fact we have to prove *) |
|
448 |
(* * WO6 ==> NN(y) \<noteq> 0 *) |
|
449 |
(* * reverse induction which lets us infer that 1 \<in> NN(y) *) |
|
450 |
(* * 1 \<in> NN(y) ==> y can be well-ordered *) |
|
451 |
(* ********************************************************************** *) |
|
452 |
||
453 |
(* ********************************************************************** *) |
|
454 |
(* WO6 ==> NN(y) \<noteq> 0 *) |
|
455 |
(* ********************************************************************** *) |
|
456 |
||
457 |
lemma WO6_imp_NN_not_empty: "WO6 ==> NN(y) \<noteq> 0" |
|
12820 | 458 |
by (unfold WO6_def NN_def, clarify, blast) |
12776 | 459 |
|
460 |
(* ********************************************************************** *) |
|
461 |
(* 1 \<in> NN(y) ==> y can be well-ordered *) |
|
462 |
(* ********************************************************************** *) |
|
463 |
||
464 |
lemma lemma1: |
|
465 |
"[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |] ==> \<exists>c<a. f`c = {x}" |
|
466 |
by (fast elim!: lepoll_1_is_sing) |
|
467 |
||
468 |
lemma lemma2: |
|
469 |
"[| (\<Union>b<a. f`b)=y; x \<in> y; \<forall>b<a. f`b \<lesssim> 1; Ord(a) |] |
|
470 |
==> f` (LEAST i. f`i = {x}) = {x}" |
|
471 |
apply (drule lemma1, assumption+) |
|
472 |
apply (fast elim!: lt_Ord intro: LeastI) |
|
473 |
done |
|
1019 | 474 |
|
12776 | 475 |
lemma NN_imp_ex_inj: "1 \<in> NN(y) ==> \<exists>a f. Ord(a) & f \<in> inj(y, a)" |
476 |
apply (unfold NN_def) |
|
477 |
apply (elim CollectE exE conjE) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
478 |
apply (rule_tac x = a in exI) |
12776 | 479 |
apply (rule_tac x = "\<lambda>x \<in> y. LEAST i. f`i = {x}" in exI) |
480 |
apply (rule conjI, assumption) |
|
481 |
apply (rule_tac d = "%i. THE x. x \<in> f`i" in lam_injective) |
|
482 |
apply (drule lemma1, assumption+) |
|
483 |
apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord) |
|
484 |
apply (rule lemma2 [THEN ssubst], assumption+, blast) |
|
485 |
done |
|
486 |
||
487 |
lemma y_well_ord: "[| y*y \<subseteq> y; 1 \<in> NN(y) |] ==> \<exists>r. well_ord(y, r)" |
|
488 |
apply (drule NN_imp_ex_inj) |
|
489 |
apply (fast elim!: well_ord_rvimage [OF _ well_ord_Memrel]) |
|
490 |
done |
|
491 |
||
492 |
(* ********************************************************************** *) |
|
493 |
(* reverse induction which lets us infer that 1 \<in> NN(y) *) |
|
494 |
(* ********************************************************************** *) |
|
495 |
||
496 |
lemma rev_induct_lemma [rule_format]: |
|
497 |
"[| n \<in> nat; !!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |] |
|
498 |
==> n\<noteq>0 --> P(n) --> P(1)" |
|
499 |
by (erule nat_induct, blast+) |
|
500 |
||
501 |
lemma rev_induct: |
|
502 |
"[| n \<in> nat; P(n); n\<noteq>0; |
|
503 |
!!m. [| m \<in> nat; m\<noteq>0; P(succ(m)) |] ==> P(m) |] |
|
504 |
==> P(1)" |
|
505 |
by (rule rev_induct_lemma, blast+) |
|
1019 | 506 |
|
12776 | 507 |
lemma NN_into_nat: "n \<in> NN(y) ==> n \<in> nat" |
508 |
by (simp add: NN_def) |
|
509 |
||
510 |
lemma lemma3: "[| n \<in> NN(y); y*y \<subseteq> y; n\<noteq>0 |] ==> 1 \<in> NN(y)" |
|
511 |
apply (rule rev_induct [OF NN_into_nat], assumption+) |
|
512 |
apply (rule lemma_ii, assumption+) |
|
513 |
done |
|
514 |
||
515 |
(* ********************************************************************** *) |
|
516 |
(* Main theorem "WO6 ==> WO1" *) |
|
517 |
(* ********************************************************************** *) |
|
518 |
||
519 |
(* another helpful lemma *) |
|
520 |
lemma NN_y_0: "0 \<in> NN(y) ==> y=0" |
|
521 |
apply (unfold NN_def) |
|
522 |
apply (fast intro!: equalityI dest!: lepoll_0_is_0 elim: subst) |
|
523 |
done |
|
524 |
||
525 |
lemma WO6_imp_WO1: "WO6 ==> WO1" |
|
526 |
apply (unfold WO1_def) |
|
527 |
apply (rule allI) |
|
528 |
apply (case_tac "A=0") |
|
529 |
apply (fast intro!: well_ord_Memrel nat_0I [THEN nat_into_Ord]) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
530 |
apply (rule_tac x1 = A in lemma_iv [THEN revcut_rl]) |
12776 | 531 |
apply (erule exE) |
532 |
apply (drule WO6_imp_NN_not_empty) |
|
533 |
apply (erule Un_subset_iff [THEN iffD1, THEN conjE]) |
|
534 |
apply (erule_tac A = "NN (y) " in not_emptyE) |
|
535 |
apply (frule y_well_ord) |
|
536 |
apply (fast intro!: lemma3 dest!: NN_y_0 elim!: not_emptyE) |
|
537 |
apply (fast elim: well_ord_subset) |
|
538 |
done |
|
539 |
||
1019 | 540 |
end |