author | lcp |
Thu, 08 Dec 1994 14:06:16 +0100 | |
changeset 767 | a4fce3b94065 |
parent 760 | f0200e91b272 |
child 782 | 200a16083201 |
permissions | -rw-r--r-- |
437 | 1 |
(* Title: ZF/CardinalArith.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1994 University of Cambridge |
|
5 |
||
6 |
Cardinal arithmetic -- WITHOUT the Axiom of Choice |
|
7 |
*) |
|
8 |
||
9 |
open CardinalArith; |
|
10 |
||
484 | 11 |
(*** Elementary properties ***) |
467 | 12 |
|
437 | 13 |
(*Use AC to discharge first premise*) |
14 |
goal CardinalArith.thy |
|
15 |
"!!A B. [| well_ord(B,r); A lepoll B |] ==> |A| le |B|"; |
|
16 |
by (res_inst_tac [("i","|A|"),("j","|B|")] Ord_linear_le 1); |
|
17 |
by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI])); |
|
18 |
by (rtac (eqpollI RS cardinal_cong) 1 THEN assume_tac 1); |
|
19 |
by (rtac lepoll_trans 1); |
|
20 |
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll) 1); |
|
21 |
by (assume_tac 1); |
|
22 |
by (etac (le_imp_subset RS subset_imp_lepoll RS lepoll_trans) 1); |
|
23 |
by (rtac eqpoll_imp_lepoll 1); |
|
24 |
by (rewtac lepoll_def); |
|
25 |
by (etac exE 1); |
|
26 |
by (rtac well_ord_cardinal_eqpoll 1); |
|
27 |
by (etac well_ord_rvimage 1); |
|
28 |
by (assume_tac 1); |
|
767 | 29 |
qed "well_ord_lepoll_imp_Card_le"; |
437 | 30 |
|
484 | 31 |
(*Each element of Fin(A) is equivalent to a natural number*) |
32 |
goal CardinalArith.thy |
|
33 |
"!!X A. X: Fin(A) ==> EX n:nat. X eqpoll n"; |
|
34 |
by (eresolve_tac [Fin_induct] 1); |
|
35 |
by (fast_tac (ZF_cs addIs [eqpoll_refl, nat_0I]) 1); |
|
36 |
by (fast_tac (ZF_cs addSIs [cons_eqpoll_cong, |
|
37 |
rewrite_rule [succ_def] nat_succI] |
|
38 |
addSEs [mem_irrefl]) 1); |
|
760 | 39 |
qed "Fin_eqpoll"; |
484 | 40 |
|
437 | 41 |
(*** Cardinal addition ***) |
42 |
||
43 |
(** Cardinal addition is commutative **) |
|
44 |
||
45 |
goalw CardinalArith.thy [eqpoll_def] "A+B eqpoll B+A"; |
|
46 |
by (rtac exI 1); |
|
47 |
by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] |
|
48 |
lam_bijective 1); |
|
49 |
by (safe_tac (ZF_cs addSEs [sumE])); |
|
50 |
by (ALLGOALS (asm_simp_tac case_ss)); |
|
760 | 51 |
qed "sum_commute_eqpoll"; |
437 | 52 |
|
53 |
goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i"; |
|
54 |
by (rtac (sum_commute_eqpoll RS cardinal_cong) 1); |
|
760 | 55 |
qed "cadd_commute"; |
437 | 56 |
|
57 |
(** Cardinal addition is associative **) |
|
58 |
||
59 |
goalw CardinalArith.thy [eqpoll_def] "(A+B)+C eqpoll A+(B+C)"; |
|
60 |
by (rtac exI 1); |
|
61 |
by (res_inst_tac [("c", "case(case(Inl, %y.Inr(Inl(y))), %y. Inr(Inr(y)))"), |
|
62 |
("d", "case(%x.Inl(Inl(x)), case(%x.Inl(Inr(x)), Inr))")] |
|
63 |
lam_bijective 1); |
|
64 |
by (ALLGOALS (asm_simp_tac (case_ss setloop etac sumE))); |
|
760 | 65 |
qed "sum_assoc_eqpoll"; |
437 | 66 |
|
67 |
(*Unconditional version requires AC*) |
|
68 |
goalw CardinalArith.thy [cadd_def] |
|
484 | 69 |
"!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ |
437 | 70 |
\ (i |+| j) |+| k = i |+| (j |+| k)"; |
71 |
by (rtac cardinal_cong 1); |
|
72 |
br ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS sum_eqpoll_cong RS |
|
73 |
eqpoll_trans) 1; |
|
74 |
by (rtac (sum_assoc_eqpoll RS eqpoll_trans) 2); |
|
75 |
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS |
|
76 |
eqpoll_sym) 2; |
|
484 | 77 |
by (REPEAT (ares_tac [well_ord_radd] 1)); |
760 | 78 |
qed "well_ord_cadd_assoc"; |
437 | 79 |
|
80 |
(** 0 is the identity for addition **) |
|
81 |
||
82 |
goalw CardinalArith.thy [eqpoll_def] "0+A eqpoll A"; |
|
83 |
by (rtac exI 1); |
|
84 |
by (res_inst_tac [("c", "case(%x.x, %y.y)"), ("d", "Inr")] |
|
85 |
lam_bijective 1); |
|
86 |
by (ALLGOALS (asm_simp_tac (case_ss setloop eresolve_tac [sumE,emptyE]))); |
|
760 | 87 |
qed "sum_0_eqpoll"; |
437 | 88 |
|
484 | 89 |
goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K"; |
437 | 90 |
by (asm_simp_tac (ZF_ss addsimps [sum_0_eqpoll RS cardinal_cong, |
91 |
Card_cardinal_eq]) 1); |
|
760 | 92 |
qed "cadd_0"; |
437 | 93 |
|
767 | 94 |
(** Addition by another cardinal **) |
95 |
||
96 |
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A+B"; |
|
97 |
by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1); |
|
98 |
by (asm_simp_tac (sum_ss addsimps [lam_type]) 1); |
|
99 |
val sum_lepoll_self = result(); |
|
100 |
||
101 |
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) |
|
102 |
goalw CardinalArith.thy [cadd_def] |
|
103 |
"!!K. [| Card(K); Ord(L) |] ==> K le (K |+| L)"; |
|
104 |
by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1); |
|
105 |
by (rtac sum_lepoll_self 3); |
|
106 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Card_is_Ord] 1)); |
|
107 |
val cadd_le_self = result(); |
|
108 |
||
109 |
(** Monotonicity of addition **) |
|
110 |
||
111 |
goalw CardinalArith.thy [lepoll_def] |
|
112 |
"!!A B C D. [| A lepoll C; B lepoll D |] ==> A + B lepoll C + D"; |
|
113 |
by (REPEAT (etac exE 1)); |
|
114 |
by (res_inst_tac [("x", "lam z:A+B. case(%w. Inl(f`w), %y. Inr(fa`y), z)")] |
|
115 |
exI 1); |
|
116 |
by (res_inst_tac |
|
117 |
[("d", "case(%w. Inl(converse(f)`w), %y. Inr(converse(fa)`y))")] |
|
118 |
lam_injective 1); |
|
119 |
by (typechk_tac ([inj_is_fun,case_type, InlI, InrI] @ ZF_typechecks)); |
|
120 |
by (eresolve_tac [sumE] 1); |
|
121 |
by (ALLGOALS (asm_simp_tac (sum_ss addsimps [left_inverse]))); |
|
122 |
val sum_lepoll_mono = result(); |
|
123 |
||
124 |
goalw CardinalArith.thy [cadd_def] |
|
125 |
"!!K. [| K' le K; L' le L |] ==> (K' |+| L') le (K |+| L)"; |
|
126 |
by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1])); |
|
127 |
by (resolve_tac [well_ord_lepoll_imp_Card_le] 1); |
|
128 |
by (REPEAT (ares_tac [sum_lepoll_mono, subset_imp_lepoll] 2)); |
|
129 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1)); |
|
130 |
val cadd_le_mono = result(); |
|
131 |
||
437 | 132 |
(** Addition of finite cardinals is "ordinary" addition **) |
133 |
||
134 |
goalw CardinalArith.thy [eqpoll_def] "succ(A)+B eqpoll succ(A+B)"; |
|
135 |
by (rtac exI 1); |
|
136 |
by (res_inst_tac [("c", "%z.if(z=Inl(A),A+B,z)"), |
|
137 |
("d", "%z.if(z=A+B,Inl(A),z)")] |
|
138 |
lam_bijective 1); |
|
139 |
by (ALLGOALS |
|
140 |
(asm_simp_tac (case_ss addsimps [succI2, mem_imp_not_eq] |
|
141 |
setloop eresolve_tac [sumE,succE]))); |
|
760 | 142 |
qed "sum_succ_eqpoll"; |
437 | 143 |
|
144 |
(*Pulling the succ(...) outside the |...| requires m, n: nat *) |
|
145 |
(*Unconditional version requires AC*) |
|
146 |
goalw CardinalArith.thy [cadd_def] |
|
147 |
"!!m n. [| Ord(m); Ord(n) |] ==> succ(m) |+| n = |succ(m |+| n)|"; |
|
148 |
by (rtac (sum_succ_eqpoll RS cardinal_cong RS trans) 1); |
|
149 |
by (rtac (succ_eqpoll_cong RS cardinal_cong) 1); |
|
150 |
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1); |
|
151 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1)); |
|
760 | 152 |
qed "cadd_succ_lemma"; |
437 | 153 |
|
154 |
val [mnat,nnat] = goal CardinalArith.thy |
|
155 |
"[| m: nat; n: nat |] ==> m |+| n = m#+n"; |
|
156 |
by (cut_facts_tac [nnat] 1); |
|
157 |
by (nat_ind_tac "m" [mnat] 1); |
|
158 |
by (asm_simp_tac (arith_ss addsimps [nat_into_Card RS cadd_0]) 1); |
|
159 |
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cadd_succ_lemma, |
|
160 |
nat_into_Card RS Card_cardinal_eq]) 1); |
|
760 | 161 |
qed "nat_cadd_eq_add"; |
437 | 162 |
|
163 |
||
164 |
(*** Cardinal multiplication ***) |
|
165 |
||
166 |
(** Cardinal multiplication is commutative **) |
|
167 |
||
168 |
(*Easier to prove the two directions separately*) |
|
169 |
goalw CardinalArith.thy [eqpoll_def] "A*B eqpoll B*A"; |
|
170 |
by (rtac exI 1); |
|
171 |
by (res_inst_tac [("c", "split(%x y.<y,x>)"), ("d", "split(%x y.<y,x>)")] |
|
172 |
lam_bijective 1); |
|
173 |
by (safe_tac ZF_cs); |
|
174 |
by (ALLGOALS (asm_simp_tac ZF_ss)); |
|
760 | 175 |
qed "prod_commute_eqpoll"; |
437 | 176 |
|
177 |
goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i"; |
|
178 |
by (rtac (prod_commute_eqpoll RS cardinal_cong) 1); |
|
760 | 179 |
qed "cmult_commute"; |
437 | 180 |
|
181 |
(** Cardinal multiplication is associative **) |
|
182 |
||
183 |
goalw CardinalArith.thy [eqpoll_def] "(A*B)*C eqpoll A*(B*C)"; |
|
184 |
by (rtac exI 1); |
|
185 |
by (res_inst_tac [("c", "split(%w z. split(%x y. <x,<y,z>>, w))"), |
|
186 |
("d", "split(%x. split(%y z. <<x,y>, z>))")] |
|
187 |
lam_bijective 1); |
|
188 |
by (safe_tac ZF_cs); |
|
189 |
by (ALLGOALS (asm_simp_tac ZF_ss)); |
|
760 | 190 |
qed "prod_assoc_eqpoll"; |
437 | 191 |
|
192 |
(*Unconditional version requires AC*) |
|
193 |
goalw CardinalArith.thy [cmult_def] |
|
484 | 194 |
"!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ |
437 | 195 |
\ (i |*| j) |*| k = i |*| (j |*| k)"; |
196 |
by (rtac cardinal_cong 1); |
|
197 |
br ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS |
|
198 |
eqpoll_trans) 1; |
|
199 |
by (rtac (prod_assoc_eqpoll RS eqpoll_trans) 2); |
|
200 |
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS prod_eqpoll_cong RS |
|
201 |
eqpoll_sym) 2; |
|
484 | 202 |
by (REPEAT (ares_tac [well_ord_rmult] 1)); |
760 | 203 |
qed "well_ord_cmult_assoc"; |
437 | 204 |
|
205 |
(** Cardinal multiplication distributes over addition **) |
|
206 |
||
207 |
goalw CardinalArith.thy [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)"; |
|
208 |
by (rtac exI 1); |
|
209 |
by (res_inst_tac |
|
210 |
[("c", "split(%x z. case(%y.Inl(<y,z>), %y.Inr(<y,z>), x))"), |
|
211 |
("d", "case(split(%x y.<Inl(x),y>), split(%x y.<Inr(x),y>))")] |
|
212 |
lam_bijective 1); |
|
213 |
by (safe_tac (ZF_cs addSEs [sumE])); |
|
214 |
by (ALLGOALS (asm_simp_tac case_ss)); |
|
760 | 215 |
qed "sum_prod_distrib_eqpoll"; |
437 | 216 |
|
217 |
(** Multiplication by 0 yields 0 **) |
|
218 |
||
219 |
goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0"; |
|
220 |
by (rtac exI 1); |
|
221 |
by (rtac lam_bijective 1); |
|
222 |
by (safe_tac ZF_cs); |
|
760 | 223 |
qed "prod_0_eqpoll"; |
437 | 224 |
|
225 |
goalw CardinalArith.thy [cmult_def] "0 |*| i = 0"; |
|
226 |
by (asm_simp_tac (ZF_ss addsimps [prod_0_eqpoll RS cardinal_cong, |
|
227 |
Card_0 RS Card_cardinal_eq]) 1); |
|
760 | 228 |
qed "cmult_0"; |
437 | 229 |
|
230 |
(** 1 is the identity for multiplication **) |
|
231 |
||
232 |
goalw CardinalArith.thy [eqpoll_def] "{x}*A eqpoll A"; |
|
233 |
by (rtac exI 1); |
|
234 |
by (res_inst_tac [("c", "snd"), ("d", "%z.<x,z>")] lam_bijective 1); |
|
235 |
by (safe_tac ZF_cs); |
|
236 |
by (ALLGOALS (asm_simp_tac ZF_ss)); |
|
760 | 237 |
qed "prod_singleton_eqpoll"; |
437 | 238 |
|
484 | 239 |
goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K"; |
437 | 240 |
by (asm_simp_tac (ZF_ss addsimps [prod_singleton_eqpoll RS cardinal_cong, |
241 |
Card_cardinal_eq]) 1); |
|
760 | 242 |
qed "cmult_1"; |
437 | 243 |
|
767 | 244 |
(*** Some inequalities for multiplication ***) |
245 |
||
246 |
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A"; |
|
247 |
by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1); |
|
248 |
by (simp_tac (ZF_ss addsimps [lam_type]) 1); |
|
249 |
qed "prod_square_lepoll"; |
|
250 |
||
251 |
(*Could probably weaken the premise to well_ord(K,r), or removing using AC*) |
|
252 |
goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K"; |
|
253 |
by (rtac le_trans 1); |
|
254 |
by (rtac well_ord_lepoll_imp_Card_le 2); |
|
255 |
by (rtac prod_square_lepoll 3); |
|
256 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2)); |
|
257 |
by (asm_simp_tac (ZF_ss addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1); |
|
258 |
qed "cmult_square_le"; |
|
259 |
||
260 |
(** Multiplication by a non-zero cardinal **) |
|
261 |
||
262 |
goalw CardinalArith.thy [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B"; |
|
263 |
by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1); |
|
264 |
by (asm_simp_tac (ZF_ss addsimps [lam_type]) 1); |
|
265 |
val prod_lepoll_self = result(); |
|
266 |
||
267 |
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) |
|
268 |
goalw CardinalArith.thy [cmult_def] |
|
269 |
"!!K. [| Card(K); Ord(L); 0<L |] ==> K le (K |*| L)"; |
|
270 |
by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1); |
|
271 |
by (rtac prod_lepoll_self 3); |
|
272 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord, ltD] 1)); |
|
273 |
val cmult_le_self = result(); |
|
274 |
||
275 |
(** Monotonicity of multiplication **) |
|
276 |
||
277 |
(*Equivalent to KG's lepoll_SigmaI*) |
|
278 |
goalw CardinalArith.thy [lepoll_def] |
|
279 |
"!!A B C D. [| A lepoll C; B lepoll D |] ==> A * B lepoll C * D"; |
|
280 |
by (REPEAT (etac exE 1)); |
|
281 |
by (res_inst_tac [("x", "lam z:A*B. split(%w y.<f`w, fa`y>, z)")] exI 1); |
|
282 |
by (res_inst_tac [("d", "split(%w y.<converse(f)`w, converse(fa)`y>)")] |
|
283 |
lam_injective 1); |
|
284 |
by (typechk_tac (inj_is_fun::ZF_typechecks)); |
|
285 |
by (eresolve_tac [SigmaE] 1); |
|
286 |
by (asm_simp_tac (ZF_ss addsimps [left_inverse]) 1); |
|
287 |
val prod_lepoll_mono = result(); |
|
288 |
||
289 |
goalw CardinalArith.thy [cmult_def] |
|
290 |
"!!K. [| K' le K; L' le L |] ==> (K' |*| L') le (K |*| L)"; |
|
291 |
by (safe_tac (ZF_cs addSDs [le_subset_iff RS iffD1])); |
|
292 |
by (resolve_tac [well_ord_lepoll_imp_Card_le] 1); |
|
293 |
by (REPEAT (ares_tac [prod_lepoll_mono, subset_imp_lepoll] 2)); |
|
294 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
295 |
val cmult_le_mono = result(); |
|
296 |
||
297 |
(*** Multiplication of finite cardinals is "ordinary" multiplication ***) |
|
437 | 298 |
|
299 |
goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B"; |
|
300 |
by (rtac exI 1); |
|
301 |
by (res_inst_tac [("c", "split(%x y. if(x=A, Inl(y), Inr(<x,y>)))"), |
|
302 |
("d", "case(%y. <A,y>, %z.z)")] |
|
303 |
lam_bijective 1); |
|
304 |
by (safe_tac (ZF_cs addSEs [sumE])); |
|
305 |
by (ALLGOALS |
|
306 |
(asm_simp_tac (case_ss addsimps [succI2, if_type, mem_imp_not_eq]))); |
|
760 | 307 |
qed "prod_succ_eqpoll"; |
437 | 308 |
|
309 |
(*Unconditional version requires AC*) |
|
310 |
goalw CardinalArith.thy [cmult_def, cadd_def] |
|
311 |
"!!m n. [| Ord(m); Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)"; |
|
312 |
by (rtac (prod_succ_eqpoll RS cardinal_cong RS trans) 1); |
|
313 |
by (rtac (cardinal_cong RS sym) 1); |
|
314 |
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong) 1); |
|
315 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
760 | 316 |
qed "cmult_succ_lemma"; |
437 | 317 |
|
318 |
val [mnat,nnat] = goal CardinalArith.thy |
|
319 |
"[| m: nat; n: nat |] ==> m |*| n = m#*n"; |
|
320 |
by (cut_facts_tac [nnat] 1); |
|
321 |
by (nat_ind_tac "m" [mnat] 1); |
|
322 |
by (asm_simp_tac (arith_ss addsimps [cmult_0]) 1); |
|
323 |
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cmult_succ_lemma, |
|
324 |
nat_cadd_eq_add]) 1); |
|
760 | 325 |
qed "nat_cmult_eq_mult"; |
437 | 326 |
|
767 | 327 |
(*Needs Krzysztof Grabczewski's macro "2" == "succ(1)"*) |
328 |
goal CardinalArith.thy "!!m n. Card(n) ==> succ(1) |*| n = n |+| n"; |
|
329 |
by (asm_simp_tac |
|
330 |
(ZF_ss addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, Card_is_Ord, |
|
331 |
read_instantiate [("j","0")] cadd_commute, cadd_0]) 1); |
|
332 |
val cmult_2 = result(); |
|
333 |
||
437 | 334 |
|
335 |
(*** Infinite Cardinals are Limit Ordinals ***) |
|
336 |
||
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
337 |
(*This proof is modelled upon one assuming nat<=A, with injection |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
338 |
lam z:cons(u,A). if(z=u, 0, if(z : nat, succ(z), z)) and inverse |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
339 |
%y. if(y:nat, nat_case(u,%z.z,y), y). If f: inj(nat,A) then |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
340 |
range(f) behaves like the natural numbers.*) |
516 | 341 |
goalw CardinalArith.thy [lepoll_def] |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
342 |
"!!i. nat lepoll A ==> cons(u,A) lepoll A"; |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
343 |
by (etac exE 1); |
516 | 344 |
by (res_inst_tac [("x", |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
345 |
"lam z:cons(u,A). if(z=u, f`0, \ |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
346 |
\ if(z: range(f), f`succ(converse(f)`z), z))")] exI 1); |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
347 |
by (res_inst_tac [("d", "%y. if(y: range(f), \ |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
348 |
\ nat_case(u, %z.f`z, converse(f)`y), y)")] |
516 | 349 |
lam_injective 1); |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
350 |
by (fast_tac (ZF_cs addSIs [if_type, nat_0I, nat_succI, apply_type] |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
351 |
addIs [inj_is_fun, inj_converse_fun]) 1); |
516 | 352 |
by (asm_simp_tac |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
353 |
(ZF_ss addsimps [inj_is_fun RS apply_rangeI, |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
354 |
inj_converse_fun RS apply_rangeI, |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
355 |
inj_converse_fun RS apply_funtype, |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
356 |
left_inverse, right_inverse, nat_0I, nat_succI, |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
357 |
nat_case_0, nat_case_succ] |
516 | 358 |
setloop split_tac [expand_if]) 1); |
760 | 359 |
qed "nat_cons_lepoll"; |
516 | 360 |
|
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
361 |
goal CardinalArith.thy "!!i. nat lepoll A ==> cons(u,A) eqpoll A"; |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
362 |
by (etac (nat_cons_lepoll RS eqpollI) 1); |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
363 |
by (rtac (subset_consI RS subset_imp_lepoll) 1); |
760 | 364 |
qed "nat_cons_eqpoll"; |
437 | 365 |
|
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
366 |
(*Specialized version required below*) |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
367 |
goalw CardinalArith.thy [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A"; |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
368 |
by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1); |
760 | 369 |
qed "nat_succ_eqpoll"; |
437 | 370 |
|
488 | 371 |
goalw CardinalArith.thy [InfCard_def] "InfCard(nat)"; |
372 |
by (fast_tac (ZF_cs addIs [Card_nat, le_refl, Card_is_Ord]) 1); |
|
760 | 373 |
qed "InfCard_nat"; |
488 | 374 |
|
484 | 375 |
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)"; |
437 | 376 |
by (etac conjunct1 1); |
760 | 377 |
qed "InfCard_is_Card"; |
437 | 378 |
|
523 | 379 |
goalw CardinalArith.thy [InfCard_def] |
380 |
"!!K L. [| InfCard(K); Card(L) |] ==> InfCard(K Un L)"; |
|
381 |
by (asm_simp_tac (ZF_ss addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), |
|
382 |
Card_is_Ord]) 1); |
|
760 | 383 |
qed "InfCard_Un"; |
523 | 384 |
|
437 | 385 |
(*Kunen's Lemma 10.11*) |
484 | 386 |
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)"; |
437 | 387 |
by (etac conjE 1); |
388 |
by (rtac (ltI RS non_succ_LimitI) 1); |
|
389 |
by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1); |
|
390 |
by (etac Card_is_Ord 1); |
|
391 |
by (safe_tac (ZF_cs addSDs [Limit_nat RS Limit_le_succD])); |
|
392 |
by (forward_tac [Card_is_Ord RS Ord_succD] 1); |
|
393 |
by (rewtac Card_def); |
|
394 |
by (res_inst_tac [("i", "succ(y)")] lt_irrefl 1); |
|
395 |
by (dtac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1); |
|
396 |
(*Tricky combination of substitutions; backtracking needed*) |
|
397 |
by (etac ssubst 1 THEN etac ssubst 1 THEN rtac Ord_cardinal_le 1); |
|
398 |
by (assume_tac 1); |
|
760 | 399 |
qed "InfCard_is_Limit"; |
437 | 400 |
|
401 |
||
402 |
||
403 |
(*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***) |
|
404 |
||
405 |
(*A general fact about ordermap*) |
|
406 |
goalw Cardinal.thy [eqpoll_def] |
|
407 |
"!!A. [| well_ord(A,r); x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)"; |
|
408 |
by (rtac exI 1); |
|
409 |
by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, well_ord_is_wf]) 1); |
|
467 | 410 |
by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1); |
437 | 411 |
by (rtac pred_subset 1); |
760 | 412 |
qed "ordermap_eqpoll_pred"; |
437 | 413 |
|
414 |
(** Establishing the well-ordering **) |
|
415 |
||
416 |
goalw CardinalArith.thy [inj_def] |
|
484 | 417 |
"!!K. Ord(K) ==> \ |
418 |
\ (lam z:K*K. split(%x y. <x Un y, <x, y>>, z)) : inj(K*K, K*K*K)"; |
|
437 | 419 |
by (safe_tac ZF_cs); |
420 |
by (fast_tac (ZF_cs addIs [lam_type, Un_least_lt RS ltD, ltI] |
|
421 |
addSEs [split_type]) 1); |
|
422 |
by (asm_full_simp_tac ZF_ss 1); |
|
760 | 423 |
qed "csquare_lam_inj"; |
437 | 424 |
|
425 |
goalw CardinalArith.thy [csquare_rel_def] |
|
484 | 426 |
"!!K. Ord(K) ==> well_ord(K*K, csquare_rel(K))"; |
437 | 427 |
by (rtac (csquare_lam_inj RS well_ord_rvimage) 1); |
428 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
760 | 429 |
qed "well_ord_csquare"; |
437 | 430 |
|
431 |
(** Characterising initial segments of the well-ordering **) |
|
432 |
||
433 |
goalw CardinalArith.thy [csquare_rel_def] |
|
484 | 434 |
"!!K. [| x<K; y<K; z<K |] ==> \ |
435 |
\ <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z"; |
|
437 | 436 |
by (REPEAT (etac ltE 1)); |
437 |
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff, |
|
438 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
|
439 |
by (safe_tac (ZF_cs addSEs [mem_irrefl] |
|
440 |
addSIs [Un_upper1_le, Un_upper2_le])); |
|
441 |
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [lt_def, succI2, Ord_succ]))); |
|
760 | 442 |
qed "csquareD_lemma"; |
437 | 443 |
val csquareD = csquareD_lemma RS mp |> standard; |
444 |
||
445 |
goalw CardinalArith.thy [pred_def] |
|
484 | 446 |
"!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)"; |
437 | 447 |
by (safe_tac (lemmas_cs addSEs [SigmaE])); (*avoids using succCI*) |
448 |
by (rtac (csquareD RS conjE) 1); |
|
449 |
by (rewtac lt_def); |
|
450 |
by (assume_tac 4); |
|
451 |
by (ALLGOALS (fast_tac ZF_cs)); |
|
760 | 452 |
qed "pred_csquare_subset"; |
437 | 453 |
|
454 |
goalw CardinalArith.thy [csquare_rel_def] |
|
484 | 455 |
"!!K. [| x<z; y<z; z<K |] ==> \ |
456 |
\ <<x,y>, <z,z>> : csquare_rel(K)"; |
|
457 |
by (subgoals_tac ["x<K", "y<K"] 1); |
|
437 | 458 |
by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2)); |
459 |
by (REPEAT (etac ltE 1)); |
|
460 |
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff, |
|
461 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
|
760 | 462 |
qed "csquare_ltI"; |
437 | 463 |
|
464 |
(*Part of the traditional proof. UNUSED since it's harder to prove & apply *) |
|
465 |
goalw CardinalArith.thy [csquare_rel_def] |
|
484 | 466 |
"!!K. [| x le z; y le z; z<K |] ==> \ |
467 |
\ <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z"; |
|
468 |
by (subgoals_tac ["x<K", "y<K"] 1); |
|
437 | 469 |
by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2)); |
470 |
by (REPEAT (etac ltE 1)); |
|
471 |
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff, |
|
472 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
|
473 |
by (REPEAT_FIRST (etac succE)); |
|
474 |
by (ALLGOALS |
|
475 |
(asm_simp_tac (ZF_ss addsimps [subset_Un_iff RS iff_sym, |
|
476 |
subset_Un_iff2 RS iff_sym, OrdmemD]))); |
|
760 | 477 |
qed "csquare_or_eqI"; |
437 | 478 |
|
479 |
(** The cardinality of initial segments **) |
|
480 |
||
481 |
goal CardinalArith.thy |
|
484 | 482 |
"!!K. [| InfCard(K); x<K; y<K; z=succ(x Un y) |] ==> \ |
483 |
\ ordermap(K*K, csquare_rel(K)) ` <x,y> lepoll \ |
|
484 |
\ ordermap(K*K, csquare_rel(K)) ` <z,z>"; |
|
485 |
by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1); |
|
437 | 486 |
by (etac (InfCard_is_Card RS Card_is_Ord RS well_ord_csquare) 2); |
487 |
by (fast_tac (ZF_cs addSIs [Un_least_lt, InfCard_is_Limit, Limit_has_succ]) 2); |
|
488 |
by (rtac (OrdmemD RS subset_imp_lepoll) 1); |
|
467 | 489 |
by (res_inst_tac [("z1","z")] (csquare_ltI RS ordermap_mono) 1); |
437 | 490 |
by (etac well_ord_is_wf 4); |
491 |
by (ALLGOALS |
|
492 |
(fast_tac (ZF_cs addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap] |
|
493 |
addSEs [ltE]))); |
|
760 | 494 |
qed "ordermap_z_lepoll"; |
437 | 495 |
|
484 | 496 |
(*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *) |
437 | 497 |
goalw CardinalArith.thy [cmult_def] |
484 | 498 |
"!!K. [| InfCard(K); x<K; y<K; z=succ(x Un y) |] ==> \ |
499 |
\ | ordermap(K*K, csquare_rel(K)) ` <x,y> | le |succ(z)| |*| |succ(z)|"; |
|
767 | 500 |
by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1); |
437 | 501 |
by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1)); |
484 | 502 |
by (subgoals_tac ["z<K"] 1); |
437 | 503 |
by (fast_tac (ZF_cs addSIs [Un_least_lt, InfCard_is_Limit, |
504 |
Limit_has_succ]) 2); |
|
505 |
by (rtac (ordermap_z_lepoll RS lepoll_trans) 1); |
|
506 |
by (REPEAT_SOME assume_tac); |
|
507 |
by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1); |
|
508 |
by (etac (InfCard_is_Card RS Card_is_Ord RS well_ord_csquare) 1); |
|
509 |
by (fast_tac (ZF_cs addIs [ltD]) 1); |
|
510 |
by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN |
|
511 |
assume_tac 1); |
|
512 |
by (REPEAT_FIRST (etac ltE)); |
|
513 |
by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1); |
|
514 |
by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll))); |
|
760 | 515 |
qed "ordermap_csquare_le"; |
437 | 516 |
|
484 | 517 |
(*Kunen: "... so the order type <= K" *) |
437 | 518 |
goal CardinalArith.thy |
484 | 519 |
"!!K. [| InfCard(K); ALL y:K. InfCard(y) --> y |*| y = y |] ==> \ |
520 |
\ ordertype(K*K, csquare_rel(K)) le K"; |
|
437 | 521 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); |
522 |
by (rtac all_lt_imp_le 1); |
|
523 |
by (assume_tac 1); |
|
524 |
by (etac (well_ord_csquare RS Ord_ordertype) 1); |
|
525 |
by (rtac Card_lt_imp_lt 1); |
|
526 |
by (etac InfCard_is_Card 3); |
|
527 |
by (etac ltE 2 THEN assume_tac 2); |
|
528 |
by (asm_full_simp_tac (ZF_ss addsimps [ordertype_unfold]) 1); |
|
529 |
by (safe_tac (ZF_cs addSEs [ltE])); |
|
530 |
by (subgoals_tac ["Ord(xb)", "Ord(y)"] 1); |
|
531 |
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2)); |
|
532 |
by (rtac (ordermap_csquare_le RS lt_trans1) 1 THEN |
|
533 |
REPEAT (ares_tac [refl] 1 ORELSE etac ltI 1)); |
|
534 |
by (res_inst_tac [("i","xb Un y"), ("j","nat")] Ord_linear2 1 THEN |
|
535 |
REPEAT (ares_tac [Ord_Un, Ord_nat] 1)); |
|
536 |
(*the finite case: xb Un y < nat *) |
|
537 |
by (res_inst_tac [("j", "nat")] lt_trans2 1); |
|
538 |
by (asm_full_simp_tac (FOL_ss addsimps [InfCard_def]) 2); |
|
539 |
by (asm_full_simp_tac |
|
540 |
(ZF_ss addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type, |
|
541 |
nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1); |
|
542 |
(*case nat le (xb Un y), equivalently InfCard(xb Un y) *) |
|
543 |
by (asm_full_simp_tac |
|
544 |
(ZF_ss addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong, |
|
545 |
le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt, |
|
546 |
Ord_Un, ltI, nat_le_cardinal, |
|
547 |
Ord_cardinal_le RS lt_trans1 RS ltD]) 1); |
|
760 | 548 |
qed "ordertype_csquare_le"; |
437 | 549 |
|
550 |
(*This lemma can easily be generalized to premise well_ord(A*A,r) *) |
|
551 |
goalw CardinalArith.thy [cmult_def] |
|
484 | 552 |
"!!K. Ord(K) ==> K |*| K = |ordertype(K*K, csquare_rel(K))|"; |
437 | 553 |
by (rtac cardinal_cong 1); |
554 |
by (rewtac eqpoll_def); |
|
555 |
by (rtac exI 1); |
|
467 | 556 |
by (etac (well_ord_csquare RS ordermap_bij) 1); |
760 | 557 |
qed "csquare_eq_ordertype"; |
437 | 558 |
|
559 |
(*Main result: Kunen's Theorem 10.12*) |
|
484 | 560 |
goal CardinalArith.thy "!!K. InfCard(K) ==> K |*| K = K"; |
437 | 561 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); |
562 |
by (etac rev_mp 1); |
|
484 | 563 |
by (trans_ind_tac "K" [] 1); |
437 | 564 |
by (rtac impI 1); |
565 |
by (rtac le_anti_sym 1); |
|
566 |
by (etac (InfCard_is_Card RS cmult_square_le) 2); |
|
567 |
by (rtac (ordertype_csquare_le RSN (2, le_trans)) 1); |
|
568 |
by (assume_tac 2); |
|
569 |
by (assume_tac 2); |
|
570 |
by (asm_simp_tac |
|
571 |
(ZF_ss addsimps [csquare_eq_ordertype, Ord_cardinal_le, |
|
572 |
well_ord_csquare RS Ord_ordertype]) 1); |
|
760 | 573 |
qed "InfCard_csquare_eq"; |
484 | 574 |
|
767 | 575 |
(*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*) |
484 | 576 |
goal CardinalArith.thy |
577 |
"!!A. [| well_ord(A,r); InfCard(|A|) |] ==> A*A eqpoll A"; |
|
578 |
by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1); |
|
579 |
by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1)); |
|
580 |
by (resolve_tac [well_ord_cardinal_eqE] 1); |
|
581 |
by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1)); |
|
582 |
by (asm_simp_tac (ZF_ss addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1); |
|
760 | 583 |
qed "well_ord_InfCard_square_eq"; |
484 | 584 |
|
767 | 585 |
(** Toward's Kunen's Corollary 10.13 (1) **) |
586 |
||
587 |
goal CardinalArith.thy "!!K. [| InfCard(K); L le K; 0<L |] ==> K |*| L = K"; |
|
588 |
by (resolve_tac [le_anti_sym] 1); |
|
589 |
by (eresolve_tac [ltE] 2 THEN |
|
590 |
REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2)); |
|
591 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1); |
|
592 |
by (resolve_tac [cmult_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1)); |
|
593 |
by (asm_simp_tac (ZF_ss addsimps [InfCard_csquare_eq]) 1); |
|
594 |
val InfCard_le_cmult_eq = result(); |
|
595 |
||
596 |
(*Corollary 10.13 (1), for cardinal multiplication*) |
|
597 |
goal CardinalArith.thy |
|
598 |
"!!K. [| InfCard(K); InfCard(L) |] ==> K |*| L = K Un L"; |
|
599 |
by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1); |
|
600 |
by (typechk_tac [InfCard_is_Card, Card_is_Ord]); |
|
601 |
by (resolve_tac [cmult_commute RS ssubst] 1); |
|
602 |
by (resolve_tac [Un_commute RS ssubst] 1); |
|
603 |
by (ALLGOALS |
|
604 |
(asm_simp_tac |
|
605 |
(ZF_ss addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq, |
|
606 |
subset_Un_iff2 RS iffD1, le_imp_subset]))); |
|
607 |
val InfCard_cmult_eq = result(); |
|
608 |
||
609 |
(*This proof appear to be the simplest!*) |
|
610 |
goal CardinalArith.thy "!!K. InfCard(K) ==> K |+| K = K"; |
|
611 |
by (asm_simp_tac |
|
612 |
(ZF_ss addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1); |
|
613 |
by (resolve_tac [InfCard_le_cmult_eq] 1); |
|
614 |
by (typechk_tac [Ord_0, le_refl, leI]); |
|
615 |
by (typechk_tac [InfCard_is_Limit, Limit_has_0, Limit_has_succ]); |
|
616 |
val InfCard_cdouble_eq = result(); |
|
617 |
||
618 |
(*Corollary 10.13 (1), for cardinal addition*) |
|
619 |
goal CardinalArith.thy "!!K. [| InfCard(K); L le K |] ==> K |+| L = K"; |
|
620 |
by (resolve_tac [le_anti_sym] 1); |
|
621 |
by (eresolve_tac [ltE] 2 THEN |
|
622 |
REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2)); |
|
623 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1); |
|
624 |
by (resolve_tac [cadd_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1)); |
|
625 |
by (asm_simp_tac (ZF_ss addsimps [InfCard_cdouble_eq]) 1); |
|
626 |
val InfCard_le_cadd_eq = result(); |
|
627 |
||
628 |
goal CardinalArith.thy |
|
629 |
"!!K. [| InfCard(K); InfCard(L) |] ==> K |+| L = K Un L"; |
|
630 |
by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1); |
|
631 |
by (typechk_tac [InfCard_is_Card, Card_is_Ord]); |
|
632 |
by (resolve_tac [cadd_commute RS ssubst] 1); |
|
633 |
by (resolve_tac [Un_commute RS ssubst] 1); |
|
634 |
by (ALLGOALS |
|
635 |
(asm_simp_tac |
|
636 |
(ZF_ss addsimps [InfCard_le_cadd_eq, |
|
637 |
subset_Un_iff2 RS iffD1, le_imp_subset]))); |
|
638 |
val InfCard_cadd_eq = result(); |
|
639 |
||
640 |
(*The other part, Corollary 10.13 (2), refers to the cardinality of the set |
|
641 |
of all n-tuples of elements of K. A better version for the Isabelle theory |
|
642 |
might be InfCard(K) ==> |list(K)| = K. |
|
643 |
*) |
|
484 | 644 |
|
645 |
(*** For every cardinal number there exists a greater one |
|
646 |
[Kunen's Theorem 10.16, which would be trivial using AC] ***) |
|
647 |
||
648 |
goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))"; |
|
649 |
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1); |
|
650 |
by (safe_tac (ZF_cs addSIs [Ord_ordertype])); |
|
651 |
bw Transset_def; |
|
652 |
by (safe_tac ZF_cs); |
|
653 |
by (rtac (ordertype_subset RS exE) 1 THEN REPEAT (assume_tac 1)); |
|
654 |
by (resolve_tac [UN_I] 1); |
|
655 |
by (resolve_tac [ReplaceI] 2); |
|
656 |
by (ALLGOALS (fast_tac (ZF_cs addSEs [well_ord_subset]))); |
|
760 | 657 |
qed "Ord_jump_cardinal"; |
484 | 658 |
|
659 |
(*Allows selective unfolding. Less work than deriving intro/elim rules*) |
|
660 |
goalw CardinalArith.thy [jump_cardinal_def] |
|
661 |
"i : jump_cardinal(K) <-> \ |
|
662 |
\ (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))"; |
|
663 |
by (fast_tac subset_cs 1); (*It's vital to avoid reasoning about <=*) |
|
760 | 664 |
qed "jump_cardinal_iff"; |
484 | 665 |
|
666 |
(*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) |
|
667 |
goal CardinalArith.thy "!!K. Ord(K) ==> K < jump_cardinal(K)"; |
|
668 |
by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1); |
|
669 |
by (resolve_tac [jump_cardinal_iff RS iffD2] 1); |
|
670 |
by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel])); |
|
671 |
by (resolve_tac [subset_refl] 2); |
|
672 |
by (asm_simp_tac (ZF_ss addsimps [Memrel_def, subset_iff]) 1); |
|
673 |
by (asm_simp_tac (ZF_ss addsimps [ordertype_Memrel]) 1); |
|
760 | 674 |
qed "K_lt_jump_cardinal"; |
484 | 675 |
|
676 |
(*The proof by contradiction: the bijection f yields a wellordering of X |
|
677 |
whose ordertype is jump_cardinal(K). *) |
|
678 |
goal CardinalArith.thy |
|
679 |
"!!K. [| well_ord(X,r); r <= K * K; X <= K; \ |
|
680 |
\ f : bij(ordertype(X,r), jump_cardinal(K)) \ |
|
681 |
\ |] ==> jump_cardinal(K) : jump_cardinal(K)"; |
|
682 |
by (subgoal_tac "f O ordermap(X,r): bij(X, jump_cardinal(K))" 1); |
|
683 |
by (REPEAT (ares_tac [comp_bij, ordermap_bij] 2)); |
|
684 |
by (resolve_tac [jump_cardinal_iff RS iffD2] 1); |
|
685 |
by (REPEAT_FIRST (resolve_tac [exI, conjI])); |
|
686 |
by (rtac ([rvimage_type, Sigma_mono] MRS subset_trans) 1); |
|
687 |
by (REPEAT (assume_tac 1)); |
|
688 |
by (etac (bij_is_inj RS well_ord_rvimage) 1); |
|
689 |
by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1); |
|
690 |
by (asm_simp_tac |
|
691 |
(ZF_ss addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), |
|
692 |
ordertype_Memrel, Ord_jump_cardinal]) 1); |
|
760 | 693 |
qed "Card_jump_cardinal_lemma"; |
484 | 694 |
|
695 |
(*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*) |
|
696 |
goal CardinalArith.thy "Card(jump_cardinal(K))"; |
|
697 |
by (rtac (Ord_jump_cardinal RS CardI) 1); |
|
698 |
by (rewrite_goals_tac [eqpoll_def]); |
|
699 |
by (safe_tac (ZF_cs addSDs [ltD, jump_cardinal_iff RS iffD1])); |
|
700 |
by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1)); |
|
760 | 701 |
qed "Card_jump_cardinal"; |
484 | 702 |
|
703 |
(*** Basic properties of successor cardinals ***) |
|
704 |
||
705 |
goalw CardinalArith.thy [csucc_def] |
|
706 |
"!!K. Ord(K) ==> Card(csucc(K)) & K < csucc(K)"; |
|
707 |
by (rtac LeastI 1); |
|
708 |
by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal, |
|
709 |
Ord_jump_cardinal] 1)); |
|
760 | 710 |
qed "csucc_basic"; |
484 | 711 |
|
712 |
val Card_csucc = csucc_basic RS conjunct1 |> standard; |
|
713 |
||
714 |
val lt_csucc = csucc_basic RS conjunct2 |> standard; |
|
715 |
||
517 | 716 |
goal CardinalArith.thy "!!K. Ord(K) ==> 0 < csucc(K)"; |
717 |
by (resolve_tac [[Ord_0_le, lt_csucc] MRS lt_trans1] 1); |
|
718 |
by (REPEAT (assume_tac 1)); |
|
760 | 719 |
qed "Ord_0_lt_csucc"; |
517 | 720 |
|
484 | 721 |
goalw CardinalArith.thy [csucc_def] |
722 |
"!!K L. [| Card(L); K<L |] ==> csucc(K) le L"; |
|
723 |
by (rtac Least_le 1); |
|
724 |
by (REPEAT (ares_tac [conjI, Card_is_Ord] 1)); |
|
760 | 725 |
qed "csucc_le"; |
484 | 726 |
|
727 |
goal CardinalArith.thy |
|
728 |
"!!K. [| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K"; |
|
729 |
by (resolve_tac [iffI] 1); |
|
730 |
by (resolve_tac [Card_lt_imp_lt] 2); |
|
731 |
by (eresolve_tac [lt_trans1] 2); |
|
732 |
by (REPEAT (ares_tac [lt_csucc, Card_csucc, Card_is_Ord] 2)); |
|
733 |
by (resolve_tac [notI RS not_lt_imp_le] 1); |
|
734 |
by (resolve_tac [Card_cardinal RS csucc_le RS lt_trans1 RS lt_irrefl] 1); |
|
735 |
by (assume_tac 1); |
|
736 |
by (resolve_tac [Ord_cardinal_le RS lt_trans1] 1); |
|
737 |
by (REPEAT (ares_tac [Ord_cardinal] 1 |
|
738 |
ORELSE eresolve_tac [ltE, Card_is_Ord] 1)); |
|
760 | 739 |
qed "lt_csucc_iff"; |
484 | 740 |
|
741 |
goal CardinalArith.thy |
|
742 |
"!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K"; |
|
743 |
by (asm_simp_tac |
|
744 |
(ZF_ss addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1); |
|
760 | 745 |
qed "Card_lt_csucc_iff"; |
488 | 746 |
|
747 |
goalw CardinalArith.thy [InfCard_def] |
|
748 |
"!!K. InfCard(K) ==> InfCard(csucc(K))"; |
|
749 |
by (asm_simp_tac (ZF_ss addsimps [Card_csucc, Card_is_Ord, |
|
750 |
lt_csucc RS leI RSN (2,le_trans)]) 1); |
|
760 | 751 |
qed "InfCard_csucc"; |
517 | 752 |
|
753 |
val Limit_csucc = InfCard_csucc RS InfCard_is_Limit |> standard; |