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