| author | paulson | 
| Wed, 01 Dec 2004 10:14:10 +0100 | |
| changeset 15351 | bdcd0f321df0 | 
| parent 14171 | 0cab06e3bbd0 | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 12776 | 1  | 
(* Title: ZF/AC/AC15_WO6.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Krzysztof Grabczewski  | 
|
4  | 
||
5  | 
The proofs needed to state that AC10, ..., AC15 are equivalent to the rest.  | 
|
6  | 
We need the following:  | 
|
7  | 
||
8  | 
WO1 ==> AC10(n) ==> AC11 ==> AC12 ==> AC15 ==> WO6  | 
|
9  | 
||
10  | 
In order to add the formulations AC13 and AC14 we need:  | 
|
11  | 
||
12  | 
AC10(succ(n)) ==> AC13(n) ==> AC14 ==> AC15  | 
|
13  | 
||
14  | 
or  | 
|
15  | 
||
16  | 
AC1 ==> AC13(1); AC13(m) ==> AC13(n) ==> AC14 ==> AC15 (m\<le>n)  | 
|
17  | 
||
18  | 
So we don't have to prove all implications of both cases.  | 
|
19  | 
Moreover we don't need to prove AC13(1) ==> AC1 and AC11 ==> AC14 as  | 
|
20  | 
Rubin & Rubin do.  | 
|
21  | 
*)  | 
|
22  | 
||
23  | 
theory AC15_WO6 = HH + Cardinal_aux:  | 
|
24  | 
||
25  | 
||
26  | 
(* ********************************************************************** *)  | 
|
27  | 
(* Lemmas used in the proofs in which the conclusion is AC13, AC14 *)  | 
|
28  | 
(* or AC15 *)  | 
|
29  | 
(* - cons_times_nat_not_Finite *)  | 
|
30  | 
(* - ex_fun_AC13_AC15 *)  | 
|
31  | 
(* ********************************************************************** *)  | 
|
32  | 
||
33  | 
lemma lepoll_Sigma: "A\<noteq>0 ==> B \<lesssim> A*B"  | 
|
34  | 
apply (unfold lepoll_def)  | 
|
35  | 
apply (erule not_emptyE)  | 
|
36  | 
apply (rule_tac x = "\<lambda>z \<in> B. <x,z>" in exI)  | 
|
37  | 
apply (fast intro!: snd_conv lam_injective)  | 
|
38  | 
done  | 
|
39  | 
||
40  | 
lemma cons_times_nat_not_Finite:  | 
|
41  | 
     "0\<notin>A ==> \<forall>B \<in> {cons(0,x*nat). x \<in> A}. ~Finite(B)"
 | 
|
42  | 
apply clarify  | 
|
43  | 
apply (rule nat_not_Finite [THEN notE] )  | 
|
44  | 
apply (subgoal_tac "x ~= 0")  | 
|
| 13245 | 45  | 
apply (blast intro: lepoll_Sigma [THEN lepoll_Finite])+  | 
| 12776 | 46  | 
done  | 
47  | 
||
48  | 
lemma lemma1: "[| Union(C)=A; a \<in> A |] ==> \<exists>B \<in> C. a \<in> B & B \<subseteq> A"  | 
|
49  | 
by fast  | 
|
50  | 
||
51  | 
lemma lemma2:  | 
|
52  | 
"[| pairwise_disjoint(A); B \<in> A; C \<in> A; a \<in> B; a \<in> C |] ==> B=C"  | 
|
53  | 
by (unfold pairwise_disjoint_def, blast)  | 
|
54  | 
||
55  | 
lemma lemma3:  | 
|
| 12788 | 56  | 
     "\<forall>B \<in> {cons(0, x*nat). x \<in> A}. pairwise_disjoint(f`B) &   
 | 
57  | 
sets_of_size_between(f`B, 2, n) & Union(f`B)=B  | 
|
58  | 
==> \<forall>B \<in> A. \<exists>! u. u \<in> f`cons(0, B*nat) & u \<subseteq> cons(0, B*nat) &  | 
|
59  | 
0 \<in> u & 2 \<lesssim> u & u \<lesssim> n"  | 
|
| 12776 | 60  | 
apply (unfold sets_of_size_between_def)  | 
61  | 
apply (rule ballI)  | 
|
| 12788 | 62  | 
apply (erule_tac x="cons(0, B*nat)" in ballE)  | 
| 12820 | 63  | 
apply (blast dest: lemma1 intro!: lemma2, blast)  | 
| 12776 | 64  | 
done  | 
65  | 
||
66  | 
lemma lemma4: "[| A \<lesssim> i; Ord(i) |] ==> {P(a). a \<in> A} \<lesssim> i"
 | 
|
67  | 
apply (unfold lepoll_def)  | 
|
68  | 
apply (erule exE)  | 
|
69  | 
apply (rule_tac x = "\<lambda>x \<in> RepFun(A,P). LEAST j. \<exists>a\<in>A. x=P(a) & f`a=j"  | 
|
70  | 
in exI)  | 
|
71  | 
apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective)  | 
|
72  | 
apply (erule RepFunE)  | 
|
73  | 
apply (frule inj_is_fun [THEN apply_type], assumption)  | 
|
74  | 
apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type])  | 
|
75  | 
apply (erule RepFunE)  | 
|
76  | 
apply (rule LeastI2)  | 
|
77  | 
apply fast  | 
|
78  | 
apply (fast elim!: Ord_in_Ord inj_is_fun [THEN apply_type])  | 
|
79  | 
apply (fast elim: sym left_inverse [THEN ssubst])  | 
|
80  | 
done  | 
|
81  | 
||
82  | 
lemma lemma5_1:  | 
|
83  | 
     "[| B \<in> A; 2 \<lesssim> u(B) |] ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<noteq> 0"
 | 
|
84  | 
apply simp  | 
|
85  | 
apply (fast dest: lepoll_Diff_sing  | 
|
86  | 
elim: lepoll_trans [THEN succ_lepoll_natE] ssubst  | 
|
87  | 
intro!: lepoll_refl)  | 
|
88  | 
done  | 
|
89  | 
||
90  | 
lemma lemma5_2:  | 
|
91  | 
"[| B \<in> A; u(B) \<subseteq> cons(0, B*nat) |]  | 
|
92  | 
      ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<subseteq> B"
 | 
|
93  | 
apply auto  | 
|
94  | 
done  | 
|
95  | 
||
96  | 
lemma lemma5_3:  | 
|
97  | 
"[| n \<in> nat; B \<in> A; 0 \<in> u(B); u(B) \<lesssim> succ(n) |]  | 
|
98  | 
      ==> (\<lambda>x \<in> A. {fst(x). x \<in> u(x)-{0}})`B \<lesssim> n"
 | 
|
99  | 
apply simp  | 
|
100  | 
apply (fast elim!: Diff_lepoll [THEN lemma4 [OF _ nat_into_Ord]])  | 
|
101  | 
done  | 
|
102  | 
||
103  | 
lemma ex_fun_AC13_AC15:  | 
|
104  | 
     "[| \<forall>B \<in> {cons(0, x*nat). x \<in> A}.   
 | 
|
105  | 
pairwise_disjoint(f`B) &  | 
|
106  | 
sets_of_size_between(f`B, 2, succ(n)) & Union(f`B)=B;  | 
|
107  | 
n \<in> nat |]  | 
|
108  | 
==> \<exists>f. \<forall>B \<in> A. f`B \<noteq> 0 & f`B \<subseteq> B & f`B \<lesssim> n"  | 
|
109  | 
by (fast del: subsetI notI  | 
|
110  | 
dest!: lemma3 theI intro!: lemma5_1 lemma5_2 lemma5_3)  | 
|
111  | 
||
112  | 
||
113  | 
(* ********************************************************************** *)  | 
|
114  | 
(* The target proofs *)  | 
|
115  | 
(* ********************************************************************** *)  | 
|
116  | 
||
117  | 
(* ********************************************************************** *)  | 
|
118  | 
(* AC10(n) ==> AC11 *)  | 
|
119  | 
(* ********************************************************************** *)  | 
|
120  | 
||
| 12788 | 121  | 
theorem AC10_AC11: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC11"  | 
| 12776 | 122  | 
by (unfold AC10_def AC11_def, blast)  | 
123  | 
||
124  | 
(* ********************************************************************** *)  | 
|
125  | 
(* AC11 ==> AC12 *)  | 
|
126  | 
(* ********************************************************************** *)  | 
|
127  | 
||
| 12788 | 128  | 
theorem AC11_AC12: "AC11 ==> AC12"  | 
| 12776 | 129  | 
by (unfold AC10_def AC11_def AC11_def AC12_def, blast)  | 
130  | 
||
131  | 
(* ********************************************************************** *)  | 
|
132  | 
(* AC12 ==> AC15 *)  | 
|
133  | 
(* ********************************************************************** *)  | 
|
134  | 
||
| 12788 | 135  | 
theorem AC12_AC15: "AC12 ==> AC15"  | 
| 12776 | 136  | 
apply (unfold AC12_def AC15_def)  | 
137  | 
apply (blast del: ballI  | 
|
138  | 
intro!: cons_times_nat_not_Finite ex_fun_AC13_AC15)  | 
|
139  | 
done  | 
|
| 1123 | 140  | 
|
| 12776 | 141  | 
(* ********************************************************************** *)  | 
142  | 
(* AC15 ==> WO6 *)  | 
|
143  | 
(* ********************************************************************** *)  | 
|
144  | 
||
145  | 
lemma OUN_eq_UN: "Ord(x) ==> (\<Union>a<x. F(a)) = (\<Union>a \<in> x. F(a))"  | 
|
146  | 
by (fast intro!: ltI dest!: ltD)  | 
|
147  | 
||
| 13535 | 148  | 
lemma AC15_WO6_aux1:  | 
| 12776 | 149  | 
     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
 | 
150  | 
      ==> (\<Union>i<LEAST x. HH(f,A,x)={A}. HH(f,A,i)) = A"
 | 
|
151  | 
apply (simp add: Ord_Least [THEN OUN_eq_UN])  | 
|
152  | 
apply (rule equalityI)  | 
|
153  | 
apply (fast dest!: less_Least_subset_x)  | 
|
154  | 
apply (blast del: subsetI  | 
|
155  | 
intro!: f_subsets_imp_UN_HH_eq_x [THEN Diff_eq_0_iff [THEN iffD1]])  | 
|
156  | 
done  | 
|
157  | 
||
| 13535 | 158  | 
lemma AC15_WO6_aux2:  | 
| 12776 | 159  | 
     "\<forall>x \<in> Pow(A)-{0}. f`x\<noteq>0 & f`x \<subseteq> x & f`x \<lesssim> m 
 | 
160  | 
      ==> \<forall>x < (LEAST x. HH(f,A,x)={A}). HH(f,A,x) \<lesssim> m"
 | 
|
161  | 
apply (rule oallI)  | 
|
162  | 
apply (drule ltD [THEN less_Least_subset_x])  | 
|
163  | 
apply (frule HH_subset_imp_eq)  | 
|
164  | 
apply (erule ssubst)  | 
|
165  | 
apply (blast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2])  | 
|
| 13535 | 166  | 
(*but can't use del: DiffE despite the obvious conflict*)  | 
| 12776 | 167  | 
done  | 
168  | 
||
| 12788 | 169  | 
theorem AC15_WO6: "AC15 ==> WO6"  | 
| 12776 | 170  | 
apply (unfold AC15_def WO6_def)  | 
171  | 
apply (rule allI)  | 
|
172  | 
apply (erule_tac x = "Pow (A) -{0}" in allE)
 | 
|
173  | 
apply (erule impE, fast)  | 
|
174  | 
apply (elim bexE conjE exE)  | 
|
175  | 
apply (rule bexI)  | 
|
| 
13175
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
176  | 
apply (rule conjI, assumption)  | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
177  | 
 apply (rule_tac x = "LEAST i. HH (f,A,i) ={A}" in exI)
 | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
178  | 
 apply (rule_tac x = "\<lambda>j \<in> (LEAST i. HH (f,A,i) ={A}) . HH (f,A,j) " in exI)
 | 
| 
 
81082cfa5618
new definition of "apply" and new simprule "beta_if"
 
paulson 
parents: 
12820 
diff
changeset
 | 
179  | 
apply (simp_all add: ltD)  | 
| 12776 | 180  | 
apply (fast intro!: Ord_Least lam_type [THEN domain_of_fun]  | 
| 13535 | 181  | 
elim!: less_Least_subset_x AC15_WO6_aux1 AC15_WO6_aux2)  | 
| 12776 | 182  | 
done  | 
183  | 
||
184  | 
||
185  | 
(* ********************************************************************** *)  | 
|
186  | 
(* The proof needed in the first case, not in the second *)  | 
|
187  | 
(* ********************************************************************** *)  | 
|
188  | 
||
189  | 
(* ********************************************************************** *)  | 
|
190  | 
(* AC10(n) ==> AC13(n-1) if 2\<le>n *)  | 
|
191  | 
(* *)  | 
|
192  | 
(* Because of the change to the formal definition of AC10(n) we prove *)  | 
|
193  | 
(* the following obviously equivalent theorem \<in> *)  | 
|
194  | 
(* AC10(n) implies AC13(n) for (1\<le>n) *)  | 
|
195  | 
(* ********************************************************************** *)  | 
|
196  | 
||
| 12788 | 197  | 
theorem AC10_AC13: "[| n \<in> nat; 1\<le>n; AC10(n) |] ==> AC13(n)"  | 
| 12776 | 198  | 
apply (unfold AC10_def AC13_def, safe)  | 
199  | 
apply (erule allE)  | 
|
| 12820 | 200  | 
apply (erule impE [OF _ cons_times_nat_not_Finite], assumption)  | 
| 12776 | 201  | 
apply (fast elim!: impE [OF _ cons_times_nat_not_Finite]  | 
202  | 
dest!: ex_fun_AC13_AC15)  | 
|
203  | 
done  | 
|
204  | 
||
205  | 
(* ********************************************************************** *)  | 
|
206  | 
(* The proofs needed in the second case, not in the first *)  | 
|
207  | 
(* ********************************************************************** *)  | 
|
208  | 
||
209  | 
(* ********************************************************************** *)  | 
|
210  | 
(* AC1 ==> AC13(1) *)  | 
|
211  | 
(* ********************************************************************** *)  | 
|
212  | 
||
213  | 
lemma AC1_AC13: "AC1 ==> AC13(1)"  | 
|
214  | 
apply (unfold AC1_def AC13_def)  | 
|
215  | 
apply (rule allI)  | 
|
216  | 
apply (erule allE)  | 
|
217  | 
apply (rule impI)  | 
|
218  | 
apply (drule mp, assumption)  | 
|
219  | 
apply (elim exE)  | 
|
220  | 
apply (rule_tac x = "\<lambda>x \<in> A. {f`x}" in exI)
 | 
|
221  | 
apply (simp add: singleton_eqpoll_1 [THEN eqpoll_imp_lepoll])  | 
|
222  | 
done  | 
|
223  | 
||
224  | 
(* ********************************************************************** *)  | 
|
225  | 
(* AC13(m) ==> AC13(n) for m \<subseteq> n *)  | 
|
226  | 
(* ********************************************************************** *)  | 
|
227  | 
||
228  | 
lemma AC13_mono: "[| m\<le>n; AC13(m) |] ==> AC13(n)"  | 
|
229  | 
apply (unfold AC13_def)  | 
|
230  | 
apply (drule le_imp_lepoll)  | 
|
231  | 
apply (fast elim!: lepoll_trans)  | 
|
232  | 
done  | 
|
233  | 
||
234  | 
(* ********************************************************************** *)  | 
|
235  | 
(* The proofs necessary for both cases *)  | 
|
236  | 
(* ********************************************************************** *)  | 
|
237  | 
||
238  | 
(* ********************************************************************** *)  | 
|
239  | 
(* AC13(n) ==> AC14 if 1 \<subseteq> n *)  | 
|
240  | 
(* ********************************************************************** *)  | 
|
241  | 
||
| 12788 | 242  | 
theorem AC13_AC14: "[| n \<in> nat; 1\<le>n; AC13(n) |] ==> AC14"  | 
| 12776 | 243  | 
by (unfold AC13_def AC14_def, auto)  | 
244  | 
||
245  | 
(* ********************************************************************** *)  | 
|
246  | 
(* AC14 ==> AC15 *)  | 
|
247  | 
(* ********************************************************************** *)  | 
|
248  | 
||
| 12788 | 249  | 
theorem AC14_AC15: "AC14 ==> AC15"  | 
| 12776 | 250  | 
by (unfold AC13_def AC14_def AC15_def, fast)  | 
251  | 
||
252  | 
(* ********************************************************************** *)  | 
|
253  | 
(* The redundant proofs; however cited by Rubin & Rubin *)  | 
|
254  | 
(* ********************************************************************** *)  | 
|
255  | 
||
256  | 
(* ********************************************************************** *)  | 
|
257  | 
(* AC13(1) ==> AC1 *)  | 
|
258  | 
(* ********************************************************************** *)  | 
|
259  | 
||
260  | 
lemma lemma_aux: "[| A\<noteq>0; A \<lesssim> 1 |] ==> \<exists>a. A={a}"
 | 
|
261  | 
by (fast elim!: not_emptyE lepoll_1_is_sing)  | 
|
262  | 
||
263  | 
lemma AC13_AC1_lemma:  | 
|
264  | 
"\<forall>B \<in> A. f(B)\<noteq>0 & f(B)<=B & f(B) \<lesssim> 1  | 
|
| 
14171
 
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
 
skalberg 
parents: 
13535 
diff
changeset
 | 
265  | 
      ==> (\<lambda>x \<in> A. THE y. f(x)={y}) \<in> (\<Pi> X \<in> A. X)"
 | 
| 12776 | 266  | 
apply (rule lam_type)  | 
267  | 
apply (drule bspec, assumption)  | 
|
268  | 
apply (elim conjE)  | 
|
269  | 
apply (erule lemma_aux [THEN exE], assumption)  | 
|
| 12814 | 270  | 
apply (simp add: the_equality)  | 
| 12776 | 271  | 
done  | 
272  | 
||
| 12788 | 273  | 
theorem AC13_AC1: "AC13(1) ==> AC1"  | 
| 12776 | 274  | 
apply (unfold AC13_def AC1_def)  | 
275  | 
apply (fast elim!: AC13_AC1_lemma)  | 
|
276  | 
done  | 
|
277  | 
||
278  | 
(* ********************************************************************** *)  | 
|
279  | 
(* AC11 ==> AC14 *)  | 
|
280  | 
(* ********************************************************************** *)  | 
|
281  | 
||
| 12788 | 282  | 
theorem AC11_AC14: "AC11 ==> AC14"  | 
| 12776 | 283  | 
apply (unfold AC11_def AC14_def)  | 
284  | 
apply (fast intro!: AC10_AC13)  | 
|
285  | 
done  | 
|
286  | 
||
287  | 
end  | 
|
288  |