| author | haftmann | 
| Wed, 29 Apr 2009 14:20:26 +0200 | |
| changeset 31021 | 53642251a04f | 
| parent 27517 | c055e1d49285 | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/CardinalArith.thy | 
| 437 | 2 | ID: $Id$ | 
| 1478 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 437 | 4 | Copyright 1994 University of Cambridge | 
| 5 | ||
| 13328 | 6 | *) | 
| 13216 | 7 | |
| 13328 | 8 | header{*Cardinal Arithmetic Without the Axiom of Choice*}
 | 
| 437 | 9 | |
| 16417 | 10 | theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin | 
| 467 | 11 | |
| 24893 | 12 | definition | 
| 13 | InfCard :: "i=>o" where | |
| 12667 | 14 | "InfCard(i) == Card(i) & nat le i" | 
| 437 | 15 | |
| 24893 | 16 | definition | 
| 17 | cmult :: "[i,i]=>i" (infixl "|*|" 70) where | |
| 12667 | 18 | "i |*| j == |i*j|" | 
| 19 | ||
| 24893 | 20 | definition | 
| 21 | cadd :: "[i,i]=>i" (infixl "|+|" 65) where | |
| 12667 | 22 | "i |+| j == |i+j|" | 
| 437 | 23 | |
| 24893 | 24 | definition | 
| 25 | csquare_rel :: "i=>i" where | |
| 12667 | 26 | "csquare_rel(K) == | 
| 27 | rvimage(K*K, | |
| 28 | lam <x,y>:K*K. <x Un y, x, y>, | |
| 29 | rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))" | |
| 437 | 30 | |
| 24893 | 31 | definition | 
| 32 | jump_cardinal :: "i=>i" where | |
| 14883 | 33 |     --{*This def is more complex than Kunen's but it more easily proved to
 | 
| 34 | be a cardinal*} | |
| 12667 | 35 | "jump_cardinal(K) == | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 36 |          \<Union>X\<in>Pow(K). {z. r: Pow(K*K), well_ord(X,r) & z = ordertype(X,r)}"
 | 
| 12667 | 37 | |
| 24893 | 38 | definition | 
| 39 | csucc :: "i=>i" where | |
| 14883 | 40 |     --{*needed because @{term "jump_cardinal(K)"} might not be the successor
 | 
| 41 |         of @{term K}*}
 | |
| 12667 | 42 | "csucc(K) == LEAST L. Card(L) & K<L" | 
| 484 | 43 | |
| 24893 | 44 | notation (xsymbols output) | 
| 45 | cadd (infixl "\<oplus>" 65) and | |
| 46 | cmult (infixl "\<otimes>" 70) | |
| 47 | ||
| 48 | notation (HTML output) | |
| 49 | cadd (infixl "\<oplus>" 65) and | |
| 50 | cmult (infixl "\<otimes>" 70) | |
| 12667 | 51 | |
| 52 | ||
| 53 | lemma Card_Union [simp,intro,TC]: "(ALL x:A. Card(x)) ==> Card(Union(A))" | |
| 54 | apply (rule CardI) | |
| 55 | apply (simp add: Card_is_Ord) | |
| 56 | apply (clarify dest!: ltD) | |
| 57 | apply (drule bspec, assumption) | |
| 58 | apply (frule lt_Card_imp_lesspoll, blast intro: ltI Card_is_Ord) | |
| 59 | apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll]) | |
| 60 | apply (drule lesspoll_trans1, assumption) | |
| 13216 | 61 | apply (subgoal_tac "B \<lesssim> \<Union>A") | 
| 12667 | 62 | apply (drule lesspoll_trans1, assumption, blast) | 
| 63 | apply (blast intro: subset_imp_lepoll) | |
| 64 | done | |
| 65 | ||
| 14883 | 66 | lemma Card_UN: "(!!x. x:A ==> Card(K(x))) ==> Card(\<Union>x\<in>A. K(x))" | 
| 12667 | 67 | by (blast intro: Card_Union) | 
| 68 | ||
| 69 | lemma Card_OUN [simp,intro,TC]: | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 70 | "(!!x. x:A ==> Card(K(x))) ==> Card(\<Union>x<A. K(x))" | 
| 12667 | 71 | by (simp add: OUnion_def Card_0) | 
| 9654 
9754ba005b64
X-symbols for ordinal, cardinal, integer arithmetic
 paulson parents: 
9548diff
changeset | 72 | |
| 12776 | 73 | lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat" | 
| 74 | apply (unfold lesspoll_def) | |
| 75 | apply (rule conjI) | |
| 76 | apply (erule OrdmemD [THEN subset_imp_lepoll], rule Ord_nat) | |
| 77 | apply (rule notI) | |
| 78 | apply (erule eqpollE) | |
| 79 | apply (rule succ_lepoll_natE) | |
| 80 | apply (blast intro: nat_succI [THEN OrdmemD, THEN subset_imp_lepoll] | |
| 12820 | 81 | lepoll_trans, assumption) | 
| 12776 | 82 | done | 
| 83 | ||
| 84 | lemma in_Card_imp_lesspoll: "[| Card(K); b \<in> K |] ==> b \<prec> K" | |
| 85 | apply (unfold lesspoll_def) | |
| 86 | apply (simp add: Card_iff_initial) | |
| 87 | apply (fast intro!: le_imp_lepoll ltI leI) | |
| 88 | done | |
| 89 | ||
| 14883 | 90 | lemma lesspoll_lemma: "[| ~ A \<prec> B; C \<prec> B |] ==> A - C \<noteq> 0" | 
| 12776 | 91 | apply (unfold lesspoll_def) | 
| 92 | apply (fast dest!: Diff_eq_0_iff [THEN iffD1, THEN subset_imp_lepoll] | |
| 93 | intro!: eqpollI elim: notE | |
| 94 | elim!: eqpollE lepoll_trans) | |
| 95 | done | |
| 96 | ||
| 13216 | 97 | |
| 13356 | 98 | subsection{*Cardinal addition*}
 | 
| 13216 | 99 | |
| 13328 | 100 | text{*Note: Could omit proving the algebraic laws for cardinal addition and
 | 
| 101 | multiplication. On finite cardinals these operations coincide with | |
| 102 | addition and multiplication of natural numbers; on infinite cardinals they | |
| 103 | coincide with union (maximum). Either way we get most laws for free.*} | |
| 104 | ||
| 14883 | 105 | subsubsection{*Cardinal addition is commutative*}
 | 
| 13216 | 106 | |
| 107 | lemma sum_commute_eqpoll: "A+B \<approx> B+A" | |
| 108 | apply (unfold eqpoll_def) | |
| 109 | apply (rule exI) | |
| 110 | apply (rule_tac c = "case(Inr,Inl)" and d = "case(Inr,Inl)" in lam_bijective) | |
| 111 | apply auto | |
| 112 | done | |
| 113 | ||
| 114 | lemma cadd_commute: "i |+| j = j |+| i" | |
| 115 | apply (unfold cadd_def) | |
| 116 | apply (rule sum_commute_eqpoll [THEN cardinal_cong]) | |
| 117 | done | |
| 118 | ||
| 14883 | 119 | subsubsection{*Cardinal addition is associative*}
 | 
| 13216 | 120 | |
| 121 | lemma sum_assoc_eqpoll: "(A+B)+C \<approx> A+(B+C)" | |
| 122 | apply (unfold eqpoll_def) | |
| 123 | apply (rule exI) | |
| 124 | apply (rule sum_assoc_bij) | |
| 125 | done | |
| 126 | ||
| 127 | (*Unconditional version requires AC*) | |
| 128 | lemma well_ord_cadd_assoc: | |
| 129 | "[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] | |
| 130 | ==> (i |+| j) |+| k = i |+| (j |+| k)" | |
| 131 | apply (unfold cadd_def) | |
| 132 | apply (rule cardinal_cong) | |
| 133 | apply (rule eqpoll_trans) | |
| 134 | apply (rule sum_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl]) | |
| 13221 | 135 | apply (blast intro: well_ord_radd ) | 
| 13216 | 136 | apply (rule sum_assoc_eqpoll [THEN eqpoll_trans]) | 
| 137 | apply (rule eqpoll_sym) | |
| 138 | apply (rule sum_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll]) | |
| 13221 | 139 | apply (blast intro: well_ord_radd ) | 
| 13216 | 140 | done | 
| 141 | ||
| 14883 | 142 | subsubsection{*0 is the identity for addition*}
 | 
| 13216 | 143 | |
| 144 | lemma sum_0_eqpoll: "0+A \<approx> A" | |
| 145 | apply (unfold eqpoll_def) | |
| 146 | apply (rule exI) | |
| 147 | apply (rule bij_0_sum) | |
| 148 | done | |
| 149 | ||
| 150 | lemma cadd_0 [simp]: "Card(K) ==> 0 |+| K = K" | |
| 151 | apply (unfold cadd_def) | |
| 152 | apply (simp add: sum_0_eqpoll [THEN cardinal_cong] Card_cardinal_eq) | |
| 153 | done | |
| 154 | ||
| 14883 | 155 | subsubsection{*Addition by another cardinal*}
 | 
| 13216 | 156 | |
| 157 | lemma sum_lepoll_self: "A \<lesssim> A+B" | |
| 158 | apply (unfold lepoll_def inj_def) | |
| 159 | apply (rule_tac x = "lam x:A. Inl (x) " in exI) | |
| 13221 | 160 | apply simp | 
| 13216 | 161 | done | 
| 162 | ||
| 163 | (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) | |
| 164 | ||
| 165 | lemma cadd_le_self: | |
| 166 | "[| Card(K); Ord(L) |] ==> K le (K |+| L)" | |
| 167 | apply (unfold cadd_def) | |
| 13221 | 168 | apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_Card_le], | 
| 169 | assumption) | |
| 13216 | 170 | apply (rule_tac [2] sum_lepoll_self) | 
| 171 | apply (blast intro: well_ord_radd well_ord_Memrel Card_is_Ord) | |
| 172 | done | |
| 173 | ||
| 14883 | 174 | subsubsection{*Monotonicity of addition*}
 | 
| 13216 | 175 | |
| 176 | lemma sum_lepoll_mono: | |
| 13221 | 177 | "[| A \<lesssim> C; B \<lesssim> D |] ==> A + B \<lesssim> C + D" | 
| 13216 | 178 | apply (unfold lepoll_def) | 
| 13221 | 179 | apply (elim exE) | 
| 13216 | 180 | apply (rule_tac x = "lam z:A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI) | 
| 13221 | 181 | apply (rule_tac d = "case (%w. Inl(converse(f) `w), %y. Inr(converse(fa) ` y))" | 
| 13216 | 182 | in lam_injective) | 
| 13221 | 183 | apply (typecheck add: inj_is_fun, auto) | 
| 13216 | 184 | done | 
| 185 | ||
| 186 | lemma cadd_le_mono: | |
| 187 | "[| K' le K; L' le L |] ==> (K' |+| L') le (K |+| L)" | |
| 188 | apply (unfold cadd_def) | |
| 189 | apply (safe dest!: le_subset_iff [THEN iffD1]) | |
| 190 | apply (rule well_ord_lepoll_imp_Card_le) | |
| 191 | apply (blast intro: well_ord_radd well_ord_Memrel) | |
| 192 | apply (blast intro: sum_lepoll_mono subset_imp_lepoll) | |
| 193 | done | |
| 194 | ||
| 14883 | 195 | subsubsection{*Addition of finite cardinals is "ordinary" addition*}
 | 
| 13216 | 196 | |
| 197 | lemma sum_succ_eqpoll: "succ(A)+B \<approx> succ(A+B)" | |
| 198 | apply (unfold eqpoll_def) | |
| 199 | apply (rule exI) | |
| 200 | apply (rule_tac c = "%z. if z=Inl (A) then A+B else z" | |
| 201 | and d = "%z. if z=A+B then Inl (A) else z" in lam_bijective) | |
| 13221 | 202 | apply simp_all | 
| 13216 | 203 | apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+ | 
| 204 | done | |
| 205 | ||
| 206 | (*Pulling the succ(...) outside the |...| requires m, n: nat *) | |
| 207 | (*Unconditional version requires AC*) | |
| 208 | lemma cadd_succ_lemma: | |
| 209 | "[| Ord(m); Ord(n) |] ==> succ(m) |+| n = |succ(m |+| n)|" | |
| 210 | apply (unfold cadd_def) | |
| 211 | apply (rule sum_succ_eqpoll [THEN cardinal_cong, THEN trans]) | |
| 212 | apply (rule succ_eqpoll_cong [THEN cardinal_cong]) | |
| 213 | apply (rule well_ord_cardinal_eqpoll [THEN eqpoll_sym]) | |
| 214 | apply (blast intro: well_ord_radd well_ord_Memrel) | |
| 215 | done | |
| 216 | ||
| 217 | lemma nat_cadd_eq_add: "[| m: nat; n: nat |] ==> m |+| n = m#+n" | |
| 13244 | 218 | apply (induct_tac m) | 
| 13216 | 219 | apply (simp add: nat_into_Card [THEN cadd_0]) | 
| 220 | apply (simp add: cadd_succ_lemma nat_into_Card [THEN Card_cardinal_eq]) | |
| 221 | done | |
| 222 | ||
| 223 | ||
| 13356 | 224 | subsection{*Cardinal multiplication*}
 | 
| 13216 | 225 | |
| 14883 | 226 | subsubsection{*Cardinal multiplication is commutative*}
 | 
| 13216 | 227 | |
| 228 | (*Easier to prove the two directions separately*) | |
| 229 | lemma prod_commute_eqpoll: "A*B \<approx> B*A" | |
| 230 | apply (unfold eqpoll_def) | |
| 231 | apply (rule exI) | |
| 13221 | 232 | apply (rule_tac c = "%<x,y>.<y,x>" and d = "%<x,y>.<y,x>" in lam_bijective, | 
| 233 | auto) | |
| 13216 | 234 | done | 
| 235 | ||
| 236 | lemma cmult_commute: "i |*| j = j |*| i" | |
| 237 | apply (unfold cmult_def) | |
| 238 | apply (rule prod_commute_eqpoll [THEN cardinal_cong]) | |
| 239 | done | |
| 240 | ||
| 14883 | 241 | subsubsection{*Cardinal multiplication is associative*}
 | 
| 13216 | 242 | |
| 243 | lemma prod_assoc_eqpoll: "(A*B)*C \<approx> A*(B*C)" | |
| 244 | apply (unfold eqpoll_def) | |
| 245 | apply (rule exI) | |
| 246 | apply (rule prod_assoc_bij) | |
| 247 | done | |
| 248 | ||
| 249 | (*Unconditional version requires AC*) | |
| 250 | lemma well_ord_cmult_assoc: | |
| 251 | "[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] | |
| 252 | ==> (i |*| j) |*| k = i |*| (j |*| k)" | |
| 253 | apply (unfold cmult_def) | |
| 254 | apply (rule cardinal_cong) | |
| 13221 | 255 | apply (rule eqpoll_trans) | 
| 13216 | 256 | apply (rule prod_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl]) | 
| 257 | apply (blast intro: well_ord_rmult) | |
| 258 | apply (rule prod_assoc_eqpoll [THEN eqpoll_trans]) | |
| 13221 | 259 | apply (rule eqpoll_sym) | 
| 13216 | 260 | apply (rule prod_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll]) | 
| 261 | apply (blast intro: well_ord_rmult) | |
| 262 | done | |
| 263 | ||
| 14883 | 264 | subsubsection{*Cardinal multiplication distributes over addition*}
 | 
| 13216 | 265 | |
| 266 | lemma sum_prod_distrib_eqpoll: "(A+B)*C \<approx> (A*C)+(B*C)" | |
| 267 | apply (unfold eqpoll_def) | |
| 268 | apply (rule exI) | |
| 269 | apply (rule sum_prod_distrib_bij) | |
| 270 | done | |
| 271 | ||
| 272 | lemma well_ord_cadd_cmult_distrib: | |
| 273 | "[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] | |
| 274 | ==> (i |+| j) |*| k = (i |*| k) |+| (j |*| k)" | |
| 275 | apply (unfold cadd_def cmult_def) | |
| 276 | apply (rule cardinal_cong) | |
| 13221 | 277 | apply (rule eqpoll_trans) | 
| 13216 | 278 | apply (rule prod_eqpoll_cong [OF well_ord_cardinal_eqpoll eqpoll_refl]) | 
| 279 | apply (blast intro: well_ord_radd) | |
| 280 | apply (rule sum_prod_distrib_eqpoll [THEN eqpoll_trans]) | |
| 13221 | 281 | apply (rule eqpoll_sym) | 
| 13216 | 282 | apply (rule sum_eqpoll_cong [OF well_ord_cardinal_eqpoll | 
| 283 | well_ord_cardinal_eqpoll]) | |
| 284 | apply (blast intro: well_ord_rmult)+ | |
| 285 | done | |
| 286 | ||
| 14883 | 287 | subsubsection{*Multiplication by 0 yields 0*}
 | 
| 13216 | 288 | |
| 289 | lemma prod_0_eqpoll: "0*A \<approx> 0" | |
| 290 | apply (unfold eqpoll_def) | |
| 291 | apply (rule exI) | |
| 13221 | 292 | apply (rule lam_bijective, safe) | 
| 13216 | 293 | done | 
| 294 | ||
| 295 | lemma cmult_0 [simp]: "0 |*| i = 0" | |
| 13221 | 296 | by (simp add: cmult_def prod_0_eqpoll [THEN cardinal_cong]) | 
| 13216 | 297 | |
| 14883 | 298 | subsubsection{*1 is the identity for multiplication*}
 | 
| 13216 | 299 | |
| 300 | lemma prod_singleton_eqpoll: "{x}*A \<approx> A"
 | |
| 301 | apply (unfold eqpoll_def) | |
| 302 | apply (rule exI) | |
| 303 | apply (rule singleton_prod_bij [THEN bij_converse_bij]) | |
| 304 | done | |
| 305 | ||
| 306 | lemma cmult_1 [simp]: "Card(K) ==> 1 |*| K = K" | |
| 307 | apply (unfold cmult_def succ_def) | |
| 308 | apply (simp add: prod_singleton_eqpoll [THEN cardinal_cong] Card_cardinal_eq) | |
| 309 | done | |
| 310 | ||
| 13356 | 311 | subsection{*Some inequalities for multiplication*}
 | 
| 13216 | 312 | |
| 313 | lemma prod_square_lepoll: "A \<lesssim> A*A" | |
| 314 | apply (unfold lepoll_def inj_def) | |
| 13221 | 315 | apply (rule_tac x = "lam x:A. <x,x>" in exI, simp) | 
| 13216 | 316 | done | 
| 317 | ||
| 318 | (*Could probably weaken the premise to well_ord(K,r), or remove using AC*) | |
| 319 | lemma cmult_square_le: "Card(K) ==> K le K |*| K" | |
| 320 | apply (unfold cmult_def) | |
| 321 | apply (rule le_trans) | |
| 322 | apply (rule_tac [2] well_ord_lepoll_imp_Card_le) | |
| 323 | apply (rule_tac [3] prod_square_lepoll) | |
| 13221 | 324 | apply (simp add: le_refl Card_is_Ord Card_cardinal_eq) | 
| 325 | apply (blast intro: well_ord_rmult well_ord_Memrel Card_is_Ord) | |
| 13216 | 326 | done | 
| 327 | ||
| 14883 | 328 | subsubsection{*Multiplication by a non-zero cardinal*}
 | 
| 13216 | 329 | |
| 330 | lemma prod_lepoll_self: "b: B ==> A \<lesssim> A*B" | |
| 331 | apply (unfold lepoll_def inj_def) | |
| 13221 | 332 | apply (rule_tac x = "lam x:A. <x,b>" in exI, simp) | 
| 13216 | 333 | done | 
| 334 | ||
| 335 | (*Could probably weaken the premises to well_ord(K,r), or removing using AC*) | |
| 336 | lemma cmult_le_self: | |
| 337 | "[| Card(K); Ord(L); 0<L |] ==> K le (K |*| L)" | |
| 338 | apply (unfold cmult_def) | |
| 339 | apply (rule le_trans [OF Card_cardinal_le well_ord_lepoll_imp_Card_le]) | |
| 13221 | 340 | apply assumption | 
| 13216 | 341 | apply (blast intro: well_ord_rmult well_ord_Memrel Card_is_Ord) | 
| 342 | apply (blast intro: prod_lepoll_self ltD) | |
| 343 | done | |
| 344 | ||
| 14883 | 345 | subsubsection{*Monotonicity of multiplication*}
 | 
| 13216 | 346 | |
| 347 | lemma prod_lepoll_mono: | |
| 348 | "[| A \<lesssim> C; B \<lesssim> D |] ==> A * B \<lesssim> C * D" | |
| 349 | apply (unfold lepoll_def) | |
| 13221 | 350 | apply (elim exE) | 
| 13216 | 351 | apply (rule_tac x = "lam <w,y>:A*B. <f`w, fa`y>" in exI) | 
| 352 | apply (rule_tac d = "%<w,y>. <converse (f) `w, converse (fa) `y>" | |
| 353 | in lam_injective) | |
| 13221 | 354 | apply (typecheck add: inj_is_fun, auto) | 
| 13216 | 355 | done | 
| 356 | ||
| 357 | lemma cmult_le_mono: | |
| 358 | "[| K' le K; L' le L |] ==> (K' |*| L') le (K |*| L)" | |
| 359 | apply (unfold cmult_def) | |
| 360 | apply (safe dest!: le_subset_iff [THEN iffD1]) | |
| 361 | apply (rule well_ord_lepoll_imp_Card_le) | |
| 362 | apply (blast intro: well_ord_rmult well_ord_Memrel) | |
| 363 | apply (blast intro: prod_lepoll_mono subset_imp_lepoll) | |
| 364 | done | |
| 365 | ||
| 13356 | 366 | subsection{*Multiplication of finite cardinals is "ordinary" multiplication*}
 | 
| 13216 | 367 | |
| 368 | lemma prod_succ_eqpoll: "succ(A)*B \<approx> B + A*B" | |
| 369 | apply (unfold eqpoll_def) | |
| 13221 | 370 | apply (rule exI) | 
| 13216 | 371 | apply (rule_tac c = "%<x,y>. if x=A then Inl (y) else Inr (<x,y>)" | 
| 372 | and d = "case (%y. <A,y>, %z. z)" in lam_bijective) | |
| 373 | apply safe | |
| 374 | apply (simp_all add: succI2 if_type mem_imp_not_eq) | |
| 375 | done | |
| 376 | ||
| 377 | (*Unconditional version requires AC*) | |
| 378 | lemma cmult_succ_lemma: | |
| 379 | "[| Ord(m); Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)" | |
| 380 | apply (unfold cmult_def cadd_def) | |
| 381 | apply (rule prod_succ_eqpoll [THEN cardinal_cong, THEN trans]) | |
| 382 | apply (rule cardinal_cong [symmetric]) | |
| 383 | apply (rule sum_eqpoll_cong [OF eqpoll_refl well_ord_cardinal_eqpoll]) | |
| 384 | apply (blast intro: well_ord_rmult well_ord_Memrel) | |
| 385 | done | |
| 386 | ||
| 387 | lemma nat_cmult_eq_mult: "[| m: nat; n: nat |] ==> m |*| n = m#*n" | |
| 13244 | 388 | apply (induct_tac m) | 
| 13221 | 389 | apply (simp_all add: cmult_succ_lemma nat_cadd_eq_add) | 
| 13216 | 390 | done | 
| 391 | ||
| 392 | lemma cmult_2: "Card(n) ==> 2 |*| n = n |+| n" | |
| 13221 | 393 | by (simp add: cmult_succ_lemma Card_is_Ord cadd_commute [of _ 0]) | 
| 13216 | 394 | |
| 395 | lemma sum_lepoll_prod: "2 \<lesssim> C ==> B+B \<lesssim> C*B" | |
| 13221 | 396 | apply (rule lepoll_trans) | 
| 13216 | 397 | apply (rule sum_eq_2_times [THEN equalityD1, THEN subset_imp_lepoll]) | 
| 398 | apply (erule prod_lepoll_mono) | |
| 13221 | 399 | apply (rule lepoll_refl) | 
| 13216 | 400 | done | 
| 401 | ||
| 402 | lemma lepoll_imp_sum_lepoll_prod: "[| A \<lesssim> B; 2 \<lesssim> A |] ==> A+B \<lesssim> A*B" | |
| 13221 | 403 | by (blast intro: sum_lepoll_mono sum_lepoll_prod lepoll_trans lepoll_refl) | 
| 13216 | 404 | |
| 405 | ||
| 13356 | 406 | subsection{*Infinite Cardinals are Limit Ordinals*}
 | 
| 13216 | 407 | |
| 408 | (*This proof is modelled upon one assuming nat<=A, with injection | |
| 409 | lam z:cons(u,A). if z=u then 0 else if z : nat then succ(z) else z | |
| 410 | and inverse %y. if y:nat then nat_case(u, %z. z, y) else y. \ | |
| 411 | If f: inj(nat,A) then range(f) behaves like the natural numbers.*) | |
| 412 | lemma nat_cons_lepoll: "nat \<lesssim> A ==> cons(u,A) \<lesssim> A" | |
| 413 | apply (unfold lepoll_def) | |
| 414 | apply (erule exE) | |
| 415 | apply (rule_tac x = | |
| 416 | "lam z:cons (u,A). | |
| 417 | if z=u then f`0 | |
| 418 | else if z: range (f) then f`succ (converse (f) `z) else z" | |
| 419 | in exI) | |
| 420 | apply (rule_tac d = | |
| 421 | "%y. if y: range(f) then nat_case (u, %z. f`z, converse(f) `y) | |
| 422 | else y" | |
| 423 | in lam_injective) | |
| 424 | apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun) | |
| 425 | apply (simp add: inj_is_fun [THEN apply_rangeI] | |
| 426 | inj_converse_fun [THEN apply_rangeI] | |
| 427 | inj_converse_fun [THEN apply_funtype]) | |
| 428 | done | |
| 429 | ||
| 430 | lemma nat_cons_eqpoll: "nat \<lesssim> A ==> cons(u,A) \<approx> A" | |
| 431 | apply (erule nat_cons_lepoll [THEN eqpollI]) | |
| 432 | apply (rule subset_consI [THEN subset_imp_lepoll]) | |
| 433 | done | |
| 434 | ||
| 435 | (*Specialized version required below*) | |
| 436 | lemma nat_succ_eqpoll: "nat <= A ==> succ(A) \<approx> A" | |
| 437 | apply (unfold succ_def) | |
| 438 | apply (erule subset_imp_lepoll [THEN nat_cons_eqpoll]) | |
| 439 | done | |
| 440 | ||
| 441 | lemma InfCard_nat: "InfCard(nat)" | |
| 442 | apply (unfold InfCard_def) | |
| 443 | apply (blast intro: Card_nat le_refl Card_is_Ord) | |
| 444 | done | |
| 445 | ||
| 446 | lemma InfCard_is_Card: "InfCard(K) ==> Card(K)" | |
| 447 | apply (unfold InfCard_def) | |
| 448 | apply (erule conjunct1) | |
| 449 | done | |
| 450 | ||
| 451 | lemma InfCard_Un: | |
| 452 | "[| InfCard(K); Card(L) |] ==> InfCard(K Un L)" | |
| 453 | apply (unfold InfCard_def) | |
| 454 | apply (simp add: Card_Un Un_upper1_le [THEN [2] le_trans] Card_is_Ord) | |
| 455 | done | |
| 456 | ||
| 457 | (*Kunen's Lemma 10.11*) | |
| 458 | lemma InfCard_is_Limit: "InfCard(K) ==> Limit(K)" | |
| 459 | apply (unfold InfCard_def) | |
| 460 | apply (erule conjE) | |
| 461 | apply (frule Card_is_Ord) | |
| 462 | apply (rule ltI [THEN non_succ_LimitI]) | |
| 463 | apply (erule le_imp_subset [THEN subsetD]) | |
| 464 | apply (safe dest!: Limit_nat [THEN Limit_le_succD]) | |
| 465 | apply (unfold Card_def) | |
| 466 | apply (drule trans) | |
| 467 | apply (erule le_imp_subset [THEN nat_succ_eqpoll, THEN cardinal_cong]) | |
| 468 | apply (erule Ord_cardinal_le [THEN lt_trans2, THEN lt_irrefl]) | |
| 13221 | 469 | apply (rule le_eqI, assumption) | 
| 13216 | 470 | apply (rule Ord_cardinal) | 
| 471 | done | |
| 472 | ||
| 473 | ||
| 474 | (*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***) | |
| 475 | ||
| 476 | (*A general fact about ordermap*) | |
| 477 | lemma ordermap_eqpoll_pred: | |
| 13269 | 478 | "[| well_ord(A,r); x:A |] ==> ordermap(A,r)`x \<approx> Order.pred(A,x,r)" | 
| 13216 | 479 | apply (unfold eqpoll_def) | 
| 480 | apply (rule exI) | |
| 13221 | 481 | apply (simp add: ordermap_eq_image well_ord_is_wf) | 
| 482 | apply (erule ordermap_bij [THEN bij_is_inj, THEN restrict_bij, | |
| 483 | THEN bij_converse_bij]) | |
| 13216 | 484 | apply (rule pred_subset) | 
| 485 | done | |
| 486 | ||
| 14883 | 487 | subsubsection{*Establishing the well-ordering*}
 | 
| 13216 | 488 | |
| 489 | lemma csquare_lam_inj: | |
| 490 | "Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)" | |
| 491 | apply (unfold inj_def) | |
| 492 | apply (force intro: lam_type Un_least_lt [THEN ltD] ltI) | |
| 493 | done | |
| 494 | ||
| 495 | lemma well_ord_csquare: "Ord(K) ==> well_ord(K*K, csquare_rel(K))" | |
| 496 | apply (unfold csquare_rel_def) | |
| 13221 | 497 | apply (rule csquare_lam_inj [THEN well_ord_rvimage], assumption) | 
| 13216 | 498 | apply (blast intro: well_ord_rmult well_ord_Memrel) | 
| 499 | done | |
| 500 | ||
| 14883 | 501 | subsubsection{*Characterising initial segments of the well-ordering*}
 | 
| 13216 | 502 | |
| 503 | lemma csquareD: | |
| 504 | "[| <<x,y>, <z,z>> : csquare_rel(K); x<K; y<K; z<K |] ==> x le z & y le z" | |
| 505 | apply (unfold csquare_rel_def) | |
| 506 | apply (erule rev_mp) | |
| 507 | apply (elim ltE) | |
| 13221 | 508 | apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD) | 
| 13216 | 509 | apply (safe elim!: mem_irrefl intro!: Un_upper1_le Un_upper2_le) | 
| 13221 | 510 | apply (simp_all add: lt_def succI2) | 
| 13216 | 511 | done | 
| 512 | ||
| 513 | lemma pred_csquare_subset: | |
| 13269 | 514 | "z<K ==> Order.pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)" | 
| 13216 | 515 | apply (unfold Order.pred_def) | 
| 516 | apply (safe del: SigmaI succCI) | |
| 517 | apply (erule csquareD [THEN conjE]) | |
| 13221 | 518 | apply (unfold lt_def, auto) | 
| 13216 | 519 | done | 
| 520 | ||
| 521 | lemma csquare_ltI: | |
| 522 | "[| x<z; y<z; z<K |] ==> <<x,y>, <z,z>> : csquare_rel(K)" | |
| 523 | apply (unfold csquare_rel_def) | |
| 524 | apply (subgoal_tac "x<K & y<K") | |
| 525 | prefer 2 apply (blast intro: lt_trans) | |
| 526 | apply (elim ltE) | |
| 13221 | 527 | apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD) | 
| 13216 | 528 | done | 
| 529 | ||
| 530 | (*Part of the traditional proof. UNUSED since it's harder to prove & apply *) | |
| 531 | lemma csquare_or_eqI: | |
| 532 | "[| x le z; y le z; z<K |] ==> <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z" | |
| 533 | apply (unfold csquare_rel_def) | |
| 534 | apply (subgoal_tac "x<K & y<K") | |
| 535 | prefer 2 apply (blast intro: lt_trans1) | |
| 536 | apply (elim ltE) | |
| 13221 | 537 | apply (simp add: rvimage_iff Un_absorb Un_least_mem_iff ltD) | 
| 13216 | 538 | apply (elim succE) | 
| 13221 | 539 | apply (simp_all add: subset_Un_iff [THEN iff_sym] | 
| 540 | subset_Un_iff2 [THEN iff_sym] OrdmemD) | |
| 13216 | 541 | done | 
| 542 | ||
| 14883 | 543 | subsubsection{*The cardinality of initial segments*}
 | 
| 13216 | 544 | |
| 545 | lemma ordermap_z_lt: | |
| 546 | "[| Limit(K); x<K; y<K; z=succ(x Un y) |] ==> | |
| 547 | ordermap(K*K, csquare_rel(K)) ` <x,y> < | |
| 548 | ordermap(K*K, csquare_rel(K)) ` <z,z>" | |
| 549 | apply (subgoal_tac "z<K & well_ord (K*K, csquare_rel (K))") | |
| 550 | prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ | |
| 13221 | 551 | Limit_is_Ord [THEN well_ord_csquare], clarify) | 
| 13216 | 552 | apply (rule csquare_ltI [THEN ordermap_mono, THEN ltI]) | 
| 553 | apply (erule_tac [4] well_ord_is_wf) | |
| 554 | apply (blast intro!: Un_upper1_le Un_upper2_le Ord_ordermap elim!: ltE)+ | |
| 555 | done | |
| 556 | ||
| 557 | (*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *) | |
| 558 | lemma ordermap_csquare_le: | |
| 13221 | 559 | "[| Limit(K); x<K; y<K; z=succ(x Un y) |] | 
| 560 | ==> | ordermap(K*K, csquare_rel(K)) ` <x,y> | le |succ(z)| |*| |succ(z)|" | |
| 13216 | 561 | apply (unfold cmult_def) | 
| 562 | apply (rule well_ord_rmult [THEN well_ord_lepoll_imp_Card_le]) | |
| 563 | apply (rule Ord_cardinal [THEN well_ord_Memrel])+ | |
| 564 | apply (subgoal_tac "z<K") | |
| 565 | prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ) | |
| 13221 | 566 | apply (rule ordermap_z_lt [THEN leI, THEN le_imp_lepoll, THEN lepoll_trans], | 
| 567 | assumption+) | |
| 13216 | 568 | apply (rule ordermap_eqpoll_pred [THEN eqpoll_imp_lepoll, THEN lepoll_trans]) | 
| 569 | apply (erule Limit_is_Ord [THEN well_ord_csquare]) | |
| 570 | apply (blast intro: ltD) | |
| 571 | apply (rule pred_csquare_subset [THEN subset_imp_lepoll, THEN lepoll_trans], | |
| 572 | assumption) | |
| 573 | apply (elim ltE) | |
| 574 | apply (rule prod_eqpoll_cong [THEN eqpoll_sym, THEN eqpoll_imp_lepoll]) | |
| 575 | apply (erule Ord_succ [THEN Ord_cardinal_eqpoll])+ | |
| 576 | done | |
| 577 | ||
| 578 | (*Kunen: "... so the order type <= K" *) | |
| 579 | lemma ordertype_csquare_le: | |
| 580 | "[| InfCard(K); ALL y:K. InfCard(y) --> y |*| y = y |] | |
| 581 | ==> ordertype(K*K, csquare_rel(K)) le K" | |
| 582 | apply (frule InfCard_is_Card [THEN Card_is_Ord]) | |
| 13221 | 583 | apply (rule all_lt_imp_le, assumption) | 
| 13216 | 584 | apply (erule well_ord_csquare [THEN Ord_ordertype]) | 
| 585 | apply (rule Card_lt_imp_lt) | |
| 586 | apply (erule_tac [3] InfCard_is_Card) | |
| 587 | apply (erule_tac [2] ltE) | |
| 588 | apply (simp add: ordertype_unfold) | |
| 589 | apply (safe elim!: ltE) | |
| 590 | apply (subgoal_tac "Ord (xa) & Ord (ya)") | |
| 13221 | 591 | prefer 2 apply (blast intro: Ord_in_Ord, clarify) | 
| 13216 | 592 | (*??WHAT A MESS!*) | 
| 593 | apply (rule InfCard_is_Limit [THEN ordermap_csquare_le, THEN lt_trans1], | |
| 594 | (assumption | rule refl | erule ltI)+) | |
| 13784 | 595 | apply (rule_tac i = "xa Un ya" and j = nat in Ord_linear2, | 
| 13216 | 596 | simp_all add: Ord_Un Ord_nat) | 
| 597 | prefer 2 (*case nat le (xa Un ya) *) | |
| 598 | apply (simp add: le_imp_subset [THEN nat_succ_eqpoll, THEN cardinal_cong] | |
| 599 | le_succ_iff InfCard_def Card_cardinal Un_least_lt Ord_Un | |
| 600 | ltI nat_le_cardinal Ord_cardinal_le [THEN lt_trans1, THEN ltD]) | |
| 601 | (*the finite case: xa Un ya < nat *) | |
| 13784 | 602 | apply (rule_tac j = nat in lt_trans2) | 
| 13216 | 603 | apply (simp add: lt_def nat_cmult_eq_mult nat_succI mult_type | 
| 604 | nat_into_Card [THEN Card_cardinal_eq] Ord_nat) | |
| 605 | apply (simp add: InfCard_def) | |
| 606 | done | |
| 607 | ||
| 608 | (*Main result: Kunen's Theorem 10.12*) | |
| 609 | lemma InfCard_csquare_eq: "InfCard(K) ==> K |*| K = K" | |
| 610 | apply (frule InfCard_is_Card [THEN Card_is_Ord]) | |
| 611 | apply (erule rev_mp) | |
| 612 | apply (erule_tac i=K in trans_induct) | |
| 613 | apply (rule impI) | |
| 614 | apply (rule le_anti_sym) | |
| 615 | apply (erule_tac [2] InfCard_is_Card [THEN cmult_square_le]) | |
| 616 | apply (rule ordertype_csquare_le [THEN [2] le_trans]) | |
| 13221 | 617 | apply (simp add: cmult_def Ord_cardinal_le | 
| 618 | well_ord_csquare [THEN Ord_ordertype] | |
| 619 | well_ord_csquare [THEN ordermap_bij, THEN bij_imp_eqpoll, | |
| 620 | THEN cardinal_cong], assumption+) | |
| 13216 | 621 | done | 
| 622 | ||
| 623 | (*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*) | |
| 624 | lemma well_ord_InfCard_square_eq: | |
| 625 | "[| well_ord(A,r); InfCard(|A|) |] ==> A*A \<approx> A" | |
| 626 | apply (rule prod_eqpoll_cong [THEN eqpoll_trans]) | |
| 627 | apply (erule well_ord_cardinal_eqpoll [THEN eqpoll_sym])+ | |
| 628 | apply (rule well_ord_cardinal_eqE) | |
| 13221 | 629 | apply (blast intro: Ord_cardinal well_ord_rmult well_ord_Memrel, assumption) | 
| 630 | apply (simp add: cmult_def [symmetric] InfCard_csquare_eq) | |
| 13216 | 631 | done | 
| 632 | ||
| 13356 | 633 | lemma InfCard_square_eqpoll: "InfCard(K) ==> K \<times> K \<approx> K" | 
| 634 | apply (rule well_ord_InfCard_square_eq) | |
| 635 | apply (erule InfCard_is_Card [THEN Card_is_Ord, THEN well_ord_Memrel]) | |
| 636 | apply (simp add: InfCard_is_Card [THEN Card_cardinal_eq]) | |
| 637 | done | |
| 638 | ||
| 639 | lemma Inf_Card_is_InfCard: "[| ~Finite(i); Card(i) |] ==> InfCard(i)" | |
| 640 | by (simp add: InfCard_def Card_is_Ord [THEN nat_le_infinite_Ord]) | |
| 641 | ||
| 14883 | 642 | subsubsection{*Toward's Kunen's Corollary 10.13 (1)*}
 | 
| 13216 | 643 | |
| 644 | lemma InfCard_le_cmult_eq: "[| InfCard(K); L le K; 0<L |] ==> K |*| L = K" | |
| 645 | apply (rule le_anti_sym) | |
| 646 | prefer 2 | |
| 647 | apply (erule ltE, blast intro: cmult_le_self InfCard_is_Card) | |
| 648 | apply (frule InfCard_is_Card [THEN Card_is_Ord, THEN le_refl]) | |
| 649 | apply (rule cmult_le_mono [THEN le_trans], assumption+) | |
| 650 | apply (simp add: InfCard_csquare_eq) | |
| 651 | done | |
| 652 | ||
| 653 | (*Corollary 10.13 (1), for cardinal multiplication*) | |
| 654 | lemma InfCard_cmult_eq: "[| InfCard(K); InfCard(L) |] ==> K |*| L = K Un L" | |
| 13784 | 655 | apply (rule_tac i = K and j = L in Ord_linear_le) | 
| 13216 | 656 | apply (typecheck add: InfCard_is_Card Card_is_Ord) | 
| 657 | apply (rule cmult_commute [THEN ssubst]) | |
| 658 | apply (rule Un_commute [THEN ssubst]) | |
| 13221 | 659 | apply (simp_all add: InfCard_is_Limit [THEN Limit_has_0] InfCard_le_cmult_eq | 
| 660 | subset_Un_iff2 [THEN iffD1] le_imp_subset) | |
| 13216 | 661 | done | 
| 662 | ||
| 663 | lemma InfCard_cdouble_eq: "InfCard(K) ==> K |+| K = K" | |
| 13221 | 664 | apply (simp add: cmult_2 [symmetric] InfCard_is_Card cmult_commute) | 
| 665 | apply (simp add: InfCard_le_cmult_eq InfCard_is_Limit Limit_has_0 Limit_has_succ) | |
| 13216 | 666 | done | 
| 667 | ||
| 668 | (*Corollary 10.13 (1), for cardinal addition*) | |
| 669 | lemma InfCard_le_cadd_eq: "[| InfCard(K); L le K |] ==> K |+| L = K" | |
| 670 | apply (rule le_anti_sym) | |
| 671 | prefer 2 | |
| 672 | apply (erule ltE, blast intro: cadd_le_self InfCard_is_Card) | |
| 673 | apply (frule InfCard_is_Card [THEN Card_is_Ord, THEN le_refl]) | |
| 674 | apply (rule cadd_le_mono [THEN le_trans], assumption+) | |
| 675 | apply (simp add: InfCard_cdouble_eq) | |
| 676 | done | |
| 677 | ||
| 678 | lemma InfCard_cadd_eq: "[| InfCard(K); InfCard(L) |] ==> K |+| L = K Un L" | |
| 13784 | 679 | apply (rule_tac i = K and j = L in Ord_linear_le) | 
| 13216 | 680 | apply (typecheck add: InfCard_is_Card Card_is_Ord) | 
| 681 | apply (rule cadd_commute [THEN ssubst]) | |
| 682 | apply (rule Un_commute [THEN ssubst]) | |
| 13221 | 683 | apply (simp_all add: InfCard_le_cadd_eq subset_Un_iff2 [THEN iffD1] le_imp_subset) | 
| 13216 | 684 | done | 
| 685 | ||
| 686 | (*The other part, Corollary 10.13 (2), refers to the cardinality of the set | |
| 687 | of all n-tuples of elements of K. A better version for the Isabelle theory | |
| 688 | might be InfCard(K) ==> |list(K)| = K. | |
| 689 | *) | |
| 690 | ||
| 27517 | 691 | subsection{*For Every Cardinal Number There Exists A Greater One*}
 | 
| 13356 | 692 | |
| 693 | text{*This result is Kunen's Theorem 10.16, which would be trivial using AC*}
 | |
| 13216 | 694 | |
| 695 | lemma Ord_jump_cardinal: "Ord(jump_cardinal(K))" | |
| 696 | apply (unfold jump_cardinal_def) | |
| 697 | apply (rule Ord_is_Transset [THEN [2] OrdI]) | |
| 698 | prefer 2 apply (blast intro!: Ord_ordertype) | |
| 699 | apply (unfold Transset_def) | |
| 700 | apply (safe del: subsetI) | |
| 13221 | 701 | apply (simp add: ordertype_pred_unfold, safe) | 
| 13216 | 702 | apply (rule UN_I) | 
| 703 | apply (rule_tac [2] ReplaceI) | |
| 704 | prefer 4 apply (blast intro: well_ord_subset elim!: predE)+ | |
| 705 | done | |
| 706 | ||
| 707 | (*Allows selective unfolding. Less work than deriving intro/elim rules*) | |
| 708 | lemma jump_cardinal_iff: | |
| 709 | "i : jump_cardinal(K) <-> | |
| 710 | (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))" | |
| 711 | apply (unfold jump_cardinal_def) | |
| 712 | apply (blast del: subsetI) | |
| 713 | done | |
| 714 | ||
| 715 | (*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) | |
| 716 | lemma K_lt_jump_cardinal: "Ord(K) ==> K < jump_cardinal(K)" | |
| 717 | apply (rule Ord_jump_cardinal [THEN [2] ltI]) | |
| 718 | apply (rule jump_cardinal_iff [THEN iffD2]) | |
| 719 | apply (rule_tac x="Memrel(K)" in exI) | |
| 720 | apply (rule_tac x=K in exI) | |
| 721 | apply (simp add: ordertype_Memrel well_ord_Memrel) | |
| 722 | apply (simp add: Memrel_def subset_iff) | |
| 723 | done | |
| 724 | ||
| 725 | (*The proof by contradiction: the bijection f yields a wellordering of X | |
| 726 | whose ordertype is jump_cardinal(K). *) | |
| 727 | lemma Card_jump_cardinal_lemma: | |
| 728 | "[| well_ord(X,r); r <= K * K; X <= K; | |
| 729 | f : bij(ordertype(X,r), jump_cardinal(K)) |] | |
| 730 | ==> jump_cardinal(K) : jump_cardinal(K)" | |
| 731 | apply (subgoal_tac "f O ordermap (X,r) : bij (X, jump_cardinal (K))") | |
| 732 | prefer 2 apply (blast intro: comp_bij ordermap_bij) | |
| 733 | apply (rule jump_cardinal_iff [THEN iffD2]) | |
| 734 | apply (intro exI conjI) | |
| 13221 | 735 | apply (rule subset_trans [OF rvimage_type Sigma_mono], assumption+) | 
| 13216 | 736 | apply (erule bij_is_inj [THEN well_ord_rvimage]) | 
| 737 | apply (rule Ord_jump_cardinal [THEN well_ord_Memrel]) | |
| 738 | apply (simp add: well_ord_Memrel [THEN [2] bij_ordertype_vimage] | |
| 739 | ordertype_Memrel Ord_jump_cardinal) | |
| 740 | done | |
| 741 | ||
| 742 | (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*) | |
| 743 | lemma Card_jump_cardinal: "Card(jump_cardinal(K))" | |
| 744 | apply (rule Ord_jump_cardinal [THEN CardI]) | |
| 745 | apply (unfold eqpoll_def) | |
| 746 | apply (safe dest!: ltD jump_cardinal_iff [THEN iffD1]) | |
| 747 | apply (blast intro: Card_jump_cardinal_lemma [THEN mem_irrefl]) | |
| 748 | done | |
| 749 | ||
| 13356 | 750 | subsection{*Basic Properties of Successor Cardinals*}
 | 
| 13216 | 751 | |
| 752 | lemma csucc_basic: "Ord(K) ==> Card(csucc(K)) & K < csucc(K)" | |
| 753 | apply (unfold csucc_def) | |
| 754 | apply (rule LeastI) | |
| 755 | apply (blast intro: Card_jump_cardinal K_lt_jump_cardinal Ord_jump_cardinal)+ | |
| 756 | done | |
| 757 | ||
| 758 | lemmas Card_csucc = csucc_basic [THEN conjunct1, standard] | |
| 759 | ||
| 760 | lemmas lt_csucc = csucc_basic [THEN conjunct2, standard] | |
| 761 | ||
| 762 | lemma Ord_0_lt_csucc: "Ord(K) ==> 0 < csucc(K)" | |
| 13221 | 763 | by (blast intro: Ord_0_le lt_csucc lt_trans1) | 
| 13216 | 764 | |
| 765 | lemma csucc_le: "[| Card(L); K<L |] ==> csucc(K) le L" | |
| 766 | apply (unfold csucc_def) | |
| 767 | apply (rule Least_le) | |
| 768 | apply (blast intro: Card_is_Ord)+ | |
| 769 | done | |
| 770 | ||
| 771 | lemma lt_csucc_iff: "[| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K" | |
| 772 | apply (rule iffI) | |
| 773 | apply (rule_tac [2] Card_lt_imp_lt) | |
| 774 | apply (erule_tac [2] lt_trans1) | |
| 775 | apply (simp_all add: lt_csucc Card_csucc Card_is_Ord) | |
| 776 | apply (rule notI [THEN not_lt_imp_le]) | |
| 13221 | 777 | apply (rule Card_cardinal [THEN csucc_le, THEN lt_trans1, THEN lt_irrefl], assumption) | 
| 13216 | 778 | apply (rule Ord_cardinal_le [THEN lt_trans1]) | 
| 779 | apply (simp_all add: Ord_cardinal Card_is_Ord) | |
| 780 | done | |
| 781 | ||
| 782 | lemma Card_lt_csucc_iff: | |
| 783 | "[| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K" | |
| 13221 | 784 | by (simp add: lt_csucc_iff Card_cardinal_eq Card_is_Ord) | 
| 13216 | 785 | |
| 786 | lemma InfCard_csucc: "InfCard(K) ==> InfCard(csucc(K))" | |
| 787 | by (simp add: InfCard_def Card_csucc Card_is_Ord | |
| 788 | lt_csucc [THEN leI, THEN [2] le_trans]) | |
| 789 | ||
| 790 | ||
| 14883 | 791 | subsubsection{*Removing elements from a finite set decreases its cardinality*}
 | 
| 13216 | 792 | |
| 793 | lemma Fin_imp_not_cons_lepoll: "A: Fin(U) ==> x~:A --> ~ cons(x,A) \<lesssim> A" | |
| 794 | apply (erule Fin_induct) | |
| 13221 | 795 | apply (simp add: lepoll_0_iff) | 
| 13216 | 796 | apply (subgoal_tac "cons (x,cons (xa,y)) = cons (xa,cons (x,y))") | 
| 13221 | 797 | apply simp | 
| 798 | apply (blast dest!: cons_lepoll_consD, blast) | |
| 13216 | 799 | done | 
| 800 | ||
| 14883 | 801 | lemma Finite_imp_cardinal_cons [simp]: | 
| 13221 | 802 | "[| Finite(A); a~:A |] ==> |cons(a,A)| = succ(|A|)" | 
| 13216 | 803 | apply (unfold cardinal_def) | 
| 804 | apply (rule Least_equality) | |
| 805 | apply (fold cardinal_def) | |
| 13221 | 806 | apply (simp add: succ_def) | 
| 13216 | 807 | apply (blast intro: cons_eqpoll_cong well_ord_cardinal_eqpoll | 
| 808 | elim!: mem_irrefl dest!: Finite_imp_well_ord) | |
| 809 | apply (blast intro: Card_cardinal Card_is_Ord) | |
| 810 | apply (rule notI) | |
| 13221 | 811 | apply (rule Finite_into_Fin [THEN Fin_imp_not_cons_lepoll, THEN mp, THEN notE], | 
| 812 | assumption, assumption) | |
| 13216 | 813 | apply (erule eqpoll_sym [THEN eqpoll_imp_lepoll, THEN lepoll_trans]) | 
| 814 | apply (erule le_imp_lepoll [THEN lepoll_trans]) | |
| 815 | apply (blast intro: well_ord_cardinal_eqpoll [THEN eqpoll_imp_lepoll] | |
| 816 | dest!: Finite_imp_well_ord) | |
| 817 | done | |
| 818 | ||
| 819 | ||
| 13221 | 820 | lemma Finite_imp_succ_cardinal_Diff: | 
| 821 |      "[| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|"
 | |
| 13784 | 822 | apply (rule_tac b = A in cons_Diff [THEN subst], assumption) | 
| 13221 | 823 | apply (simp add: Finite_imp_cardinal_cons Diff_subset [THEN subset_Finite]) | 
| 824 | apply (simp add: cons_Diff) | |
| 13216 | 825 | done | 
| 826 | ||
| 827 | lemma Finite_imp_cardinal_Diff: "[| Finite(A);  a:A |] ==> |A-{a}| < |A|"
 | |
| 828 | apply (rule succ_leE) | |
| 13221 | 829 | apply (simp add: Finite_imp_succ_cardinal_Diff) | 
| 13216 | 830 | done | 
| 831 | ||
| 14883 | 832 | lemma Finite_cardinal_in_nat [simp]: "Finite(A) ==> |A| : nat" | 
| 833 | apply (erule Finite_induct) | |
| 834 | apply (auto simp add: cardinal_0 Finite_imp_cardinal_cons) | |
| 835 | done | |
| 13216 | 836 | |
| 14883 | 837 | lemma card_Un_Int: | 
| 838 | "[|Finite(A); Finite(B)|] ==> |A| #+ |B| = |A Un B| #+ |A Int B|" | |
| 839 | apply (erule Finite_induct, simp) | |
| 840 | apply (simp add: Finite_Int cons_absorb Un_cons Int_cons_left) | |
| 841 | done | |
| 842 | ||
| 843 | lemma card_Un_disjoint: | |
| 844 | "[|Finite(A); Finite(B); A Int B = 0|] ==> |A Un B| = |A| #+ |B|" | |
| 845 | by (simp add: Finite_Un card_Un_Int) | |
| 846 | ||
| 847 | lemma card_partition [rule_format]: | |
| 848 | "Finite(C) ==> | |
| 849 | Finite (\<Union> C) --> | |
| 850 | (\<forall>c\<in>C. |c| = k) --> | |
| 851 | (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 --> c1 \<inter> c2 = 0) --> | |
| 852 | k #* |C| = |\<Union> C|" | |
| 853 | apply (erule Finite_induct, auto) | |
| 854 | apply (subgoal_tac " x \<inter> \<Union>B = 0") | |
| 855 | apply (auto simp add: card_Un_disjoint Finite_Union | |
| 856 | subset_Finite [of _ "\<Union> (cons(x,F))"]) | |
| 857 | done | |
| 858 | ||
| 859 | ||
| 860 | subsubsection{*Theorems by Krzysztof Grabczewski, proofs by lcp*}
 | |
| 13216 | 861 | |
| 862 | lemmas nat_implies_well_ord = nat_into_Ord [THEN well_ord_Memrel, standard] | |
| 863 | ||
| 864 | lemma nat_sum_eqpoll_sum: "[| m:nat; n:nat |] ==> m + n \<approx> m #+ n" | |
| 865 | apply (rule eqpoll_trans) | |
| 866 | apply (rule well_ord_radd [THEN well_ord_cardinal_eqpoll, THEN eqpoll_sym]) | |
| 867 | apply (erule nat_implies_well_ord)+ | |
| 13221 | 868 | apply (simp add: nat_cadd_eq_add [symmetric] cadd_def eqpoll_refl) | 
| 13216 | 869 | done | 
| 870 | ||
| 13221 | 871 | lemma Ord_subset_natD [rule_format]: "Ord(i) ==> i <= nat --> i : nat | i=nat" | 
| 872 | apply (erule trans_induct3, auto) | |
| 13216 | 873 | apply (blast dest!: nat_le_Limit [THEN le_imp_subset]) | 
| 874 | done | |
| 875 | ||
| 876 | lemma Ord_nat_subset_into_Card: "[| Ord(i); i <= nat |] ==> Card(i)" | |
| 13221 | 877 | by (blast dest: Ord_subset_natD intro: Card_nat nat_into_Card) | 
| 13216 | 878 | |
| 879 | lemma Finite_Diff_sing_eq_diff_1: "[| Finite(A); x:A |] ==> |A-{x}| = |A| #- 1"
 | |
| 880 | apply (rule succ_inject) | |
| 881 | apply (rule_tac b = "|A|" in trans) | |
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 882 | apply (simp add: Finite_imp_succ_cardinal_Diff) | 
| 13216 | 883 | apply (subgoal_tac "1 \<lesssim> A") | 
| 13221 | 884 | prefer 2 apply (blast intro: not_0_is_lepoll_1) | 
| 885 | apply (frule Finite_imp_well_ord, clarify) | |
| 13216 | 886 | apply (drule well_ord_lepoll_imp_Card_le) | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 887 | apply (auto simp add: cardinal_1) | 
| 13216 | 888 | apply (rule trans) | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 889 | apply (rule_tac [2] diff_succ) | 
| 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13356diff
changeset | 890 | apply (auto simp add: Finite_cardinal_in_nat) | 
| 13216 | 891 | done | 
| 892 | ||
| 13221 | 893 | lemma cardinal_lt_imp_Diff_not_0 [rule_format]: | 
| 894 | "Finite(B) ==> ALL A. |B|<|A| --> A - B ~= 0" | |
| 895 | apply (erule Finite_induct, auto) | |
| 896 | apply (case_tac "Finite (A)") | |
| 897 | apply (subgoal_tac [2] "Finite (cons (x, B))") | |
| 898 | apply (drule_tac [2] B = "cons (x, B) " in Diff_Finite) | |
| 899 | apply (auto simp add: Finite_0 Finite_cons) | |
| 13216 | 900 | apply (subgoal_tac "|B|<|A|") | 
| 13221 | 901 | prefer 2 apply (blast intro: lt_trans Ord_cardinal) | 
| 13216 | 902 | apply (case_tac "x:A") | 
| 13221 | 903 | apply (subgoal_tac [2] "A - cons (x, B) = A - B") | 
| 904 | apply auto | |
| 13216 | 905 | apply (subgoal_tac "|A| le |cons (x, B) |") | 
| 13221 | 906 | prefer 2 | 
| 13216 | 907 | apply (blast dest: Finite_cons [THEN Finite_imp_well_ord] | 
| 908 | intro: well_ord_lepoll_imp_Card_le subset_imp_lepoll) | |
| 909 | apply (auto simp add: Finite_imp_cardinal_cons) | |
| 910 | apply (auto dest!: Finite_cardinal_in_nat simp add: le_iff) | |
| 911 | apply (blast intro: lt_trans) | |
| 912 | done | |
| 913 | ||
| 914 | ||
| 915 | ML{*
 | |
| 916 | val InfCard_def = thm "InfCard_def" | |
| 917 | val cmult_def = thm "cmult_def" | |
| 918 | val cadd_def = thm "cadd_def" | |
| 919 | val jump_cardinal_def = thm "jump_cardinal_def" | |
| 920 | val csucc_def = thm "csucc_def" | |
| 921 | ||
| 922 | val sum_commute_eqpoll = thm "sum_commute_eqpoll"; | |
| 923 | val cadd_commute = thm "cadd_commute"; | |
| 924 | val sum_assoc_eqpoll = thm "sum_assoc_eqpoll"; | |
| 925 | val well_ord_cadd_assoc = thm "well_ord_cadd_assoc"; | |
| 926 | val sum_0_eqpoll = thm "sum_0_eqpoll"; | |
| 927 | val cadd_0 = thm "cadd_0"; | |
| 928 | val sum_lepoll_self = thm "sum_lepoll_self"; | |
| 929 | val cadd_le_self = thm "cadd_le_self"; | |
| 930 | val sum_lepoll_mono = thm "sum_lepoll_mono"; | |
| 931 | val cadd_le_mono = thm "cadd_le_mono"; | |
| 932 | val eq_imp_not_mem = thm "eq_imp_not_mem"; | |
| 933 | val sum_succ_eqpoll = thm "sum_succ_eqpoll"; | |
| 934 | val nat_cadd_eq_add = thm "nat_cadd_eq_add"; | |
| 935 | val prod_commute_eqpoll = thm "prod_commute_eqpoll"; | |
| 936 | val cmult_commute = thm "cmult_commute"; | |
| 937 | val prod_assoc_eqpoll = thm "prod_assoc_eqpoll"; | |
| 938 | val well_ord_cmult_assoc = thm "well_ord_cmult_assoc"; | |
| 939 | val sum_prod_distrib_eqpoll = thm "sum_prod_distrib_eqpoll"; | |
| 940 | val well_ord_cadd_cmult_distrib = thm "well_ord_cadd_cmult_distrib"; | |
| 941 | val prod_0_eqpoll = thm "prod_0_eqpoll"; | |
| 942 | val cmult_0 = thm "cmult_0"; | |
| 943 | val prod_singleton_eqpoll = thm "prod_singleton_eqpoll"; | |
| 944 | val cmult_1 = thm "cmult_1"; | |
| 945 | val prod_lepoll_self = thm "prod_lepoll_self"; | |
| 946 | val cmult_le_self = thm "cmult_le_self"; | |
| 947 | val prod_lepoll_mono = thm "prod_lepoll_mono"; | |
| 948 | val cmult_le_mono = thm "cmult_le_mono"; | |
| 949 | val prod_succ_eqpoll = thm "prod_succ_eqpoll"; | |
| 950 | val nat_cmult_eq_mult = thm "nat_cmult_eq_mult"; | |
| 951 | val cmult_2 = thm "cmult_2"; | |
| 952 | val sum_lepoll_prod = thm "sum_lepoll_prod"; | |
| 953 | val lepoll_imp_sum_lepoll_prod = thm "lepoll_imp_sum_lepoll_prod"; | |
| 954 | val nat_cons_lepoll = thm "nat_cons_lepoll"; | |
| 955 | val nat_cons_eqpoll = thm "nat_cons_eqpoll"; | |
| 956 | val nat_succ_eqpoll = thm "nat_succ_eqpoll"; | |
| 957 | val InfCard_nat = thm "InfCard_nat"; | |
| 958 | val InfCard_is_Card = thm "InfCard_is_Card"; | |
| 959 | val InfCard_Un = thm "InfCard_Un"; | |
| 960 | val InfCard_is_Limit = thm "InfCard_is_Limit"; | |
| 961 | val ordermap_eqpoll_pred = thm "ordermap_eqpoll_pred"; | |
| 962 | val ordermap_z_lt = thm "ordermap_z_lt"; | |
| 963 | val InfCard_le_cmult_eq = thm "InfCard_le_cmult_eq"; | |
| 964 | val InfCard_cmult_eq = thm "InfCard_cmult_eq"; | |
| 965 | val InfCard_cdouble_eq = thm "InfCard_cdouble_eq"; | |
| 966 | val InfCard_le_cadd_eq = thm "InfCard_le_cadd_eq"; | |
| 967 | val InfCard_cadd_eq = thm "InfCard_cadd_eq"; | |
| 968 | val Ord_jump_cardinal = thm "Ord_jump_cardinal"; | |
| 969 | val jump_cardinal_iff = thm "jump_cardinal_iff"; | |
| 970 | val K_lt_jump_cardinal = thm "K_lt_jump_cardinal"; | |
| 971 | val Card_jump_cardinal = thm "Card_jump_cardinal"; | |
| 972 | val csucc_basic = thm "csucc_basic"; | |
| 973 | val Card_csucc = thm "Card_csucc"; | |
| 974 | val lt_csucc = thm "lt_csucc"; | |
| 975 | val Ord_0_lt_csucc = thm "Ord_0_lt_csucc"; | |
| 976 | val csucc_le = thm "csucc_le"; | |
| 977 | val lt_csucc_iff = thm "lt_csucc_iff"; | |
| 978 | val Card_lt_csucc_iff = thm "Card_lt_csucc_iff"; | |
| 979 | val InfCard_csucc = thm "InfCard_csucc"; | |
| 980 | val Finite_into_Fin = thm "Finite_into_Fin"; | |
| 981 | val Fin_into_Finite = thm "Fin_into_Finite"; | |
| 982 | val Finite_Fin_iff = thm "Finite_Fin_iff"; | |
| 983 | val Finite_Un = thm "Finite_Un"; | |
| 984 | val Finite_Union = thm "Finite_Union"; | |
| 985 | val Finite_induct = thm "Finite_induct"; | |
| 986 | val Fin_imp_not_cons_lepoll = thm "Fin_imp_not_cons_lepoll"; | |
| 987 | val Finite_imp_cardinal_cons = thm "Finite_imp_cardinal_cons"; | |
| 988 | val Finite_imp_succ_cardinal_Diff = thm "Finite_imp_succ_cardinal_Diff"; | |
| 989 | val Finite_imp_cardinal_Diff = thm "Finite_imp_cardinal_Diff"; | |
| 990 | val nat_implies_well_ord = thm "nat_implies_well_ord"; | |
| 991 | val nat_sum_eqpoll_sum = thm "nat_sum_eqpoll_sum"; | |
| 992 | val Diff_sing_Finite = thm "Diff_sing_Finite"; | |
| 993 | val Diff_Finite = thm "Diff_Finite"; | |
| 994 | val Ord_subset_natD = thm "Ord_subset_natD"; | |
| 995 | val Ord_nat_subset_into_Card = thm "Ord_nat_subset_into_Card"; | |
| 996 | val Finite_cardinal_in_nat = thm "Finite_cardinal_in_nat"; | |
| 997 | val Finite_Diff_sing_eq_diff_1 = thm "Finite_Diff_sing_eq_diff_1"; | |
| 998 | val cardinal_lt_imp_Diff_not_0 = thm "cardinal_lt_imp_Diff_not_0"; | |
| 999 | *} | |
| 1000 | ||
| 437 | 1001 | end |