author | haftmann |
Fri, 25 Jul 2008 07:35:53 +0200 | |
changeset 27678 | 85ea2be46c71 |
parent 16417 | 9bc16273c2d4 |
child 32960 | 69916a850301 |
permissions | -rw-r--r-- |
12787 | 1 |
(* Title: ZF/AC/AC7_AC9.thy |
2 |
ID: $Id$ |
|
3 |
Author: Krzysztof Grabczewski |
|
4 |
||
5 |
The proofs needed to state that AC7, AC8 and AC9 are equivalent to the previous |
|
6 |
instances of AC. |
|
7 |
*) |
|
8 |
||
27678 | 9 |
theory AC7_AC9 |
10 |
imports AC_Equiv |
|
11 |
begin |
|
12787 | 12 |
|
13 |
(* ********************************************************************** *) |
|
14 |
(* Lemmas used in the proofs AC7 ==> AC6 and AC9 ==> AC1 *) |
|
15 |
(* - Sigma_fun_space_not0 *) |
|
16 |
(* - Sigma_fun_space_eqpoll *) |
|
17 |
(* ********************************************************************** *) |
|
18 |
||
19 |
lemma Sigma_fun_space_not0: "[| 0\<notin>A; B \<in> A |] ==> (nat->Union(A)) * B \<noteq> 0" |
|
20 |
by (blast dest!: Sigma_empty_iff [THEN iffD1] Union_empty_iff [THEN iffD1]) |
|
21 |
||
22 |
lemma inj_lemma: |
|
23 |
"C \<in> A ==> (\<lambda>g \<in> (nat->Union(A))*C. |
|
24 |
(\<lambda>n \<in> nat. if(n=0, snd(g), fst(g)`(n #- 1)))) |
|
25 |
\<in> inj((nat->Union(A))*C, (nat->Union(A)) ) " |
|
26 |
apply (unfold inj_def) |
|
27 |
apply (rule CollectI) |
|
28 |
apply (fast intro!: lam_type if_type apply_type fst_type snd_type, auto) |
|
29 |
apply (rule fun_extension, assumption+) |
|
30 |
apply (drule lam_eqE [OF _ nat_succI], assumption, simp) |
|
31 |
apply (drule lam_eqE [OF _ nat_0I], simp) |
|
32 |
done |
|
33 |
||
34 |
lemma Sigma_fun_space_eqpoll: |
|
35 |
"[| C \<in> A; 0\<notin>A |] ==> (nat->Union(A)) * C \<approx> (nat->Union(A))" |
|
36 |
apply (rule eqpollI) |
|
37 |
apply (simp add: lepoll_def) |
|
38 |
apply (fast intro!: inj_lemma) |
|
39 |
apply (fast elim!: prod_lepoll_self not_sym [THEN not_emptyE] subst_elem |
|
40 |
elim: swap) |
|
41 |
done |
|
42 |
||
43 |
||
44 |
(* ********************************************************************** *) |
|
45 |
(* AC6 ==> AC7 *) |
|
46 |
(* ********************************************************************** *) |
|
47 |
||
48 |
lemma AC6_AC7: "AC6 ==> AC7" |
|
49 |
by (unfold AC6_def AC7_def, blast) |
|
50 |
||
51 |
(* ********************************************************************** *) |
|
52 |
(* AC7 ==> AC6, Rubin & Rubin p. 12, Theorem 2.8 *) |
|
53 |
(* The case of the empty family of sets added in order to complete *) |
|
54 |
(* the proof. *) |
|
55 |
(* ********************************************************************** *) |
|
56 |
||
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
12787
diff
changeset
|
57 |
lemma lemma1_1: "y \<in> (\<Pi> B \<in> A. Y*B) ==> (\<lambda>B \<in> A. snd(y`B)) \<in> (\<Pi> B \<in> A. B)" |
12787 | 58 |
by (fast intro!: lam_type snd_type apply_type) |
59 |
||
60 |
lemma lemma1_2: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
12787
diff
changeset
|
61 |
"y \<in> (\<Pi> B \<in> {Y*C. C \<in> A}. B) ==> (\<lambda>B \<in> A. y`(Y*B)) \<in> (\<Pi> B \<in> A. Y*B)" |
12787 | 62 |
apply (fast intro!: lam_type apply_type) |
63 |
done |
|
64 |
||
65 |
lemma AC7_AC6_lemma1: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
12787
diff
changeset
|
66 |
"(\<Pi> B \<in> {(nat->Union(A))*C. C \<in> A}. B) \<noteq> 0 ==> (\<Pi> B \<in> A. B) \<noteq> 0" |
12787 | 67 |
by (fast intro!: equals0I lemma1_1 lemma1_2) |
68 |
||
69 |
lemma AC7_AC6_lemma2: "0 \<notin> A ==> 0 \<notin> {(nat -> Union(A)) * C. C \<in> A}" |
|
70 |
by (blast dest: Sigma_fun_space_not0) |
|
71 |
||
72 |
lemma AC7_AC6: "AC7 ==> AC6" |
|
73 |
apply (unfold AC6_def AC7_def) |
|
74 |
apply (rule allI) |
|
75 |
apply (rule impI) |
|
76 |
apply (case_tac "A=0", simp) |
|
77 |
apply (rule AC7_AC6_lemma1) |
|
78 |
apply (erule allE) |
|
79 |
apply (blast del: notI |
|
80 |
intro!: AC7_AC6_lemma2 intro: eqpoll_sym eqpoll_trans |
|
81 |
Sigma_fun_space_eqpoll) |
|
82 |
done |
|
83 |
||
84 |
||
85 |
(* ********************************************************************** *) |
|
86 |
(* AC1 ==> AC8 *) |
|
87 |
(* ********************************************************************** *) |
|
88 |
||
89 |
lemma AC1_AC8_lemma1: |
|
90 |
"\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2 |
|
91 |
==> 0 \<notin> { bij(fst(B),snd(B)). B \<in> A }" |
|
92 |
apply (unfold eqpoll_def, auto) |
|
93 |
done |
|
94 |
||
95 |
lemma AC1_AC8_lemma2: |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
12787
diff
changeset
|
96 |
"[| f \<in> (\<Pi> X \<in> RepFun(A,p). X); D \<in> A |] ==> (\<lambda>x \<in> A. f`p(x))`D \<in> p(D)" |
12787 | 97 |
apply (simp, fast elim!: apply_type) |
98 |
done |
|
99 |
||
100 |
lemma AC1_AC8: "AC1 ==> AC8" |
|
101 |
apply (unfold AC1_def AC8_def) |
|
102 |
apply (fast dest: AC1_AC8_lemma1 AC1_AC8_lemma2) |
|
103 |
done |
|
104 |
||
105 |
||
106 |
(* ********************************************************************** *) |
|
107 |
(* AC8 ==> AC9 *) |
|
108 |
(* - this proof replaces the following two from Rubin & Rubin: *) |
|
109 |
(* AC8 ==> AC1 and AC1 ==> AC9 *) |
|
110 |
(* ********************************************************************** *) |
|
111 |
||
112 |
lemma AC8_AC9_lemma: |
|
113 |
"\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1 \<approx> B2 |
|
114 |
==> \<forall>B \<in> A*A. \<exists>B1 B2. B=<B1,B2> & B1 \<approx> B2" |
|
115 |
by fast |
|
116 |
||
117 |
lemma AC8_AC9: "AC8 ==> AC9" |
|
118 |
apply (unfold AC8_def AC9_def) |
|
119 |
apply (intro allI impI) |
|
120 |
apply (erule allE) |
|
121 |
apply (erule impE, erule AC8_AC9_lemma, force) |
|
122 |
done |
|
123 |
||
124 |
||
125 |
(* ********************************************************************** *) |
|
126 |
(* AC9 ==> AC1 *) |
|
127 |
(* The idea of this proof comes from "Equivalents of the Axiom of Choice" *) |
|
128 |
(* by Rubin & Rubin. But (x * y) is not necessarily equipollent to *) |
|
129 |
(* (x * y) Un {0} when y is a set of total functions acting from nat to *) |
|
130 |
(* Union(A) -- therefore we have used the set (y * nat) instead of y. *) |
|
131 |
(* ********************************************************************** *) |
|
132 |
||
133 |
lemma snd_lepoll_SigmaI: "b \<in> B \<Longrightarrow> X \<lesssim> B \<times> X" |
|
134 |
by (blast intro: lepoll_trans prod_lepoll_self eqpoll_imp_lepoll |
|
135 |
prod_commute_eqpoll) |
|
136 |
||
137 |
lemma nat_lepoll_lemma: |
|
138 |
"[|0 \<notin> A; B \<in> A|] ==> nat \<lesssim> ((nat \<rightarrow> Union(A)) \<times> B) \<times> nat" |
|
139 |
by (blast dest: Sigma_fun_space_not0 intro: snd_lepoll_SigmaI) |
|
140 |
||
141 |
lemma AC9_AC1_lemma1: |
|
142 |
"[| 0\<notin>A; A\<noteq>0; |
|
143 |
C = {((nat->Union(A))*B)*nat. B \<in> A} Un |
|
144 |
{cons(0,((nat->Union(A))*B)*nat). B \<in> A}; |
|
145 |
B1 \<in> C; B2 \<in> C |] |
|
146 |
==> B1 \<approx> B2" |
|
147 |
by (blast intro!: nat_lepoll_lemma Sigma_fun_space_eqpoll |
|
148 |
nat_cons_eqpoll [THEN eqpoll_trans] |
|
149 |
prod_eqpoll_cong [OF _ eqpoll_refl] |
|
150 |
intro: eqpoll_trans eqpoll_sym ) |
|
151 |
||
152 |
lemma AC9_AC1_lemma2: |
|
153 |
"\<forall>B1 \<in> {(F*B)*N. B \<in> A} Un {cons(0,(F*B)*N). B \<in> A}. |
|
154 |
\<forall>B2 \<in> {(F*B)*N. B \<in> A} Un {cons(0,(F*B)*N). B \<in> A}. |
|
155 |
f`<B1,B2> \<in> bij(B1, B2) |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
12787
diff
changeset
|
156 |
==> (\<lambda>B \<in> A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) \<in> (\<Pi> X \<in> A. X)" |
12787 | 157 |
apply (intro lam_type snd_type fst_type) |
158 |
apply (rule apply_type [OF _ consI1]) |
|
159 |
apply (fast intro!: fun_weaken_type bij_is_fun) |
|
160 |
done |
|
161 |
||
162 |
lemma AC9_AC1: "AC9 ==> AC1" |
|
163 |
apply (unfold AC1_def AC9_def) |
|
164 |
apply (intro allI impI) |
|
165 |
apply (erule allE) |
|
166 |
apply (case_tac "A~=0") |
|
167 |
apply (blast dest: AC9_AC1_lemma1 AC9_AC1_lemma2, force) |
|
168 |
done |
|
169 |
||
170 |
end |