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