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