src/ZF/CardinalArith.ML
changeset 4091 771b1f6422a8
parent 3887 04f730731e43
child 4152 451104c223e2
equal deleted inserted replaced
4090:9f1eaab75e8c 4091:771b1f6422a8
    19 
    19 
    20 goalw CardinalArith.thy [eqpoll_def] "A+B eqpoll B+A";
    20 goalw CardinalArith.thy [eqpoll_def] "A+B eqpoll B+A";
    21 by (rtac exI 1);
    21 by (rtac exI 1);
    22 by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] 
    22 by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] 
    23     lam_bijective 1);
    23     lam_bijective 1);
    24 by (safe_tac (!claset addSEs [sumE]));
    24 by (safe_tac (claset() addSEs [sumE]));
    25 by (ALLGOALS (Asm_simp_tac));
    25 by (ALLGOALS (Asm_simp_tac));
    26 qed "sum_commute_eqpoll";
    26 qed "sum_commute_eqpoll";
    27 
    27 
    28 goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
    28 goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
    29 by (rtac (sum_commute_eqpoll RS cardinal_cong) 1);
    29 by (rtac (sum_commute_eqpoll RS cardinal_cong) 1);
    55 by (rtac exI 1);
    55 by (rtac exI 1);
    56 by (rtac bij_0_sum 1);
    56 by (rtac bij_0_sum 1);
    57 qed "sum_0_eqpoll";
    57 qed "sum_0_eqpoll";
    58 
    58 
    59 goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
    59 goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
    60 by (asm_simp_tac (!simpset addsimps [sum_0_eqpoll RS cardinal_cong, 
    60 by (asm_simp_tac (simpset() addsimps [sum_0_eqpoll RS cardinal_cong, 
    61                                   Card_cardinal_eq]) 1);
    61                                   Card_cardinal_eq]) 1);
    62 qed "cadd_0";
    62 qed "cadd_0";
    63 
    63 
    64 (** Addition by another cardinal **)
    64 (** Addition by another cardinal **)
    65 
    65 
    66 goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A+B";
    66 goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A+B";
    67 by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1);
    67 by (res_inst_tac [("x", "lam x:A. Inl(x)")] exI 1);
    68 by (asm_simp_tac (!simpset addsimps [lam_type]) 1);
    68 by (asm_simp_tac (simpset() addsimps [lam_type]) 1);
    69 qed "sum_lepoll_self";
    69 qed "sum_lepoll_self";
    70 
    70 
    71 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
    71 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
    72 goalw CardinalArith.thy [cadd_def]
    72 goalw CardinalArith.thy [cadd_def]
    73     "!!K. [| Card(K);  Ord(L) |] ==> K le (K |+| L)";
    73     "!!K. [| Card(K);  Ord(L) |] ==> K le (K |+| L)";
    86 by (res_inst_tac 
    86 by (res_inst_tac 
    87       [("d", "case(%w. Inl(converse(f)`w), %y. Inr(converse(fa)`y))")] 
    87       [("d", "case(%w. Inl(converse(f)`w), %y. Inr(converse(fa)`y))")] 
    88       lam_injective 1);
    88       lam_injective 1);
    89 by (typechk_tac ([inj_is_fun, case_type, InlI, InrI] @ ZF_typechecks));
    89 by (typechk_tac ([inj_is_fun, case_type, InlI, InrI] @ ZF_typechecks));
    90 by (etac sumE 1);
    90 by (etac sumE 1);
    91 by (ALLGOALS (asm_simp_tac (!simpset addsimps [left_inverse])));
    91 by (ALLGOALS (asm_simp_tac (simpset() addsimps [left_inverse])));
    92 qed "sum_lepoll_mono";
    92 qed "sum_lepoll_mono";
    93 
    93 
    94 goalw CardinalArith.thy [cadd_def]
    94 goalw CardinalArith.thy [cadd_def]
    95     "!!K. [| K' le K;  L' le L |] ==> (K' |+| L') le (K |+| L)";
    95     "!!K. [| K' le K;  L' le L |] ==> (K' |+| L') le (K |+| L)";
    96 by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
    96 by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));
    97 by (rtac well_ord_lepoll_imp_Card_le 1);
    97 by (rtac well_ord_lepoll_imp_Card_le 1);
    98 by (REPEAT (ares_tac [sum_lepoll_mono, subset_imp_lepoll] 2));
    98 by (REPEAT (ares_tac [sum_lepoll_mono, subset_imp_lepoll] 2));
    99 by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1));
    99 by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1));
   100 qed "cadd_le_mono";
   100 qed "cadd_le_mono";
   101 
   101 
   105 by (rtac exI 1);
   105 by (rtac exI 1);
   106 by (res_inst_tac [("c", "%z. if(z=Inl(A),A+B,z)"), 
   106 by (res_inst_tac [("c", "%z. if(z=Inl(A),A+B,z)"), 
   107                   ("d", "%z. if(z=A+B,Inl(A),z)")] 
   107                   ("d", "%z. if(z=A+B,Inl(A),z)")] 
   108     lam_bijective 1);
   108     lam_bijective 1);
   109 by (ALLGOALS
   109 by (ALLGOALS
   110     (asm_simp_tac (!simpset addsimps [succI2, mem_imp_not_eq]
   110     (asm_simp_tac (simpset() addsimps [succI2, mem_imp_not_eq]
   111                            setloop eresolve_tac [sumE,succE])));
   111                            setloop eresolve_tac [sumE,succE])));
   112 qed "sum_succ_eqpoll";
   112 qed "sum_succ_eqpoll";
   113 
   113 
   114 (*Pulling the  succ(...)  outside the |...| requires m, n: nat  *)
   114 (*Pulling the  succ(...)  outside the |...| requires m, n: nat  *)
   115 (*Unconditional version requires AC*)
   115 (*Unconditional version requires AC*)
   123 
   123 
   124 val [mnat,nnat] = goal CardinalArith.thy
   124 val [mnat,nnat] = goal CardinalArith.thy
   125     "[| m: nat;  n: nat |] ==> m |+| n = m#+n";
   125     "[| m: nat;  n: nat |] ==> m |+| n = m#+n";
   126 by (cut_facts_tac [nnat] 1);
   126 by (cut_facts_tac [nnat] 1);
   127 by (nat_ind_tac "m" [mnat] 1);
   127 by (nat_ind_tac "m" [mnat] 1);
   128 by (asm_simp_tac (!simpset addsimps [nat_into_Card RS cadd_0]) 1);
   128 by (asm_simp_tac (simpset() addsimps [nat_into_Card RS cadd_0]) 1);
   129 by (asm_simp_tac (!simpset addsimps [nat_into_Ord, cadd_succ_lemma,
   129 by (asm_simp_tac (simpset() addsimps [nat_into_Ord, cadd_succ_lemma,
   130                                      nat_into_Card RS Card_cardinal_eq]) 1);
   130                                      nat_into_Card RS Card_cardinal_eq]) 1);
   131 qed "nat_cadd_eq_add";
   131 qed "nat_cadd_eq_add";
   132 
   132 
   133 
   133 
   134 (*** Cardinal multiplication ***)
   134 (*** Cardinal multiplication ***)
   138 (*Easier to prove the two directions separately*)
   138 (*Easier to prove the two directions separately*)
   139 goalw CardinalArith.thy [eqpoll_def] "A*B eqpoll B*A";
   139 goalw CardinalArith.thy [eqpoll_def] "A*B eqpoll B*A";
   140 by (rtac exI 1);
   140 by (rtac exI 1);
   141 by (res_inst_tac [("c", "%<x,y>.<y,x>"), ("d", "%<x,y>.<y,x>")] 
   141 by (res_inst_tac [("c", "%<x,y>.<y,x>"), ("d", "%<x,y>.<y,x>")] 
   142     lam_bijective 1);
   142     lam_bijective 1);
   143 by (safe_tac (!claset));
   143 by (safe_tac (claset()));
   144 by (ALLGOALS (Asm_simp_tac));
   144 by (ALLGOALS (Asm_simp_tac));
   145 qed "prod_commute_eqpoll";
   145 qed "prod_commute_eqpoll";
   146 
   146 
   147 goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
   147 goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
   148 by (rtac (prod_commute_eqpoll RS cardinal_cong) 1);
   148 by (rtac (prod_commute_eqpoll RS cardinal_cong) 1);
   190 (** Multiplication by 0 yields 0 **)
   190 (** Multiplication by 0 yields 0 **)
   191 
   191 
   192 goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0";
   192 goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0";
   193 by (rtac exI 1);
   193 by (rtac exI 1);
   194 by (rtac lam_bijective 1);
   194 by (rtac lam_bijective 1);
   195 by (safe_tac (!claset));
   195 by (safe_tac (claset()));
   196 qed "prod_0_eqpoll";
   196 qed "prod_0_eqpoll";
   197 
   197 
   198 goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
   198 goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
   199 by (asm_simp_tac (!simpset addsimps [prod_0_eqpoll RS cardinal_cong, 
   199 by (asm_simp_tac (simpset() addsimps [prod_0_eqpoll RS cardinal_cong, 
   200                                   Card_0 RS Card_cardinal_eq]) 1);
   200                                   Card_0 RS Card_cardinal_eq]) 1);
   201 qed "cmult_0";
   201 qed "cmult_0";
   202 
   202 
   203 (** 1 is the identity for multiplication **)
   203 (** 1 is the identity for multiplication **)
   204 
   204 
   206 by (rtac exI 1);
   206 by (rtac exI 1);
   207 by (resolve_tac [singleton_prod_bij RS bij_converse_bij] 1);
   207 by (resolve_tac [singleton_prod_bij RS bij_converse_bij] 1);
   208 qed "prod_singleton_eqpoll";
   208 qed "prod_singleton_eqpoll";
   209 
   209 
   210 goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
   210 goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
   211 by (asm_simp_tac (!simpset addsimps [prod_singleton_eqpoll RS cardinal_cong, 
   211 by (asm_simp_tac (simpset() addsimps [prod_singleton_eqpoll RS cardinal_cong, 
   212                                   Card_cardinal_eq]) 1);
   212                                   Card_cardinal_eq]) 1);
   213 qed "cmult_1";
   213 qed "cmult_1";
   214 
   214 
   215 (*** Some inequalities for multiplication ***)
   215 (*** Some inequalities for multiplication ***)
   216 
   216 
   217 goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
   217 goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
   218 by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
   218 by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
   219 by (simp_tac (!simpset addsimps [lam_type]) 1);
   219 by (simp_tac (simpset() addsimps [lam_type]) 1);
   220 qed "prod_square_lepoll";
   220 qed "prod_square_lepoll";
   221 
   221 
   222 (*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
   222 (*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
   223 goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K";
   223 goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K";
   224 by (rtac le_trans 1);
   224 by (rtac le_trans 1);
   225 by (rtac well_ord_lepoll_imp_Card_le 2);
   225 by (rtac well_ord_lepoll_imp_Card_le 2);
   226 by (rtac prod_square_lepoll 3);
   226 by (rtac prod_square_lepoll 3);
   227 by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2));
   227 by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2));
   228 by (asm_simp_tac (!simpset addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
   228 by (asm_simp_tac (simpset() addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
   229 qed "cmult_square_le";
   229 qed "cmult_square_le";
   230 
   230 
   231 (** Multiplication by a non-zero cardinal **)
   231 (** Multiplication by a non-zero cardinal **)
   232 
   232 
   233 goalw CardinalArith.thy [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B";
   233 goalw CardinalArith.thy [lepoll_def, inj_def] "!!b. b: B ==> A lepoll A*B";
   234 by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1);
   234 by (res_inst_tac [("x", "lam x:A. <x,b>")] exI 1);
   235 by (asm_simp_tac (!simpset addsimps [lam_type]) 1);
   235 by (asm_simp_tac (simpset() addsimps [lam_type]) 1);
   236 qed "prod_lepoll_self";
   236 qed "prod_lepoll_self";
   237 
   237 
   238 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
   238 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
   239 goalw CardinalArith.thy [cmult_def]
   239 goalw CardinalArith.thy [cmult_def]
   240     "!!K. [| Card(K);  Ord(L);  0<L |] ==> K le (K |*| L)";
   240     "!!K. [| Card(K);  Ord(L);  0<L |] ==> K le (K |*| L)";
   251 by (res_inst_tac [("x", "lam <w,y>:A*B. <f`w, fa`y>")] exI 1);
   251 by (res_inst_tac [("x", "lam <w,y>:A*B. <f`w, fa`y>")] exI 1);
   252 by (res_inst_tac [("d", "%<w,y>.<converse(f)`w, converse(fa)`y>")] 
   252 by (res_inst_tac [("d", "%<w,y>.<converse(f)`w, converse(fa)`y>")] 
   253                   lam_injective 1);
   253                   lam_injective 1);
   254 by (typechk_tac (inj_is_fun::ZF_typechecks));
   254 by (typechk_tac (inj_is_fun::ZF_typechecks));
   255 by (etac SigmaE 1);
   255 by (etac SigmaE 1);
   256 by (asm_simp_tac (!simpset addsimps [left_inverse]) 1);
   256 by (asm_simp_tac (simpset() addsimps [left_inverse]) 1);
   257 qed "prod_lepoll_mono";
   257 qed "prod_lepoll_mono";
   258 
   258 
   259 goalw CardinalArith.thy [cmult_def]
   259 goalw CardinalArith.thy [cmult_def]
   260     "!!K. [| K' le K;  L' le L |] ==> (K' |*| L') le (K |*| L)";
   260     "!!K. [| K' le K;  L' le L |] ==> (K' |*| L') le (K |*| L)";
   261 by (safe_tac (!claset addSDs [le_subset_iff RS iffD1]));
   261 by (safe_tac (claset() addSDs [le_subset_iff RS iffD1]));
   262 by (rtac well_ord_lepoll_imp_Card_le 1);
   262 by (rtac well_ord_lepoll_imp_Card_le 1);
   263 by (REPEAT (ares_tac [prod_lepoll_mono, subset_imp_lepoll] 2));
   263 by (REPEAT (ares_tac [prod_lepoll_mono, subset_imp_lepoll] 2));
   264 by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
   264 by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
   265 qed "cmult_le_mono";
   265 qed "cmult_le_mono";
   266 
   266 
   269 goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B";
   269 goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B";
   270 by (rtac exI 1);
   270 by (rtac exI 1);
   271 by (res_inst_tac [("c", "%<x,y>. if(x=A, Inl(y), Inr(<x,y>))"), 
   271 by (res_inst_tac [("c", "%<x,y>. if(x=A, Inl(y), Inr(<x,y>))"), 
   272                   ("d", "case(%y. <A,y>, %z. z)")] 
   272                   ("d", "case(%y. <A,y>, %z. z)")] 
   273     lam_bijective 1);
   273     lam_bijective 1);
   274 by (safe_tac (!claset addSEs [sumE]));
   274 by (safe_tac (claset() addSEs [sumE]));
   275 by (ALLGOALS
   275 by (ALLGOALS
   276     (asm_simp_tac (!simpset addsimps [succI2, if_type, mem_imp_not_eq])));
   276     (asm_simp_tac (simpset() addsimps [succI2, if_type, mem_imp_not_eq])));
   277 qed "prod_succ_eqpoll";
   277 qed "prod_succ_eqpoll";
   278 
   278 
   279 (*Unconditional version requires AC*)
   279 (*Unconditional version requires AC*)
   280 goalw CardinalArith.thy [cmult_def, cadd_def]
   280 goalw CardinalArith.thy [cmult_def, cadd_def]
   281     "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)";
   281     "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)";
   287 
   287 
   288 val [mnat,nnat] = goal CardinalArith.thy
   288 val [mnat,nnat] = goal CardinalArith.thy
   289     "[| m: nat;  n: nat |] ==> m |*| n = m#*n";
   289     "[| m: nat;  n: nat |] ==> m |*| n = m#*n";
   290 by (cut_facts_tac [nnat] 1);
   290 by (cut_facts_tac [nnat] 1);
   291 by (nat_ind_tac "m" [mnat] 1);
   291 by (nat_ind_tac "m" [mnat] 1);
   292 by (asm_simp_tac (!simpset addsimps [cmult_0]) 1);
   292 by (asm_simp_tac (simpset() addsimps [cmult_0]) 1);
   293 by (asm_simp_tac (!simpset addsimps [nat_into_Ord, cmult_succ_lemma,
   293 by (asm_simp_tac (simpset() addsimps [nat_into_Ord, cmult_succ_lemma,
   294                                      nat_cadd_eq_add]) 1);
   294                                      nat_cadd_eq_add]) 1);
   295 qed "nat_cmult_eq_mult";
   295 qed "nat_cmult_eq_mult";
   296 
   296 
   297 goal CardinalArith.thy "!!m n. Card(n) ==> 2 |*| n = n |+| n";
   297 goal CardinalArith.thy "!!m n. Card(n) ==> 2 |*| n = n |+| n";
   298 by (asm_simp_tac 
   298 by (asm_simp_tac 
   299     (!simpset addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, 
   299     (simpset() addsimps [Ord_0, Ord_succ, cmult_0, cmult_succ_lemma, 
   300 			Card_is_Ord,
   300 			Card_is_Ord,
   301 			read_instantiate [("j","0")] cadd_commute, cadd_0]) 1);
   301 			read_instantiate [("j","0")] cadd_commute, cadd_0]) 1);
   302 qed "cmult_2";
   302 qed "cmult_2";
   303 
   303 
   304 
   304 
   315     "lam z:cons(u,A). if(z=u, f`0,      \
   315     "lam z:cons(u,A). if(z=u, f`0,      \
   316 \                        if(z: range(f), f`succ(converse(f)`z), z))")] exI 1);
   316 \                        if(z: range(f), f`succ(converse(f)`z), z))")] exI 1);
   317 by (res_inst_tac [("d", "%y. if(y: range(f),    \
   317 by (res_inst_tac [("d", "%y. if(y: range(f),    \
   318 \                               nat_case(u, %z. f`z, converse(f)`y), y)")] 
   318 \                               nat_case(u, %z. f`z, converse(f)`y), y)")] 
   319     lam_injective 1);
   319     lam_injective 1);
   320 by (fast_tac (!claset addSIs [if_type, nat_succI, apply_type]
   320 by (fast_tac (claset() addSIs [if_type, nat_succI, apply_type]
   321                       addIs  [inj_is_fun, inj_converse_fun]) 1);
   321                       addIs  [inj_is_fun, inj_converse_fun]) 1);
   322 by (asm_simp_tac 
   322 by (asm_simp_tac 
   323     (!simpset addsimps [inj_is_fun RS apply_rangeI,
   323     (simpset() addsimps [inj_is_fun RS apply_rangeI,
   324                      inj_converse_fun RS apply_rangeI,
   324                      inj_converse_fun RS apply_rangeI,
   325                      inj_converse_fun RS apply_funtype,
   325                      inj_converse_fun RS apply_funtype,
   326                      left_inverse, right_inverse, nat_0I, nat_succI, 
   326                      left_inverse, right_inverse, nat_0I, nat_succI, 
   327                      nat_case_0, nat_case_succ]
   327                      nat_case_0, nat_case_succ]
   328            setloop split_tac [expand_if]) 1);
   328            setloop split_tac [expand_if]) 1);
   337 goalw CardinalArith.thy [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
   337 goalw CardinalArith.thy [succ_def] "!!i. nat <= A ==> succ(A) eqpoll A";
   338 by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1);
   338 by (eresolve_tac [subset_imp_lepoll RS nat_cons_eqpoll] 1);
   339 qed "nat_succ_eqpoll";
   339 qed "nat_succ_eqpoll";
   340 
   340 
   341 goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
   341 goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
   342 by (blast_tac (!claset addIs [Card_nat, le_refl, Card_is_Ord]) 1);
   342 by (blast_tac (claset() addIs [Card_nat, le_refl, Card_is_Ord]) 1);
   343 qed "InfCard_nat";
   343 qed "InfCard_nat";
   344 
   344 
   345 goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
   345 goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
   346 by (etac conjunct1 1);
   346 by (etac conjunct1 1);
   347 qed "InfCard_is_Card";
   347 qed "InfCard_is_Card";
   348 
   348 
   349 goalw CardinalArith.thy [InfCard_def]
   349 goalw CardinalArith.thy [InfCard_def]
   350     "!!K L. [| InfCard(K);  Card(L) |] ==> InfCard(K Un L)";
   350     "!!K L. [| InfCard(K);  Card(L) |] ==> InfCard(K Un L)";
   351 by (asm_simp_tac (!simpset addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), 
   351 by (asm_simp_tac (simpset() addsimps [Card_Un, Un_upper1_le RSN (2,le_trans), 
   352                                   Card_is_Ord]) 1);
   352                                   Card_is_Ord]) 1);
   353 qed "InfCard_Un";
   353 qed "InfCard_Un";
   354 
   354 
   355 (*Kunen's Lemma 10.11*)
   355 (*Kunen's Lemma 10.11*)
   356 goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
   356 goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
   357 by (etac conjE 1);
   357 by (etac conjE 1);
   358 by (forward_tac [Card_is_Ord] 1);
   358 by (forward_tac [Card_is_Ord] 1);
   359 by (rtac (ltI RS non_succ_LimitI) 1);
   359 by (rtac (ltI RS non_succ_LimitI) 1);
   360 by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1);
   360 by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1);
   361 by (safe_tac (!claset addSDs [Limit_nat RS Limit_le_succD]));
   361 by (safe_tac (claset() addSDs [Limit_nat RS Limit_le_succD]));
   362 by (rewtac Card_def);
   362 by (rewtac Card_def);
   363 by (dtac trans 1);
   363 by (dtac trans 1);
   364 by (etac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1);
   364 by (etac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1);
   365 by (etac (Ord_cardinal_le RS lt_trans2 RS lt_irrefl) 1);
   365 by (etac (Ord_cardinal_le RS lt_trans2 RS lt_irrefl) 1);
   366 by (REPEAT (ares_tac [le_eqI, Ord_cardinal] 1));
   366 by (REPEAT (ares_tac [le_eqI, Ord_cardinal] 1));
   371 
   371 
   372 (*A general fact about ordermap*)
   372 (*A general fact about ordermap*)
   373 goalw Cardinal.thy [eqpoll_def]
   373 goalw Cardinal.thy [eqpoll_def]
   374     "!!A. [| well_ord(A,r);  x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)";
   374     "!!A. [| well_ord(A,r);  x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)";
   375 by (rtac exI 1);
   375 by (rtac exI 1);
   376 by (asm_simp_tac (!simpset addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
   376 by (asm_simp_tac (simpset() addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
   377 by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1);
   377 by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1);
   378 by (rtac pred_subset 1);
   378 by (rtac pred_subset 1);
   379 qed "ordermap_eqpoll_pred";
   379 qed "ordermap_eqpoll_pred";
   380 
   380 
   381 (** Establishing the well-ordering **)
   381 (** Establishing the well-ordering **)
   382 
   382 
   383 goalw CardinalArith.thy [inj_def]
   383 goalw CardinalArith.thy [inj_def]
   384  "!!K. Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)";
   384  "!!K. Ord(K) ==> (lam <x,y>:K*K. <x Un y, x, y>) : inj(K*K, K*K*K)";
   385 by (fast_tac (!claset addss (!simpset)
   385 by (fast_tac (claset() addss (simpset())
   386                     addIs [lam_type, Un_least_lt RS ltD, ltI]) 1);
   386                     addIs [lam_type, Un_least_lt RS ltD, ltI]) 1);
   387 qed "csquare_lam_inj";
   387 qed "csquare_lam_inj";
   388 
   388 
   389 goalw CardinalArith.thy [csquare_rel_def]
   389 goalw CardinalArith.thy [csquare_rel_def]
   390  "!!K. Ord(K) ==> well_ord(K*K, csquare_rel(K))";
   390  "!!K. Ord(K) ==> well_ord(K*K, csquare_rel(K))";
   396 
   396 
   397 goalw CardinalArith.thy [csquare_rel_def]
   397 goalw CardinalArith.thy [csquare_rel_def]
   398  "!!K. [| x<K;  y<K;  z<K |] ==> \
   398  "!!K. [| x<K;  y<K;  z<K |] ==> \
   399 \      <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z";
   399 \      <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z";
   400 by (REPEAT (etac ltE 1));
   400 by (REPEAT (etac ltE 1));
   401 by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   401 by (asm_simp_tac (simpset() addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   402                                   Un_absorb, Un_least_mem_iff, ltD]) 1);
   402                                   Un_absorb, Un_least_mem_iff, ltD]) 1);
   403 by (safe_tac (!claset addSEs [mem_irrefl] 
   403 by (safe_tac (claset() addSEs [mem_irrefl] 
   404                     addSIs [Un_upper1_le, Un_upper2_le]));
   404                     addSIs [Un_upper1_le, Un_upper2_le]));
   405 by (ALLGOALS (asm_simp_tac (!simpset addsimps [lt_def, succI2, Ord_succ])));
   405 by (ALLGOALS (asm_simp_tac (simpset() addsimps [lt_def, succI2, Ord_succ])));
   406 qed_spec_mp "csquareD";
   406 qed_spec_mp "csquareD";
   407 
   407 
   408 goalw CardinalArith.thy [pred_def]
   408 goalw CardinalArith.thy [pred_def]
   409  "!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)";
   409  "!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)";
   410 by (safe_tac (claset_of"ZF" addSEs [SigmaE]));  (*avoids using succCI,...*)
   410 by (safe_tac (claset_of ZF.thy addSEs [SigmaE]));  (*avoids using succCI,...*)
   411 by (rtac (csquareD RS conjE) 1);
   411 by (rtac (csquareD RS conjE) 1);
   412 by (rewtac lt_def);
   412 by (rewtac lt_def);
   413 by (assume_tac 4);
   413 by (assume_tac 4);
   414 by (ALLGOALS Blast_tac);
   414 by (ALLGOALS Blast_tac);
   415 qed "pred_csquare_subset";
   415 qed "pred_csquare_subset";
   417 goalw CardinalArith.thy [csquare_rel_def]
   417 goalw CardinalArith.thy [csquare_rel_def]
   418  "!!K. [| x<z;  y<z;  z<K |] ==>  <<x,y>, <z,z>> : csquare_rel(K)";
   418  "!!K. [| x<z;  y<z;  z<K |] ==>  <<x,y>, <z,z>> : csquare_rel(K)";
   419 by (subgoals_tac ["x<K", "y<K"] 1);
   419 by (subgoals_tac ["x<K", "y<K"] 1);
   420 by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2));
   420 by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2));
   421 by (REPEAT (etac ltE 1));
   421 by (REPEAT (etac ltE 1));
   422 by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   422 by (asm_simp_tac (simpset() addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   423                                      Un_absorb, Un_least_mem_iff, ltD]) 1);
   423                                      Un_absorb, Un_least_mem_iff, ltD]) 1);
   424 qed "csquare_ltI";
   424 qed "csquare_ltI";
   425 
   425 
   426 (*Part of the traditional proof.  UNUSED since it's harder to prove & apply *)
   426 (*Part of the traditional proof.  UNUSED since it's harder to prove & apply *)
   427 goalw CardinalArith.thy [csquare_rel_def]
   427 goalw CardinalArith.thy [csquare_rel_def]
   428  "!!K. [| x le z;  y le z;  z<K |] ==> \
   428  "!!K. [| x le z;  y le z;  z<K |] ==> \
   429 \      <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z";
   429 \      <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z";
   430 by (subgoals_tac ["x<K", "y<K"] 1);
   430 by (subgoals_tac ["x<K", "y<K"] 1);
   431 by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2));
   431 by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2));
   432 by (REPEAT (etac ltE 1));
   432 by (REPEAT (etac ltE 1));
   433 by (asm_simp_tac (!simpset addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   433 by (asm_simp_tac (simpset() addsimps [rvimage_iff, rmult_iff, Memrel_iff,
   434                                   Un_absorb, Un_least_mem_iff, ltD]) 1);
   434                                   Un_absorb, Un_least_mem_iff, ltD]) 1);
   435 by (REPEAT_FIRST (etac succE));
   435 by (REPEAT_FIRST (etac succE));
   436 by (ALLGOALS
   436 by (ALLGOALS
   437     (asm_simp_tac (!simpset addsimps [subset_Un_iff RS iff_sym, 
   437     (asm_simp_tac (simpset() addsimps [subset_Un_iff RS iff_sym, 
   438                                    subset_Un_iff2 RS iff_sym, OrdmemD])));
   438                                    subset_Un_iff2 RS iff_sym, OrdmemD])));
   439 qed "csquare_or_eqI";
   439 qed "csquare_or_eqI";
   440 
   440 
   441 (** The cardinality of initial segments **)
   441 (** The cardinality of initial segments **)
   442 
   442 
   444     "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
   444     "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
   445 \         ordermap(K*K, csquare_rel(K)) ` <x,y> <               \
   445 \         ordermap(K*K, csquare_rel(K)) ` <x,y> <               \
   446 \         ordermap(K*K, csquare_rel(K)) ` <z,z>";
   446 \         ordermap(K*K, csquare_rel(K)) ` <z,z>";
   447 by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1);
   447 by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1);
   448 by (etac (Limit_is_Ord RS well_ord_csquare) 2);
   448 by (etac (Limit_is_Ord RS well_ord_csquare) 2);
   449 by (blast_tac (!claset addSIs [Un_least_lt, Limit_has_succ]) 2);
   449 by (blast_tac (claset() addSIs [Un_least_lt, Limit_has_succ]) 2);
   450 by (rtac (csquare_ltI RS ordermap_mono RS ltI) 1);
   450 by (rtac (csquare_ltI RS ordermap_mono RS ltI) 1);
   451 by (etac well_ord_is_wf 4);
   451 by (etac well_ord_is_wf 4);
   452 by (ALLGOALS 
   452 by (ALLGOALS 
   453     (blast_tac (!claset addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap] 
   453     (blast_tac (claset() addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap] 
   454                      addSEs [ltE])));
   454                      addSEs [ltE])));
   455 qed "ordermap_z_lt";
   455 qed "ordermap_z_lt";
   456 
   456 
   457 (*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
   457 (*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
   458 goalw CardinalArith.thy [cmult_def]
   458 goalw CardinalArith.thy [cmult_def]
   459   "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
   459   "!!K. [| Limit(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
   460 \       | ordermap(K*K, csquare_rel(K)) ` <x,y> | le  |succ(z)| |*| |succ(z)|";
   460 \       | ordermap(K*K, csquare_rel(K)) ` <x,y> | le  |succ(z)| |*| |succ(z)|";
   461 by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1);
   461 by (rtac (well_ord_rmult RS well_ord_lepoll_imp_Card_le) 1);
   462 by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1));
   462 by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1));
   463 by (subgoals_tac ["z<K"] 1);
   463 by (subgoals_tac ["z<K"] 1);
   464 by (blast_tac (!claset addSIs [Un_least_lt, Limit_has_succ]) 2);
   464 by (blast_tac (claset() addSIs [Un_least_lt, Limit_has_succ]) 2);
   465 by (rtac (ordermap_z_lt RS leI RS le_imp_lepoll RS lepoll_trans) 1);
   465 by (rtac (ordermap_z_lt RS leI RS le_imp_lepoll RS lepoll_trans) 1);
   466 by (REPEAT_SOME assume_tac);
   466 by (REPEAT_SOME assume_tac);
   467 by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1);
   467 by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1);
   468 by (etac (Limit_is_Ord RS well_ord_csquare) 1);
   468 by (etac (Limit_is_Ord RS well_ord_csquare) 1);
   469 by (blast_tac (!claset addIs [ltD]) 1);
   469 by (blast_tac (claset() addIs [ltD]) 1);
   470 by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN
   470 by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN
   471     assume_tac 1);
   471     assume_tac 1);
   472 by (REPEAT_FIRST (etac ltE));
   472 by (REPEAT_FIRST (etac ltE));
   473 by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
   473 by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
   474 by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll)));
   474 by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll)));
   483 by (assume_tac 1);
   483 by (assume_tac 1);
   484 by (etac (well_ord_csquare RS Ord_ordertype) 1);
   484 by (etac (well_ord_csquare RS Ord_ordertype) 1);
   485 by (rtac Card_lt_imp_lt 1);
   485 by (rtac Card_lt_imp_lt 1);
   486 by (etac InfCard_is_Card 3);
   486 by (etac InfCard_is_Card 3);
   487 by (etac ltE 2 THEN assume_tac 2);
   487 by (etac ltE 2 THEN assume_tac 2);
   488 by (asm_full_simp_tac (!simpset addsimps [ordertype_unfold]) 1);
   488 by (asm_full_simp_tac (simpset() addsimps [ordertype_unfold]) 1);
   489 by (safe_tac (!claset addSEs [ltE]));
   489 by (safe_tac (claset() addSEs [ltE]));
   490 by (subgoals_tac ["Ord(xb)", "Ord(y)"] 1);
   490 by (subgoals_tac ["Ord(xb)", "Ord(y)"] 1);
   491 by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2));
   491 by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2));
   492 by (rtac (InfCard_is_Limit RS ordermap_csquare_le RS lt_trans1) 1  THEN
   492 by (rtac (InfCard_is_Limit RS ordermap_csquare_le RS lt_trans1) 1  THEN
   493     REPEAT (ares_tac [refl] 1 ORELSE etac ltI 1));
   493     REPEAT (ares_tac [refl] 1 ORELSE etac ltI 1));
   494 by (res_inst_tac [("i","xb Un y"), ("j","nat")] Ord_linear2 1  THEN
   494 by (res_inst_tac [("i","xb Un y"), ("j","nat")] Ord_linear2 1  THEN
   495     REPEAT (ares_tac [Ord_Un, Ord_nat] 1));
   495     REPEAT (ares_tac [Ord_Un, Ord_nat] 1));
   496 (*the finite case: xb Un y < nat *)
   496 (*the finite case: xb Un y < nat *)
   497 by (res_inst_tac [("j", "nat")] lt_trans2 1);
   497 by (res_inst_tac [("j", "nat")] lt_trans2 1);
   498 by (asm_full_simp_tac (!simpset addsimps [InfCard_def]) 2);
   498 by (asm_full_simp_tac (simpset() addsimps [InfCard_def]) 2);
   499 by (asm_full_simp_tac
   499 by (asm_full_simp_tac
   500     (!simpset addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type,
   500     (simpset() addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type,
   501                      nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1);
   501                      nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1);
   502 (*case nat le (xb Un y) *)
   502 (*case nat le (xb Un y) *)
   503 by (asm_full_simp_tac
   503 by (asm_full_simp_tac
   504     (!simpset addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong,
   504     (simpset() addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong,
   505                      le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt, 
   505                      le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt, 
   506                      Ord_Un, ltI, nat_le_cardinal,
   506                      Ord_Un, ltI, nat_le_cardinal,
   507                      Ord_cardinal_le RS lt_trans1 RS ltD]) 1);
   507                      Ord_cardinal_le RS lt_trans1 RS ltD]) 1);
   508 qed "ordertype_csquare_le";
   508 qed "ordertype_csquare_le";
   509 
   509 
   517 by (etac (InfCard_is_Card RS cmult_square_le) 2);
   517 by (etac (InfCard_is_Card RS cmult_square_le) 2);
   518 by (rtac (ordertype_csquare_le RSN (2, le_trans)) 1);
   518 by (rtac (ordertype_csquare_le RSN (2, le_trans)) 1);
   519 by (assume_tac 2);
   519 by (assume_tac 2);
   520 by (assume_tac 2);
   520 by (assume_tac 2);
   521 by (asm_simp_tac 
   521 by (asm_simp_tac 
   522     (!simpset addsimps [cmult_def, Ord_cardinal_le,
   522     (simpset() addsimps [cmult_def, Ord_cardinal_le,
   523                      well_ord_csquare RS ordermap_bij RS 
   523                      well_ord_csquare RS ordermap_bij RS 
   524                           bij_imp_eqpoll RS cardinal_cong,
   524                           bij_imp_eqpoll RS cardinal_cong,
   525                      well_ord_csquare RS Ord_ordertype]) 1);
   525                      well_ord_csquare RS Ord_ordertype]) 1);
   526 qed "InfCard_csquare_eq";
   526 qed "InfCard_csquare_eq";
   527 
   527 
   530     "!!A. [| well_ord(A,r);  InfCard(|A|) |] ==> A*A eqpoll A";
   530     "!!A. [| well_ord(A,r);  InfCard(|A|) |] ==> A*A eqpoll A";
   531 by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1);
   531 by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1);
   532 by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1));
   532 by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1));
   533 by (rtac well_ord_cardinal_eqE 1);
   533 by (rtac well_ord_cardinal_eqE 1);
   534 by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1));
   534 by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1));
   535 by (asm_simp_tac (!simpset addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
   535 by (asm_simp_tac (simpset() addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
   536 qed "well_ord_InfCard_square_eq";
   536 qed "well_ord_InfCard_square_eq";
   537 
   537 
   538 (** Toward's Kunen's Corollary 10.13 (1) **)
   538 (** Toward's Kunen's Corollary 10.13 (1) **)
   539 
   539 
   540 goal CardinalArith.thy "!!K. [| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
   540 goal CardinalArith.thy "!!K. [| InfCard(K);  L le K;  0<L |] ==> K |*| L = K";
   541 by (rtac le_anti_sym 1);
   541 by (rtac le_anti_sym 1);
   542 by (etac ltE 2 THEN
   542 by (etac ltE 2 THEN
   543     REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2));
   543     REPEAT (ares_tac [cmult_le_self, InfCard_is_Card] 2));
   544 by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
   544 by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
   545 by (resolve_tac [cmult_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
   545 by (resolve_tac [cmult_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
   546 by (asm_simp_tac (!simpset addsimps [InfCard_csquare_eq]) 1);
   546 by (asm_simp_tac (simpset() addsimps [InfCard_csquare_eq]) 1);
   547 qed "InfCard_le_cmult_eq";
   547 qed "InfCard_le_cmult_eq";
   548 
   548 
   549 (*Corollary 10.13 (1), for cardinal multiplication*)
   549 (*Corollary 10.13 (1), for cardinal multiplication*)
   550 goal CardinalArith.thy
   550 goal CardinalArith.thy
   551     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |*| L = K Un L";
   551     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |*| L = K Un L";
   553 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
   553 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
   554 by (resolve_tac [cmult_commute RS ssubst] 1);
   554 by (resolve_tac [cmult_commute RS ssubst] 1);
   555 by (resolve_tac [Un_commute RS ssubst] 1);
   555 by (resolve_tac [Un_commute RS ssubst] 1);
   556 by (ALLGOALS
   556 by (ALLGOALS
   557     (asm_simp_tac 
   557     (asm_simp_tac 
   558      (!simpset addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq,
   558      (simpset() addsimps [InfCard_is_Limit RS Limit_has_0, InfCard_le_cmult_eq,
   559                       subset_Un_iff2 RS iffD1, le_imp_subset])));
   559                       subset_Un_iff2 RS iffD1, le_imp_subset])));
   560 qed "InfCard_cmult_eq";
   560 qed "InfCard_cmult_eq";
   561 
   561 
   562 (*This proof appear to be the simplest!*)
   562 (*This proof appear to be the simplest!*)
   563 goal CardinalArith.thy "!!K. InfCard(K) ==> K |+| K = K";
   563 goal CardinalArith.thy "!!K. InfCard(K) ==> K |+| K = K";
   564 by (asm_simp_tac
   564 by (asm_simp_tac
   565     (!simpset addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
   565     (simpset() addsimps [cmult_2 RS sym, InfCard_is_Card, cmult_commute]) 1);
   566 by (rtac InfCard_le_cmult_eq 1);
   566 by (rtac InfCard_le_cmult_eq 1);
   567 by (typechk_tac [Ord_0, le_refl, leI]);
   567 by (typechk_tac [Ord_0, le_refl, leI]);
   568 by (typechk_tac [InfCard_is_Limit, Limit_has_0, Limit_has_succ]);
   568 by (typechk_tac [InfCard_is_Limit, Limit_has_0, Limit_has_succ]);
   569 qed "InfCard_cdouble_eq";
   569 qed "InfCard_cdouble_eq";
   570 
   570 
   573 by (rtac le_anti_sym 1);
   573 by (rtac le_anti_sym 1);
   574 by (etac ltE 2 THEN
   574 by (etac ltE 2 THEN
   575     REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2));
   575     REPEAT (ares_tac [cadd_le_self, InfCard_is_Card] 2));
   576 by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
   576 by (forward_tac [InfCard_is_Card RS Card_is_Ord RS le_refl] 1);
   577 by (resolve_tac [cadd_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
   577 by (resolve_tac [cadd_le_mono RS le_trans] 1 THEN REPEAT (assume_tac 1));
   578 by (asm_simp_tac (!simpset addsimps [InfCard_cdouble_eq]) 1);
   578 by (asm_simp_tac (simpset() addsimps [InfCard_cdouble_eq]) 1);
   579 qed "InfCard_le_cadd_eq";
   579 qed "InfCard_le_cadd_eq";
   580 
   580 
   581 goal CardinalArith.thy
   581 goal CardinalArith.thy
   582     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |+| L = K Un L";
   582     "!!K. [| InfCard(K);  InfCard(L) |] ==> K |+| L = K Un L";
   583 by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1);
   583 by (res_inst_tac [("i","K"),("j","L")] Ord_linear_le 1);
   584 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
   584 by (typechk_tac [InfCard_is_Card, Card_is_Ord]);
   585 by (resolve_tac [cadd_commute RS ssubst] 1);
   585 by (resolve_tac [cadd_commute RS ssubst] 1);
   586 by (resolve_tac [Un_commute RS ssubst] 1);
   586 by (resolve_tac [Un_commute RS ssubst] 1);
   587 by (ALLGOALS
   587 by (ALLGOALS
   588     (asm_simp_tac 
   588     (asm_simp_tac 
   589      (!simpset addsimps [InfCard_le_cadd_eq,
   589      (simpset() addsimps [InfCard_le_cadd_eq,
   590                       subset_Un_iff2 RS iffD1, le_imp_subset])));
   590                       subset_Un_iff2 RS iffD1, le_imp_subset])));
   591 qed "InfCard_cadd_eq";
   591 qed "InfCard_cadd_eq";
   592 
   592 
   593 (*The other part, Corollary 10.13 (2), refers to the cardinality of the set
   593 (*The other part, Corollary 10.13 (2), refers to the cardinality of the set
   594   of all n-tuples of elements of K.  A better version for the Isabelle theory
   594   of all n-tuples of elements of K.  A better version for the Isabelle theory
   598 (*** For every cardinal number there exists a greater one
   598 (*** For every cardinal number there exists a greater one
   599      [Kunen's Theorem 10.16, which would be trivial using AC] ***)
   599      [Kunen's Theorem 10.16, which would be trivial using AC] ***)
   600 
   600 
   601 goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))";
   601 goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))";
   602 by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
   602 by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
   603 by (blast_tac (!claset addSIs [Ord_ordertype]) 2);
   603 by (blast_tac (claset() addSIs [Ord_ordertype]) 2);
   604 by (rewtac Transset_def);
   604 by (rewtac Transset_def);
   605 by (safe_tac subset_cs);
   605 by (safe_tac subset_cs);
   606 by (asm_full_simp_tac (!simpset addsimps [ordertype_pred_unfold]) 1);
   606 by (asm_full_simp_tac (simpset() addsimps [ordertype_pred_unfold]) 1);
   607 by (safe_tac (!claset));
   607 by (safe_tac (claset()));
   608 by (rtac UN_I 1);
   608 by (rtac UN_I 1);
   609 by (rtac ReplaceI 2);
   609 by (rtac ReplaceI 2);
   610 by (ALLGOALS (blast_tac (!claset addIs [well_ord_subset] addSEs [predE])));
   610 by (ALLGOALS (blast_tac (claset() addIs [well_ord_subset] addSEs [predE])));
   611 qed "Ord_jump_cardinal";
   611 qed "Ord_jump_cardinal";
   612 
   612 
   613 (*Allows selective unfolding.  Less work than deriving intro/elim rules*)
   613 (*Allows selective unfolding.  Less work than deriving intro/elim rules*)
   614 goalw CardinalArith.thy [jump_cardinal_def]
   614 goalw CardinalArith.thy [jump_cardinal_def]
   615      "i : jump_cardinal(K) <->   \
   615      "i : jump_cardinal(K) <->   \
   621 goal CardinalArith.thy "!!K. Ord(K) ==> K < jump_cardinal(K)";
   621 goal CardinalArith.thy "!!K. Ord(K) ==> K < jump_cardinal(K)";
   622 by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1);
   622 by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1);
   623 by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
   623 by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
   624 by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
   624 by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
   625 by (rtac subset_refl 2);
   625 by (rtac subset_refl 2);
   626 by (asm_simp_tac (!simpset addsimps [Memrel_def, subset_iff]) 1);
   626 by (asm_simp_tac (simpset() addsimps [Memrel_def, subset_iff]) 1);
   627 by (asm_simp_tac (!simpset addsimps [ordertype_Memrel]) 1);
   627 by (asm_simp_tac (simpset() addsimps [ordertype_Memrel]) 1);
   628 qed "K_lt_jump_cardinal";
   628 qed "K_lt_jump_cardinal";
   629 
   629 
   630 (*The proof by contradiction: the bijection f yields a wellordering of X
   630 (*The proof by contradiction: the bijection f yields a wellordering of X
   631   whose ordertype is jump_cardinal(K).  *)
   631   whose ordertype is jump_cardinal(K).  *)
   632 goal CardinalArith.thy
   632 goal CardinalArith.thy
   640 by (rtac ([rvimage_type, Sigma_mono] MRS subset_trans) 1);
   640 by (rtac ([rvimage_type, Sigma_mono] MRS subset_trans) 1);
   641 by (REPEAT (assume_tac 1));
   641 by (REPEAT (assume_tac 1));
   642 by (etac (bij_is_inj RS well_ord_rvimage) 1);
   642 by (etac (bij_is_inj RS well_ord_rvimage) 1);
   643 by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1);
   643 by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1);
   644 by (asm_simp_tac
   644 by (asm_simp_tac
   645     (!simpset addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), 
   645     (simpset() addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), 
   646                      ordertype_Memrel, Ord_jump_cardinal]) 1);
   646                      ordertype_Memrel, Ord_jump_cardinal]) 1);
   647 qed "Card_jump_cardinal_lemma";
   647 qed "Card_jump_cardinal_lemma";
   648 
   648 
   649 (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*)
   649 (*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*)
   650 goal CardinalArith.thy "Card(jump_cardinal(K))";
   650 goal CardinalArith.thy "Card(jump_cardinal(K))";
   651 by (rtac (Ord_jump_cardinal RS CardI) 1);
   651 by (rtac (Ord_jump_cardinal RS CardI) 1);
   652 by (rewtac eqpoll_def);
   652 by (rewtac eqpoll_def);
   653 by (safe_tac (!claset addSDs [ltD, jump_cardinal_iff RS iffD1]));
   653 by (safe_tac (claset() addSDs [ltD, jump_cardinal_iff RS iffD1]));
   654 by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1));
   654 by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1));
   655 qed "Card_jump_cardinal";
   655 qed "Card_jump_cardinal";
   656 
   656 
   657 (*** Basic properties of successor cardinals ***)
   657 (*** Basic properties of successor cardinals ***)
   658 
   658 
   693 qed "lt_csucc_iff";
   693 qed "lt_csucc_iff";
   694 
   694 
   695 goal CardinalArith.thy
   695 goal CardinalArith.thy
   696     "!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
   696     "!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
   697 by (asm_simp_tac 
   697 by (asm_simp_tac 
   698     (!simpset addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
   698     (simpset() addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
   699 qed "Card_lt_csucc_iff";
   699 qed "Card_lt_csucc_iff";
   700 
   700 
   701 goalw CardinalArith.thy [InfCard_def]
   701 goalw CardinalArith.thy [InfCard_def]
   702     "!!K. InfCard(K) ==> InfCard(csucc(K))";
   702     "!!K. InfCard(K) ==> InfCard(csucc(K))";
   703 by (asm_simp_tac (!simpset addsimps [Card_csucc, Card_is_Ord, 
   703 by (asm_simp_tac (simpset() addsimps [Card_csucc, Card_is_Ord, 
   704                                   lt_csucc RS leI RSN (2,le_trans)]) 1);
   704                                   lt_csucc RS leI RSN (2,le_trans)]) 1);
   705 qed "InfCard_csucc";
   705 qed "InfCard_csucc";
   706 
   706 
   707 
   707 
   708 (*** Finite sets ***)
   708 (*** Finite sets ***)
   709 
   709 
   710 goal CardinalArith.thy
   710 goal CardinalArith.thy
   711     "!!n. n: nat ==> ALL A. A eqpoll n --> A : Fin(A)";
   711     "!!n. n: nat ==> ALL A. A eqpoll n --> A : Fin(A)";
   712 by (etac nat_induct 1);
   712 by (etac nat_induct 1);
   713 by (simp_tac (!simpset addsimps (eqpoll_0_iff::Fin.intrs)) 1);
   713 by (simp_tac (simpset() addsimps (eqpoll_0_iff::Fin.intrs)) 1);
   714 by (Clarify_tac 1);
   714 by (Clarify_tac 1);
   715 by (subgoal_tac "EX u. u:A" 1);
   715 by (subgoal_tac "EX u. u:A" 1);
   716 by (etac exE 1);
   716 by (etac exE 1);
   717 by (resolve_tac [Diff_sing_eqpoll RS revcut_rl] 1);
   717 by (resolve_tac [Diff_sing_eqpoll RS revcut_rl] 1);
   718 by (assume_tac 2);
   718 by (assume_tac 2);
   719 by (assume_tac 1);
   719 by (assume_tac 1);
   720 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
   720 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
   721 by (assume_tac 1);
   721 by (assume_tac 1);
   722 by (resolve_tac [Fin.consI] 1);
   722 by (resolve_tac [Fin.consI] 1);
   723 by (Blast_tac 1);
   723 by (Blast_tac 1);
   724 by (blast_tac (!claset addIs [subset_consI  RS Fin_mono RS subsetD]) 1); 
   724 by (blast_tac (claset() addIs [subset_consI  RS Fin_mono RS subsetD]) 1); 
   725 (*Now for the lemma assumed above*)
   725 (*Now for the lemma assumed above*)
   726 by (rewtac eqpoll_def);
   726 by (rewtac eqpoll_def);
   727 by (blast_tac (!claset addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
   727 by (blast_tac (claset() addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
   728 val lemma = result();
   728 val lemma = result();
   729 
   729 
   730 goalw CardinalArith.thy [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
   730 goalw CardinalArith.thy [Finite_def] "!!A. Finite(A) ==> A : Fin(A)";
   731 by (blast_tac (!claset addIs [lemma RS spec RS mp]) 1);
   731 by (blast_tac (claset() addIs [lemma RS spec RS mp]) 1);
   732 qed "Finite_into_Fin";
   732 qed "Finite_into_Fin";
   733 
   733 
   734 goal CardinalArith.thy "!!A. A : Fin(U) ==> Finite(A)";
   734 goal CardinalArith.thy "!!A. A : Fin(U) ==> Finite(A)";
   735 by (fast_tac (!claset addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
   735 by (fast_tac (claset() addSIs [Finite_0, Finite_cons] addEs [Fin.induct]) 1);
   736 qed "Fin_into_Finite";
   736 qed "Fin_into_Finite";
   737 
   737 
   738 goal CardinalArith.thy "Finite(A) <-> A : Fin(A)";
   738 goal CardinalArith.thy "Finite(A) <-> A : Fin(A)";
   739 by (blast_tac (!claset addIs [Finite_into_Fin, Fin_into_Finite]) 1);
   739 by (blast_tac (claset() addIs [Finite_into_Fin, Fin_into_Finite]) 1);
   740 qed "Finite_Fin_iff";
   740 qed "Finite_Fin_iff";
   741 
   741 
   742 goal CardinalArith.thy
   742 goal CardinalArith.thy
   743     "!!A. [| Finite(A); Finite(B) |] ==> Finite(A Un B)";
   743     "!!A. [| Finite(A); Finite(B) |] ==> Finite(A Un B)";
   744 by (blast_tac (!claset addSIs [Fin_into_Finite, Fin_UnI] 
   744 by (blast_tac (claset() addSIs [Fin_into_Finite, Fin_UnI] 
   745                        addSDs [Finite_into_Fin]
   745                        addSDs [Finite_into_Fin]
   746                        addIs [Un_upper1 RS Fin_mono RS subsetD,
   746                        addIs [Un_upper1 RS Fin_mono RS subsetD,
   747 			      Un_upper2 RS Fin_mono RS subsetD]) 1);
   747 			      Un_upper2 RS Fin_mono RS subsetD]) 1);
   748 qed "Finite_Un";
   748 qed "Finite_Un";
   749 
   749 
   751 (** Removing elements from a finite set decreases its cardinality **)
   751 (** Removing elements from a finite set decreases its cardinality **)
   752 
   752 
   753 goal CardinalArith.thy
   753 goal CardinalArith.thy
   754     "!!A. A: Fin(U) ==> x~:A --> ~ cons(x,A) lepoll A";
   754     "!!A. A: Fin(U) ==> x~:A --> ~ cons(x,A) lepoll A";
   755 by (etac Fin_induct 1);
   755 by (etac Fin_induct 1);
   756 by (simp_tac (!simpset addsimps [lepoll_0_iff]) 1);
   756 by (simp_tac (simpset() addsimps [lepoll_0_iff]) 1);
   757 by (subgoal_tac "cons(x,cons(xa,y)) = cons(xa,cons(x,y))" 1);
   757 by (subgoal_tac "cons(x,cons(xa,y)) = cons(xa,cons(x,y))" 1);
   758 by (Asm_simp_tac 1);
   758 by (Asm_simp_tac 1);
   759 by (blast_tac (!claset addSDs [cons_lepoll_consD]) 1);
   759 by (blast_tac (claset() addSDs [cons_lepoll_consD]) 1);
   760 by (Blast_tac 1);
   760 by (Blast_tac 1);
   761 qed "Fin_imp_not_cons_lepoll";
   761 qed "Fin_imp_not_cons_lepoll";
   762 
   762 
   763 goal CardinalArith.thy
   763 goal CardinalArith.thy
   764     "!!a A. [| Finite(A);  a~:A |] ==> |cons(a,A)| = succ(|A|)";
   764     "!!a A. [| Finite(A);  a~:A |] ==> |cons(a,A)| = succ(|A|)";
   765 by (rewtac cardinal_def);
   765 by (rewtac cardinal_def);
   766 by (rtac Least_equality 1);
   766 by (rtac Least_equality 1);
   767 by (fold_tac [cardinal_def]);
   767 by (fold_tac [cardinal_def]);
   768 by (simp_tac (!simpset addsimps [succ_def]) 1);
   768 by (simp_tac (simpset() addsimps [succ_def]) 1);
   769 by (blast_tac (!claset addIs [cons_eqpoll_cong, well_ord_cardinal_eqpoll] 
   769 by (blast_tac (claset() addIs [cons_eqpoll_cong, well_ord_cardinal_eqpoll] 
   770                     addSEs [mem_irrefl]
   770                     addSEs [mem_irrefl]
   771                     addSDs [Finite_imp_well_ord]) 1);
   771                     addSDs [Finite_imp_well_ord]) 1);
   772 by (blast_tac (!claset addIs [Ord_succ, Card_cardinal, Card_is_Ord]) 1);
   772 by (blast_tac (claset() addIs [Ord_succ, Card_cardinal, Card_is_Ord]) 1);
   773 by (rtac notI 1);
   773 by (rtac notI 1);
   774 by (resolve_tac [Finite_into_Fin RS Fin_imp_not_cons_lepoll RS mp RS notE] 1);
   774 by (resolve_tac [Finite_into_Fin RS Fin_imp_not_cons_lepoll RS mp RS notE] 1);
   775 by (assume_tac 1);
   775 by (assume_tac 1);
   776 by (assume_tac 1);
   776 by (assume_tac 1);
   777 by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] 1);
   777 by (eresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll RS lepoll_trans] 1);
   778 by (eresolve_tac [le_imp_lepoll RS lepoll_trans] 1);
   778 by (eresolve_tac [le_imp_lepoll RS lepoll_trans] 1);
   779 by (blast_tac (!claset addIs [well_ord_cardinal_eqpoll RS eqpoll_imp_lepoll] 
   779 by (blast_tac (claset() addIs [well_ord_cardinal_eqpoll RS eqpoll_imp_lepoll] 
   780                     addSDs [Finite_imp_well_ord]) 1);
   780                     addSDs [Finite_imp_well_ord]) 1);
   781 qed "Finite_imp_cardinal_cons";
   781 qed "Finite_imp_cardinal_cons";
   782 
   782 
   783 
   783 
   784 goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
   784 goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> succ(|A-{a}|) = |A|";
   785 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
   785 by (res_inst_tac [("b", "A")] (cons_Diff RS subst) 1);
   786 by (assume_tac 1);
   786 by (assume_tac 1);
   787 by (asm_simp_tac (!simpset addsimps [Finite_imp_cardinal_cons,
   787 by (asm_simp_tac (simpset() addsimps [Finite_imp_cardinal_cons,
   788                                   Diff_subset RS subset_Finite]) 1);
   788                                   Diff_subset RS subset_Finite]) 1);
   789 by (asm_simp_tac (!simpset addsimps [cons_Diff]) 1);
   789 by (asm_simp_tac (simpset() addsimps [cons_Diff]) 1);
   790 qed "Finite_imp_succ_cardinal_Diff";
   790 qed "Finite_imp_succ_cardinal_Diff";
   791 
   791 
   792 goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> |A-{a}| < |A|";
   792 goal CardinalArith.thy "!!a A. [| Finite(A);  a:A |] ==> |A-{a}| < |A|";
   793 by (rtac succ_leE 1);
   793 by (rtac succ_leE 1);
   794 by (asm_simp_tac (!simpset addsimps [Finite_imp_succ_cardinal_Diff, 
   794 by (asm_simp_tac (simpset() addsimps [Finite_imp_succ_cardinal_Diff, 
   795                                   Ord_cardinal RS le_refl]) 1);
   795                                   Ord_cardinal RS le_refl]) 1);
   796 qed "Finite_imp_cardinal_Diff";
   796 qed "Finite_imp_cardinal_Diff";
   797 
   797 
   798 
   798 
   799 (** Thanks to Krzysztof Grabczewski **)
   799 (** Thanks to Krzysztof Grabczewski **)
   806 by (eresolve_tac [nat_implies_well_ord RS (
   806 by (eresolve_tac [nat_implies_well_ord RS (
   807                   nat_implies_well_ord RSN (2,
   807                   nat_implies_well_ord RSN (2,
   808                   well_ord_radd RS well_ord_cardinal_eqpoll)) RS eqpoll_sym] 1 
   808                   well_ord_radd RS well_ord_cardinal_eqpoll)) RS eqpoll_sym] 1 
   809     THEN (assume_tac 1));
   809     THEN (assume_tac 1));
   810 by (eresolve_tac [nat_cadd_eq_add RS subst] 1 THEN (assume_tac 1));
   810 by (eresolve_tac [nat_cadd_eq_add RS subst] 1 THEN (assume_tac 1));
   811 by (asm_full_simp_tac (!simpset addsimps [cadd_def, eqpoll_refl]) 1);
   811 by (asm_full_simp_tac (simpset() addsimps [cadd_def, eqpoll_refl]) 1);
   812 qed "nat_sum_eqpoll_sum";
   812 qed "nat_sum_eqpoll_sum";
   813 
   813 
   814 goal Nat.thy "!!m. [| m le n; n:nat |] ==> m:nat";
   814 goal Nat.thy "!!m. [| m le n; n:nat |] ==> m:nat";
   815 by (blast_tac (!claset addSDs [nat_succI RS (Ord_nat RSN (2, OrdmemD))]
   815 by (blast_tac (claset() addSDs [nat_succI RS (Ord_nat RSN (2, OrdmemD))]
   816         addSEs [ltE]) 1);
   816         addSEs [ltE]) 1);
   817 qed "le_in_nat";
   817 qed "le_in_nat";
   818 
   818