|
1478
|
1 |
(* Title: ZF/AC/HH.thy
|
|
1123
|
2 |
ID: $Id$
|
|
1478
|
3 |
Author: Krzysztof Grabczewski
|
|
1123
|
4 |
|
|
12776
|
5 |
Some properties of the recursive definition of HH used in the proofs of
|
|
1123
|
6 |
AC17 ==> AC1
|
|
|
7 |
AC1 ==> WO2
|
|
|
8 |
AC15 ==> WO6
|
|
|
9 |
*)
|
|
|
10 |
|
|
12776
|
11 |
theory HH = AC_Equiv + Hartog:
|
|
1123
|
12 |
|
|
12776
|
13 |
constdefs
|
|
1123
|
14 |
|
|
12776
|
15 |
HH :: "[i, i, i] => i"
|
|
|
16 |
"HH(f,x,a) == transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
|
|
|
17 |
in if f`z \<in> Pow(z)-{0} then f`z else {x})"
|
|
|
18 |
|
|
|
19 |
|
|
|
20 |
(* ********************************************************************** *)
|
|
|
21 |
(* Lemmas useful in each of the three proofs *)
|
|
|
22 |
(* ********************************************************************** *)
|
|
|
23 |
|
|
|
24 |
lemma HH_def_satisfies_eq:
|
|
|
25 |
"HH(f,x,a) = (let z = x - (\<Union>b \<in> a. HH(f,x,b))
|
|
|
26 |
in if f`z \<in> Pow(z)-{0} then f`z else {x})"
|
|
|
27 |
by (rule HH_def [THEN def_transrec, THEN trans], simp)
|
|
|
28 |
|
|
|
29 |
lemma HH_values: "HH(f,x,a) \<in> Pow(x)-{0} | HH(f,x,a)={x}"
|
|
|
30 |
apply (rule HH_def_satisfies_eq [THEN ssubst])
|
|
|
31 |
apply (simp add: Let_def Diff_subset [THEN PowI], fast)
|
|
|
32 |
done
|
|
|
33 |
|
|
|
34 |
lemma subset_imp_Diff_eq:
|
|
|
35 |
"B \<subseteq> A ==> X-(\<Union>a \<in> A. P(a)) = X-(\<Union>a \<in> A-B. P(a))-(\<Union>b \<in> B. P(b))"
|
|
|
36 |
by fast
|
|
|
37 |
|
|
|
38 |
lemma Ord_DiffE: "[| c \<in> a-b; b<a |] ==> c=b | b<c & c<a"
|
|
|
39 |
apply (erule ltE)
|
|
|
40 |
apply (drule Ord_linear [of _ c])
|
|
|
41 |
apply (fast elim: Ord_in_Ord)
|
|
|
42 |
apply (fast intro!: ltI intro: Ord_in_Ord)
|
|
|
43 |
done
|
|
|
44 |
|
|
|
45 |
lemma Diff_UN_eq_self:
|
|
|
46 |
"(!!y. y \<in> A ==> P(y) = {x}) ==> x - (\<Union>y \<in> A. P(y)) = x"
|
|
|
47 |
apply (simp, fast elim!: mem_irrefl)
|
|
|
48 |
done
|
|
|
49 |
|
|
|
50 |
lemma HH_eq: "x - (\<Union>b \<in> a. HH(f,x,b)) = x - (\<Union>b \<in> a1. HH(f,x,b))
|
|
|
51 |
==> HH(f,x,a) = HH(f,x,a1)"
|
|
|
52 |
apply (subst HH_def_satisfies_eq)
|
|
|
53 |
apply (rule HH_def_satisfies_eq [THEN trans], simp)
|
|
|
54 |
done
|
|
|
55 |
|
|
|
56 |
lemma HH_is_x_gt_too: "[| HH(f,x,b)={x}; b<a |] ==> HH(f,x,a)={x}"
|
|
|
57 |
apply (rule_tac P = "b<a" in impE)
|
|
|
58 |
prefer 2 apply assumption+
|
|
|
59 |
apply (erule lt_Ord2 [THEN trans_induct])
|
|
|
60 |
apply (rule impI)
|
|
|
61 |
apply (rule HH_eq [THEN trans])
|
|
|
62 |
prefer 2 apply assumption+
|
|
|
63 |
apply (rule leI [THEN le_imp_subset, THEN subset_imp_Diff_eq, THEN ssubst],
|
|
|
64 |
assumption)
|
|
|
65 |
apply (rule_tac t = "%z. z-?X" in subst_context)
|
|
|
66 |
apply (rule Diff_UN_eq_self)
|
|
|
67 |
apply (drule Ord_DiffE, assumption)
|
|
|
68 |
apply (fast elim: ltE, auto)
|
|
|
69 |
done
|
|
|
70 |
|
|
|
71 |
lemma HH_subset_x_lt_too:
|
|
|
72 |
"[| HH(f,x,a) \<in> Pow(x)-{0}; b<a |] ==> HH(f,x,b) \<in> Pow(x)-{0}"
|
|
|
73 |
apply (rule HH_values [THEN disjE], assumption)
|
|
|
74 |
apply (drule HH_is_x_gt_too, assumption)
|
|
|
75 |
apply (drule subst, assumption)
|
|
|
76 |
apply (fast elim!: mem_irrefl)
|
|
|
77 |
done
|
|
|
78 |
|
|
|
79 |
lemma HH_subset_x_imp_subset_Diff_UN:
|
|
|
80 |
"HH(f,x,a) \<in> Pow(x)-{0} ==> HH(f,x,a) \<in> Pow(x - (\<Union>b \<in> a. HH(f,x,b)))-{0}"
|
|
|
81 |
apply (drule HH_def_satisfies_eq [THEN subst])
|
|
|
82 |
apply (rule HH_def_satisfies_eq [THEN ssubst])
|
|
|
83 |
apply (simp add: Let_def Diff_subset [THEN PowI])
|
|
|
84 |
apply (drule split_if [THEN iffD1])
|
|
|
85 |
apply (fast elim!: mem_irrefl)
|
|
|
86 |
done
|
|
|
87 |
|
|
|
88 |
lemma HH_eq_arg_lt:
|
|
|
89 |
"[| HH(f,x,v)=HH(f,x,w); HH(f,x,v) \<in> Pow(x)-{0}; v \<in> w |] ==> P"
|
|
|
90 |
apply (frule_tac P = "%y. y \<in> Pow (x) -{0}" in subst, assumption)
|
|
|
91 |
apply (drule_tac a = "w" in HH_subset_x_imp_subset_Diff_UN)
|
|
|
92 |
apply (drule subst_elem, assumption)
|
|
|
93 |
apply (fast intro!: singleton_iff [THEN iffD2] equals0I)
|
|
|
94 |
done
|
|
|
95 |
|
|
|
96 |
lemma HH_eq_imp_arg_eq:
|
|
|
97 |
"[| HH(f,x,v)=HH(f,x,w); HH(f,x,w) \<in> Pow(x)-{0}; Ord(v); Ord(w) |] ==> v=w"
|
|
|
98 |
apply (rule_tac j = "w" in Ord_linear_lt)
|
|
|
99 |
apply (simp_all (no_asm_simp))
|
|
|
100 |
apply (drule subst_elem, assumption)
|
|
|
101 |
apply (blast dest: ltD HH_eq_arg_lt)
|
|
|
102 |
apply (blast dest: HH_eq_arg_lt [OF sym] ltD)
|
|
|
103 |
done
|
|
|
104 |
|
|
|
105 |
lemma HH_subset_x_imp_lepoll:
|
|
|
106 |
"[| HH(f, x, i) \<in> Pow(x)-{0}; Ord(i) |] ==> i lepoll Pow(x)-{0}"
|
|
|
107 |
apply (unfold lepoll_def inj_def)
|
|
|
108 |
apply (rule_tac x = "\<lambda>j \<in> i. HH (f, x, j) " in exI)
|
|
|
109 |
apply (simp (no_asm_simp))
|
|
|
110 |
apply (fast del: DiffE
|
|
|
111 |
elim!: HH_eq_imp_arg_eq Ord_in_Ord HH_subset_x_lt_too
|
|
|
112 |
intro!: lam_type ballI ltI intro: bexI)
|
|
|
113 |
done
|
|
|
114 |
|
|
|
115 |
lemma HH_Hartog_is_x: "HH(f, x, Hartog(Pow(x)-{0})) = {x}"
|
|
|
116 |
apply (rule HH_values [THEN disjE])
|
|
|
117 |
prefer 2 apply assumption
|
|
|
118 |
apply (fast del: DiffE
|
|
|
119 |
intro!: Ord_Hartog
|
|
|
120 |
dest!: HH_subset_x_imp_lepoll
|
|
|
121 |
elim!: Hartog_lepoll_selfE)
|
|
|
122 |
done
|
|
|
123 |
|
|
|
124 |
lemma HH_Least_eq_x: "HH(f, x, LEAST i. HH(f, x, i) = {x}) = {x}"
|
|
|
125 |
by (fast intro!: Ord_Hartog HH_Hartog_is_x LeastI)
|
|
|
126 |
|
|
|
127 |
lemma less_Least_subset_x:
|
|
|
128 |
"a \<in> (LEAST i. HH(f,x,i)={x}) ==> HH(f,x,a) \<in> Pow(x)-{0}"
|
|
|
129 |
apply (rule HH_values [THEN disjE], assumption)
|
|
|
130 |
apply (rule less_LeastE)
|
|
|
131 |
apply (erule_tac [2] ltI [OF _ Ord_Least], assumption)
|
|
|
132 |
done
|
|
1123
|
133 |
|
|
12776
|
134 |
(* ********************************************************************** *)
|
|
|
135 |
(* Lemmas used in the proofs of AC1 ==> WO2 and AC17 ==> AC1 *)
|
|
|
136 |
(* ********************************************************************** *)
|
|
|
137 |
|
|
|
138 |
lemma lam_Least_HH_inj_Pow:
|
|
|
139 |
"(\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))
|
|
|
140 |
\<in> inj(LEAST i. HH(f,x,i)={x}, Pow(x)-{0})"
|
|
|
141 |
apply (unfold inj_def, simp)
|
|
|
142 |
apply (fast intro!: lam_type dest: less_Least_subset_x
|
|
|
143 |
elim!: HH_eq_imp_arg_eq Ord_Least [THEN Ord_in_Ord])
|
|
|
144 |
done
|
|
|
145 |
|
|
|
146 |
lemma lam_Least_HH_inj:
|
|
|
147 |
"\<forall>a \<in> (LEAST i. HH(f,x,i)={x}). \<exists>z \<in> x. HH(f,x,a) = {z}
|
|
|
148 |
==> (\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))
|
|
|
149 |
\<in> inj(LEAST i. HH(f,x,i)={x}, {{y}. y \<in> x})"
|
|
|
150 |
by (rule lam_Least_HH_inj_Pow [THEN inj_strengthen_type], simp)
|
|
|
151 |
|
|
|
152 |
lemma lam_surj_sing:
|
|
|
153 |
"[| x - (\<Union>a \<in> A. F(a)) = 0; \<forall>a \<in> A. \<exists>z \<in> x. F(a) = {z} |]
|
|
|
154 |
==> (\<lambda>a \<in> A. F(a)) \<in> surj(A, {{y}. y \<in> x})"
|
|
|
155 |
apply (simp add: surj_def lam_type Diff_eq_0_iff)
|
|
|
156 |
apply (blast elim: equalityE)
|
|
|
157 |
done
|
|
|
158 |
|
|
|
159 |
lemma not_emptyI2: "y \<in> Pow(x)-{0} ==> x \<noteq> 0"
|
|
|
160 |
by auto
|
|
|
161 |
|
|
|
162 |
lemma f_subset_imp_HH_subset:
|
|
|
163 |
"f`(x - (\<Union>j \<in> i. HH(f,x,j))) \<in> Pow(x - (\<Union>j \<in> i. HH(f,x,j)))-{0}
|
|
|
164 |
==> HH(f, x, i) \<in> Pow(x) - {0}"
|
|
|
165 |
apply (rule HH_def_satisfies_eq [THEN ssubst])
|
|
|
166 |
apply (simp add: Let_def Diff_subset [THEN PowI] not_emptyI2 [THEN if_P], fast)
|
|
|
167 |
done
|
|
|
168 |
|
|
|
169 |
lemma f_subsets_imp_UN_HH_eq_x:
|
|
|
170 |
"\<forall>z \<in> Pow(x)-{0}. f`z \<in> Pow(z)-{0}
|
|
|
171 |
==> x - (\<Union>j \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0"
|
|
|
172 |
apply (case_tac "?P \<in> {0}", fast)
|
|
|
173 |
apply (drule Diff_subset [THEN PowI, THEN DiffI])
|
|
|
174 |
apply (drule bspec, assumption)
|
|
|
175 |
apply (drule f_subset_imp_HH_subset)
|
|
|
176 |
apply (blast dest!: subst_elem [OF _ HH_Least_eq_x [symmetric]]
|
|
|
177 |
elim!: mem_irrefl)
|
|
|
178 |
done
|
|
|
179 |
|
|
|
180 |
lemma HH_values2: "HH(f,x,i) = f`(x - (\<Union>j \<in> i. HH(f,x,j))) | HH(f,x,i)={x}"
|
|
|
181 |
apply (rule HH_def_satisfies_eq [THEN ssubst])
|
|
|
182 |
apply (simp add: Let_def Diff_subset [THEN PowI])
|
|
|
183 |
done
|
|
|
184 |
|
|
|
185 |
lemma HH_subset_imp_eq:
|
|
|
186 |
"HH(f,x,i): Pow(x)-{0} ==> HH(f,x,i)=f`(x - (\<Union>j \<in> i. HH(f,x,j)))"
|
|
|
187 |
apply (rule HH_values2 [THEN disjE], assumption)
|
|
|
188 |
apply (fast elim!: equalityE mem_irrefl dest!: singleton_subsetD)
|
|
|
189 |
done
|
|
1123
|
190 |
|
|
12776
|
191 |
lemma f_sing_imp_HH_sing:
|
|
|
192 |
"[| f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x};
|
|
|
193 |
a \<in> (LEAST i. HH(f,x,i)={x}) |] ==> \<exists>z \<in> x. HH(f,x,a) = {z}"
|
|
|
194 |
apply (drule less_Least_subset_x)
|
|
|
195 |
apply (frule HH_subset_imp_eq)
|
|
|
196 |
apply (drule apply_type)
|
|
|
197 |
apply (rule Diff_subset [THEN PowI, THEN DiffI])
|
|
|
198 |
apply (fast dest!: HH_subset_x_imp_subset_Diff_UN [THEN not_emptyI2], force)
|
|
|
199 |
done
|
|
|
200 |
|
|
|
201 |
lemma f_sing_lam_bij:
|
|
|
202 |
"[| x - (\<Union>j \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,j)) = 0;
|
|
|
203 |
f \<in> (Pow(x)-{0}) -> {{z}. z \<in> x} |]
|
|
|
204 |
==> (\<lambda>a \<in> (LEAST i. HH(f,x,i)={x}). HH(f,x,a))
|
|
|
205 |
\<in> bij(LEAST i. HH(f,x,i)={x}, {{y}. y \<in> x})"
|
|
|
206 |
apply (unfold bij_def)
|
|
|
207 |
apply (fast intro!: lam_Least_HH_inj lam_surj_sing f_sing_imp_HH_sing)
|
|
|
208 |
done
|
|
|
209 |
|
|
|
210 |
lemma lam_singI:
|
|
|
211 |
"f \<in> (\<Pi>X \<in> Pow(x)-{0}. F(X))
|
|
|
212 |
==> (\<lambda>X \<in> Pow(x)-{0}. {f`X}) \<in> (\<Pi>X \<in> Pow(x)-{0}. {{z}. z \<in> F(X)})"
|
|
|
213 |
by (fast del: DiffI DiffE
|
|
|
214 |
intro!: lam_type singleton_eq_iff [THEN iffD2] dest: apply_type)
|
|
|
215 |
|
|
|
216 |
(*FIXME: both uses have the form ...[THEN bij_converse_bij], so
|
|
|
217 |
simplification is needed!*)
|
|
|
218 |
lemmas bij_Least_HH_x =
|
|
|
219 |
comp_bij [OF f_sing_lam_bij [OF _ lam_singI]
|
|
|
220 |
lam_sing_bij [THEN bij_converse_bij], standard]
|
|
|
221 |
|
|
|
222 |
|
|
|
223 |
(* ********************************************************************** *)
|
|
|
224 |
(* The proof of AC1 ==> WO2 *)
|
|
|
225 |
(* ********************************************************************** *)
|
|
|
226 |
|
|
|
227 |
(*Establishing the existence of a bijection, namely
|
|
|
228 |
converse
|
|
|
229 |
(converse(\<lambda>x\<in>x. {x}) O
|
|
|
230 |
Lambda
|
|
|
231 |
(LEAST i. HH(\<lambda>X\<in>Pow(x) - {0}. {f ` X}, x, i) = {x},
|
|
|
232 |
HH(\<lambda>X\<in>Pow(x) - {0}. {f ` X}, x)))
|
|
|
233 |
Perhaps it could be simplified. *)
|
|
|
234 |
|
|
|
235 |
lemma bijection:
|
|
|
236 |
"f \<in> (\<Pi>X \<in> Pow(x) - {0}. X)
|
|
|
237 |
==> \<exists>g. g \<in> bij(x, LEAST i. HH(\<lambda>X \<in> Pow(x)-{0}. {f`X}, x, i) = {x})"
|
|
|
238 |
apply (rule exI)
|
|
|
239 |
apply (rule bij_Least_HH_x [THEN bij_converse_bij])
|
|
|
240 |
apply (rule f_subsets_imp_UN_HH_eq_x)
|
|
|
241 |
apply (intro ballI apply_type)
|
|
12820
|
242 |
apply (fast intro: lam_type apply_type del: DiffE, assumption)
|
|
12776
|
243 |
apply (fast intro: Pi_weaken_type)
|
|
|
244 |
done
|
|
|
245 |
|
|
|
246 |
lemma AC1_WO2: "AC1 ==> WO2"
|
|
|
247 |
apply (unfold AC1_def WO2_def eqpoll_def)
|
|
|
248 |
apply (intro allI)
|
|
|
249 |
apply (drule_tac x = "Pow(A) - {0}" in spec)
|
|
|
250 |
apply (blast dest: bijection)
|
|
|
251 |
done
|
|
1123
|
252 |
|
|
|
253 |
end
|
|
|
254 |
|
|
12776
|
255 |
|