author | haftmann |
Thu, 22 Jul 2010 17:26:22 +0200 | |
changeset 37934 | 440114da2488 |
parent 32960 | 69916a850301 |
child 45602 | 2a858377c3d2 |
permissions | -rw-r--r-- |
12776 | 1 |
(* Title: ZF/AC/WO2_AC16.thy |
2 |
Author: Krzysztof Grabczewski |
|
3 |
||
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
4 |
The proof of WO2 ==> AC16(k #+ m, k) |
12776 | 5 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
6 |
The main part of the proof is the inductive reasoning concerning |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
7 |
properties of constructed family T_gamma. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
8 |
The proof deals with three cases for ordinals: 0, succ and limit ordinal. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
9 |
The first instance is trivial, the third not difficult, but the second |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
10 |
is very complicated requiring many lemmas. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
11 |
We also need to prove that at any stage gamma the set |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
12 |
(s - Union(...) - k_gamma) (Rubin & Rubin page 15) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
13 |
contains m distinct elements (in fact is equipollent to s) |
12776 | 14 |
*) |
15 |
||
16417 | 16 |
theory WO2_AC16 imports AC_Equiv AC16_lemmas Cardinal_aux begin |
12776 | 17 |
|
18 |
(**** A recursive definition used in the proof of WO2 ==> AC16 ****) |
|
19 |
||
24893 | 20 |
definition |
21 |
recfunAC16 :: "[i,i,i,i] => i" where |
|
12776 | 22 |
"recfunAC16(f,h,i,a) == |
23 |
transrec2(i, 0, |
|
24 |
%g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r |
|
25 |
else r Un {f`(LEAST i. h`g \<subseteq> f`i & |
|
26 |
(\<forall>b<a. (h`b \<subseteq> f`i --> (\<forall>t \<in> r. ~ h`b \<subseteq> t))))})" |
|
27 |
||
28 |
(* ********************************************************************** *) |
|
29 |
(* Basic properties of recfunAC16 *) |
|
30 |
(* ********************************************************************** *) |
|
31 |
||
32 |
lemma recfunAC16_0: "recfunAC16(f,h,0,a) = 0" |
|
12820 | 33 |
by (simp add: recfunAC16_def) |
12776 | 34 |
|
35 |
lemma recfunAC16_succ: |
|
36 |
"recfunAC16(f,h,succ(i),a) = |
|
37 |
(if (\<exists>y \<in> recfunAC16(f,h,i,a). h ` i \<subseteq> y) then recfunAC16(f,h,i,a) |
|
38 |
else recfunAC16(f,h,i,a) Un |
|
39 |
{f ` (LEAST j. h ` i \<subseteq> f ` j & |
|
40 |
(\<forall>b<a. (h`b \<subseteq> f`j |
|
41 |
--> (\<forall>t \<in> recfunAC16(f,h,i,a). ~ h`b \<subseteq> t))))})" |
|
12820 | 42 |
apply (simp add: recfunAC16_def) |
12776 | 43 |
done |
44 |
||
45 |
lemma recfunAC16_Limit: "Limit(i) |
|
46 |
==> recfunAC16(f,h,i,a) = (\<Union>j<i. recfunAC16(f,h,j,a))" |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
47 |
by (simp add: recfunAC16_def transrec2_Limit) |
12776 | 48 |
|
49 |
(* ********************************************************************** *) |
|
50 |
(* Monotonicity of transrec2 *) |
|
51 |
(* ********************************************************************** *) |
|
52 |
||
53 |
lemma transrec2_mono_lemma [rule_format]: |
|
54 |
"[| !!g r. r \<subseteq> B(g,r); Ord(i) |] |
|
55 |
==> j<i --> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)" |
|
12820 | 56 |
apply (erule trans_induct) |
57 |
apply (rule Ord_cases, assumption+, fast) |
|
12776 | 58 |
apply (simp (no_asm_simp)) |
12820 | 59 |
apply (blast elim!: leE) |
60 |
apply (simp add: transrec2_Limit) |
|
12776 | 61 |
apply (blast intro: OUN_I ltI Ord_in_Ord [THEN le_refl] |
62 |
elim!: Limit_has_succ [THEN ltE]) |
|
63 |
done |
|
64 |
||
65 |
lemma transrec2_mono: |
|
66 |
"[| !!g r. r \<subseteq> B(g,r); j\<le>i |] |
|
67 |
==> transrec2(j, 0, B) \<subseteq> transrec2(i, 0, B)" |
|
12820 | 68 |
apply (erule leE) |
12776 | 69 |
apply (rule transrec2_mono_lemma) |
12820 | 70 |
apply (auto intro: lt_Ord2 ) |
12776 | 71 |
done |
72 |
||
73 |
(* ********************************************************************** *) |
|
74 |
(* Monotonicity of recfunAC16 *) |
|
75 |
(* ********************************************************************** *) |
|
76 |
||
77 |
lemma recfunAC16_mono: |
|
78 |
"i\<le>j ==> recfunAC16(f, g, i, a) \<subseteq> recfunAC16(f, g, j, a)" |
|
79 |
apply (unfold recfunAC16_def) |
|
12820 | 80 |
apply (rule transrec2_mono, auto) |
12776 | 81 |
done |
82 |
||
83 |
(* ********************************************************************** *) |
|
84 |
(* case of limit ordinal *) |
|
85 |
(* ********************************************************************** *) |
|
86 |
||
87 |
||
88 |
lemma lemma3_1: |
|
89 |
"[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) --> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y); |
|
90 |
\<forall>i j. i\<le>j --> F(i) \<subseteq> F(j); j\<le>i; i<x; z<a; |
|
91 |
V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |] |
|
92 |
==> V = W" |
|
93 |
apply (erule asm_rl allE impE)+ |
|
12820 | 94 |
apply (drule subsetD, assumption, blast) |
12776 | 95 |
done |
96 |
||
97 |
||
98 |
lemma lemma3: |
|
99 |
"[| \<forall>y<x. \<forall>z<a. z<y | (\<exists>Y \<in> F(y). f(z)<=Y) --> (\<exists>! Y. Y \<in> F(y) & f(z)<=Y); |
|
100 |
\<forall>i j. i\<le>j --> F(i) \<subseteq> F(j); i<x; j<x; z<a; |
|
101 |
V \<in> F(i); f(z)<=V; W \<in> F(j); f(z)<=W |] |
|
102 |
==> V = W" |
|
103 |
apply (rule_tac j=j in Ord_linear_le [OF lt_Ord lt_Ord], assumption+) |
|
104 |
apply (erule lemma3_1 [symmetric], assumption+) |
|
105 |
apply (erule lemma3_1, assumption+) |
|
106 |
done |
|
107 |
||
108 |
lemma lemma4: |
|
109 |
"[| \<forall>y<x. F(y) \<subseteq> X & |
|
110 |
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) --> |
|
111 |
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y)); |
|
112 |
x < a |] |
|
113 |
==> \<forall>y<x. \<forall>z<a. z < y | (\<exists>Y \<in> F(y). h(z) \<subseteq> Y) --> |
|
114 |
(\<exists>! Y. Y \<in> F(y) & h(z) \<subseteq> Y)" |
|
115 |
apply (intro oallI impI) |
|
12820 | 116 |
apply (drule ospec, assumption, clarify) |
12776 | 117 |
apply (blast elim!: oallE ) |
118 |
done |
|
119 |
||
120 |
lemma lemma5: |
|
121 |
"[| \<forall>y<x. F(y) \<subseteq> X & |
|
122 |
(\<forall>x<a. x < y | (\<exists>Y \<in> F(y). h(x) \<subseteq> Y) --> |
|
123 |
(\<exists>! Y. Y \<in> F(y) & h(x) \<subseteq> Y)); |
|
124 |
x < a; Limit(x); \<forall>i j. i\<le>j --> F(i) \<subseteq> F(j) |] |
|
125 |
==> (\<Union>x<x. F(x)) \<subseteq> X & |
|
126 |
(\<forall>xa<a. xa < x | (\<exists>x \<in> \<Union>x<x. F(x). h(xa) \<subseteq> x) |
|
127 |
--> (\<exists>! Y. Y \<in> (\<Union>x<x. F(x)) & h(xa) \<subseteq> Y))" |
|
128 |
apply (rule conjI) |
|
129 |
apply (rule subsetI) |
|
130 |
apply (erule OUN_E) |
|
12820 | 131 |
apply (drule ospec, assumption, fast) |
12776 | 132 |
apply (drule lemma4, assumption) |
133 |
apply (rule oallI) |
|
134 |
apply (rule impI) |
|
135 |
apply (erule disjE) |
|
136 |
apply (frule ospec, erule Limit_has_succ, assumption) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
137 |
apply (drule_tac A = a and x = xa in ospec, assumption) |
12776 | 138 |
apply (erule impE, rule le_refl [THEN disjI1], erule lt_Ord) |
139 |
apply (blast intro: lemma3 Limit_has_succ) |
|
140 |
apply (blast intro: lemma3) |
|
141 |
done |
|
142 |
||
143 |
(* ********************************************************************** *) |
|
144 |
(* case of successor ordinal *) |
|
145 |
(* ********************************************************************** *) |
|
146 |
||
147 |
(* |
|
148 |
First quite complicated proof of the fact used in the recursive construction |
|
149 |
of the family T_gamma (WO2 ==> AC16(k #+ m, k)) - the fact that at any stage |
|
150 |
gamma the set (s - Union(...) - k_gamma) is equipollent to s |
|
151 |
(Rubin & Rubin page 15). |
|
152 |
*) |
|
153 |
||
154 |
(* ********************************************************************** *) |
|
155 |
(* dbl_Diff_eqpoll_Card *) |
|
156 |
(* ********************************************************************** *) |
|
157 |
||
158 |
||
159 |
lemma dbl_Diff_eqpoll_Card: |
|
160 |
"[| A\<approx>a; Card(a); ~Finite(a); B\<prec>a; C\<prec>a |] ==> A - B - C\<approx>a" |
|
161 |
by (blast intro: Diff_lesspoll_eqpoll_Card) |
|
162 |
||
163 |
(* ********************************************************************** *) |
|
164 |
(* Case of finite ordinals *) |
|
165 |
(* ********************************************************************** *) |
|
166 |
||
167 |
||
168 |
lemma Finite_lesspoll_infinite_Ord: |
|
169 |
"[| Finite(X); ~Finite(a); Ord(a) |] ==> X\<prec>a" |
|
170 |
apply (unfold lesspoll_def) |
|
171 |
apply (rule conjI) |
|
172 |
apply (drule nat_le_infinite_Ord [THEN le_imp_lepoll], assumption) |
|
173 |
apply (unfold Finite_def) |
|
174 |
apply (blast intro: leI [THEN le_imp_subset, THEN subset_imp_lepoll] |
|
175 |
ltI eqpoll_imp_lepoll lepoll_trans) |
|
176 |
apply (blast intro: eqpoll_sym [THEN eqpoll_trans]) |
|
177 |
done |
|
178 |
||
179 |
lemma Union_lesspoll: |
|
180 |
"[| \<forall>x \<in> X. x lepoll n & x \<subseteq> T; well_ord(T, R); X lepoll b; |
|
181 |
b<a; ~Finite(a); Card(a); n \<in> nat |] |
|
182 |
==> Union(X)\<prec>a" |
|
183 |
apply (case_tac "Finite (X)") |
|
184 |
apply (blast intro: Card_is_Ord Finite_lesspoll_infinite_Ord |
|
185 |
lepoll_nat_imp_Finite Finite_Union) |
|
186 |
apply (drule lepoll_imp_ex_le_eqpoll) |
|
187 |
apply (erule lt_Ord) |
|
188 |
apply (elim exE conjE) |
|
189 |
apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption) |
|
190 |
apply (erule eqpoll_sym [THEN eqpoll_def [THEN def_imp_iff, THEN iffD1], |
|
191 |
THEN exE]) |
|
192 |
apply (frule bij_is_surj [THEN surj_image_eq]) |
|
193 |
apply (drule image_fun [OF bij_is_fun subset_refl]) |
|
194 |
apply (drule sym [THEN trans], assumption) |
|
195 |
apply (blast intro: lt_Ord UN_lepoll lt_Card_imp_lesspoll |
|
196 |
lt_trans1 lesspoll_trans1) |
|
197 |
done |
|
198 |
||
199 |
(* ********************************************************************** *) |
|
200 |
(* recfunAC16_lepoll_index *) |
|
201 |
(* ********************************************************************** *) |
|
202 |
||
203 |
lemma Un_sing_eq_cons: "A Un {a} = cons(a, A)" |
|
204 |
by fast |
|
205 |
||
206 |
lemma Un_lepoll_succ: "A lepoll B ==> A Un {a} lepoll succ(B)" |
|
207 |
apply (simp add: Un_sing_eq_cons succ_def) |
|
208 |
apply (blast elim!: mem_irrefl intro: cons_lepoll_cong) |
|
209 |
done |
|
210 |
||
211 |
lemma Diff_UN_succ_empty: "Ord(a) ==> F(a) - (\<Union>b<succ(a). F(b)) = 0" |
|
212 |
by (fast intro!: le_refl) |
|
213 |
||
214 |
lemma Diff_UN_succ_subset: "Ord(a) ==> F(a) Un X - (\<Union>b<succ(a). F(b)) \<subseteq> X" |
|
215 |
by blast |
|
216 |
||
217 |
lemma recfunAC16_Diff_lepoll_1: |
|
218 |
"Ord(x) |
|
219 |
==> recfunAC16(f, g, x, a) - (\<Union>i<x. recfunAC16(f, g, i, a)) lepoll 1" |
|
220 |
apply (erule Ord_cases) |
|
221 |
apply (simp add: recfunAC16_0 empty_subsetI [THEN subset_imp_lepoll]) |
|
222 |
(*Limit case*) |
|
223 |
prefer 2 apply (simp add: recfunAC16_Limit Diff_cancel |
|
224 |
empty_subsetI [THEN subset_imp_lepoll]) |
|
225 |
(*succ case*) |
|
226 |
apply (simp add: recfunAC16_succ |
|
227 |
Diff_UN_succ_empty [of _ "%j. recfunAC16(f,g,j,a)"] |
|
228 |
empty_subsetI [THEN subset_imp_lepoll]) |
|
229 |
apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll] |
|
230 |
singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans) |
|
231 |
done |
|
232 |
||
233 |
lemma in_Least_Diff: |
|
234 |
"[| z \<in> F(x); Ord(x) |] |
|
235 |
==> z \<in> F(LEAST i. z \<in> F(i)) - (\<Union>j<(LEAST i. z \<in> F(i)). F(j))" |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
236 |
by (fast elim: less_LeastE elim!: LeastI) |
12776 | 237 |
|
238 |
lemma Least_eq_imp_ex: |
|
239 |
"[| (LEAST i. w \<in> F(i)) = (LEAST i. z \<in> F(i)); |
|
240 |
w \<in> (\<Union>i<a. F(i)); z \<in> (\<Union>i<a. F(i)) |] |
|
241 |
==> \<exists>b<a. w \<in> (F(b) - (\<Union>c<b. F(c))) & z \<in> (F(b) - (\<Union>c<b. F(c)))" |
|
242 |
apply (elim OUN_E) |
|
243 |
apply (drule in_Least_Diff, erule lt_Ord) |
|
244 |
apply (frule in_Least_Diff, erule lt_Ord) |
|
245 |
apply (rule oexI, force) |
|
246 |
apply (blast intro: lt_Ord Least_le [THEN lt_trans1]) |
|
247 |
done |
|
248 |
||
249 |
||
250 |
lemma two_in_lepoll_1: "[| A lepoll 1; a \<in> A; b \<in> A |] ==> a=b" |
|
251 |
by (fast dest!: lepoll_1_is_sing) |
|
252 |
||
253 |
||
254 |
lemma UN_lepoll_index: |
|
255 |
"[| \<forall>i<a. F(i)-(\<Union>j<i. F(j)) lepoll 1; Limit(a) |] |
|
256 |
==> (\<Union>x<a. F(x)) lepoll a" |
|
257 |
apply (rule lepoll_def [THEN def_imp_iff [THEN iffD2]]) |
|
258 |
apply (rule_tac x = "\<lambda>z \<in> (\<Union>x<a. F (x)). LEAST i. z \<in> F (i) " in exI) |
|
259 |
apply (unfold inj_def) |
|
260 |
apply (rule CollectI) |
|
261 |
apply (rule lam_type) |
|
262 |
apply (erule OUN_E) |
|
263 |
apply (erule Least_in_Ord) |
|
264 |
apply (erule ltD) |
|
265 |
apply (erule lt_Ord2) |
|
266 |
apply (intro ballI) |
|
267 |
apply (simp (no_asm_simp)) |
|
268 |
apply (rule impI) |
|
269 |
apply (drule Least_eq_imp_ex, assumption+) |
|
270 |
apply (fast elim!: two_in_lepoll_1) |
|
271 |
done |
|
272 |
||
273 |
||
274 |
lemma recfunAC16_lepoll_index: "Ord(y) ==> recfunAC16(f, h, y, a) lepoll y" |
|
275 |
apply (erule trans_induct3) |
|
276 |
(*0 case*) |
|
277 |
apply (simp (no_asm_simp) add: recfunAC16_0 lepoll_refl) |
|
278 |
(*succ case*) |
|
279 |
apply (simp (no_asm_simp) add: recfunAC16_succ) |
|
280 |
apply (blast dest!: succI1 [THEN rev_bspec] |
|
281 |
intro: subset_succI [THEN subset_imp_lepoll] Un_lepoll_succ |
|
282 |
lepoll_trans) |
|
283 |
apply (simp (no_asm_simp) add: recfunAC16_Limit) |
|
284 |
apply (blast intro: lt_Ord [THEN recfunAC16_Diff_lepoll_1] UN_lepoll_index) |
|
285 |
done |
|
286 |
||
1196 | 287 |
|
12776 | 288 |
lemma Union_recfunAC16_lesspoll: |
289 |
"[| recfunAC16(f,g,y,a) \<subseteq> {X \<in> Pow(A). X\<approx>n}; |
|
290 |
A\<approx>a; y<a; ~Finite(a); Card(a); n \<in> nat |] |
|
291 |
==> Union(recfunAC16(f,g,y,a))\<prec>a" |
|
292 |
apply (erule eqpoll_def [THEN def_imp_iff, THEN iffD1, THEN exE]) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
293 |
apply (rule_tac T=A in Union_lesspoll, simp_all) |
12776 | 294 |
apply (blast intro!: eqpoll_imp_lepoll) |
295 |
apply (blast intro: bij_is_inj Card_is_Ord [THEN well_ord_Memrel] |
|
296 |
well_ord_rvimage) |
|
297 |
apply (erule lt_Ord [THEN recfunAC16_lepoll_index]) |
|
298 |
done |
|
299 |
||
300 |
lemma dbl_Diff_eqpoll: |
|
301 |
"[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)}; |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
302 |
Card(a); ~ Finite(a); A\<approx>a; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
303 |
k \<in> nat; y<a; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
304 |
h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}) |] |
12776 | 305 |
==> A - Union(recfunAC16(f, h, y, a)) - h`y\<approx>a" |
306 |
apply (rule dbl_Diff_eqpoll_Card, simp_all) |
|
307 |
apply (simp add: Union_recfunAC16_lesspoll) |
|
308 |
apply (rule Finite_lesspoll_infinite_Ord) |
|
309 |
apply (rule Finite_def [THEN def_imp_iff, THEN iffD2]) |
|
12820 | 310 |
apply (blast dest: ltD bij_is_fun [THEN apply_type], assumption) |
12776 | 311 |
apply (blast intro: Card_is_Ord) |
312 |
done; |
|
313 |
||
314 |
(* back to the proof *) |
|
315 |
||
316 |
lemmas disj_Un_eqpoll_nat_sum = |
|
317 |
eqpoll_trans [THEN eqpoll_trans, |
|
318 |
OF disj_Un_eqpoll_sum sum_eqpoll_cong nat_sum_eqpoll_sum, |
|
319 |
standard]; |
|
320 |
||
321 |
||
322 |
lemma Un_in_Collect: "[| x \<in> Pow(A - B - h`i); x\<approx>m; |
|
323 |
h \<in> bij(a, {x \<in> Pow(A) . x\<approx>k}); i<a; k \<in> nat; m \<in> nat |] |
|
324 |
==> h ` i Un x \<in> {x \<in> Pow(A) . x\<approx>k #+ m}" |
|
325 |
by (blast intro: disj_Un_eqpoll_nat_sum |
|
326 |
dest: ltD bij_is_fun [THEN apply_type]) |
|
327 |
||
328 |
||
329 |
(* ********************************************************************** *) |
|
330 |
(* Lemmas simplifying assumptions *) |
|
331 |
(* ********************************************************************** *) |
|
332 |
||
333 |
lemma lemma6: |
|
334 |
"[| \<forall>y<succ(j). F(y)<=X & (\<forall>x<a. x<y | P(x,y) --> Q(x,y)); succ(j)<a |] |
|
335 |
==> F(j)<=X & (\<forall>x<a. x<j | P(x,j) --> Q(x,j))" |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
336 |
by (blast intro!: lt_Ord succI1 [THEN ltI, THEN lt_Ord, THEN le_refl]) |
12776 | 337 |
|
338 |
||
339 |
lemma lemma7: |
|
340 |
"[| \<forall>x<a. x<j | P(x,j) --> Q(x,j); succ(j)<a |] |
|
341 |
==> P(j,j) --> (\<forall>x<a. x\<le>j | P(x,j) --> Q(x,j))" |
|
342 |
by (fast elim!: leE) |
|
343 |
||
344 |
(* ********************************************************************** *) |
|
345 |
(* Lemmas needed to prove ex_next_set, which means that for any successor *) |
|
346 |
(* ordinal there is a set satisfying certain properties *) |
|
347 |
(* ********************************************************************** *) |
|
348 |
||
349 |
lemma ex_subset_eqpoll: |
|
350 |
"[| A\<approx>a; ~ Finite(a); Ord(a); m \<in> nat |] ==> \<exists>X \<in> Pow(A). X\<approx>m" |
|
351 |
apply (rule lepoll_imp_eqpoll_subset [of m A, THEN exE]) |
|
352 |
apply (rule lepoll_trans, rule leI [THEN le_imp_lepoll]) |
|
353 |
apply (blast intro: lt_trans2 [OF ltI nat_le_infinite_Ord] Ord_nat) |
|
354 |
apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll]) |
|
355 |
apply (fast elim!: eqpoll_sym) |
|
356 |
done |
|
357 |
||
358 |
lemma subset_Un_disjoint: "[| A \<subseteq> B Un C; A Int C = 0 |] ==> A \<subseteq> B" |
|
359 |
by blast |
|
360 |
||
361 |
||
362 |
lemma Int_empty: |
|
363 |
"[| X \<in> Pow(A - Union(B) -C); T \<in> B; F \<subseteq> T |] ==> F Int X = 0" |
|
364 |
by blast |
|
365 |
||
366 |
||
367 |
(* ********************************************************************** *) |
|
368 |
(* equipollent subset (and finite) is the whole set *) |
|
369 |
(* ********************************************************************** *) |
|
370 |
||
371 |
||
372 |
lemma subset_imp_eq_lemma: |
|
373 |
"m \<in> nat ==> \<forall>A B. A \<subseteq> B & m lepoll A & B lepoll m --> A=B" |
|
374 |
apply (induct_tac "m") |
|
375 |
apply (fast dest!: lepoll_0_is_0) |
|
376 |
apply (intro allI impI) |
|
377 |
apply (elim conjE) |
|
378 |
apply (rule succ_lepoll_imp_not_empty [THEN not_emptyE], assumption) |
|
379 |
apply (frule subsetD [THEN Diff_sing_lepoll], assumption+) |
|
380 |
apply (frule lepoll_Diff_sing) |
|
381 |
apply (erule allE impE)+ |
|
382 |
apply (rule conjI) |
|
383 |
prefer 2 apply fast |
|
384 |
apply fast |
|
385 |
apply (blast elim: equalityE) |
|
386 |
done |
|
387 |
||
388 |
||
389 |
lemma subset_imp_eq: "[| A \<subseteq> B; m lepoll A; B lepoll m; m \<in> nat |] ==> A=B" |
|
390 |
by (blast dest!: subset_imp_eq_lemma) |
|
391 |
||
392 |
||
393 |
lemma bij_imp_arg_eq: |
|
394 |
"[| f \<in> bij(a, {Y \<in> X. Y\<approx>succ(k)}); k \<in> nat; f`b \<subseteq> f`y; b<a; y<a |] |
|
395 |
==> b=y" |
|
396 |
apply (drule subset_imp_eq) |
|
397 |
apply (erule_tac [3] nat_succI) |
|
398 |
apply (unfold bij_def inj_def) |
|
399 |
apply (blast intro: eqpoll_sym eqpoll_imp_lepoll |
|
400 |
dest: ltD apply_type)+ |
|
401 |
done |
|
402 |
||
403 |
||
404 |
lemma ex_next_set: |
|
405 |
"[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)}; |
|
406 |
Card(a); ~ Finite(a); A\<approx>a; |
|
407 |
k \<in> nat; m \<in> nat; y<a; |
|
408 |
h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}); |
|
409 |
~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |] |
|
410 |
==> \<exists>X \<in> {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}. h`y \<subseteq> X & |
|
411 |
(\<forall>b<a. h`b \<subseteq> X --> |
|
412 |
(\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))" |
|
413 |
apply (erule_tac m1=m in dbl_Diff_eqpoll [THEN ex_subset_eqpoll, THEN bexE], |
|
414 |
assumption+) |
|
415 |
apply (erule Card_is_Ord, assumption) |
|
416 |
apply (frule Un_in_Collect, (erule asm_rl nat_succI)+) |
|
417 |
apply (erule CollectE) |
|
418 |
apply (rule rev_bexI, simp) |
|
419 |
apply (rule conjI, blast) |
|
420 |
apply (intro ballI impI oallI notI) |
|
421 |
apply (drule subset_Un_disjoint, rule Int_empty, assumption+) |
|
422 |
apply (blast dest: bij_imp_arg_eq) |
|
423 |
done |
|
424 |
||
425 |
(* ********************************************************************** *) |
|
426 |
(* Lemma ex_next_Ord states that for any successor *) |
|
427 |
(* ordinal there is a number of the set satisfying certain properties *) |
|
428 |
(* ********************************************************************** *) |
|
429 |
||
430 |
lemma ex_next_Ord: |
|
431 |
"[| recfunAC16(f, h, y, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)}; |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
432 |
Card(a); ~ Finite(a); A\<approx>a; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
433 |
k \<in> nat; m \<in> nat; y<a; |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
434 |
h \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k)}); |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
435 |
f \<in> bij(a, {Y \<in> Pow(A). Y\<approx>succ(k #+ m)}); |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
436 |
~ (\<exists>Y \<in> recfunAC16(f, h, y, a). h`y \<subseteq> Y) |] |
12776 | 437 |
==> \<exists>c<a. h`y \<subseteq> f`c & |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
438 |
(\<forall>b<a. h`b \<subseteq> f`c --> |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
439 |
(\<forall>T \<in> recfunAC16(f, h, y, a). ~ h`b \<subseteq> T))" |
12776 | 440 |
apply (drule ex_next_set, assumption+) |
441 |
apply (erule bexE) |
|
15481 | 442 |
apply (rule_tac x="converse(f)`X" in oexI) |
443 |
apply (simp add: right_inverse_bij) |
|
12776 | 444 |
apply (blast intro: bij_converse_bij bij_is_fun [THEN apply_type] ltI |
445 |
Card_is_Ord) |
|
446 |
done |
|
447 |
||
448 |
||
449 |
(* ********************************************************************** *) |
|
450 |
(* Lemma simplifying assumptions *) |
|
451 |
(* ********************************************************************** *) |
|
452 |
||
453 |
lemma lemma8: |
|
454 |
"[| \<forall>x<a. x<j | (\<exists>xa \<in> F(j). P(x, xa)) |
|
455 |
--> (\<exists>! Y. Y \<in> F(j) & P(x, Y)); F(j) \<subseteq> X; |
|
456 |
L \<in> X; P(j, L) & (\<forall>x<a. P(x, L) --> (\<forall>xa \<in> F(j). ~P(x, xa))) |] |
|
457 |
==> F(j) Un {L} \<subseteq> X & |
|
458 |
(\<forall>x<a. x\<le>j | (\<exists>xa \<in> (F(j) Un {L}). P(x, xa)) --> |
|
459 |
(\<exists>! Y. Y \<in> (F(j) Un {L}) & P(x, Y)))" |
|
460 |
apply (rule conjI) |
|
461 |
apply (fast intro!: singleton_subsetI) |
|
462 |
apply (rule oallI) |
|
463 |
apply (blast elim!: leE oallE) |
|
464 |
done |
|
465 |
||
466 |
(* ********************************************************************** *) |
|
467 |
(* The main part of the proof: inductive proof of the property of T_gamma *) |
|
468 |
(* lemma main_induct *) |
|
469 |
(* ********************************************************************** *) |
|
470 |
||
471 |
lemma main_induct: |
|
472 |
"[| b < a; f \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k #+ m)}); |
|
473 |
h \<in> bij(a, {Y \<in> Pow(A) . Y\<approx>succ(k)}); |
|
474 |
~Finite(a); Card(a); A\<approx>a; k \<in> nat; m \<in> nat |] |
|
475 |
==> recfunAC16(f, h, b, a) \<subseteq> {X \<in> Pow(A) . X\<approx>succ(k #+ m)} & |
|
476 |
(\<forall>x<a. x < b | (\<exists>Y \<in> recfunAC16(f, h, b, a). h ` x \<subseteq> Y) --> |
|
477 |
(\<exists>! Y. Y \<in> recfunAC16(f, h, b, a) & h ` x \<subseteq> Y))" |
|
478 |
apply (erule lt_induct) |
|
479 |
apply (frule lt_Ord) |
|
480 |
apply (erule Ord_cases) |
|
481 |
(* case 0 *) |
|
482 |
apply (simp add: recfunAC16_0) |
|
483 |
(* case Limit *) |
|
484 |
prefer 2 apply (simp add: recfunAC16_Limit) |
|
485 |
apply (rule lemma5, assumption+) |
|
486 |
apply (blast dest!: recfunAC16_mono) |
|
487 |
(* case succ *) |
|
488 |
apply clarify |
|
489 |
apply (erule lemma6 [THEN conjE], assumption) |
|
490 |
apply (simp (no_asm_simp) split del: split_if add: recfunAC16_succ) |
|
491 |
apply (rule conjI [THEN split_if [THEN iffD2]]) |
|
492 |
apply (simp, erule lemma7, assumption) |
|
493 |
apply (rule impI) |
|
494 |
apply (rule ex_next_Ord [THEN oexE], |
|
495 |
assumption+, rule le_refl [THEN lt_trans], assumption+) |
|
496 |
apply (erule lemma8, assumption) |
|
497 |
apply (rule bij_is_fun [THEN apply_type], assumption) |
|
498 |
apply (erule Least_le [THEN lt_trans2, THEN ltD]) |
|
499 |
apply (erule lt_Ord) |
|
500 |
apply (erule succ_leI) |
|
501 |
apply (erule LeastI) |
|
502 |
apply (erule lt_Ord) |
|
503 |
done |
|
504 |
||
505 |
(* ********************************************************************** *) |
|
506 |
(* Lemma to simplify the inductive proof *) |
|
507 |
(* - the desired property is a consequence of the inductive assumption *) |
|
508 |
(* ********************************************************************** *) |
|
509 |
||
510 |
lemma lemma_simp_induct: |
|
511 |
"[| \<forall>b. b<a --> F(b) \<subseteq> S & (\<forall>x<a. (x<b | (\<exists>Y \<in> F(b). f`x \<subseteq> Y)) |
|
512 |
--> (\<exists>! Y. Y \<in> F(b) & f`x \<subseteq> Y)); |
|
513 |
f \<in> a->f``(a); Limit(a); |
|
514 |
\<forall>i j. i\<le>j --> F(i) \<subseteq> F(j) |] |
|
515 |
==> (\<Union>j<a. F(j)) \<subseteq> S & |
|
516 |
(\<forall>x \<in> f``a. \<exists>! Y. Y \<in> (\<Union>j<a. F(j)) & x \<subseteq> Y)" |
|
517 |
apply (rule conjI) |
|
518 |
apply (rule subsetI) |
|
519 |
apply (erule OUN_E, blast) |
|
520 |
apply (rule ballI) |
|
521 |
apply (erule imageE) |
|
522 |
apply (drule ltI, erule Limit_is_Ord) |
|
523 |
apply (drule Limit_has_succ, assumption) |
|
12820 | 524 |
apply (frule_tac x1="succ(xa)" in spec [THEN mp], assumption) |
12776 | 525 |
apply (erule conjE) |
526 |
apply (drule ospec) |
|
527 |
(** LEVEL 10 **) |
|
528 |
apply (erule leI [THEN succ_leE]) |
|
529 |
apply (erule impE) |
|
530 |
apply (fast elim!: leI [THEN succ_leE, THEN lt_Ord, THEN le_refl]) |
|
531 |
apply (drule apply_equality, assumption) |
|
532 |
apply (elim conjE ex1E) |
|
533 |
(** LEVEL 15 **) |
|
534 |
apply (rule ex1I, blast) |
|
535 |
apply (elim conjE OUN_E) |
|
536 |
apply (erule_tac i="succ(xa)" and j=aa |
|
537 |
in Ord_linear_le [OF lt_Ord lt_Ord], assumption) |
|
538 |
prefer 2 |
|
539 |
apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) |
|
540 |
(** LEVEL 20 **) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
12820
diff
changeset
|
541 |
apply (drule_tac x1=aa in spec [THEN mp], assumption) |
12776 | 542 |
apply (frule succ_leE) |
543 |
apply (drule spec [THEN spec, THEN mp, THEN subsetD], assumption+, blast) |
|
544 |
done |
|
545 |
||
546 |
(* ********************************************************************** *) |
|
547 |
(* The target theorem *) |
|
548 |
(* ********************************************************************** *) |
|
549 |
||
550 |
theorem WO2_AC16: "[| WO2; 0<m; k \<in> nat; m \<in> nat |] ==> AC16(k #+ m,k)" |
|
551 |
apply (unfold AC16_def) |
|
552 |
apply (rule allI) |
|
553 |
apply (rule impI) |
|
554 |
apply (frule WO2_infinite_subsets_eqpoll_X, assumption+) |
|
12820 | 555 |
apply (frule_tac n="k #+ m" in WO2_infinite_subsets_eqpoll_X, simp, simp) |
12776 | 556 |
apply (frule WO2_imp_ex_Card) |
557 |
apply (elim exE conjE) |
|
558 |
apply (drule eqpoll_trans [THEN eqpoll_sym, |
|
559 |
THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], |
|
560 |
assumption) |
|
561 |
apply (drule eqpoll_trans [THEN eqpoll_sym, |
|
562 |
THEN eqpoll_def [THEN def_imp_iff, THEN iffD1]], |
|
563 |
assumption+) |
|
564 |
apply (elim exE) |
|
565 |
apply (rename_tac h) |
|
566 |
apply (rule_tac x = "\<Union>j<a. recfunAC16 (h,f,j,a) " in exI) |
|
567 |
apply (rule_tac P="%z. ?Y & (\<forall>x \<in> z. ?Z (x))" |
|
568 |
in bij_is_surj [THEN surj_image_eq, THEN subst], assumption) |
|
569 |
apply (rule lemma_simp_induct) |
|
570 |
apply (blast del: conjI notI |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
24893
diff
changeset
|
571 |
intro!: main_induct eqpoll_imp_lepoll [THEN lepoll_infinite] ) |
12776 | 572 |
apply (blast intro: bij_is_fun [THEN surj_image, THEN surj_is_fun]) |
573 |
apply (erule eqpoll_imp_lepoll [THEN lepoll_infinite, |
|
574 |
THEN infinite_Card_is_InfCard, |
|
575 |
THEN InfCard_is_Limit], |
|
576 |
assumption+) |
|
577 |
apply (blast dest!: recfunAC16_mono) |
|
578 |
done |
|
579 |
||
580 |
end |