| author | blanchet | 
| Tue, 07 May 2013 10:18:59 +0200 | |
| changeset 51888 | 1cbcc0cc6bdf | 
| parent 46822 | 95f1e700b712 | 
| child 61394 | 6142b282b164 | 
| permissions | -rw-r--r-- | 
| 12776 | 1  | 
(* Title: ZF/AC/AC16_lemmas.thy  | 
2  | 
Author: Krzysztof Grabczewski  | 
|
3  | 
||
4  | 
Lemmas used in the proofs concerning AC16  | 
|
5  | 
*)  | 
|
6  | 
||
| 27678 | 7  | 
theory AC16_lemmas  | 
8  | 
imports AC_Equiv Hartog Cardinal_aux  | 
|
9  | 
begin  | 
|
| 12776 | 10  | 
|
11  | 
lemma cons_Diff_eq: "a\<notin>A ==> cons(a,A)-{a}=A"
 | 
|
12  | 
by fast  | 
|
13  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
14  | 
lemma nat_1_lepoll_iff: "1\<lesssim>X \<longleftrightarrow> (\<exists>x. x \<in> X)"  | 
| 12776 | 15  | 
apply (unfold lepoll_def)  | 
16  | 
apply (rule iffI)  | 
|
17  | 
apply (fast intro: inj_is_fun [THEN apply_type])  | 
|
18  | 
apply (erule exE)  | 
|
19  | 
apply (rule_tac x = "\<lambda>a \<in> 1. x" in exI)  | 
|
20  | 
apply (fast intro!: lam_injective)  | 
|
21  | 
done  | 
|
22  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
23  | 
lemma eqpoll_1_iff_singleton: "X\<approx>1 \<longleftrightarrow> (\<exists>x. X={x})"
 | 
| 12776 | 24  | 
apply (rule iffI)  | 
25  | 
apply (erule eqpollE)  | 
|
26  | 
apply (drule nat_1_lepoll_iff [THEN iffD1])  | 
|
27  | 
apply (fast intro!: lepoll_1_is_sing)  | 
|
28  | 
apply (fast intro!: singleton_eqpoll_1)  | 
|
29  | 
done  | 
|
30  | 
||
31  | 
lemma cons_eqpoll_succ: "[| x\<approx>n; y\<notin>x |] ==> cons(y,x)\<approx>succ(n)"  | 
|
32  | 
apply (unfold succ_def)  | 
|
33  | 
apply (fast elim!: cons_eqpoll_cong mem_irrefl)  | 
|
34  | 
done  | 
|
35  | 
||
36  | 
lemma subsets_eqpoll_1_eq: "{Y \<in> Pow(X). Y\<approx>1} = {{x}. x \<in> X}"
 | 
|
37  | 
apply (rule equalityI)  | 
|
38  | 
apply (rule subsetI)  | 
|
39  | 
apply (erule CollectE)  | 
|
40  | 
apply (drule eqpoll_1_iff_singleton [THEN iffD1])  | 
|
41  | 
apply (fast intro!: RepFunI)  | 
|
42  | 
apply (rule subsetI)  | 
|
43  | 
apply (erule RepFunE)  | 
|
| 12820 | 44  | 
apply (rule CollectI, fast)  | 
| 12776 | 45  | 
apply (fast intro!: singleton_eqpoll_1)  | 
46  | 
done  | 
|
47  | 
||
48  | 
lemma eqpoll_RepFun_sing: "X\<approx>{{x}. x \<in> X}"
 | 
|
49  | 
apply (unfold eqpoll_def bij_def)  | 
|
50  | 
apply (rule_tac x = "\<lambda>x \<in> X. {x}" in exI)
 | 
|
51  | 
apply (rule IntI)  | 
|
| 12820 | 52  | 
apply (unfold inj_def surj_def, simp)  | 
53  | 
apply (fast intro!: lam_type RepFunI intro: singleton_eq_iff [THEN iffD1], simp)  | 
|
| 12776 | 54  | 
apply (fast intro!: lam_type)  | 
55  | 
done  | 
|
56  | 
||
57  | 
lemma subsets_eqpoll_1_eqpoll: "{Y \<in> Pow(X). Y\<approx>1}\<approx>X"
 | 
|
58  | 
apply (rule subsets_eqpoll_1_eq [THEN ssubst])  | 
|
59  | 
apply (rule eqpoll_RepFun_sing [THEN eqpoll_sym])  | 
|
60  | 
done  | 
|
61  | 
||
62  | 
lemma InfCard_Least_in:  | 
|
63  | 
"[| InfCard(x); y \<subseteq> x; y \<approx> succ(z) |] ==> (LEAST i. i \<in> y) \<in> y"  | 
|
64  | 
apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll,  | 
|
65  | 
THEN succ_lepoll_imp_not_empty, THEN not_emptyE])  | 
|
66  | 
apply (fast intro: LeastI  | 
|
67  | 
dest!: InfCard_is_Card [THEN Card_is_Ord]  | 
|
68  | 
elim: Ord_in_Ord)  | 
|
69  | 
done  | 
|
70  | 
||
71  | 
lemma subsets_lepoll_lemma1:  | 
|
72  | 
"[| InfCard(x); n \<in> nat |]  | 
|
73  | 
      ==> {y \<in> Pow(x). y\<approx>succ(succ(n))} \<lesssim> x*{y \<in> Pow(x). y\<approx>succ(n)}"
 | 
|
74  | 
apply (unfold lepoll_def)  | 
|
75  | 
apply (rule_tac x = "\<lambda>y \<in> {y \<in> Pow(x) . y\<approx>succ (succ (n))}. 
 | 
|
76  | 
                      <LEAST i. i \<in> y, y-{LEAST i. i \<in> y}>" in exI)
 | 
|
| 12820 | 77  | 
apply (rule_tac d = "%z. cons (fst(z), snd(z))" in lam_injective)  | 
78  | 
apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in)  | 
|
79  | 
apply (simp, blast intro: InfCard_Least_in)  | 
|
| 12776 | 80  | 
done  | 
81  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
82  | 
lemma set_of_Ord_succ_Union: "(\<forall>y \<in> z. Ord(y)) ==> z \<subseteq> succ(\<Union>(z))"  | 
| 12776 | 83  | 
apply (rule subsetI)  | 
| 12820 | 84  | 
apply (case_tac "\<forall>y \<in> z. y \<subseteq> x", blast )  | 
85  | 
apply (simp, erule bexE)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12820 
diff
changeset
 | 
86  | 
apply (rule_tac i=y and j=x in Ord_linear_le)  | 
| 12776 | 87  | 
apply (blast dest: le_imp_subset elim: leE ltE)+  | 
88  | 
done  | 
|
89  | 
||
90  | 
lemma subset_not_mem: "j \<subseteq> i ==> i \<notin> j"  | 
|
91  | 
by (fast elim!: mem_irrefl)  | 
|
92  | 
||
93  | 
lemma succ_Union_not_mem:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
94  | 
"(!!y. y \<in> z ==> Ord(y)) ==> succ(\<Union>(z)) \<notin> z"  | 
| 12820 | 95  | 
apply (rule set_of_Ord_succ_Union [THEN subset_not_mem], blast)  | 
| 12776 | 96  | 
done  | 
97  | 
||
98  | 
lemma Union_cons_eq_succ_Union:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
99  | 
"\<Union>(cons(succ(\<Union>(z)),z)) = succ(\<Union>(z))"  | 
| 12776 | 100  | 
by fast  | 
101  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
102  | 
lemma Un_Ord_disj: "[| Ord(i); Ord(j) |] ==> i \<union> j = i | i \<union> j = j"  | 
| 12776 | 103  | 
by (fast dest!: le_imp_subset elim: Ord_linear_le)  | 
104  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
105  | 
lemma Union_eq_Un: "x \<in> X ==> \<Union>(X) = x \<union> \<Union>(X-{x})"
 | 
| 12776 | 106  | 
by fast  | 
| 1196 | 107  | 
|
| 12776 | 108  | 
lemma Union_in_lemma [rule_format]:  | 
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
109  | 
"n \<in> nat ==> \<forall>z. (\<forall>y \<in> z. Ord(y)) & z\<approx>n & z\<noteq>0 \<longrightarrow> \<Union>(z) \<in> z"  | 
| 12776 | 110  | 
apply (induct_tac "n")  | 
111  | 
apply (fast dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])  | 
|
112  | 
apply (intro allI impI)  | 
|
113  | 
apply (erule natE)  | 
|
114  | 
apply (fast dest!: eqpoll_1_iff_singleton [THEN iffD1]  | 
|
| 12820 | 115  | 
intro!: Union_singleton, clarify)  | 
| 12776 | 116  | 
apply (elim not_emptyE)  | 
117  | 
apply (erule_tac x = "z-{xb}" in allE)
 | 
|
118  | 
apply (erule impE)  | 
|
119  | 
apply (fast elim!: Diff_sing_eqpoll  | 
|
120  | 
Diff_sing_eqpoll [THEN eqpoll_succ_imp_not_empty])  | 
|
| 12820 | 121  | 
apply (subgoal_tac "xb \<union> \<Union>(z - {xb}) \<in> z")
 | 
122  | 
apply (simp add: Union_eq_Un [symmetric])  | 
|
| 12776 | 123  | 
apply (frule bspec, assumption)  | 
| 12820 | 124  | 
apply (drule bspec)  | 
125  | 
apply (erule Diff_subset [THEN subsetD])  | 
|
126  | 
apply (drule Un_Ord_disj, assumption, auto)  | 
|
| 12776 | 127  | 
done  | 
128  | 
||
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
129  | 
lemma Union_in: "[| \<forall>x \<in> z. Ord(x); z\<approx>n; z\<noteq>0; n \<in> nat |] ==> \<Union>(z) \<in> z"  | 
| 12820 | 130  | 
by (blast intro: Union_in_lemma)  | 
| 12776 | 131  | 
|
132  | 
lemma succ_Union_in_x:  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
133  | 
"[| InfCard(x); z \<in> Pow(x); z\<approx>n; n \<in> nat |] ==> succ(\<Union>(z)) \<in> x"  | 
| 12820 | 134  | 
apply (rule Limit_has_succ [THEN ltE])  | 
| 12776 | 135  | 
prefer 3 apply assumption  | 
136  | 
apply (erule InfCard_is_Limit)  | 
|
| 12820 | 137  | 
apply (case_tac "z=0")  | 
138  | 
apply (simp, fast intro!: InfCard_is_Limit [THEN Limit_has_0])  | 
|
139  | 
apply (rule ltI [OF PowD [THEN subsetD] InfCard_is_Card [THEN Card_is_Ord]], assumption)  | 
|
| 12776 | 140  | 
apply (blast intro: Union_in  | 
141  | 
InfCard_is_Card [THEN Card_is_Ord, THEN Ord_in_Ord])+  | 
|
142  | 
done  | 
|
143  | 
||
144  | 
lemma succ_lepoll_succ_succ:  | 
|
145  | 
"[| InfCard(x); n \<in> nat |]  | 
|
146  | 
      ==> {y \<in> Pow(x). y\<approx>succ(n)} \<lesssim> {y \<in> Pow(x). y\<approx>succ(succ(n))}"
 | 
|
| 12820 | 147  | 
apply (unfold lepoll_def)  | 
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
148  | 
apply (rule_tac x = "\<lambda>z \<in> {y\<in>Pow(x). y\<approx>succ(n)}. cons(succ(\<Union>(z)), z)" 
 | 
| 12776 | 149  | 
in exI)  | 
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
32960 
diff
changeset
 | 
150  | 
apply (rule_tac d = "%z. z-{\<Union>(z) }" in lam_injective)
 | 
| 12776 | 151  | 
apply (blast intro!: succ_Union_in_x succ_Union_not_mem  | 
152  | 
intro: cons_eqpoll_succ Ord_in_Ord  | 
|
153  | 
dest!: InfCard_is_Card [THEN Card_is_Ord])  | 
|
| 12820 | 154  | 
apply (simp only: Union_cons_eq_succ_Union)  | 
155  | 
apply (rule cons_Diff_eq)  | 
|
| 12776 | 156  | 
apply (fast dest!: InfCard_is_Card [THEN Card_is_Ord]  | 
157  | 
elim: Ord_in_Ord  | 
|
| 12820 | 158  | 
intro!: succ_Union_not_mem)  | 
| 12776 | 159  | 
done  | 
160  | 
||
161  | 
lemma subsets_eqpoll_X:  | 
|
162  | 
     "[| InfCard(X); n \<in> nat |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)} \<approx> X"
 | 
|
163  | 
apply (induct_tac "n")  | 
|
164  | 
apply (rule subsets_eqpoll_1_eqpoll)  | 
|
165  | 
apply (rule eqpollI)  | 
|
| 12820 | 166  | 
apply (rule subsets_lepoll_lemma1 [THEN lepoll_trans], assumption+)  | 
167  | 
apply (rule eqpoll_trans [THEN eqpoll_imp_lepoll])  | 
|
168  | 
apply (erule eqpoll_refl [THEN prod_eqpoll_cong])  | 
|
| 12776 | 169  | 
apply (erule InfCard_square_eqpoll)  | 
170  | 
apply (fast elim: eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]  | 
|
171  | 
intro!: succ_lepoll_succ_succ)  | 
|
172  | 
done  | 
|
173  | 
||
174  | 
lemma image_vimage_eq:  | 
|
175  | 
"[| f \<in> surj(A,B); y \<subseteq> B |] ==> f``(converse(f)``y) = y"  | 
|
176  | 
apply (unfold surj_def)  | 
|
177  | 
apply (fast dest: apply_equality2 elim: apply_iff [THEN iffD2])  | 
|
178  | 
done  | 
|
179  | 
||
180  | 
lemma vimage_image_eq: "[| f \<in> inj(A,B); y \<subseteq> A |] ==> converse(f)``(f``y) = y"  | 
|
| 12820 | 181  | 
by (fast elim!: inj_is_fun [THEN apply_Pair] dest: inj_equality)  | 
| 12776 | 182  | 
|
183  | 
lemma subsets_eqpoll:  | 
|
184  | 
     "A\<approx>B ==> {Y \<in> Pow(A). Y\<approx>n}\<approx>{Y \<in> Pow(B). Y\<approx>n}"
 | 
|
185  | 
apply (unfold eqpoll_def)  | 
|
186  | 
apply (erule exE)  | 
|
187  | 
apply (rule_tac x = "\<lambda>X \<in> {Y \<in> Pow (A) . \<exists>f. f \<in> bij (Y, n) }. f``X" in exI)
 | 
|
188  | 
apply (rule_tac d = "%Z. converse (f) ``Z" in lam_bijective)  | 
|
189  | 
apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij,  | 
|
190  | 
THEN comp_bij]  | 
|
191  | 
elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset])  | 
|
192  | 
apply (blast intro!: bij_is_inj [THEN restrict_bij]  | 
|
193  | 
comp_bij bij_converse_bij  | 
|
194  | 
bij_is_fun [THEN fun_is_rel, THEN image_subset])  | 
|
195  | 
apply (fast elim!: bij_is_inj [THEN vimage_image_eq])  | 
|
196  | 
apply (fast elim!: bij_is_surj [THEN image_vimage_eq])  | 
|
197  | 
done  | 
|
198  | 
||
199  | 
lemma WO2_imp_ex_Card: "WO2 ==> \<exists>a. Card(a) & X\<approx>a"  | 
|
200  | 
apply (unfold WO2_def)  | 
|
201  | 
apply (drule spec [of _ X])  | 
|
202  | 
apply (blast intro: Card_cardinal eqpoll_trans  | 
|
203  | 
well_ord_Memrel [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym])  | 
|
204  | 
done  | 
|
205  | 
||
206  | 
lemma lepoll_infinite: "[| X\<lesssim>Y; ~Finite(X) |] ==> ~Finite(Y)"  | 
|
207  | 
by (blast intro: lepoll_Finite)  | 
|
208  | 
||
209  | 
lemma infinite_Card_is_InfCard: "[| ~Finite(X); Card(X) |] ==> InfCard(X)"  | 
|
210  | 
apply (unfold InfCard_def)  | 
|
211  | 
apply (fast elim!: Card_is_Ord [THEN nat_le_infinite_Ord])  | 
|
212  | 
done  | 
|
213  | 
||
214  | 
lemma WO2_infinite_subsets_eqpoll_X: "[| WO2; n \<in> nat; ~Finite(X) |]  | 
|
215  | 
        ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
 | 
|
216  | 
apply (drule WO2_imp_ex_Card)  | 
|
| 12820 | 217  | 
apply (elim allE exE conjE)  | 
| 12776 | 218  | 
apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)  | 
219  | 
apply (drule infinite_Card_is_InfCard, assumption)  | 
|
| 12820 | 220  | 
apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans)  | 
| 12776 | 221  | 
done  | 
222  | 
||
223  | 
lemma well_ord_imp_ex_Card: "well_ord(X,R) ==> \<exists>a. Card(a) & X\<approx>a"  | 
|
224  | 
by (fast elim!: well_ord_cardinal_eqpoll [THEN eqpoll_sym]  | 
|
225  | 
intro!: Card_cardinal)  | 
|
226  | 
||
227  | 
lemma well_ord_infinite_subsets_eqpoll_X:  | 
|
228  | 
     "[| well_ord(X,R); n \<in> nat; ~Finite(X) |] ==> {Y \<in> Pow(X). Y\<approx>succ(n)}\<approx>X"
 | 
|
229  | 
apply (drule well_ord_imp_ex_Card)  | 
|
230  | 
apply (elim allE exE conjE)  | 
|
231  | 
apply (frule eqpoll_imp_lepoll [THEN lepoll_infinite], assumption)  | 
|
232  | 
apply (drule infinite_Card_is_InfCard, assumption)  | 
|
| 12820 | 233  | 
apply (blast intro: subsets_eqpoll subsets_eqpoll_X eqpoll_sym eqpoll_trans)  | 
| 12776 | 234  | 
done  | 
235  | 
||
236  | 
end  |