author | nipkow |
Sat, 08 Aug 2020 18:20:09 +0200 | |
changeset 72122 | 2dcb6523f6af |
parent 68847 | 511d163ab623 |
child 76213 | e44d86131648 |
permissions | -rw-r--r-- |
1478 | 1 |
(* Title: ZF/AC/DC.thy |
2 |
Author: Krzysztof Grabczewski |
|
1196 | 3 |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
4 |
The proofs concerning the Axiom of Dependent Choice. |
1196 | 5 |
*) |
6 |
||
27678 | 7 |
theory DC |
8 |
imports AC_Equiv Hartog Cardinal_aux |
|
9 |
begin |
|
12776 | 10 |
|
11 |
lemma RepFun_lepoll: "Ord(a) ==> {P(b). b \<in> a} \<lesssim> a" |
|
12 |
apply (unfold lepoll_def) |
|
61394 | 13 |
apply (rule_tac x = "\<lambda>z \<in> RepFun (a,P) . \<mu> i. z=P (i) " in exI) |
12776 | 14 |
apply (rule_tac d="%z. P (z)" in lam_injective) |
15 |
apply (fast intro!: Least_in_Ord) |
|
16 |
apply (rule sym) |
|
17 |
apply (fast intro: LeastI Ord_in_Ord) |
|
18 |
done |
|
1196 | 19 |
|
60770 | 20 |
text\<open>Trivial in the presence of AC, but here we need a wellordering of X\<close> |
12776 | 21 |
lemma image_Ord_lepoll: "[| f \<in> X->Y; Ord(X) |] ==> f``X \<lesssim> X" |
22 |
apply (unfold lepoll_def) |
|
61394 | 23 |
apply (rule_tac x = "\<lambda>x \<in> f``X. \<mu> y. f`y = x" in exI) |
12776 | 24 |
apply (rule_tac d = "%z. f`z" in lam_injective) |
25 |
apply (fast intro!: Least_in_Ord apply_equality, clarify) |
|
26 |
apply (rule LeastI) |
|
27 |
apply (erule apply_equality, assumption+) |
|
28 |
apply (blast intro: Ord_in_Ord) |
|
29 |
done |
|
1196 | 30 |
|
12776 | 31 |
lemma range_subset_domain: |
32 |
"[| R \<subseteq> X*X; !!g. g \<in> X ==> \<exists>u. <g,u> \<in> R |] |
|
33 |
==> range(R) \<subseteq> domain(R)" |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12891
diff
changeset
|
34 |
by blast |
12776 | 35 |
|
36 |
lemma cons_fun_type: "g \<in> n->X ==> cons(<n,x>, g) \<in> succ(n) -> cons(x, X)" |
|
37 |
apply (unfold succ_def) |
|
38 |
apply (fast intro!: fun_extend elim!: mem_irrefl) |
|
39 |
done |
|
1196 | 40 |
|
12776 | 41 |
lemma cons_fun_type2: |
42 |
"[| g \<in> n->X; x \<in> X |] ==> cons(<n,x>, g) \<in> succ(n) -> X" |
|
43 |
by (erule cons_absorb [THEN subst], erule cons_fun_type) |
|
44 |
||
45 |
lemma cons_image_n: "n \<in> nat ==> cons(<n,x>, g)``n = g``n" |
|
46 |
by (fast elim!: mem_irrefl) |
|
47 |
||
48 |
lemma cons_val_n: "g \<in> n->X ==> cons(<n,x>, g)`n = x" |
|
49 |
by (fast intro!: apply_equality elim!: cons_fun_type) |
|
50 |
||
51 |
lemma cons_image_k: "k \<in> n ==> cons(<n,x>, g)``k = g``k" |
|
52 |
by (fast elim: mem_asym) |
|
1196 | 53 |
|
12776 | 54 |
lemma cons_val_k: "[| k \<in> n; g \<in> n->X |] ==> cons(<n,x>, g)`k = g`k" |
55 |
by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair) |
|
56 |
||
57 |
lemma domain_cons_eq_succ: "domain(f)=x ==> domain(cons(<x,y>, f)) = succ(x)" |
|
58 |
by (simp add: domain_cons succ_def) |
|
59 |
||
60 |
lemma restrict_cons_eq: "g \<in> n->X ==> restrict(cons(<n,x>, g), n) = g" |
|
12891 | 61 |
apply (simp add: restrict_def Pi_iff) |
62 |
apply (blast intro: elim: mem_irrefl) |
|
12776 | 63 |
done |
64 |
||
65 |
lemma succ_in_succ: "[| Ord(k); i \<in> k |] ==> succ(i) \<in> succ(k)" |
|
66 |
apply (rule Ord_linear [of "succ(i)" "succ(k)", THEN disjE]) |
|
67 |
apply (fast elim: Ord_in_Ord mem_irrefl mem_asym)+ |
|
68 |
done |
|
1196 | 69 |
|
12776 | 70 |
lemma restrict_eq_imp_val_eq: |
12891 | 71 |
"[|restrict(f, domain(g)) = g; x \<in> domain(g)|] |
72 |
==> f`x = g`x" |
|
73 |
by (erule subst, simp add: restrict) |
|
12776 | 74 |
|
75 |
lemma domain_eq_imp_fun_type: "[| domain(f)=A; f \<in> B->C |] ==> f \<in> A->C" |
|
76 |
by (frule domain_of_fun, fast) |
|
77 |
||
78 |
lemma ex_in_domain: "[| R \<subseteq> A * B; R \<noteq> 0 |] ==> \<exists>x. x \<in> domain(R)" |
|
79 |
by (fast elim!: not_emptyE) |
|
80 |
||
5482 | 81 |
|
24893 | 82 |
definition |
83 |
DC :: "i => o" where |
|
12776 | 84 |
"DC(a) == \<forall>X R. R \<subseteq> Pow(X)*X & |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
85 |
(\<forall>Y \<in> Pow(X). Y \<prec> a \<longrightarrow> (\<exists>x \<in> X. <Y,x> \<in> R)) |
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
86 |
\<longrightarrow> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)" |
12776 | 87 |
|
24893 | 88 |
definition |
89 |
DC0 :: o where |
|
12776 | 90 |
"DC0 == \<forall>A B R. R \<subseteq> A*B & R\<noteq>0 & range(R) \<subseteq> domain(R) |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
91 |
\<longrightarrow> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)" |
12776 | 92 |
|
24893 | 93 |
definition |
94 |
ff :: "[i, i, i, i] => i" where |
|
12776 | 95 |
"ff(b, X, Q, R) == |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
96 |
transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))" |
12776 | 97 |
|
5482 | 98 |
|
27678 | 99 |
locale DC0_imp = |
12776 | 100 |
fixes XX and RR and X and R |
101 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
102 |
assumes all_ex: "\<forall>Y \<in> Pow(X). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R)" |
12776 | 103 |
|
104 |
defines XX_def: "XX == (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})" |
|
105 |
and RR_def: "RR == {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1)) |
|
106 |
& restrict(z2, domain(z1)) = z1}" |
|
68847 | 107 |
begin |
12776 | 108 |
|
109 |
(* ********************************************************************** *) |
|
110 |
(* DC ==> DC(omega) *) |
|
111 |
(* *) |
|
112 |
(* The scheme of the proof: *) |
|
113 |
(* *) |
|
114 |
(* Assume DC. Let R and X satisfy the premise of DC(omega). *) |
|
115 |
(* *) |
|
116 |
(* Define XX and RR as follows: *) |
|
117 |
(* *) |
|
118 |
(* XX = (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R}) *) |
|
119 |
(* f RR g iff domain(g)=succ(domain(f)) & *) |
|
120 |
(* restrict(g, domain(f)) = f *) |
|
121 |
(* *) |
|
122 |
(* Then RR satisfies the hypotheses of DC. *) |
|
123 |
(* So applying DC: *) |
|
124 |
(* *) |
|
125 |
(* \<exists>f \<in> nat->XX. \<forall>n \<in> nat. f`n RR f`succ(n) *) |
|
126 |
(* *) |
|
127 |
(* Thence *) |
|
128 |
(* *) |
|
129 |
(* ff = {<n, f`succ(n)`n>. n \<in> nat} *) |
|
130 |
(* *) |
|
131 |
(* is the desired function. *) |
|
132 |
(* *) |
|
133 |
(* ********************************************************************** *) |
|
134 |
||
68847 | 135 |
lemma lemma1_1: "RR \<subseteq> XX*XX" |
12776 | 136 |
by (unfold RR_def, fast) |
137 |
||
68847 | 138 |
lemma lemma1_2: "RR \<noteq> 0" |
12776 | 139 |
apply (unfold RR_def XX_def) |
140 |
apply (rule all_ex [THEN ballE]) |
|
141 |
apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]]) |
|
142 |
apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]]) |
|
143 |
apply (erule bexE) |
|
144 |
apply (rule_tac a = "<0, {<0, x>}>" in not_emptyI) |
|
145 |
apply (rule CollectI) |
|
146 |
apply (rule SigmaI) |
|
147 |
apply (rule nat_0I [THEN UN_I]) |
|
148 |
apply (simp (no_asm_simp) add: nat_0I [THEN UN_I]) |
|
149 |
apply (rule nat_1I [THEN UN_I]) |
|
150 |
apply (force intro!: singleton_fun [THEN Pi_type] |
|
151 |
simp add: singleton_0 [symmetric]) |
|
152 |
apply (simp add: singleton_0) |
|
153 |
done |
|
154 |
||
68847 | 155 |
lemma lemma1_3: "range(RR) \<subseteq> domain(RR)" |
12776 | 156 |
apply (unfold RR_def XX_def) |
157 |
apply (rule range_subset_domain, blast, clarify) |
|
158 |
apply (frule fun_is_rel [THEN image_subset, THEN PowI, |
|
159 |
THEN all_ex [THEN bspec]]) |
|
160 |
apply (erule impE[OF _ lesspoll_trans1[OF image_Ord_lepoll |
|
161 |
[OF _ nat_into_Ord] n_lesspoll_nat]], |
|
162 |
assumption+) |
|
163 |
apply (erule bexE) |
|
164 |
apply (rule_tac x = "cons (<n,x>, g) " in exI) |
|
165 |
apply (rule CollectI) |
|
166 |
apply (force elim!: cons_fun_type2 |
|
167 |
simp add: cons_image_n cons_val_n cons_image_k cons_val_k) |
|
168 |
apply (simp add: domain_of_fun succ_def restrict_cons_eq) |
|
169 |
done |
|
5482 | 170 |
|
68847 | 171 |
lemma lemma2: |
12776 | 172 |
"[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; n \<in> nat |] |
173 |
==> \<exists>k \<in> nat. f`succ(n) \<in> k -> X & n \<in> k |
|
174 |
& <f`succ(n)``n, f`succ(n)`n> \<in> R" |
|
175 |
apply (induct_tac "n") |
|
176 |
apply (drule apply_type [OF _ nat_1I]) |
|
177 |
apply (drule bspec [OF _ nat_0I]) |
|
178 |
apply (simp add: XX_def, safe) |
|
179 |
apply (rule rev_bexI, assumption) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
180 |
apply (subgoal_tac "0 \<in> y", force) |
12776 | 181 |
apply (force simp add: RR_def |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
182 |
intro: ltD elim!: nat_0_le [THEN leE]) |
12776 | 183 |
(** LEVEL 7, other subgoal **) |
184 |
apply (drule bspec [OF _ nat_succI], assumption) |
|
185 |
apply (subgoal_tac "f ` succ (succ (x)) \<in> succ (k) ->X") |
|
186 |
apply (drule apply_type [OF _ nat_succI [THEN nat_succI]], assumption) |
|
187 |
apply (simp (no_asm_use) add: XX_def RR_def) |
|
188 |
apply safe |
|
189 |
apply (frule_tac a="succ(k)" in domain_of_fun [symmetric, THEN trans], |
|
190 |
assumption) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
191 |
apply (frule_tac a=y in domain_of_fun [symmetric, THEN trans], |
12776 | 192 |
assumption) |
193 |
apply (fast elim!: nat_into_Ord [THEN succ_in_succ] |
|
194 |
dest!: bspec [OF _ nat_into_Ord [THEN succ_in_succ]]) |
|
195 |
apply (drule domain_of_fun) |
|
196 |
apply (simp add: XX_def RR_def, clarify) |
|
197 |
apply (blast dest: domain_of_fun [symmetric, THEN trans] ) |
|
198 |
done |
|
199 |
||
68847 | 200 |
lemma lemma3_1: |
12776 | 201 |
"[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat |] |
202 |
==> {f`succ(x)`x. x \<in> m} = {f`succ(m)`x. x \<in> m}" |
|
203 |
apply (subgoal_tac "\<forall>x \<in> m. f`succ (m) `x = f`succ (x) `x") |
|
204 |
apply simp |
|
205 |
apply (induct_tac "m", blast) |
|
206 |
apply (rule ballI) |
|
207 |
apply (erule succE) |
|
208 |
apply (rule restrict_eq_imp_val_eq) |
|
209 |
apply (drule bspec [OF _ nat_succI], assumption) |
|
210 |
apply (simp add: RR_def) |
|
211 |
apply (drule lemma2, assumption+) |
|
212 |
apply (fast dest!: domain_of_fun) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
213 |
apply (drule_tac x = xa in bspec, assumption) |
12776 | 214 |
apply (erule sym [THEN trans, symmetric]) |
215 |
apply (rule restrict_eq_imp_val_eq [symmetric]) |
|
216 |
apply (drule bspec [OF _ nat_succI], assumption) |
|
217 |
apply (simp add: RR_def) |
|
218 |
apply (drule lemma2, assumption+) |
|
219 |
apply (blast dest!: domain_of_fun |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
220 |
intro: nat_into_Ord OrdmemD [THEN subsetD]) |
12776 | 221 |
done |
5482 | 222 |
|
68847 | 223 |
lemma lemma3: |
12776 | 224 |
"[| \<forall>n \<in> nat. <f`n, f`succ(n)> \<in> RR; f \<in> nat -> XX; m \<in> nat |] |
225 |
==> (\<lambda>x \<in> nat. f`succ(x)`x) `` m = f`succ(m)``m" |
|
226 |
apply (erule natE, simp) |
|
227 |
apply (subst image_lam) |
|
228 |
apply (fast elim!: OrdmemD [OF nat_succI Ord_nat]) |
|
229 |
apply (subst lemma3_1, assumption+) |
|
230 |
apply fast |
|
231 |
apply (fast dest!: lemma2 |
|
232 |
elim!: image_fun [symmetric, OF _ OrdmemD [OF _ nat_into_Ord]]) |
|
233 |
done |
|
234 |
||
68847 | 235 |
end |
12776 | 236 |
|
237 |
theorem DC0_imp_DC_nat: "DC0 ==> DC(nat)" |
|
238 |
apply (unfold DC_def DC0_def, clarify) |
|
239 |
apply (elim allE) |
|
240 |
apply (erule impE) |
|
241 |
(*these three results comprise Lemma 1*) |
|
27678 | 242 |
apply (blast intro!: DC0_imp.lemma1_1 [OF DC0_imp.intro] DC0_imp.lemma1_2 [OF DC0_imp.intro] DC0_imp.lemma1_3 [OF DC0_imp.intro]) |
12776 | 243 |
apply (erule bexE) |
244 |
apply (rule_tac x = "\<lambda>n \<in> nat. f`succ (n) `n" in rev_bexI) |
|
27678 | 245 |
apply (rule lam_type, blast dest!: DC0_imp.lemma2 [OF DC0_imp.intro] intro: fun_weaken_type) |
12776 | 246 |
apply (rule oallI) |
27678 | 247 |
apply (frule DC0_imp.lemma2 [OF DC0_imp.intro], assumption) |
12776 | 248 |
apply (blast intro: fun_weaken_type) |
249 |
apply (erule ltD) |
|
250 |
(** LEVEL 11: last subgoal **) |
|
27678 | 251 |
apply (subst DC0_imp.lemma3 [OF DC0_imp.intro], assumption+) |
12776 | 252 |
apply (fast elim!: fun_weaken_type) |
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12891
diff
changeset
|
253 |
apply (erule ltD) |
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12891
diff
changeset
|
254 |
apply (force simp add: lt_def) |
12776 | 255 |
done |
256 |
||
257 |
||
258 |
(* ************************************************************************ |
|
259 |
DC(omega) ==> DC |
|
260 |
||
261 |
The scheme of the proof: |
|
262 |
||
263 |
Assume DC(omega). Let R and x satisfy the premise of DC. |
|
264 |
||
265 |
Define XX and RR as follows: |
|
266 |
||
267 |
XX = (\<Union>n \<in> nat. {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R}) |
|
268 |
||
269 |
RR = {<z1,z2>:Fin(XX)*XX. |
|
270 |
(domain(z2)=succ(\<Union>f \<in> z1. domain(f)) & |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
271 |
(\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) | |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
272 |
(~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) & |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
273 |
(\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
274 |
z2={<0,x>})} |
12776 | 275 |
|
276 |
Then XX and RR satisfy the hypotheses of DC(omega). |
|
277 |
So applying DC: |
|
278 |
||
279 |
\<exists>f \<in> nat->XX. \<forall>n \<in> nat. f``n RR f`n |
|
280 |
||
281 |
Thence |
|
282 |
||
283 |
ff = {<n, f`n`n>. n \<in> nat} |
|
284 |
||
285 |
is the desired function. |
|
286 |
||
287 |
************************************************************************* *) |
|
288 |
||
289 |
lemma singleton_in_funs: |
|
290 |
"x \<in> X ==> {<0,x>} \<in> |
|
291 |
(\<Union>n \<in> nat. {f \<in> succ(n)->X. \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})" |
|
292 |
apply (rule nat_0I [THEN UN_I]) |
|
293 |
apply (force simp add: singleton_0 [symmetric] |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
294 |
intro!: singleton_fun [THEN Pi_type]) |
12776 | 295 |
done |
5482 | 296 |
|
1196 | 297 |
|
27678 | 298 |
locale imp_DC0 = |
12776 | 299 |
fixes XX and RR and x and R and f and allRR |
300 |
defines XX_def: "XX == (\<Union>n \<in> nat. |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
301 |
{f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})" |
12776 | 302 |
and RR_def: |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
303 |
"RR == {<z1,z2>:Fin(XX)*XX. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
304 |
(domain(z2)=succ(\<Union>f \<in> z1. domain(f)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
305 |
& (\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
306 |
| (~ (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
307 |
& (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) & z2={<0,x>})}" |
12776 | 308 |
and allRR_def: |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
309 |
"allRR == \<forall>b<nat. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
310 |
<f``b, f`b> \<in> |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
311 |
{<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f)) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
312 |
& (\<Union>f \<in> z1. domain(f)) = b |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
313 |
& (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}" |
68847 | 314 |
begin |
12776 | 315 |
|
68847 | 316 |
lemma lemma4: |
12776 | 317 |
"[| range(R) \<subseteq> domain(R); x \<in> domain(R) |] |
318 |
==> RR \<subseteq> Pow(XX)*XX & |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
319 |
(\<forall>Y \<in> Pow(XX). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> XX. <Y,x>:RR))" |
12776 | 320 |
apply (rule conjI) |
321 |
apply (force dest!: FinD [THEN PowI] simp add: RR_def) |
|
322 |
apply (rule impI [THEN ballI]) |
|
323 |
apply (drule Finite_Fin [OF lesspoll_nat_is_Finite PowD], assumption) |
|
324 |
apply (case_tac |
|
325 |
"\<exists>g \<in> XX. domain (g) = |
|
326 |
succ(\<Union>f \<in> Y. domain(f)) & (\<forall>f\<in>Y. restrict(g, domain(f)) = f)") |
|
327 |
apply (simp add: RR_def, blast) |
|
328 |
apply (safe del: domainE) |
|
329 |
apply (unfold XX_def RR_def) |
|
330 |
apply (rule rev_bexI, erule singleton_in_funs) |
|
331 |
apply (simp add: nat_0I [THEN rev_bexI] cons_fun_type2) |
|
332 |
done |
|
333 |
||
68847 | 334 |
lemma UN_image_succ_eq: |
12776 | 335 |
"[| f \<in> nat->X; n \<in> nat |] |
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
336 |
==> (\<Union>x \<in> f``succ(n). P(x)) = P(f`n) \<union> (\<Union>x \<in> f``n. P(x))" |
12776 | 337 |
by (simp add: image_fun OrdmemD) |
338 |
||
68847 | 339 |
lemma UN_image_succ_eq_succ: |
12776 | 340 |
"[| (\<Union>x \<in> f``n. P(x)) = y; P(f`n) = succ(y); |
341 |
f \<in> nat -> X; n \<in> nat |] ==> (\<Union>x \<in> f``succ(n). P(x)) = succ(y)" |
|
342 |
by (simp add: UN_image_succ_eq, blast) |
|
343 |
||
68847 | 344 |
lemma apply_domain_type: |
12776 | 345 |
"[| h \<in> succ(n) -> D; n \<in> nat; domain(h)=succ(y) |] ==> h`y \<in> D" |
346 |
by (fast elim: apply_type dest!: trans [OF sym domain_of_fun]) |
|
347 |
||
68847 | 348 |
lemma image_fun_succ: |
12776 | 349 |
"[| h \<in> nat -> X; n \<in> nat |] ==> h``succ(n) = cons(h`n, h``n)" |
350 |
by (simp add: image_fun OrdmemD) |
|
351 |
||
68847 | 352 |
lemma f_n_type: |
12776 | 353 |
"[| domain(f`n) = succ(k); f \<in> nat -> XX; n \<in> nat |] |
354 |
==> f`n \<in> succ(k) -> domain(R)" |
|
355 |
apply (unfold XX_def) |
|
356 |
apply (drule apply_type, assumption) |
|
357 |
apply (fast elim: domain_eq_imp_fun_type) |
|
358 |
done |
|
359 |
||
68847 | 360 |
lemma f_n_pairs_in_R [rule_format]: |
12776 | 361 |
"[| h \<in> nat -> XX; domain(h`n) = succ(k); n \<in> nat |] |
362 |
==> \<forall>i \<in> k. <h`n`i, h`n`succ(i)> \<in> R" |
|
363 |
apply (unfold XX_def) |
|
364 |
apply (drule apply_type, assumption) |
|
365 |
apply (elim UN_E CollectE) |
|
12820 | 366 |
apply (drule domain_of_fun [symmetric, THEN trans], assumption, simp) |
12776 | 367 |
done |
368 |
||
68847 | 369 |
lemma restrict_cons_eq_restrict: |
12776 | 370 |
"[| restrict(h, domain(u))=u; h \<in> n->X; domain(u) \<subseteq> n |] |
371 |
==> restrict(cons(<n, y>, h), domain(u)) = u" |
|
372 |
apply (unfold restrict_def) |
|
12891 | 373 |
apply (simp add: restrict_def Pi_iff) |
12776 | 374 |
apply (erule sym [THEN trans, symmetric]) |
12891 | 375 |
apply (blast elim: mem_irrefl) |
12776 | 376 |
done |
377 |
||
68847 | 378 |
lemma all_in_image_restrict_eq: |
12776 | 379 |
"[| \<forall>x \<in> f``n. restrict(f`n, domain(x))=x; |
380 |
f \<in> nat -> XX; |
|
381 |
n \<in> nat; domain(f`n) = succ(n); |
|
382 |
(\<Union>x \<in> f``n. domain(x)) \<subseteq> n |] |
|
383 |
==> \<forall>x \<in> f``succ(n). restrict(cons(<succ(n),y>, f`n), domain(x)) = x" |
|
384 |
apply (rule ballI) |
|
385 |
apply (simp add: image_fun_succ) |
|
386 |
apply (drule f_n_type, assumption+) |
|
387 |
apply (erule disjE) |
|
388 |
apply (simp add: domain_of_fun restrict_cons_eq) |
|
389 |
apply (blast intro!: restrict_cons_eq_restrict) |
|
390 |
done |
|
391 |
||
68847 | 392 |
lemma simplify_recursion: |
12776 | 393 |
"[| \<forall>b<nat. <f``b, f`b> \<in> RR; |
394 |
f \<in> nat -> XX; range(R) \<subseteq> domain(R); x \<in> domain(R)|] |
|
395 |
==> allRR" |
|
396 |
apply (unfold RR_def allRR_def) |
|
397 |
apply (rule oallI, drule ltD) |
|
398 |
apply (erule nat_induct) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
399 |
apply (drule_tac x=0 in ospec, blast intro: Limit_has_0) |
12776 | 400 |
apply (force simp add: singleton_fun [THEN domain_of_fun] singleton_in_funs) |
401 |
(*induction step*) (** LEVEL 5 **) |
|
402 |
(*prevent simplification of ~\<exists> to \<forall>~ *) |
|
403 |
apply (simp only: separation split) |
|
12820 | 404 |
apply (drule_tac x="succ(xa)" in ospec, blast intro: ltI) |
12776 | 405 |
apply (elim conjE disjE) |
406 |
apply (force elim!: trans subst_context |
|
407 |
intro!: UN_image_succ_eq_succ) |
|
408 |
apply (erule notE) |
|
409 |
apply (simp add: XX_def UN_image_succ_eq_succ) |
|
410 |
apply (elim conjE bexE) |
|
411 |
apply (drule apply_domain_type, assumption+) |
|
412 |
apply (erule domainE)+ |
|
413 |
apply (frule f_n_type) |
|
414 |
apply (simp add: XX_def, assumption+) |
|
415 |
apply (rule rev_bexI, erule nat_succI) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
416 |
apply (rename_tac m i j y z) |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
417 |
apply (rule_tac x = "cons(<succ(m), z>, f`m)" in bexI) |
12776 | 418 |
prefer 2 apply (blast intro: cons_fun_type2) |
419 |
apply (rule conjI) |
|
420 |
prefer 2 apply (fast del: ballI subsetI |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
421 |
elim: trans [OF _ subst_context, THEN domain_cons_eq_succ] |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
422 |
subst_context |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
423 |
all_in_image_restrict_eq [simplified XX_def] |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
27678
diff
changeset
|
424 |
trans equalityD1) |
12776 | 425 |
(*one remaining subgoal*) |
426 |
apply (rule ballI) |
|
427 |
apply (erule succE) |
|
428 |
(** LEVEL 25 **) |
|
429 |
apply (simp add: cons_val_n cons_val_k) |
|
430 |
(*assumption+ will not perform the required backtracking!*) |
|
431 |
apply (drule f_n_pairs_in_R [simplified XX_def, OF _ domain_of_fun], |
|
432 |
assumption, assumption, assumption) |
|
433 |
apply (simp add: nat_into_Ord [THEN succ_in_succ] succI2 cons_val_k) |
|
434 |
done |
|
435 |
||
5482 | 436 |
|
68847 | 437 |
lemma lemma2: |
12776 | 438 |
"[| allRR; f \<in> nat->XX; range(R) \<subseteq> domain(R); x \<in> domain(R); n \<in> nat |] |
439 |
==> f`n \<in> succ(n) -> domain(R) & (\<forall>i \<in> n. <f`n`i, f`n`succ(i)>:R)" |
|
440 |
apply (unfold allRR_def) |
|
441 |
apply (drule ospec) |
|
442 |
apply (erule ltI [OF _ Ord_nat]) |
|
443 |
apply (erule CollectE, simp) |
|
444 |
apply (rule conjI) |
|
445 |
prefer 2 apply (fast elim!: f_n_pairs_in_R trans subst_context) |
|
446 |
apply (unfold XX_def) |
|
447 |
apply (fast elim!: trans [THEN domain_eq_imp_fun_type] subst_context) |
|
448 |
done |
|
449 |
||
68847 | 450 |
lemma lemma3: |
12776 | 451 |
"[| allRR; f \<in> nat->XX; n\<in>nat; range(R) \<subseteq> domain(R); x \<in> domain(R) |] |
452 |
==> f`n`n = f`succ(n)`n" |
|
453 |
apply (frule lemma2 [THEN conjunct1, THEN domain_of_fun], assumption+) |
|
454 |
apply (unfold allRR_def) |
|
455 |
apply (drule ospec) |
|
12820 | 456 |
apply (drule ltI [OF nat_succI Ord_nat], assumption, simp) |
12776 | 457 |
apply (elim conjE ballE) |
458 |
apply (erule restrict_eq_imp_val_eq [symmetric], force) |
|
459 |
apply (simp add: image_fun OrdmemD) |
|
460 |
done |
|
461 |
||
68847 | 462 |
end |
463 |
||
12776 | 464 |
|
465 |
theorem DC_nat_imp_DC0: "DC(nat) ==> DC0" |
|
466 |
apply (unfold DC_def DC0_def) |
|
467 |
apply (intro allI impI) |
|
468 |
apply (erule asm_rl conjE ex_in_domain [THEN exE] allE)+ |
|
469 |
apply (erule impE [OF _ imp_DC0.lemma4], assumption+) |
|
470 |
apply (erule bexE) |
|
471 |
apply (drule imp_DC0.simplify_recursion, assumption+) |
|
472 |
apply (rule_tac x = "\<lambda>n \<in> nat. f`n`n" in bexI) |
|
473 |
apply (rule_tac [2] lam_type) |
|
474 |
apply (erule_tac [2] apply_type [OF imp_DC0.lemma2 [THEN conjunct1] succI1]) |
|
475 |
apply (rule ballI) |
|
476 |
apply (frule_tac n="succ(n)" in imp_DC0.lemma2, |
|
477 |
(assumption|erule nat_succI)+) |
|
478 |
apply (drule imp_DC0.lemma3, auto) |
|
479 |
done |
|
480 |
||
481 |
(* ********************************************************************** *) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
482 |
(* \<forall>K. Card(K) \<longrightarrow> DC(K) ==> WO3 *) |
12776 | 483 |
(* ********************************************************************** *) |
484 |
||
485 |
lemma fun_Ord_inj: |
|
486 |
"[| f \<in> a->X; Ord(a); |
|
487 |
!!b c. [| b<c; c \<in> a |] ==> f`b\<noteq>f`c |] |
|
488 |
==> f \<in> inj(a, X)" |
|
489 |
apply (unfold inj_def, simp) |
|
490 |
apply (intro ballI impI) |
|
491 |
apply (rule_tac j=x in Ord_in_Ord [THEN Ord_linear_lt], assumption+) |
|
492 |
apply (blast intro: Ord_in_Ord, auto) |
|
493 |
apply (atomize, blast dest: not_sym) |
|
494 |
done |
|
495 |
||
496 |
lemma value_in_image: "[| f \<in> X->Y; A \<subseteq> X; a \<in> A |] ==> f`a \<in> f``A" |
|
13175
81082cfa5618
new definition of "apply" and new simprule "beta_if"
paulson
parents:
12891
diff
changeset
|
497 |
by (fast elim!: image_fun [THEN ssubst]) |
12776 | 498 |
|
46901 | 499 |
lemma lesspoll_lemma: "[| ~ A \<prec> B; C \<prec> B |] ==> A - C \<noteq> 0" |
500 |
apply (unfold lesspoll_def) |
|
501 |
apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll] |
|
502 |
intro!: eqpollI elim: notE |
|
503 |
elim!: eqpollE lepoll_trans) |
|
504 |
done |
|
505 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
506 |
theorem DC_WO3: "(\<forall>K. Card(K) \<longrightarrow> DC(K)) ==> WO3" |
12776 | 507 |
apply (unfold DC_def WO3_def) |
508 |
apply (rule allI) |
|
12820 | 509 |
apply (case_tac "A \<prec> Hartog (A)") |
12776 | 510 |
apply (fast dest!: lesspoll_imp_ex_lt_eqpoll |
511 |
intro!: Ord_Hartog leI [THEN le_imp_subset]) |
|
512 |
apply (erule allE impE)+ |
|
513 |
apply (rule Card_Hartog) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13175
diff
changeset
|
514 |
apply (erule_tac x = A in allE) |
12776 | 515 |
apply (erule_tac x = "{<z1,z2> \<in> Pow (A) *A . z1 \<prec> Hartog (A) & z2 \<notin> z1}" |
516 |
in allE) |
|
517 |
apply simp |
|
518 |
apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE]) |
|
519 |
apply (erule bexE) |
|
520 |
apply (rule Hartog_lepoll_selfE) |
|
521 |
apply (rule lepoll_def [THEN def_imp_iff, THEN iffD2]) |
|
522 |
apply (rule exI, rule fun_Ord_inj, assumption, rule Ord_Hartog) |
|
523 |
apply (drule value_in_image) |
|
524 |
apply (drule OrdmemD, rule Ord_Hartog, assumption+, erule ltD) |
|
525 |
apply (drule ospec) |
|
526 |
apply (blast intro: ltI Ord_Hartog, force) |
|
527 |
done |
|
528 |
||
529 |
(* ********************************************************************** *) |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
530 |
(* WO1 ==> \<forall>K. Card(K) \<longrightarrow> DC(K) *) |
12776 | 531 |
(* ********************************************************************** *) |
532 |
||
533 |
lemma images_eq: |
|
534 |
"[| \<forall>x \<in> A. f`x=g`x; f \<in> Df->Cf; g \<in> Dg->Cg; A \<subseteq> Df; A \<subseteq> Dg |] |
|
535 |
==> f``A = g``A" |
|
536 |
apply (simp (no_asm_simp) add: image_fun) |
|
537 |
done |
|
538 |
||
539 |
lemma lam_images_eq: |
|
540 |
"[| Ord(a); b \<in> a |] ==> (\<lambda>x \<in> a. h(x))``b = (\<lambda>x \<in> b. h(x))``b" |
|
541 |
apply (rule images_eq) |
|
542 |
apply (rule ballI) |
|
543 |
apply (drule OrdmemD [THEN subsetD], assumption+) |
|
544 |
apply simp |
|
545 |
apply (fast elim!: RepFunI OrdmemD intro!: lam_type)+ |
|
546 |
done |
|
547 |
||
548 |
lemma lam_type_RepFun: "(\<lambda>b \<in> a. h(b)) \<in> a -> {h(b). b \<in> a}" |
|
549 |
by (fast intro!: lam_type RepFunI) |
|
550 |
||
551 |
lemma lemmaX: |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
552 |
"[| \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R); |
12776 | 553 |
b \<in> K; Z \<in> Pow(X); Z \<prec> K |] |
554 |
==> {x \<in> X. <Z,x> \<in> R} \<noteq> 0" |
|
555 |
by blast |
|
556 |
||
557 |
||
558 |
lemma WO1_DC_lemma: |
|
559 |
"[| Card(K); well_ord(X,Q); |
|
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
560 |
\<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R); b \<in> K |] |
12776 | 561 |
==> ff(b, X, Q, R) \<in> {x \<in> X. <(\<lambda>c \<in> b. ff(c, X, Q, R))``b, x> \<in> R}" |
562 |
apply (rule_tac P = "b \<in> K" in impE, (erule_tac [2] asm_rl)+) |
|
563 |
apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], |
|
564 |
assumption+) |
|
565 |
apply (rule impI) |
|
566 |
apply (rule ff_def [THEN def_transrec, THEN ssubst]) |
|
567 |
apply (erule the_first_in, fast) |
|
568 |
apply (simp add: image_fun [OF lam_type_RepFun subset_refl]) |
|
569 |
apply (erule lemmaX, assumption) |
|
570 |
apply (blast intro: Card_is_Ord OrdmemD [THEN subsetD]) |
|
571 |
apply (blast intro: lesspoll_trans1 in_Card_imp_lesspoll RepFun_lepoll) |
|
572 |
done |
|
573 |
||
46822
95f1e700b712
mathematical symbols for Isabelle/ZF example theories
paulson
parents:
32960
diff
changeset
|
574 |
theorem WO1_DC_Card: "WO1 ==> \<forall>K. Card(K) \<longrightarrow> DC(K)" |
12776 | 575 |
apply (unfold DC_def WO1_def) |
576 |
apply (rule allI impI)+ |
|
577 |
apply (erule allE exE conjE)+ |
|
578 |
apply (rule_tac x = "\<lambda>b \<in> K. ff (b, X, Ra, R) " in bexI) |
|
579 |
apply (simp add: lam_images_eq [OF Card_is_Ord ltD]) |
|
580 |
apply (fast elim!: ltE WO1_DC_lemma [THEN CollectD2]) |
|
581 |
apply (rule_tac lam_type) |
|
582 |
apply (rule WO1_DC_lemma [THEN CollectD1], assumption+) |
|
583 |
done |
|
584 |
||
1196 | 585 |
end |