author | wenzelm |
Tue, 08 Jan 2002 21:02:15 +0100 | |
changeset 12678 | 4d36d8df29fa |
parent 12667 | 7e6eaaa125f2 |
child 13140 | 6d97dbb189a9 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/CardinalArith.ML |
437 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
437 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Cardinal arithmetic -- WITHOUT the Axiom of Choice |
|
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
7 |
|
846 | 8 |
Note: Could omit proving the algebraic laws for cardinal addition and |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
9 |
multiplication. On finite cardinals these operations coincide with |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
10 |
addition and multiplication of natural numbers; on infinite cardinals they |
846 | 11 |
coincide with union (maximum). Either way we get most laws for free. |
437 | 12 |
*) |
13 |
||
12667 | 14 |
val InfCard_def = thm "InfCard_def"; |
15 |
val cmult_def = thm "cmult_def"; |
|
16 |
val cadd_def = thm "cadd_def"; |
|
17 |
val csquare_rel_def = thm "csquare_rel_def"; |
|
18 |
val jump_cardinal_def = thm "jump_cardinal_def"; |
|
19 |
val csucc_def = thm "csucc_def"; |
|
20 |
||
21 |
||
437 | 22 |
(*** Cardinal addition ***) |
23 |
||
24 |
(** Cardinal addition is commutative **) |
|
25 |
||
5067 | 26 |
Goalw [eqpoll_def] "A+B eqpoll B+A"; |
437 | 27 |
by (rtac exI 1); |
28 |
by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] |
|
29 |
lam_bijective 1); |
|
5488 | 30 |
by Auto_tac; |
760 | 31 |
qed "sum_commute_eqpoll"; |
437 | 32 |
|
5067 | 33 |
Goalw [cadd_def] "i |+| j = j |+| i"; |
437 | 34 |
by (rtac (sum_commute_eqpoll RS cardinal_cong) 1); |
760 | 35 |
qed "cadd_commute"; |
437 | 36 |
|
37 |
(** Cardinal addition is associative **) |
|
38 |
||
5067 | 39 |
Goalw [eqpoll_def] "(A+B)+C eqpoll A+(B+C)"; |
437 | 40 |
by (rtac exI 1); |
1461 | 41 |
by (rtac sum_assoc_bij 1); |
760 | 42 |
qed "sum_assoc_eqpoll"; |
437 | 43 |
|
44 |
(*Unconditional version requires AC*) |
|
5067 | 45 |
Goalw [cadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
46 |
"[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ |
437 | 47 |
\ (i |+| j) |+| k = i |+| (j |+| k)"; |
48 |
by (rtac cardinal_cong 1); |
|
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
49 |
by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS sum_eqpoll_cong RS |
1461 | 50 |
eqpoll_trans) 1); |
437 | 51 |
by (rtac (sum_assoc_eqpoll RS eqpoll_trans) 2); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
52 |
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS |
1461 | 53 |
eqpoll_sym) 2); |
484 | 54 |
by (REPEAT (ares_tac [well_ord_radd] 1)); |
760 | 55 |
qed "well_ord_cadd_assoc"; |
437 | 56 |
|
57 |
(** 0 is the identity for addition **) |
|
58 |
||
5067 | 59 |
Goalw [eqpoll_def] "0+A eqpoll A"; |
437 | 60 |
by (rtac exI 1); |
846 | 61 |
by (rtac bij_0_sum 1); |
760 | 62 |
qed "sum_0_eqpoll"; |
437 | 63 |
|
5137 | 64 |
Goalw [cadd_def] "Card(K) ==> 0 |+| K = K"; |
4091 | 65 |
by (asm_simp_tac (simpset() addsimps [sum_0_eqpoll RS cardinal_cong, |
4312 | 66 |
Card_cardinal_eq]) 1); |
760 | 67 |
qed "cadd_0"; |
6070 | 68 |
Addsimps [cadd_0]; |
437 | 69 |
|
767 | 70 |
(** Addition by another cardinal **) |
71 |
||
5067 | 72 |
Goalw [lepoll_def, inj_def] "A lepoll A+B"; |
767 | 73 |
by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1); |
6153 | 74 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
75 |
qed "sum_lepoll_self"; |
767 | 76 |
|
77 |
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) |
|
5067 | 78 |
Goalw [cadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
79 |
"[| Card(K); Ord(L) |] ==> K le (K |+| L)"; |
767 | 80 |
by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1); |
81 |
by (rtac sum_lepoll_self 3); |
|
82 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel, Card_is_Ord] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
83 |
qed "cadd_le_self"; |
767 | 84 |
|
85 |
(** Monotonicity of addition **) |
|
86 |
||
5067 | 87 |
Goalw [lepoll_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
88 |
"[| A lepoll C; B lepoll D |] ==> A + B lepoll C + D"; |
767 | 89 |
by (REPEAT (etac exE 1)); |
90 |
by (res_inst_tac [("x", "lam z:A+B. case(%w. Inl(f`w), %y. Inr(fa`y), z)")] |
|
91 |
exI 1); |
|
92 |
by (res_inst_tac |
|
93 |
[("d", "case(%w. Inl(converse(f)`w), %y. Inr(converse(fa)`y))")] |
|
94 |
lam_injective 1); |
|
6153 | 95 |
by (typecheck_tac (tcset() addTCs [inj_is_fun])); |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
96 |
by Auto_tac; |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
97 |
qed "sum_lepoll_mono"; |
767 | 98 |
|
5067 | 99 |
Goalw [cadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
100 |
"[| K' le K; L' le L |] ==> (K' |+| L') le (K |+| L)"; |
4091 | 101 |
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1])); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
102 |
by (rtac well_ord_lepoll_imp_Card_le 1); |
767 | 103 |
by (REPEAT (ares_tac [sum_lepoll_mono, subset_imp_lepoll] 2)); |
104 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
105 |
qed "cadd_le_mono"; |
767 | 106 |
|
437 | 107 |
(** Addition of finite cardinals is "ordinary" addition **) |
108 |
||
5067 | 109 |
Goalw [eqpoll_def] "succ(A)+B eqpoll succ(A+B)"; |
437 | 110 |
by (rtac exI 1); |
6068 | 111 |
by (res_inst_tac [("c", "%z. if z=Inl(A) then A+B else z"), |
112 |
("d", "%z. if z=A+B then Inl(A) else z")] |
|
437 | 113 |
lam_bijective 1); |
114 |
by (ALLGOALS |
|
4091 | 115 |
(asm_simp_tac (simpset() addsimps [succI2, mem_imp_not_eq] |
4312 | 116 |
setloop eresolve_tac [sumE,succE]))); |
760 | 117 |
qed "sum_succ_eqpoll"; |
437 | 118 |
|
119 |
(*Pulling the succ(...) outside the |...| requires m, n: nat *) |
|
120 |
(*Unconditional version requires AC*) |
|
5067 | 121 |
Goalw [cadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
122 |
"[| Ord(m); Ord(n) |] ==> succ(m) |+| n = |succ(m |+| n)|"; |
437 | 123 |
by (rtac (sum_succ_eqpoll RS cardinal_cong RS trans) 1); |
124 |
by (rtac (succ_eqpoll_cong RS cardinal_cong) 1); |
|
125 |
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1); |
|
126 |
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1)); |
|
760 | 127 |
qed "cadd_succ_lemma"; |
437 | 128 |
|
6070 | 129 |
Goal "[| m: nat; n: nat |] ==> m |+| n = m#+n"; |
130 |
by (induct_tac "m" 1); |
|
4091 | 131 |
by (asm_simp_tac (simpset() addsimps [nat_into_Card RS cadd_0]) 1); |
8201 | 132 |
by (asm_simp_tac (simpset() addsimps [cadd_succ_lemma, |
4312 | 133 |
nat_into_Card RS Card_cardinal_eq]) 1); |
760 | 134 |
qed "nat_cadd_eq_add"; |
437 | 135 |
|
136 |
||
137 |
(*** Cardinal multiplication ***) |
|
138 |
||
139 |
(** Cardinal multiplication is commutative **) |
|
140 |
||
141 |
(*Easier to prove the two directions separately*) |
|
5067 | 142 |
Goalw [eqpoll_def] "A*B eqpoll B*A"; |
437 | 143 |
by (rtac exI 1); |
1090
8ab69b3e396b
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1075
diff
changeset
|
144 |
by (res_inst_tac [("c", "%<x,y>.<y,x>"), ("d", "%<x,y>.<y,x>")] |
437 | 145 |
lam_bijective 1); |
4152 | 146 |
by Safe_tac; |
2469 | 147 |
by (ALLGOALS (Asm_simp_tac)); |
760 | 148 |
qed "prod_commute_eqpoll"; |
437 | 149 |
|
5067 | 150 |
Goalw [cmult_def] "i |*| j = j |*| i"; |
437 | 151 |
by (rtac (prod_commute_eqpoll RS cardinal_cong) 1); |
760 | 152 |
qed "cmult_commute"; |
437 | 153 |
|
154 |
(** Cardinal multiplication is associative **) |
|
155 |
||
5067 | 156 |
Goalw [eqpoll_def] "(A*B)*C eqpoll A*(B*C)"; |
437 | 157 |
by (rtac exI 1); |
1461 | 158 |
by (rtac prod_assoc_bij 1); |
760 | 159 |
qed "prod_assoc_eqpoll"; |
437 | 160 |
|
161 |
(*Unconditional version requires AC*) |
|
5067 | 162 |
Goalw [cmult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
163 |
"[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ |
437 | 164 |
\ (i |*| j) |*| k = i |*| (j |*| k)"; |
165 |
by (rtac cardinal_cong 1); |
|
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
166 |
by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS |
1461 | 167 |
eqpoll_trans) 1); |
437 | 168 |
by (rtac (prod_assoc_eqpoll RS eqpoll_trans) 2); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
169 |
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS prod_eqpoll_cong RS |
1461 | 170 |
eqpoll_sym) 2); |
484 | 171 |
by (REPEAT (ares_tac [well_ord_rmult] 1)); |
760 | 172 |
qed "well_ord_cmult_assoc"; |
437 | 173 |
|
174 |
(** Cardinal multiplication distributes over addition **) |
|
175 |
||
5067 | 176 |
Goalw [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)"; |
437 | 177 |
by (rtac exI 1); |
1461 | 178 |
by (rtac sum_prod_distrib_bij 1); |
760 | 179 |
qed "sum_prod_distrib_eqpoll"; |
437 | 180 |
|
5067 | 181 |
Goalw [cadd_def, cmult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
182 |
"[| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==> \ |
846 | 183 |
\ (i |+| j) |*| k = (i |*| k) |+| (j |*| k)"; |
184 |
by (rtac cardinal_cong 1); |
|
185 |
by (rtac ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS |
|
1461 | 186 |
eqpoll_trans) 1); |
846 | 187 |
by (rtac (sum_prod_distrib_eqpoll RS eqpoll_trans) 2); |
4312 | 188 |
by (rtac ([well_ord_cardinal_eqpoll, well_ord_cardinal_eqpoll] MRS |
189 |
sum_eqpoll_cong RS eqpoll_sym) 2); |
|
846 | 190 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_radd] 1)); |
191 |
qed "well_ord_cadd_cmult_distrib"; |
|
192 |
||
437 | 193 |
(** Multiplication by 0 yields 0 **) |
194 |
||
5067 | 195 |
Goalw [eqpoll_def] "0*A eqpoll 0"; |
437 | 196 |
by (rtac exI 1); |
197 |
by (rtac lam_bijective 1); |
|
4152 | 198 |
by Safe_tac; |
760 | 199 |
qed "prod_0_eqpoll"; |
437 | 200 |
|
5067 | 201 |
Goalw [cmult_def] "0 |*| i = 0"; |
4091 | 202 |
by (asm_simp_tac (simpset() addsimps [prod_0_eqpoll RS cardinal_cong, |
4312 | 203 |
Card_0 RS Card_cardinal_eq]) 1); |
760 | 204 |
qed "cmult_0"; |
6070 | 205 |
Addsimps [cmult_0]; |
437 | 206 |
|
207 |
(** 1 is the identity for multiplication **) |
|
208 |
||
5067 | 209 |
Goalw [eqpoll_def] "{x}*A eqpoll A"; |
437 | 210 |
by (rtac exI 1); |
846 | 211 |
by (resolve_tac [singleton_prod_bij RS bij_converse_bij] 1); |
760 | 212 |
qed "prod_singleton_eqpoll"; |
437 | 213 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
214 |
Goalw [cmult_def, succ_def] "Card(K) ==> 1 |*| K = K"; |
4091 | 215 |
by (asm_simp_tac (simpset() addsimps [prod_singleton_eqpoll RS cardinal_cong, |
4312 | 216 |
Card_cardinal_eq]) 1); |
760 | 217 |
qed "cmult_1"; |
6070 | 218 |
Addsimps [cmult_1]; |
437 | 219 |
|
767 | 220 |
(*** Some inequalities for multiplication ***) |
221 |
||
5067 | 222 |
Goalw [lepoll_def, inj_def] "A lepoll A*A"; |
767 | 223 |
by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1); |
6153 | 224 |
by (Simp_tac 1); |
767 | 225 |
qed "prod_square_lepoll"; |
226 |
||
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
227 |
(*Could probably weaken the premise to well_ord(K,r), or remove using AC*) |
5137 | 228 |
Goalw [cmult_def] "Card(K) ==> K le K |*| K"; |
767 | 229 |
by (rtac le_trans 1); |
230 |
by (rtac well_ord_lepoll_imp_Card_le 2); |
|
231 |
by (rtac prod_square_lepoll 3); |
|
232 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2)); |
|
4312 | 233 |
by (asm_simp_tac (simpset() |
234 |
addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1); |
|
767 | 235 |
qed "cmult_square_le"; |
236 |
||
237 |
(** Multiplication by a non-zero cardinal **) |
|
238 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
239 |
Goalw [lepoll_def, inj_def] "b: B ==> A lepoll A*B"; |
767 | 240 |
by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1); |
6153 | 241 |
by (Asm_simp_tac 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
242 |
qed "prod_lepoll_self"; |
767 | 243 |
|
244 |
(*Could probably weaken the premises to well_ord(K,r), or removing using AC*) |
|
5067 | 245 |
Goalw [cmult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
246 |
"[| Card(K); Ord(L); 0<L |] ==> K le (K |*| L)"; |
767 | 247 |
by (rtac ([Card_cardinal_le, well_ord_lepoll_imp_Card_le] MRS le_trans) 1); |
248 |
by (rtac prod_lepoll_self 3); |
|
249 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord, ltD] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
250 |
qed "cmult_le_self"; |
767 | 251 |
|
252 |
(** Monotonicity of multiplication **) |
|
253 |
||
5067 | 254 |
Goalw [lepoll_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
255 |
"[| A lepoll C; B lepoll D |] ==> A * B lepoll C * D"; |
767 | 256 |
by (REPEAT (etac exE 1)); |
1090
8ab69b3e396b
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1075
diff
changeset
|
257 |
by (res_inst_tac [("x", "lam <w,y>:A*B. <f`w, fa`y>")] exI 1); |
8ab69b3e396b
Changed some definitions and proofs to use pattern-matching.
lcp
parents:
1075
diff
changeset
|
258 |
by (res_inst_tac [("d", "%<w,y>.<converse(f)`w, converse(fa)`y>")] |
1461 | 259 |
lam_injective 1); |
6153 | 260 |
by (typecheck_tac (tcset() addTCs [inj_is_fun])); |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
261 |
by Auto_tac; |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
262 |
qed "prod_lepoll_mono"; |
767 | 263 |
|
5067 | 264 |
Goalw [cmult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
265 |
"[| K' le K; L' le L |] ==> (K' |*| L') le (K |*| L)"; |
4091 | 266 |
by (safe_tac (claset() addSDs [le_subset_iff RS iffD1])); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
267 |
by (rtac well_ord_lepoll_imp_Card_le 1); |
767 | 268 |
by (REPEAT (ares_tac [prod_lepoll_mono, subset_imp_lepoll] 2)); |
269 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
270 |
qed "cmult_le_mono"; |
767 | 271 |
|
272 |
(*** Multiplication of finite cardinals is "ordinary" multiplication ***) |
|
437 | 273 |
|
5067 | 274 |
Goalw [eqpoll_def] "succ(A)*B eqpoll B + A*B"; |
437 | 275 |
by (rtac exI 1); |
6068 | 276 |
by (res_inst_tac [("c", "%<x,y>. if x=A then Inl(y) else Inr(<x,y>)"), |
3840 | 277 |
("d", "case(%y. <A,y>, %z. z)")] |
437 | 278 |
lam_bijective 1); |
5488 | 279 |
by Safe_tac; |
437 | 280 |
by (ALLGOALS |
4091 | 281 |
(asm_simp_tac (simpset() addsimps [succI2, if_type, mem_imp_not_eq]))); |
760 | 282 |
qed "prod_succ_eqpoll"; |
437 | 283 |
|
284 |
(*Unconditional version requires AC*) |
|
5067 | 285 |
Goalw [cmult_def, cadd_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
286 |
"[| Ord(m); Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)"; |
437 | 287 |
by (rtac (prod_succ_eqpoll RS cardinal_cong RS trans) 1); |
288 |
by (rtac (cardinal_cong RS sym) 1); |
|
289 |
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong) 1); |
|
290 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
760 | 291 |
qed "cmult_succ_lemma"; |
437 | 292 |
|
6070 | 293 |
Goal "[| m: nat; n: nat |] ==> m |*| n = m#*n"; |
294 |
by (induct_tac "m" 1); |
|
295 |
by (Asm_simp_tac 1); |
|
8201 | 296 |
by (asm_simp_tac (simpset() addsimps [cmult_succ_lemma, nat_cadd_eq_add]) 1); |
760 | 297 |
qed "nat_cmult_eq_mult"; |
437 | 298 |
|
5137 | 299 |
Goal "Card(n) ==> 2 |*| n = n |+| n"; |
767 | 300 |
by (asm_simp_tac |
6153 | 301 |
(simpset() addsimps [cmult_succ_lemma, Card_is_Ord, |
8551 | 302 |
inst "j" "0" cadd_commute]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
303 |
qed "cmult_2"; |
767 | 304 |
|
437 | 305 |
|
9907 | 306 |
bind_thm ("sum_lepoll_prod", [sum_eq_2_times RS equalityD1 RS subset_imp_lepoll, |
5284
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
307 |
asm_rl, lepoll_refl] MRS (prod_lepoll_mono RSN (2, lepoll_trans)) |
9907 | 308 |
|> standard); |
5284
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
309 |
|
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
310 |
Goal "[| A lepoll B; 2 lepoll A |] ==> A+B lepoll A*B"; |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
311 |
by (REPEAT (ares_tac [[sum_lepoll_mono, sum_lepoll_prod] |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
312 |
MRS lepoll_trans, lepoll_refl] 1)); |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
313 |
qed "lepoll_imp_sum_lepoll_prod"; |
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
314 |
|
c77e9dd9bafc
Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents:
5147
diff
changeset
|
315 |
|
437 | 316 |
(*** Infinite Cardinals are Limit Ordinals ***) |
317 |
||
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
318 |
(*This proof is modelled upon one assuming nat<=A, with injection |
6068 | 319 |
lam z:cons(u,A). if z=u then 0 else if z : nat then succ(z) else z |
320 |
and inverse %y. if y:nat then nat_case(u, %z. z, y) else y. \ |
|
321 |
If f: inj(nat,A) then range(f) behaves like the natural numbers.*) |
|
322 |
Goalw [lepoll_def] "nat lepoll A ==> cons(u,A) lepoll A"; |
|
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
323 |
by (etac exE 1); |
516 | 324 |
by (res_inst_tac [("x", |
6068 | 325 |
"lam z:cons(u,A). if z=u then f`0 \ |
326 |
\ else if z: range(f) then f`succ(converse(f)`z) \ |
|
327 |
\ else z")] exI 1); |
|
328 |
by (res_inst_tac [("d", "%y. if y: range(f) \ |
|
329 |
\ then nat_case(u, %z. f`z, converse(f)`y) \ |
|
330 |
\ else y")] |
|
516 | 331 |
lam_injective 1); |
5137 | 332 |
by (fast_tac (claset() addSIs [if_type, apply_type] |
333 |
addIs [inj_is_fun, inj_converse_fun]) 1); |
|
516 | 334 |
by (asm_simp_tac |
4091 | 335 |
(simpset() addsimps [inj_is_fun RS apply_rangeI, |
4312 | 336 |
inj_converse_fun RS apply_rangeI, |
6176
707b6f9859d2
tidied, with left_inverse & right_inverse as default simprules
paulson
parents:
6153
diff
changeset
|
337 |
inj_converse_fun RS apply_funtype]) 1); |
760 | 338 |
qed "nat_cons_lepoll"; |
516 | 339 |
|
5137 | 340 |
Goal "nat lepoll A ==> cons(u,A) eqpoll A"; |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
341 |
by (etac (nat_cons_lepoll RS eqpollI) 1); |
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
342 |
by (rtac (subset_consI RS subset_imp_lepoll) 1); |
760 | 343 |
qed "nat_cons_eqpoll"; |
437 | 344 |
|
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
345 |
(*Specialized version required below*) |
5137 | 346 |
Goalw [succ_def] "nat <= A ==> succ(A) eqpoll A"; |
571
0b03ce5b62f7
ZF/Cardinal: some results moved here from CardinalArith
lcp
parents:
523
diff
changeset
|
347 |
by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1); |
760 | 348 |
qed "nat_succ_eqpoll"; |
437 | 349 |
|
5067 | 350 |
Goalw [InfCard_def] "InfCard(nat)"; |
4091 | 351 |
by (blast_tac (claset() addIs [Card_nat, le_refl, Card_is_Ord]) 1); |
760 | 352 |
qed "InfCard_nat"; |
488 | 353 |
|
5137 | 354 |
Goalw [InfCard_def] "InfCard(K) ==> Card(K)"; |
437 | 355 |
by (etac conjunct1 1); |
760 | 356 |
qed "InfCard_is_Card"; |
437 | 357 |
|
5067 | 358 |
Goalw [InfCard_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
359 |
"[| InfCard(K); Card(L) |] ==> InfCard(K Un L)"; |
4091 | 360 |
by (asm_simp_tac (simpset() addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), |
4312 | 361 |
Card_is_Ord]) 1); |
760 | 362 |
qed "InfCard_Un"; |
523 | 363 |
|
437 | 364 |
(*Kunen's Lemma 10.11*) |
5137 | 365 |
Goalw [InfCard_def] "InfCard(K) ==> Limit(K)"; |
437 | 366 |
by (etac conjE 1); |
7499 | 367 |
by (ftac Card_is_Ord 1); |
437 | 368 |
by (rtac (ltI RS non_succ_LimitI) 1); |
369 |
by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1); |
|
4091 | 370 |
by (safe_tac (claset() addSDs [Limit_nat RS Limit_le_succD])); |
437 | 371 |
by (rewtac Card_def); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
372 |
by (dtac trans 1); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
373 |
by (etac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1); |
3016 | 374 |
by (etac (Ord_cardinal_le RS lt_trans2 RS lt_irrefl) 1); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
375 |
by (REPEAT (ares_tac [le_eqI, Ord_cardinal] 1)); |
760 | 376 |
qed "InfCard_is_Limit"; |
437 | 377 |
|
378 |
||
379 |
(*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***) |
|
380 |
||
381 |
(*A general fact about ordermap*) |
|
5325
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5284
diff
changeset
|
382 |
Goalw [eqpoll_def] |
f7a5e06adea1
Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents:
5284
diff
changeset
|
383 |
"[| well_ord(A,r); x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)"; |
437 | 384 |
by (rtac exI 1); |
4091 | 385 |
by (asm_simp_tac (simpset() addsimps [ordermap_eq_image, well_ord_is_wf]) 1); |
467 | 386 |
by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1); |
437 | 387 |
by (rtac pred_subset 1); |
760 | 388 |
qed "ordermap_eqpoll_pred"; |
437 | 389 |
|
390 |
(** Establishing the well-ordering **) |
|
391 |
||
5488 | 392 |
Goalw [inj_def] "Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)"; |
393 |
by (force_tac (claset() addIs [lam_type, Un_least_lt RS ltD, ltI], |
|
394 |
simpset()) 1); |
|
760 | 395 |
qed "csquare_lam_inj"; |
437 | 396 |
|
5488 | 397 |
Goalw [csquare_rel_def] "Ord(K) ==> well_ord(K*K, csquare_rel(K))"; |
437 | 398 |
by (rtac (csquare_lam_inj RS well_ord_rvimage) 1); |
399 |
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1)); |
|
760 | 400 |
qed "well_ord_csquare"; |
437 | 401 |
|
402 |
(** Characterising initial segments of the well-ordering **) |
|
403 |
||
5067 | 404 |
Goalw [csquare_rel_def] |
5488 | 405 |
"[| <<x,y>, <z,z>> : csquare_rel(K); x<K; y<K; z<K |] ==> x le z & y le z"; |
406 |
by (etac rev_mp 1); |
|
437 | 407 |
by (REPEAT (etac ltE 1)); |
9842 | 408 |
by (asm_simp_tac (simpset() addsimps [rvimage_iff, |
4312 | 409 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
4091 | 410 |
by (safe_tac (claset() addSEs [mem_irrefl] |
4312 | 411 |
addSIs [Un_upper1_le, Un_upper2_le])); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
412 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_def, succI2]))); |
5488 | 413 |
qed "csquareD"; |
437 | 414 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9180
diff
changeset
|
415 |
Goalw [Order.pred_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
416 |
"z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)"; |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9180
diff
changeset
|
417 |
by (safe_tac (claset() delrules [SigmaI, succCI])); |
5488 | 418 |
by (etac (csquareD RS conjE) 1); |
437 | 419 |
by (rewtac lt_def); |
2925 | 420 |
by (ALLGOALS Blast_tac); |
760 | 421 |
qed "pred_csquare_subset"; |
437 | 422 |
|
5067 | 423 |
Goalw [csquare_rel_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
424 |
"[| x<z; y<z; z<K |] ==> <<x,y>, <z,z>> : csquare_rel(K)"; |
484 | 425 |
by (subgoals_tac ["x<K", "y<K"] 1); |
437 | 426 |
by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2)); |
427 |
by (REPEAT (etac ltE 1)); |
|
9842 | 428 |
by (asm_simp_tac (simpset() addsimps [rvimage_iff, |
4312 | 429 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
760 | 430 |
qed "csquare_ltI"; |
437 | 431 |
|
432 |
(*Part of the traditional proof. UNUSED since it's harder to prove & apply *) |
|
5067 | 433 |
Goalw [csquare_rel_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
434 |
"[| x le z; y le z; z<K |] ==> \ |
484 | 435 |
\ <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z"; |
436 |
by (subgoals_tac ["x<K", "y<K"] 1); |
|
437 | 437 |
by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2)); |
438 |
by (REPEAT (etac ltE 1)); |
|
9842 | 439 |
by (asm_simp_tac (simpset() addsimps [rvimage_iff, |
4312 | 440 |
Un_absorb, Un_least_mem_iff, ltD]) 1); |
437 | 441 |
by (REPEAT_FIRST (etac succE)); |
442 |
by (ALLGOALS |
|
4091 | 443 |
(asm_simp_tac (simpset() addsimps [subset_Un_iff RS iff_sym, |
4312 | 444 |
subset_Un_iff2 RS iff_sym, OrdmemD]))); |
760 | 445 |
qed "csquare_or_eqI"; |
437 | 446 |
|
447 |
(** The cardinality of initial segments **) |
|
448 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
449 |
Goal "[| Limit(K); x<K; y<K; z=succ(x Un y) |] ==> \ |
1461 | 450 |
\ ordermap(K*K, csquare_rel(K)) ` <x,y> < \ |
484 | 451 |
\ ordermap(K*K, csquare_rel(K)) ` <z,z>"; |
452 |
by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1); |
|
846 | 453 |
by (etac (Limit_is_Ord RS well_ord_csquare) 2); |
4091 | 454 |
by (blast_tac (claset() addSIs [Un_least_lt, Limit_has_succ]) 2); |
870 | 455 |
by (rtac (csquare_ltI RS ordermap_mono RS ltI) 1); |
437 | 456 |
by (etac well_ord_is_wf 4); |
457 |
by (ALLGOALS |
|
4091 | 458 |
(blast_tac (claset() addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap] |
4312 | 459 |
addSEs [ltE]))); |
870 | 460 |
qed "ordermap_z_lt"; |
437 | 461 |
|
484 | 462 |
(*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *) |
5067 | 463 |
Goalw [cmult_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
464 |
"[| Limit(K); x<K; y<K; z=succ(x Un y) |] ==> \ |
484 | 465 |
\ | ordermap(K*K, csquare_rel(K)) ` <x,y> | le |succ(z)| |*| |succ(z)|"; |
767 | 466 |
by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1); |
437 | 467 |
by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1)); |
484 | 468 |
by (subgoals_tac ["z<K"] 1); |
4091 | 469 |
by (blast_tac (claset() addSIs [Un_least_lt, Limit_has_succ]) 2); |
1609 | 470 |
by (rtac (ordermap_z_lt RS leI RS le_imp_lepoll RS lepoll_trans) 1); |
437 | 471 |
by (REPEAT_SOME assume_tac); |
472 |
by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1); |
|
846 | 473 |
by (etac (Limit_is_Ord RS well_ord_csquare) 1); |
4091 | 474 |
by (blast_tac (claset() addIs [ltD]) 1); |
437 | 475 |
by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN |
476 |
assume_tac 1); |
|
477 |
by (REPEAT_FIRST (etac ltE)); |
|
478 |
by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1); |
|
479 |
by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll))); |
|
760 | 480 |
qed "ordermap_csquare_le"; |
437 | 481 |
|
484 | 482 |
(*Kunen: "... so the order type <= K" *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
483 |
Goal "[| InfCard(K); ALL y:K. InfCard(y) --> y |*| y = y |] ==> \ |
484 | 484 |
\ ordertype(K*K, csquare_rel(K)) le K"; |
437 | 485 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); |
486 |
by (rtac all_lt_imp_le 1); |
|
487 |
by (assume_tac 1); |
|
488 |
by (etac (well_ord_csquare RS Ord_ordertype) 1); |
|
489 |
by (rtac Card_lt_imp_lt 1); |
|
490 |
by (etac InfCard_is_Card 3); |
|
491 |
by (etac ltE 2 THEN assume_tac 2); |
|
4091 | 492 |
by (asm_full_simp_tac (simpset() addsimps [ordertype_unfold]) 1); |
493 |
by (safe_tac (claset() addSEs [ltE])); |
|
12152
46f128d8133c
Renamed some bound variables due to changes in simplifier.
berghofe
parents:
12089
diff
changeset
|
494 |
by (subgoals_tac ["Ord(xa)", "Ord(ya)"] 1); |
437 | 495 |
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2)); |
846 | 496 |
by (rtac (InfCard_is_Limit RS ordermap_csquare_le RS lt_trans1) 1 THEN |
437 | 497 |
REPEAT (ares_tac [refl] 1 ORELSE etac ltI 1)); |
12152
46f128d8133c
Renamed some bound variables due to changes in simplifier.
berghofe
parents:
12089
diff
changeset
|
498 |
by (res_inst_tac [("i","xa Un ya"), ("j","nat")] Ord_linear2 1 THEN |
437 | 499 |
REPEAT (ares_tac [Ord_Un, Ord_nat] 1)); |
12152
46f128d8133c
Renamed some bound variables due to changes in simplifier.
berghofe
parents:
12089
diff
changeset
|
500 |
(*the finite case: xa Un ya < nat *) |
437 | 501 |
by (res_inst_tac [("j", "nat")] lt_trans2 1); |
4091 | 502 |
by (asm_full_simp_tac (simpset() addsimps [InfCard_def]) 2); |
437 | 503 |
by (asm_full_simp_tac |
4091 | 504 |
(simpset() addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type, |
4312 | 505 |
nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1); |
12152
46f128d8133c
Renamed some bound variables due to changes in simplifier.
berghofe
parents:
12089
diff
changeset
|
506 |
(*case nat le (xa Un ya) *) |
437 | 507 |
by (asm_full_simp_tac |
4091 | 508 |
(simpset() addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong, |
4312 | 509 |
le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt, |
510 |
Ord_Un, ltI, nat_le_cardinal, |
|
511 |
Ord_cardinal_le RS lt_trans1 RS ltD]) 1); |
|
760 | 512 |
qed "ordertype_csquare_le"; |
437 | 513 |
|
514 |
(*Main result: Kunen's Theorem 10.12*) |
|
5137 | 515 |
Goal "InfCard(K) ==> K |*| K = K"; |
437 | 516 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1); |
517 |
by (etac rev_mp 1); |
|
484 | 518 |
by (trans_ind_tac "K" [] 1); |
437 | 519 |
by (rtac impI 1); |
520 |
by (rtac le_anti_sym 1); |
|
521 |
by (etac (InfCard_is_Card RS cmult_square_le) 2); |
|
522 |
by (rtac (ordertype_csquare_le RSN (2, le_trans)) 1); |
|
523 |
by (assume_tac 2); |
|
524 |
by (assume_tac 2); |
|
525 |
by (asm_simp_tac |
|
4091 | 526 |
(simpset() addsimps [cmult_def, Ord_cardinal_le, |
4312 | 527 |
well_ord_csquare RS ordermap_bij RS |
528 |
bij_imp_eqpoll RS cardinal_cong, |
|
529 |
well_ord_csquare RS Ord_ordertype]) 1); |
|
760 | 530 |
qed "InfCard_csquare_eq"; |
484 | 531 |
|
767 | 532 |
(*Corollary for arbitrary well-ordered sets (all sets, assuming AC)*) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
533 |
Goal "[| well_ord(A,r); InfCard(|A|) |] ==> A*A eqpoll A"; |
484 | 534 |
by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1); |
535 |
by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1)); |
|
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
536 |
by (rtac well_ord_cardinal_eqE 1); |
484 | 537 |
by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1)); |
4312 | 538 |
by (asm_simp_tac (simpset() |
539 |
addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1); |
|
760 | 540 |
qed "well_ord_InfCard_square_eq"; |
484 | 541 |
|
767 | 542 |
(** Toward's Kunen's Corollary 10.13 (1) **) |
543 |
||
5137 | 544 |
Goal "[| InfCard(K); L le K; 0<L |] ==> K |*| L = K"; |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
545 |
by (rtac le_anti_sym 1); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
546 |
by (etac ltE 2 THEN |
767 | 547 |
REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2)); |
548 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1); |
|
549 |
by (resolve_tac [cmult_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1)); |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
550 |
by (asm_full_simp_tac (simpset() addsimps [InfCard_csquare_eq]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
551 |
qed "InfCard_le_cmult_eq"; |
767 | 552 |
|
553 |
(*Corollary 10.13 (1), for cardinal multiplication*) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
554 |
Goal "[| InfCard(K); InfCard(L) |] ==> K |*| L = K Un L"; |
767 | 555 |
by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1); |
6153 | 556 |
by (typecheck_tac (tcset() addTCs [InfCard_is_Card, Card_is_Ord])); |
767 | 557 |
by (resolve_tac [cmult_commute RS ssubst] 1); |
558 |
by (resolve_tac [Un_commute RS ssubst] 1); |
|
559 |
by (ALLGOALS |
|
560 |
(asm_simp_tac |
|
4091 | 561 |
(simpset() addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq, |
4312 | 562 |
subset_Un_iff2 RS iffD1, le_imp_subset]))); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
563 |
qed "InfCard_cmult_eq"; |
767 | 564 |
|
5137 | 565 |
Goal "InfCard(K) ==> K |+| K = K"; |
767 | 566 |
by (asm_simp_tac |
4091 | 567 |
(simpset() addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1); |
6153 | 568 |
by (asm_simp_tac |
569 |
(simpset() addsimps [InfCard_le_cmult_eq, InfCard_is_Limit, |
|
570 |
Limit_has_0, Limit_has_succ]) 1); |
|
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
571 |
qed "InfCard_cdouble_eq"; |
767 | 572 |
|
573 |
(*Corollary 10.13 (1), for cardinal addition*) |
|
5137 | 574 |
Goal "[| InfCard(K); L le K |] ==> K |+| L = K"; |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
575 |
by (rtac le_anti_sym 1); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
576 |
by (etac ltE 2 THEN |
767 | 577 |
REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2)); |
578 |
by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1); |
|
579 |
by (resolve_tac [cadd_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1)); |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
580 |
by (asm_full_simp_tac (simpset() addsimps [InfCard_cdouble_eq]) 1); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
581 |
qed "InfCard_le_cadd_eq"; |
767 | 582 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
583 |
Goal "[| InfCard(K); InfCard(L) |] ==> K |+| L = K Un L"; |
767 | 584 |
by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1); |
6153 | 585 |
by (typecheck_tac (tcset() addTCs [InfCard_is_Card, Card_is_Ord])); |
767 | 586 |
by (resolve_tac [cadd_commute RS ssubst] 1); |
587 |
by (resolve_tac [Un_commute RS ssubst] 1); |
|
588 |
by (ALLGOALS |
|
589 |
(asm_simp_tac |
|
4091 | 590 |
(simpset() addsimps [InfCard_le_cadd_eq, |
4312 | 591 |
subset_Un_iff2 RS iffD1, le_imp_subset]))); |
782
200a16083201
added bind_thm for theorems defined by "standard ..."
clasohm
parents:
767
diff
changeset
|
592 |
qed "InfCard_cadd_eq"; |
767 | 593 |
|
594 |
(*The other part, Corollary 10.13 (2), refers to the cardinality of the set |
|
595 |
of all n-tuples of elements of K. A better version for the Isabelle theory |
|
596 |
might be InfCard(K) ==> |list(K)| = K. |
|
597 |
*) |
|
484 | 598 |
|
599 |
(*** For every cardinal number there exists a greater one |
|
600 |
[Kunen's Theorem 10.16, which would be trivial using AC] ***) |
|
601 |
||
5067 | 602 |
Goalw [jump_cardinal_def] "Ord(jump_cardinal(K))"; |
484 | 603 |
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1); |
4091 | 604 |
by (blast_tac (claset() addSIs [Ord_ordertype]) 2); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
605 |
by (rewtac Transset_def); |
1075
848bf2e18dff
Modified proofs for new claset primitives. The problem is that they enforce
lcp
parents:
989
diff
changeset
|
606 |
by (safe_tac subset_cs); |
4091 | 607 |
by (asm_full_simp_tac (simpset() addsimps [ordertype_pred_unfold]) 1); |
4152 | 608 |
by Safe_tac; |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
609 |
by (rtac UN_I 1); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
610 |
by (rtac ReplaceI 2); |
4091 | 611 |
by (ALLGOALS (blast_tac (claset() addIs [well_ord_subset] addSEs [predE]))); |
760 | 612 |
qed "Ord_jump_cardinal"; |
484 | 613 |
|
614 |
(*Allows selective unfolding. Less work than deriving intro/elim rules*) |
|
5067 | 615 |
Goalw [jump_cardinal_def] |
484 | 616 |
"i : jump_cardinal(K) <-> \ |
617 |
\ (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))"; |
|
1461 | 618 |
by (fast_tac subset_cs 1); (*It's vital to avoid reasoning about <=*) |
760 | 619 |
qed "jump_cardinal_iff"; |
484 | 620 |
|
621 |
(*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*) |
|
5137 | 622 |
Goal "Ord(K) ==> K < jump_cardinal(K)"; |
484 | 623 |
by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1); |
624 |
by (resolve_tac [jump_cardinal_iff RS iffD2] 1); |
|
625 |
by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel])); |
|
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
626 |
by (rtac subset_refl 2); |
4091 | 627 |
by (asm_simp_tac (simpset() addsimps [Memrel_def, subset_iff]) 1); |
628 |
by (asm_simp_tac (simpset() addsimps [ordertype_Memrel]) 1); |
|
760 | 629 |
qed "K_lt_jump_cardinal"; |
484 | 630 |
|
631 |
(*The proof by contradiction: the bijection f yields a wellordering of X |
|
632 |
whose ordertype is jump_cardinal(K). *) |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
633 |
Goal "[| well_ord(X,r); r <= K * K; X <= K; \ |
1461 | 634 |
\ f : bij(ordertype(X,r), jump_cardinal(K)) \ |
635 |
\ |] ==> jump_cardinal(K) : jump_cardinal(K)"; |
|
484 | 636 |
by (subgoal_tac "f O ordermap(X,r): bij(X, jump_cardinal(K))" 1); |
637 |
by (REPEAT (ares_tac [comp_bij, ordermap_bij] 2)); |
|
638 |
by (resolve_tac [jump_cardinal_iff RS iffD2] 1); |
|
639 |
by (REPEAT_FIRST (resolve_tac [exI, conjI])); |
|
640 |
by (rtac ([rvimage_type, Sigma_mono] MRS subset_trans) 1); |
|
641 |
by (REPEAT (assume_tac 1)); |
|
642 |
by (etac (bij_is_inj RS well_ord_rvimage) 1); |
|
643 |
by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1); |
|
644 |
by (asm_simp_tac |
|
4091 | 645 |
(simpset() addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), |
4312 | 646 |
ordertype_Memrel, Ord_jump_cardinal]) 1); |
760 | 647 |
qed "Card_jump_cardinal_lemma"; |
484 | 648 |
|
649 |
(*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*) |
|
5067 | 650 |
Goal "Card(jump_cardinal(K))"; |
484 | 651 |
by (rtac (Ord_jump_cardinal RS CardI) 1); |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
652 |
by (rewtac eqpoll_def); |
4091 | 653 |
by (safe_tac (claset() addSDs [ltD, jump_cardinal_iff RS iffD1])); |
484 | 654 |
by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1)); |
760 | 655 |
qed "Card_jump_cardinal"; |
484 | 656 |
|
657 |
(*** Basic properties of successor cardinals ***) |
|
658 |
||
5067 | 659 |
Goalw [csucc_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
660 |
"Ord(K) ==> Card(csucc(K)) & K < csucc(K)"; |
484 | 661 |
by (rtac LeastI 1); |
662 |
by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal, |
|
1461 | 663 |
Ord_jump_cardinal] 1)); |
760 | 664 |
qed "csucc_basic"; |
484 | 665 |
|
800
23f55b829ccb
Limit_csucc: moved to InfDatatype and proved explicitly in
lcp
parents:
782
diff
changeset
|
666 |
bind_thm ("Card_csucc", csucc_basic RS conjunct1); |
484 | 667 |
|
800
23f55b829ccb
Limit_csucc: moved to InfDatatype and proved explicitly in
lcp
parents:
782
diff
changeset
|
668 |
bind_thm ("lt_csucc", csucc_basic RS conjunct2); |
484 | 669 |
|
5137 | 670 |
Goal "Ord(K) ==> 0 < csucc(K)"; |
517 | 671 |
by (resolve_tac [[Ord_0_le, lt_csucc] MRS lt_trans1] 1); |
672 |
by (REPEAT (assume_tac 1)); |
|
760 | 673 |
qed "Ord_0_lt_csucc"; |
517 | 674 |
|
5067 | 675 |
Goalw [csucc_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
676 |
"[| Card(L); K<L |] ==> csucc(K) le L"; |
484 | 677 |
by (rtac Least_le 1); |
678 |
by (REPEAT (ares_tac [conjI, Card_is_Ord] 1)); |
|
760 | 679 |
qed "csucc_le"; |
484 | 680 |
|
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
681 |
Goal "[| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K"; |
823
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
682 |
by (rtac iffI 1); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
683 |
by (rtac Card_lt_imp_lt 2); |
33dc37d46296
Changed succ(1) to 2 in cmult_2; Simplified proof of InfCard_is_Limit
lcp
parents:
800
diff
changeset
|
684 |
by (etac lt_trans1 2); |
484 | 685 |
by (REPEAT (ares_tac [lt_csucc, Card_csucc, Card_is_Ord] 2)); |
686 |
by (resolve_tac [notI RS not_lt_imp_le] 1); |
|
687 |
by (resolve_tac [Card_cardinal RS csucc_le RS lt_trans1 RS lt_irrefl] 1); |
|
688 |
by (assume_tac 1); |
|
689 |
by (resolve_tac [Ord_cardinal_le RS lt_trans1] 1); |
|
690 |
by (REPEAT (ares_tac [Ord_cardinal] 1 |
|
691 |
ORELSE eresolve_tac [ltE, Card_is_Ord] 1)); |
|
760 | 692 |
qed "lt_csucc_iff"; |
484 | 693 |
|
9180 | 694 |
Goal "[| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K"; |
484 | 695 |
by (asm_simp_tac |
4091 | 696 |
(simpset() addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1); |
760 | 697 |
qed "Card_lt_csucc_iff"; |
488 | 698 |
|
5067 | 699 |
Goalw [InfCard_def] |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
700 |
"InfCard(K) ==> InfCard(csucc(K))"; |
4091 | 701 |
by (asm_simp_tac (simpset() addsimps [Card_csucc, Card_is_Ord, |
4312 | 702 |
lt_csucc RS leI RSN (2,le_trans)]) 1); |
760 | 703 |
qed "InfCard_csucc"; |
517 | 704 |
|
1609 | 705 |
|
706 |
(*** Finite sets ***) |
|
707 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
708 |
Goal "n: nat ==> ALL A. A eqpoll n --> A : Fin(A)"; |
6070 | 709 |
by (induct_tac "n" 1); |
5529 | 710 |
by (simp_tac (simpset() addsimps eqpoll_0_iff::Fin.intrs) 1); |
3736
39ee3d31cfbc
Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents:
3016
diff
changeset
|
711 |
by (Clarify_tac 1); |
1609 | 712 |
by (subgoal_tac "EX u. u:A" 1); |
1622 | 713 |
by (etac exE 1); |
1609 | 714 |
by (resolve_tac [Diff_sing_eqpoll RS revcut_rl] 1); |
715 |
by (assume_tac 2); |
|
716 |
by (assume_tac 1); |
|
717 |
by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1); |
|
718 |
by (assume_tac 1); |
|
719 |
by (resolve_tac [Fin.consI] 1); |
|
2925 | 720 |
by (Blast_tac 1); |
4091 | 721 |
by (blast_tac (claset() addIs [subset_consI RS Fin_mono RS subsetD]) 1); |
1609 | 722 |
(*Now for the lemma assumed above*) |
1622 | 723 |
by (rewtac eqpoll_def); |
4091 | 724 |
by (blast_tac (claset() addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1); |
1609 | 725 |
val lemma = result(); |
726 |
||
5137 | 727 |
Goalw [Finite_def] "Finite(A) ==> A : Fin(A)"; |
4091 | 728 |
by (blast_tac (claset() addIs [lemma RS spec RS mp]) 1); |
1609 | 729 |
qed "Finite_into_Fin"; |
730 |
||
5137 | 731 |
Goal "A : Fin(U) ==> Finite(A)"; |
4091 | 732 |
by (fast_tac (claset() addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1); |
1609 | 733 |
qed "Fin_into_Finite"; |
734 |
||
5067 | 735 |
Goal "Finite(A) <-> A : Fin(A)"; |
4091 | 736 |
by (blast_tac (claset() addIs [Finite_into_Fin, Fin_into_Finite]) 1); |
1609 | 737 |
qed "Finite_Fin_iff"; |
738 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
739 |
Goal "[| Finite(A); Finite(B) |] ==> Finite(A Un B)"; |
4091 | 740 |
by (blast_tac (claset() addSIs [Fin_into_Finite, Fin_UnI] |
4312 | 741 |
addSDs [Finite_into_Fin] |
742 |
addIs [Un_upper1 RS Fin_mono RS subsetD, |
|
743 |
Un_upper2 RS Fin_mono RS subsetD]) 1); |
|
1609 | 744 |
qed "Finite_Un"; |
745 |
||
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
746 |
Goal "[| ALL y:X. Finite(y); Finite(X) |] ==> Finite(Union(X))"; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
747 |
by (asm_full_simp_tac (simpset() addsimps [Finite_Fin_iff]) 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
748 |
by (rtac Fin_UnionI 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
749 |
by (etac Fin.induct 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
750 |
by (Simp_tac 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
751 |
by (blast_tac (claset() addIs [Fin.consI, impOfSubs Fin_mono]) 1); |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
752 |
qed "Finite_Union"; |
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
753 |
|
12089
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
754 |
(* Induction principle for Finite(A), by Sidi Ehmety *) |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
755 |
val major::prems = Goal |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
756 |
"[| Finite(A); P(0); \ |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
757 |
\ !! x B. [| Finite(B); x ~: B; P(B) |] \ |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
758 |
\ ==> P(cons(x, B)) |] \ |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
759 |
\ ==> P(A)"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
760 |
by (resolve_tac [major RS Finite_into_Fin RS Fin_induct] 1); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
761 |
by (ALLGOALS(resolve_tac prems)); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
762 |
by (ALLGOALS(asm_simp_tac (simpset() addsimps [Fin_into_Finite]))); |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
763 |
qed "Finite_induct"; |
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
paulson
parents:
9907
diff
changeset
|
764 |
|
1609 | 765 |
|
766 |
(** Removing elements from a finite set decreases its cardinality **) |
|
767 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
768 |
Goal "A: Fin(U) ==> x~:A --> ~ cons(x,A) lepoll A"; |
1622 | 769 |
by (etac Fin_induct 1); |
4091 | 770 |
by (simp_tac (simpset() addsimps [lepoll_0_iff]) 1); |
1609 | 771 |
by (subgoal_tac "cons(x,cons(xa,y)) = cons(xa,cons(x,y))" 1); |
2469 | 772 |
by (Asm_simp_tac 1); |
4091 | 773 |
by (blast_tac (claset() addSDs [cons_lepoll_consD]) 1); |
2925 | 774 |
by (Blast_tac 1); |
1609 | 775 |
qed "Fin_imp_not_cons_lepoll"; |
776 |
||
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5137
diff
changeset
|
777 |
Goal "[| Finite(A); a~:A |] ==> |cons(a,A)| = succ(|A|)"; |
1622 | 778 |
by (rewtac cardinal_def); |
779 |
by (rtac Least_equality 1); |
|
1609 | 780 |
by (fold_tac [cardinal_def]); |
4091 | 781 |
by (simp_tac (simpset() addsimps [succ_def]) 1); |
782 |
by (blast_tac (claset() addIs [cons_eqpoll_cong, well_ord_cardinal_eqpoll] |
|
4312 | 783 |
addSEs [mem_irrefl] |
784 |
addSDs [Finite_imp_well_ord]) 1); |
|
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
785 |
by (blast_tac (claset() addIs [Card_cardinal, Card_is_Ord]) 1); |
1622 | 786 |
by (rtac notI 1); |
1609 | 787 |
by (resolve_tac [Finite_into_Fin RS Fin_imp_not_cons_lepoll RS mp RS notE] 1); |
788 |
by (assume_tac 1); |
|
789 |
by (assume_tac 1); |
|
790 |
by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] 1); |
|
791 |
by (eresolve_tac [le_imp_lepoll RS lepoll_trans] 1); |
|
4091 | 792 |
by (blast_tac (claset() addIs [well_ord_cardinal_eqpoll RS eqpoll_imp_lepoll] |
1609 | 793 |
addSDs [Finite_imp_well_ord]) 1); |
794 |
qed "Finite_imp_cardinal_cons"; |
|
795 |
||
796 |
||
5137 | 797 |
Goal "[| Finite(A); a:A |] ==> succ(|A-{a}|) = |A|"; |
1609 | 798 |
by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1); |
799 |
by (assume_tac 1); |
|
4091 | 800 |
by (asm_simp_tac (simpset() addsimps [Finite_imp_cardinal_cons, |
1622 | 801 |
Diff_subset RS subset_Finite]) 1); |
4091 | 802 |
by (asm_simp_tac (simpset() addsimps [cons_Diff]) 1); |
1622 | 803 |
qed "Finite_imp_succ_cardinal_Diff"; |
804 |
||
5137 | 805 |
Goal "[| Finite(A); a:A |] ==> |A-{a}| < |A|"; |
1622 | 806 |
by (rtac succ_leE 1); |
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
807 |
by (asm_simp_tac (simpset() addsimps [Finite_imp_succ_cardinal_Diff]) 1); |
1609 | 808 |
qed "Finite_imp_cardinal_Diff"; |
809 |
||
810 |
||
4312 | 811 |
(** Theorems by Krzysztof Grabczewski, proofs by lcp **) |
1609 | 812 |
|
9907 | 813 |
bind_thm ("nat_implies_well_ord", |
814 |
(transfer (the_context ()) nat_into_Ord) RS well_ord_Memrel); |
|
1609 | 815 |
|
5137 | 816 |
Goal "[| m:nat; n:nat |] ==> m + n eqpoll m #+ n"; |
1609 | 817 |
by (rtac eqpoll_trans 1); |
4312 | 818 |
by (resolve_tac [well_ord_radd RS well_ord_cardinal_eqpoll RS eqpoll_sym] 1); |
4477
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
paulson
parents:
4312
diff
changeset
|
819 |
by (REPEAT (etac nat_implies_well_ord 1)); |
4312 | 820 |
by (asm_simp_tac (simpset() |
821 |
addsimps [nat_cadd_eq_add RS sym, cadd_def, eqpoll_refl]) 1); |
|
1609 | 822 |
qed "nat_sum_eqpoll_sum"; |
823 |
||
12614 | 824 |
|
825 |
(*** Theorems by Sidi Ehmety ***) |
|
826 |
||
827 |
(*The contrapositive says ~Finite(A) ==> ~Finite(A-{a}) *) |
|
828 |
Goalw [Finite_def] "Finite(A - {a}) ==> Finite(A)"; |
|
829 |
by (case_tac "a:A" 1); |
|
830 |
by (subgoal_tac "A-{a}=A" 2); |
|
831 |
by Auto_tac; |
|
832 |
by (res_inst_tac [("x", "succ(n)")] bexI 1); |
|
833 |
by (subgoal_tac "cons(a, A - {a}) = A & cons(n, n) = succ(n)" 1); |
|
834 |
by (dres_inst_tac [("a", "a"), ("b", "n")] cons_eqpoll_cong 1); |
|
835 |
by (auto_tac (claset() addDs [mem_irrefl], simpset())); |
|
836 |
qed "Diff_sing_Finite"; |
|
837 |
||
838 |
(*And the contrapositive of this says |
|
839 |
[| ~Finite(A); Finite(B) |] ==> ~Finite(A-B) *) |
|
840 |
Goal "Finite(B) ==> Finite(A-B) --> Finite(A)"; |
|
841 |
by (etac Finite_induct 1); |
|
842 |
by Auto_tac; |
|
843 |
by (case_tac "x:A" 1); |
|
844 |
by (subgoal_tac "A-cons(x, B) = A - B" 2); |
|
845 |
by (subgoal_tac "A - cons(x, B) = (A - B) - {x}" 1); |
|
846 |
by (rotate_tac ~1 1); |
|
847 |
by (Asm_full_simp_tac 1); |
|
848 |
by (dtac Diff_sing_Finite 1); |
|
849 |
by Auto_tac; |
|
850 |
qed_spec_mp "Diff_Finite"; |
|
851 |
||
852 |
Goal "Ord(i) ==> i <= nat --> i : nat | i=nat"; |
|
853 |
by (etac trans_induct3 1); |
|
854 |
by Auto_tac; |
|
855 |
by (blast_tac (claset() addSDs [nat_le_Limit RS le_imp_subset]) 1); |
|
856 |
qed_spec_mp "Ord_subset_natD"; |
|
857 |
||
858 |
Goal "[| Ord(i); i <= nat |] ==> Card(i)"; |
|
859 |
by (blast_tac (claset() addDs [Ord_subset_natD] |
|
860 |
addIs [Card_nat, nat_into_Card]) 1); |
|
861 |
qed "Ord_nat_subset_into_Card"; |
|
862 |
||
863 |
Goal "Finite(A) ==> |A| : nat"; |
|
864 |
by (etac Finite_induct 1); |
|
865 |
by (auto_tac (claset(), |
|
866 |
simpset() addsimps |
|
867 |
[cardinal_0, Finite_imp_cardinal_cons])); |
|
868 |
qed "Finite_cardinal_in_nat"; |
|
869 |
Addsimps [Finite_cardinal_in_nat]; |
|
870 |
||
871 |
Goal "[| Finite(A); x:A |] ==> |A-{x}| = |A| #- 1"; |
|
872 |
by (rtac succ_inject 1); |
|
873 |
by (res_inst_tac [("b", "|A|")] trans 1); |
|
874 |
by (asm_simp_tac (simpset() addsimps |
|
875 |
[Finite_imp_succ_cardinal_Diff]) 1); |
|
876 |
by (subgoal_tac "1 lepoll A" 1); |
|
877 |
by (blast_tac (claset() addIs [not_0_is_lepoll_1]) 2); |
|
878 |
by (forward_tac [Finite_imp_well_ord] 1); |
|
879 |
by (Clarify_tac 1); |
|
880 |
by (rotate_tac ~1 1); |
|
881 |
by (dtac well_ord_lepoll_imp_Card_le 1); |
|
882 |
by (auto_tac (claset(), simpset() addsimps [cardinal_1])); |
|
883 |
by (rtac trans 1); |
|
884 |
by (rtac diff_succ 2); |
|
885 |
by (auto_tac (claset(), simpset() addsimps [Finite_cardinal_in_nat])); |
|
886 |
qed "Finite_Diff_sing_eq_diff_1"; |
|
887 |
||
888 |
Goal "Finite(B) ==> ALL A. |B|<|A| --> A - B ~= 0"; |
|
889 |
by (etac Finite_induct 1); |
|
890 |
by (ALLGOALS(Clarify_tac)); |
|
891 |
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps |
|
892 |
[cardinal_0,Finite_imp_cardinal_cons]))); |
|
893 |
by (case_tac "Finite(A)" 1); |
|
894 |
by (subgoal_tac "Finite(cons(x, B))" 2); |
|
895 |
by (dres_inst_tac [("B", "cons(x, B)")] Diff_Finite 2); |
|
896 |
by (auto_tac (claset(), simpset() addsimps [Finite_0, Finite_cons])); |
|
897 |
by (subgoal_tac "|B|<|A|" 1); |
|
898 |
by (blast_tac (claset() addIs [lt_trans, Ord_cardinal]) 2); |
|
899 |
by (case_tac "x:A" 1); |
|
900 |
by (subgoal_tac "A - cons(x, B)= A - B" 2); |
|
901 |
by Auto_tac; |
|
902 |
by (subgoal_tac "|A| le |cons(x, B)|" 1); |
|
903 |
by (blast_tac (claset() addDs [Finite_cons RS Finite_imp_well_ord] |
|
904 |
addIs [well_ord_lepoll_imp_Card_le, subset_imp_lepoll]) 2); |
|
905 |
by (auto_tac (claset(), simpset() addsimps [Finite_imp_cardinal_cons])); |
|
906 |
by (auto_tac (claset() addSDs [Finite_cardinal_in_nat], |
|
907 |
simpset() addsimps [le_iff])); |
|
908 |
by (blast_tac (claset() addIs [lt_trans]) 1); |
|
909 |
qed_spec_mp "cardinal_lt_imp_Diff_not_0"; |