src/ZF/CardinalArith.ML
author lcp
Fri, 12 Aug 1994 12:51:34 +0200
changeset 516 1957113f0d7d
parent 488 52f7447d4f1b
child 517 a9f93400f307
permissions -rw-r--r--
installation of new inductive/datatype sections
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/CardinalArith.ML
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     2
    ID:         $Id$
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     5
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     6
Cardinal arithmetic -- WITHOUT the Axiom of Choice
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     7
*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     8
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
     9
open CardinalArith;
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    10
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    11
(*** Elementary properties ***)
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
    12
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    13
(*Use AC to discharge first premise*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    14
goal CardinalArith.thy
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    15
    "!!A B. [| well_ord(B,r);  A lepoll B |] ==> |A| le |B|";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    16
by (res_inst_tac [("i","|A|"),("j","|B|")] Ord_linear_le 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    17
by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    18
by (rtac (eqpollI RS cardinal_cong) 1 THEN assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    19
by (rtac lepoll_trans 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    20
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    21
by (assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    22
by (etac (le_imp_subset RS subset_imp_lepoll RS lepoll_trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    23
by (rtac eqpoll_imp_lepoll 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    24
by (rewtac lepoll_def);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    25
by (etac exE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    26
by (rtac well_ord_cardinal_eqpoll 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    27
by (etac well_ord_rvimage 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    28
by (assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    29
val well_ord_lepoll_imp_le = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    30
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    31
val case_ss = ZF_ss addsimps [Inl_iff, Inl_Inr_iff, Inr_iff, Inr_Inl_iff,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    32
			      case_Inl, case_Inr, InlI, InrI];
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    33
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    34
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    35
(** Congruence laws for successor, cardinal addition and multiplication **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    36
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    37
val bij_inverse_ss =
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    38
    case_ss addsimps [bij_is_fun RS apply_type,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    39
		      bij_converse_bij RS bij_is_fun RS apply_type,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    40
		      left_inverse_bij, right_inverse_bij];
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    41
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    42
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    43
(*Congruence law for  cons  under equipollence*)
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    44
goalw CardinalArith.thy [eqpoll_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    45
    "!!A B. [| A eqpoll B;  a ~: A;  b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    46
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    47
by (rtac exI 1);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    48
by (res_inst_tac [("c", "%z.if(z=a,b,f`z)"), 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    49
                  ("d", "%z.if(z=b,a,converse(f)`z)")] lam_bijective 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    50
by (ALLGOALS
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    51
    (asm_simp_tac (bij_inverse_ss addsimps [consI2]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    52
 		                  setloop (etac consE ORELSE' 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    53
				           split_tac [expand_if]))));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    54
by (fast_tac (ZF_cs addIs [bij_is_fun RS apply_type]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    55
by (fast_tac (ZF_cs addIs [bij_converse_bij RS bij_is_fun RS apply_type]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    56
val cons_eqpoll_cong = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    57
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    58
(*Congruence law for  succ  under equipollence*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    59
goalw CardinalArith.thy [succ_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    60
    "!!A B. A eqpoll B ==> succ(A) eqpoll succ(B)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    61
by (REPEAT (ares_tac [cons_eqpoll_cong, mem_not_refl] 1));
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    62
val succ_eqpoll_cong = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    63
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    64
(*Each element of Fin(A) is equivalent to a natural number*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    65
goal CardinalArith.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    66
    "!!X A. X: Fin(A) ==> EX n:nat. X eqpoll n";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    67
by (eresolve_tac [Fin_induct] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    68
by (fast_tac (ZF_cs addIs [eqpoll_refl, nat_0I]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    69
by (fast_tac (ZF_cs addSIs [cons_eqpoll_cong, 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    70
			    rewrite_rule [succ_def] nat_succI] 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    71
                            addSEs [mem_irrefl]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    72
val Fin_eqpoll = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    73
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    74
(*Congruence law for + under equipollence*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    75
goalw CardinalArith.thy [eqpoll_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    76
    "!!A B C D. [| A eqpoll C;  B eqpoll D |] ==> A+B eqpoll C+D";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    77
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    78
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    79
by (res_inst_tac [("c", "case(%x. Inl(f`x), %y. Inr(fa`y))"),
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    80
	 ("d", "case(%x. Inl(converse(f)`x), %y. Inr(converse(fa)`y))")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    81
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    82
by (safe_tac (ZF_cs addSEs [sumE]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    83
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    84
val sum_eqpoll_cong = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    85
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    86
(*Congruence law for * under equipollence*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    87
goalw CardinalArith.thy [eqpoll_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    88
    "!!A B C D. [| A eqpoll C;  B eqpoll D |] ==> A*B eqpoll C*D";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    89
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    90
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    91
by (res_inst_tac [("c", "split(%x y. <f`x, fa`y>)"),
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    92
		  ("d", "split(%x y. <converse(f)`x, converse(fa)`y>)")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    93
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    94
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    95
by (ALLGOALS (asm_simp_tac bij_inverse_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    96
val prod_eqpoll_cong = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    97
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    98
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
    99
(*** Cardinal addition ***)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   100
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   101
(** Cardinal addition is commutative **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   102
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   103
(*Easier to prove the two directions separately*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   104
goalw CardinalArith.thy [eqpoll_def] "A+B eqpoll B+A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   105
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   106
by (res_inst_tac [("c", "case(Inr, Inl)"), ("d", "case(Inr, Inl)")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   107
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   108
by (safe_tac (ZF_cs addSEs [sumE]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   109
by (ALLGOALS (asm_simp_tac case_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   110
val sum_commute_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   111
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   112
goalw CardinalArith.thy [cadd_def] "i |+| j = j |+| i";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   113
by (rtac (sum_commute_eqpoll RS cardinal_cong) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   114
val cadd_commute = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   115
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   116
(** Cardinal addition is associative **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   117
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   118
goalw CardinalArith.thy [eqpoll_def] "(A+B)+C eqpoll A+(B+C)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   119
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   120
by (res_inst_tac [("c", "case(case(Inl, %y.Inr(Inl(y))), %y. Inr(Inr(y)))"),
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   121
		  ("d", "case(%x.Inl(Inl(x)), case(%x.Inl(Inr(x)), Inr))")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   122
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   123
by (ALLGOALS (asm_simp_tac (case_ss setloop etac sumE)));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   124
val sum_assoc_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   125
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   126
(*Unconditional version requires AC*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   127
goalw CardinalArith.thy [cadd_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   128
    "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==>	\
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   129
\             (i |+| j) |+| k = i |+| (j |+| k)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   130
by (rtac cardinal_cong 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   131
br ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS sum_eqpoll_cong RS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   132
    eqpoll_trans) 1;
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   133
by (rtac (sum_assoc_eqpoll RS eqpoll_trans) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   134
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong RS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   135
    eqpoll_sym) 2;
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   136
by (REPEAT (ares_tac [well_ord_radd] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   137
val well_ord_cadd_assoc = result();
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   138
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   139
(** 0 is the identity for addition **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   140
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   141
goalw CardinalArith.thy [eqpoll_def] "0+A eqpoll A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   142
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   143
by (res_inst_tac [("c", "case(%x.x, %y.y)"), ("d", "Inr")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   144
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   145
by (ALLGOALS (asm_simp_tac (case_ss setloop eresolve_tac [sumE,emptyE])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   146
val sum_0_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   147
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   148
goalw CardinalArith.thy [cadd_def] "!!K. Card(K) ==> 0 |+| K = K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   149
by (asm_simp_tac (ZF_ss addsimps [sum_0_eqpoll RS cardinal_cong, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   150
				  Card_cardinal_eq]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   151
val cadd_0 = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   152
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   153
(** Addition of finite cardinals is "ordinary" addition **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   154
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   155
goalw CardinalArith.thy [eqpoll_def] "succ(A)+B eqpoll succ(A+B)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   156
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   157
by (res_inst_tac [("c", "%z.if(z=Inl(A),A+B,z)"), 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   158
		  ("d", "%z.if(z=A+B,Inl(A),z)")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   159
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   160
by (ALLGOALS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   161
    (asm_simp_tac (case_ss addsimps [succI2, mem_imp_not_eq]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   162
		           setloop eresolve_tac [sumE,succE])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   163
val sum_succ_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   164
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   165
(*Pulling the  succ(...)  outside the |...| requires m, n: nat  *)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   166
(*Unconditional version requires AC*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   167
goalw CardinalArith.thy [cadd_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   168
    "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |+| n = |succ(m |+| n)|";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   169
by (rtac (sum_succ_eqpoll RS cardinal_cong RS trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   170
by (rtac (succ_eqpoll_cong RS cardinal_cong) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   171
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   172
by (REPEAT (ares_tac [well_ord_radd, well_ord_Memrel] 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   173
val cadd_succ_lemma = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   174
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   175
val [mnat,nnat] = goal CardinalArith.thy
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   176
    "[| m: nat;  n: nat |] ==> m |+| n = m#+n";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   177
by (cut_facts_tac [nnat] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   178
by (nat_ind_tac "m" [mnat] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   179
by (asm_simp_tac (arith_ss addsimps [nat_into_Card RS cadd_0]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   180
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cadd_succ_lemma,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   181
				     nat_into_Card RS Card_cardinal_eq]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   182
val nat_cadd_eq_add = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   183
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   184
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   185
(*** Cardinal multiplication ***)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   186
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   187
(** Cardinal multiplication is commutative **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   188
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   189
(*Easier to prove the two directions separately*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   190
goalw CardinalArith.thy [eqpoll_def] "A*B eqpoll B*A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   191
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   192
by (res_inst_tac [("c", "split(%x y.<y,x>)"), ("d", "split(%x y.<y,x>)")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   193
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   194
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   195
by (ALLGOALS (asm_simp_tac ZF_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   196
val prod_commute_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   197
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   198
goalw CardinalArith.thy [cmult_def] "i |*| j = j |*| i";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   199
by (rtac (prod_commute_eqpoll RS cardinal_cong) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   200
val cmult_commute = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   201
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   202
(** Cardinal multiplication is associative **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   203
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   204
goalw CardinalArith.thy [eqpoll_def] "(A*B)*C eqpoll A*(B*C)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   205
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   206
by (res_inst_tac [("c", "split(%w z. split(%x y. <x,<y,z>>, w))"),
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   207
		  ("d", "split(%x.   split(%y z. <<x,y>, z>))")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   208
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   209
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   210
by (ALLGOALS (asm_simp_tac ZF_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   211
val prod_assoc_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   212
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   213
(*Unconditional version requires AC*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   214
goalw CardinalArith.thy [cmult_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   215
    "!!i j k. [| well_ord(i,ri); well_ord(j,rj); well_ord(k,rk) |] ==>	\
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   216
\             (i |*| j) |*| k = i |*| (j |*| k)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   217
by (rtac cardinal_cong 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   218
br ([well_ord_cardinal_eqpoll, eqpoll_refl] MRS prod_eqpoll_cong RS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   219
    eqpoll_trans) 1;
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   220
by (rtac (prod_assoc_eqpoll RS eqpoll_trans) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   221
br ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS prod_eqpoll_cong RS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   222
    eqpoll_sym) 2;
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   223
by (REPEAT (ares_tac [well_ord_rmult] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   224
val well_ord_cmult_assoc = result();
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   225
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   226
(** Cardinal multiplication distributes over addition **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   227
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   228
goalw CardinalArith.thy [eqpoll_def] "(A+B)*C eqpoll (A*C)+(B*C)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   229
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   230
by (res_inst_tac
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   231
    [("c", "split(%x z. case(%y.Inl(<y,z>), %y.Inr(<y,z>), x))"),
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   232
     ("d", "case(split(%x y.<Inl(x),y>), split(%x y.<Inr(x),y>))")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   233
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   234
by (safe_tac (ZF_cs addSEs [sumE]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   235
by (ALLGOALS (asm_simp_tac case_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   236
val sum_prod_distrib_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   237
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   238
goalw CardinalArith.thy [lepoll_def, inj_def] "A lepoll A*A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   239
by (res_inst_tac [("x", "lam x:A. <x,x>")] exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   240
by (simp_tac (ZF_ss addsimps [lam_type]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   241
val prod_square_lepoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   242
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   243
goalw CardinalArith.thy [cmult_def] "!!K. Card(K) ==> K le K |*| K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   244
by (rtac le_trans 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   245
by (rtac well_ord_lepoll_imp_le 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   246
by (rtac prod_square_lepoll 3);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   247
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel, Card_is_Ord] 2));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   248
by (asm_simp_tac (ZF_ss addsimps [le_refl, Card_is_Ord, Card_cardinal_eq]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   249
val cmult_square_le = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   250
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   251
(** Multiplication by 0 yields 0 **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   252
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   253
goalw CardinalArith.thy [eqpoll_def] "0*A eqpoll 0";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   254
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   255
by (rtac lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   256
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   257
val prod_0_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   258
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   259
goalw CardinalArith.thy [cmult_def] "0 |*| i = 0";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   260
by (asm_simp_tac (ZF_ss addsimps [prod_0_eqpoll RS cardinal_cong, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   261
				  Card_0 RS Card_cardinal_eq]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   262
val cmult_0 = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   263
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   264
(** 1 is the identity for multiplication **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   265
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   266
goalw CardinalArith.thy [eqpoll_def] "{x}*A eqpoll A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   267
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   268
by (res_inst_tac [("c", "snd"), ("d", "%z.<x,z>")] lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   269
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   270
by (ALLGOALS (asm_simp_tac ZF_ss));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   271
val prod_singleton_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   272
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   273
goalw CardinalArith.thy [cmult_def, succ_def] "!!K. Card(K) ==> 1 |*| K = K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   274
by (asm_simp_tac (ZF_ss addsimps [prod_singleton_eqpoll RS cardinal_cong, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   275
				  Card_cardinal_eq]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   276
val cmult_1 = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   277
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   278
(** Multiplication of finite cardinals is "ordinary" multiplication **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   279
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   280
goalw CardinalArith.thy [eqpoll_def] "succ(A)*B eqpoll B + A*B";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   281
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   282
by (res_inst_tac [("c", "split(%x y. if(x=A, Inl(y), Inr(<x,y>)))"), 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   283
		  ("d", "case(%y. <A,y>, %z.z)")] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   284
    lam_bijective 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   285
by (safe_tac (ZF_cs addSEs [sumE]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   286
by (ALLGOALS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   287
    (asm_simp_tac (case_ss addsimps [succI2, if_type, mem_imp_not_eq])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   288
val prod_succ_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   289
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   290
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   291
(*Unconditional version requires AC*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   292
goalw CardinalArith.thy [cmult_def, cadd_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   293
    "!!m n. [| Ord(m);  Ord(n) |] ==> succ(m) |*| n = n |+| (m |*| n)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   294
by (rtac (prod_succ_eqpoll RS cardinal_cong RS trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   295
by (rtac (cardinal_cong RS sym) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   296
by (rtac ([eqpoll_refl, well_ord_cardinal_eqpoll] MRS sum_eqpoll_cong) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   297
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   298
val cmult_succ_lemma = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   299
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   300
val [mnat,nnat] = goal CardinalArith.thy
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   301
    "[| m: nat;  n: nat |] ==> m |*| n = m#*n";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   302
by (cut_facts_tac [nnat] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   303
by (nat_ind_tac "m" [mnat] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   304
by (asm_simp_tac (arith_ss addsimps [cmult_0]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   305
by (asm_simp_tac (arith_ss addsimps [nat_into_Ord, cmult_succ_lemma,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   306
				     nat_cadd_eq_add]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   307
val nat_cmult_eq_mult = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   308
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   309
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   310
(*** Infinite Cardinals are Limit Ordinals ***)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   311
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   312
goalw CardinalArith.thy [lepoll_def]
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   313
    "!!i. nat <= A ==> succ(A) lepoll A";
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   314
by (res_inst_tac [("x",
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   315
    "lam z:succ(A). if(z=A, 0, if(z:nat, succ(z), z))")] exI 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   316
by (res_inst_tac [("d", "%y. if(y:nat, nat_case(A,%z.z,y), y)")] 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   317
    lam_injective 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   318
by (fast_tac (ZF_cs addSIs [if_type, nat_0I, nat_succI]) 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   319
by (asm_simp_tac 
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   320
    (ZF_ss addsimps [nat_case_0, nat_case_succ, nat_0I, nat_succI]
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   321
           setloop split_tac [expand_if]) 1);
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   322
val nat_succ_lepoll = result();
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   323
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   324
goalw CardinalArith.thy [lepoll_def, inj_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   325
    "!!i. nat <= A ==> succ(A) lepoll A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   326
by (res_inst_tac [("x",
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   327
   "lam z:succ(A). if(z=A, 0, if(z:nat, succ(z), z))")] exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   328
by (rtac (lam_type RS CollectI) 1);
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   329
by (fast_tac (ZF_cs addSIs [if_type, nat_0I, nat_succI]) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   330
by (REPEAT (rtac ballI 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   331
by (asm_simp_tac 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   332
    (ZF_ss addsimps [succ_inject_iff, succ_not_0, succ_not_0 RS not_sym]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   333
           setloop split_tac [expand_if]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   334
by (safe_tac (ZF_cs addSIs [nat_0I, nat_succI]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   335
val nat_succ_lepoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   336
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   337
goal CardinalArith.thy "!!i. nat <= A ==> succ(A) eqpoll A";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   338
by (etac (nat_succ_lepoll RS eqpollI) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   339
by (rtac (subset_succI RS subset_imp_lepoll) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   340
val nat_succ_eqpoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   341
488
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   342
goalw CardinalArith.thy [InfCard_def] "InfCard(nat)";
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   343
by (fast_tac (ZF_cs addIs [Card_nat, le_refl, Card_is_Ord]) 1);
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   344
val InfCard_nat = result();
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   345
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   346
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Card(K)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   347
by (etac conjunct1 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   348
val InfCard_is_Card = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   349
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   350
(*Kunen's Lemma 10.11*)
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   351
goalw CardinalArith.thy [InfCard_def] "!!K. InfCard(K) ==> Limit(K)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   352
by (etac conjE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   353
by (rtac (ltI RS non_succ_LimitI) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   354
by (etac ([asm_rl, nat_0I] MRS (le_imp_subset RS subsetD)) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   355
by (etac Card_is_Ord 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   356
by (safe_tac (ZF_cs addSDs [Limit_nat RS Limit_le_succD]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   357
by (forward_tac [Card_is_Ord RS Ord_succD] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   358
by (rewtac Card_def);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   359
by (res_inst_tac [("i", "succ(y)")] lt_irrefl 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   360
by (dtac (le_imp_subset RS nat_succ_eqpoll RS cardinal_cong) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   361
(*Tricky combination of substitutions; backtracking needed*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   362
by (etac ssubst 1 THEN etac ssubst 1 THEN rtac Ord_cardinal_le 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   363
by (assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   364
val InfCard_is_Limit = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   365
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   366
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   367
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   368
(*** An infinite cardinal equals its square (Kunen, Thm 10.12, page 29) ***)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   369
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   370
(*A general fact about ordermap*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   371
goalw Cardinal.thy [eqpoll_def]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   372
    "!!A. [| well_ord(A,r);  x:A |] ==> ordermap(A,r)`x eqpoll pred(A,x,r)";
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   373
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   374
by (asm_simp_tac (ZF_ss addsimps [ordermap_eq_image, well_ord_is_wf]) 1);
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   375
by (etac (ordermap_bij RS bij_is_inj RS restrict_bij RS bij_converse_bij) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   376
by (rtac pred_subset 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   377
val ordermap_eqpoll_pred = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   378
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   379
(** Establishing the well-ordering **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   380
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   381
goalw CardinalArith.thy [inj_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   382
 "!!K. Ord(K) ==>	\
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   383
\ (lam z:K*K. split(%x y. <x Un y, <x, y>>, z)) : inj(K*K, K*K*K)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   384
by (safe_tac ZF_cs);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   385
by (fast_tac (ZF_cs addIs [lam_type, Un_least_lt RS ltD, ltI]
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   386
                    addSEs [split_type]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   387
by (asm_full_simp_tac ZF_ss 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   388
val csquare_lam_inj = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   389
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   390
goalw CardinalArith.thy [csquare_rel_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   391
 "!!K. Ord(K) ==> well_ord(K*K, csquare_rel(K))";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   392
by (rtac (csquare_lam_inj RS well_ord_rvimage) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   393
by (REPEAT (ares_tac [well_ord_rmult, well_ord_Memrel] 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   394
val well_ord_csquare = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   395
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   396
(** Characterising initial segments of the well-ordering **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   397
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   398
goalw CardinalArith.thy [csquare_rel_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   399
 "!!K. [| x<K;  y<K;  z<K |] ==> \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   400
\      <<x,y>, <z,z>> : csquare_rel(K) --> x le z & y le z";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   401
by (REPEAT (etac ltE 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   402
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   403
                                  Un_absorb, Un_least_mem_iff, ltD]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   404
by (safe_tac (ZF_cs addSEs [mem_irrefl] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   405
                    addSIs [Un_upper1_le, Un_upper2_le]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   406
by (ALLGOALS (asm_simp_tac (ZF_ss addsimps [lt_def, succI2, Ord_succ])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   407
val csquareD_lemma = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   408
val csquareD = csquareD_lemma RS mp |> standard;
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   409
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   410
goalw CardinalArith.thy [pred_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   411
 "!!K. z<K ==> pred(K*K, <z,z>, csquare_rel(K)) <= succ(z)*succ(z)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   412
by (safe_tac (lemmas_cs addSEs [SigmaE]));	(*avoids using succCI*)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   413
by (rtac (csquareD RS conjE) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   414
by (rewtac lt_def);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   415
by (assume_tac 4);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   416
by (ALLGOALS (fast_tac ZF_cs));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   417
val pred_csquare_subset = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   418
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   419
goalw CardinalArith.thy [csquare_rel_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   420
 "!!K. [| x<z;  y<z;  z<K |] ==> \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   421
\      <<x,y>, <z,z>> : csquare_rel(K)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   422
by (subgoals_tac ["x<K", "y<K"] 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   423
by (REPEAT (eresolve_tac [asm_rl, lt_trans] 2));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   424
by (REPEAT (etac ltE 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   425
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   426
                                  Un_absorb, Un_least_mem_iff, ltD]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   427
val csquare_ltI = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   428
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   429
(*Part of the traditional proof.  UNUSED since it's harder to prove & apply *)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   430
goalw CardinalArith.thy [csquare_rel_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   431
 "!!K. [| x le z;  y le z;  z<K |] ==> \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   432
\      <<x,y>, <z,z>> : csquare_rel(K) | x=z & y=z";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   433
by (subgoals_tac ["x<K", "y<K"] 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   434
by (REPEAT (eresolve_tac [asm_rl, lt_trans1] 2));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   435
by (REPEAT (etac ltE 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   436
by (asm_simp_tac (ZF_ss addsimps [rvimage_iff, rmult_iff, Memrel_iff,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   437
                                  Un_absorb, Un_least_mem_iff, ltD]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   438
by (REPEAT_FIRST (etac succE));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   439
by (ALLGOALS
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   440
    (asm_simp_tac (ZF_ss addsimps [subset_Un_iff RS iff_sym, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   441
				   subset_Un_iff2 RS iff_sym, OrdmemD])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   442
val csquare_or_eqI = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   443
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   444
(** The cardinality of initial segments **)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   445
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   446
goal CardinalArith.thy
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   447
    "!!K. [| InfCard(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   448
\         ordermap(K*K, csquare_rel(K)) ` <x,y> lepoll 		\
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   449
\         ordermap(K*K, csquare_rel(K)) ` <z,z>";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   450
by (subgoals_tac ["z<K", "well_ord(K*K, csquare_rel(K))"] 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   451
by (etac (InfCard_is_Card RS Card_is_Ord RS well_ord_csquare) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   452
by (fast_tac (ZF_cs addSIs [Un_least_lt, InfCard_is_Limit, Limit_has_succ]) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   453
by (rtac (OrdmemD RS subset_imp_lepoll) 1);
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   454
by (res_inst_tac [("z1","z")] (csquare_ltI RS ordermap_mono) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   455
by (etac well_ord_is_wf 4);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   456
by (ALLGOALS 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   457
    (fast_tac (ZF_cs addSIs [Un_upper1_le, Un_upper2_le, Ord_ordermap] 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   458
                     addSEs [ltE])));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   459
val ordermap_z_lepoll = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   460
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   461
(*Kunen: "each <x,y>: K*K has no more than z*z predecessors..." (page 29) *)
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   462
goalw CardinalArith.thy [cmult_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   463
  "!!K. [| InfCard(K);  x<K;  y<K;  z=succ(x Un y) |] ==> \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   464
\       | ordermap(K*K, csquare_rel(K)) ` <x,y> | le  |succ(z)| |*| |succ(z)|";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   465
by (rtac (well_ord_rmult RS well_ord_lepoll_imp_le) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   466
by (REPEAT (ares_tac [Ord_cardinal, well_ord_Memrel] 1));
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   467
by (subgoals_tac ["z<K"] 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   468
by (fast_tac (ZF_cs addSIs [Un_least_lt, InfCard_is_Limit, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   469
                            Limit_has_succ]) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   470
by (rtac (ordermap_z_lepoll RS lepoll_trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   471
by (REPEAT_SOME assume_tac);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   472
by (rtac (ordermap_eqpoll_pred RS eqpoll_imp_lepoll RS lepoll_trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   473
by (etac (InfCard_is_Card RS Card_is_Ord RS well_ord_csquare) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   474
by (fast_tac (ZF_cs addIs [ltD]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   475
by (rtac (pred_csquare_subset RS subset_imp_lepoll RS lepoll_trans) 1 THEN
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   476
    assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   477
by (REPEAT_FIRST (etac ltE));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   478
by (rtac (prod_eqpoll_cong RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   479
by (REPEAT_FIRST (etac (Ord_succ RS Ord_cardinal_eqpoll)));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   480
val ordermap_csquare_le = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   481
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   482
(*Kunen: "... so the order type <= K" *)
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   483
goal CardinalArith.thy
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   484
    "!!K. [| InfCard(K);  ALL y:K. InfCard(y) --> y |*| y = y |]  ==>  \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   485
\         ordertype(K*K, csquare_rel(K)) le K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   486
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   487
by (rtac all_lt_imp_le 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   488
by (assume_tac 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   489
by (etac (well_ord_csquare RS Ord_ordertype) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   490
by (rtac Card_lt_imp_lt 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   491
by (etac InfCard_is_Card 3);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   492
by (etac ltE 2 THEN assume_tac 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   493
by (asm_full_simp_tac (ZF_ss addsimps [ordertype_unfold]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   494
by (safe_tac (ZF_cs addSEs [ltE]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   495
by (subgoals_tac ["Ord(xb)", "Ord(y)"] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   496
by (REPEAT (eresolve_tac [asm_rl, Ord_in_Ord] 2));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   497
by (rtac (ordermap_csquare_le RS lt_trans1) 1  THEN
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   498
    REPEAT (ares_tac [refl] 1 ORELSE etac ltI 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   499
by (res_inst_tac [("i","xb Un y"), ("j","nat")] Ord_linear2 1  THEN
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   500
    REPEAT (ares_tac [Ord_Un, Ord_nat] 1));
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   501
(*the finite case: xb Un y < nat *)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   502
by (res_inst_tac [("j", "nat")] lt_trans2 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   503
by (asm_full_simp_tac (FOL_ss addsimps [InfCard_def]) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   504
by (asm_full_simp_tac
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   505
    (ZF_ss addsimps [lt_def, nat_cmult_eq_mult, nat_succI, mult_type,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   506
		     nat_into_Card RS Card_cardinal_eq, Ord_nat]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   507
(*case nat le (xb Un y), equivalently InfCard(xb Un y)  *)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   508
by (asm_full_simp_tac
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   509
    (ZF_ss addsimps [le_imp_subset RS nat_succ_eqpoll RS cardinal_cong,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   510
		     le_succ_iff, InfCard_def, Card_cardinal, Un_least_lt, 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   511
		     Ord_Un, ltI, nat_le_cardinal,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   512
		     Ord_cardinal_le RS lt_trans1 RS ltD]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   513
val ordertype_csquare_le = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   514
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   515
(*This lemma can easily be generalized to premise well_ord(A*A,r) *)
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   516
goalw CardinalArith.thy [cmult_def]
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   517
    "!!K. Ord(K) ==> K |*| K  =  |ordertype(K*K, csquare_rel(K))|";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   518
by (rtac cardinal_cong 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   519
by (rewtac eqpoll_def);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   520
by (rtac exI 1);
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   521
by (etac (well_ord_csquare RS ordermap_bij) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   522
val csquare_eq_ordertype = result();
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   523
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   524
(*Main result: Kunen's Theorem 10.12*)
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   525
goal CardinalArith.thy "!!K. InfCard(K) ==> K |*| K = K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   526
by (forward_tac [InfCard_is_Card RS Card_is_Ord] 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   527
by (etac rev_mp 1);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   528
by (trans_ind_tac "K" [] 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   529
by (rtac impI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   530
by (rtac le_anti_sym 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   531
by (etac (InfCard_is_Card RS cmult_square_le) 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   532
by (rtac (ordertype_csquare_le RSN (2, le_trans)) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   533
by (assume_tac 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   534
by (assume_tac 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   535
by (asm_simp_tac 
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   536
    (ZF_ss addsimps [csquare_eq_ordertype, Ord_cardinal_le,
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   537
                     well_ord_csquare RS Ord_ordertype]) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents:
diff changeset
   538
val InfCard_csquare_eq = result();
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   539
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   540
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   541
goal CardinalArith.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   542
    "!!A. [| well_ord(A,r);  InfCard(|A|) |] ==> A*A eqpoll A";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   543
by (resolve_tac [prod_eqpoll_cong RS eqpoll_trans] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   544
by (REPEAT (etac (well_ord_cardinal_eqpoll RS eqpoll_sym) 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   545
by (resolve_tac [well_ord_cardinal_eqE] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   546
by (REPEAT (ares_tac [Ord_cardinal, well_ord_rmult, well_ord_Memrel] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   547
by (asm_simp_tac (ZF_ss addsimps [symmetric cmult_def, InfCard_csquare_eq]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   548
val well_ord_InfCard_square_eq = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   549
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   550
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   551
(*** For every cardinal number there exists a greater one
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   552
     [Kunen's Theorem 10.16, which would be trivial using AC] ***)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   553
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   554
goalw CardinalArith.thy [jump_cardinal_def] "Ord(jump_cardinal(K))";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   555
by (rtac (Ord_is_Transset RSN (2,OrdI)) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   556
by (safe_tac (ZF_cs addSIs [Ord_ordertype]));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   557
bw Transset_def;
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   558
by (safe_tac ZF_cs);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   559
by (rtac (ordertype_subset RS exE) 1 THEN REPEAT (assume_tac 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   560
by (resolve_tac [UN_I] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   561
by (resolve_tac [ReplaceI] 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   562
by (ALLGOALS (fast_tac (ZF_cs addSEs [well_ord_subset])));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   563
val Ord_jump_cardinal = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   564
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   565
(*Allows selective unfolding.  Less work than deriving intro/elim rules*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   566
goalw CardinalArith.thy [jump_cardinal_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   567
     "i : jump_cardinal(K) <->   \
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   568
\         (EX r X. r <= K*K & X <= K & well_ord(X,r) & i = ordertype(X,r))";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   569
by (fast_tac subset_cs 1);	(*It's vital to avoid reasoning about <=*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   570
val jump_cardinal_iff = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   571
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   572
(*The easy part of Theorem 10.16: jump_cardinal(K) exceeds K*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   573
goal CardinalArith.thy "!!K. Ord(K) ==> K < jump_cardinal(K)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   574
by (resolve_tac [Ord_jump_cardinal RSN (2,ltI)] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   575
by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   576
by (REPEAT_FIRST (ares_tac [exI, conjI, well_ord_Memrel]));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   577
by (resolve_tac [subset_refl] 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   578
by (asm_simp_tac (ZF_ss addsimps [Memrel_def, subset_iff]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   579
by (asm_simp_tac (ZF_ss addsimps [ordertype_Memrel]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   580
val K_lt_jump_cardinal = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   581
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   582
(*The proof by contradiction: the bijection f yields a wellordering of X
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   583
  whose ordertype is jump_cardinal(K).  *)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   584
goal CardinalArith.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   585
    "!!K. [| well_ord(X,r);  r <= K * K;  X <= K;	\
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   586
\            f : bij(ordertype(X,r), jump_cardinal(K)) 	\
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   587
\	  |] ==> jump_cardinal(K) : jump_cardinal(K)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   588
by (subgoal_tac "f O ordermap(X,r): bij(X, jump_cardinal(K))" 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   589
by (REPEAT (ares_tac [comp_bij, ordermap_bij] 2));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   590
by (resolve_tac [jump_cardinal_iff RS iffD2] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   591
by (REPEAT_FIRST (resolve_tac [exI, conjI]));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   592
by (rtac ([rvimage_type, Sigma_mono] MRS subset_trans) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   593
by (REPEAT (assume_tac 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   594
by (etac (bij_is_inj RS well_ord_rvimage) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   595
by (rtac (Ord_jump_cardinal RS well_ord_Memrel) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   596
by (asm_simp_tac
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   597
    (ZF_ss addsimps [well_ord_Memrel RSN (2, bij_ordertype_vimage), 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   598
		     ordertype_Memrel, Ord_jump_cardinal]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   599
val Card_jump_cardinal_lemma = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   600
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   601
(*The hard part of Theorem 10.16: jump_cardinal(K) is itself a cardinal*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   602
goal CardinalArith.thy "Card(jump_cardinal(K))";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   603
by (rtac (Ord_jump_cardinal RS CardI) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   604
by (rewrite_goals_tac [eqpoll_def]);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   605
by (safe_tac (ZF_cs addSDs [ltD, jump_cardinal_iff RS iffD1]));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   606
by (REPEAT (ares_tac [Card_jump_cardinal_lemma RS mem_irrefl] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   607
val Card_jump_cardinal = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   608
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   609
(*** Basic properties of successor cardinals ***)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   610
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   611
goalw CardinalArith.thy [csucc_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   612
    "!!K. Ord(K) ==> Card(csucc(K)) & K < csucc(K)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   613
by (rtac LeastI 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   614
by (REPEAT (ares_tac [conjI, Card_jump_cardinal, K_lt_jump_cardinal,
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   615
		      Ord_jump_cardinal] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   616
val csucc_basic = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   617
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   618
val Card_csucc = csucc_basic RS conjunct1 |> standard;
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   619
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   620
val lt_csucc = csucc_basic RS conjunct2 |> standard;
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   621
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   622
goalw CardinalArith.thy [csucc_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   623
    "!!K L. [| Card(L);  K<L |] ==> csucc(K) le L";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   624
by (rtac Least_le 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   625
by (REPEAT (ares_tac [conjI, Card_is_Ord] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   626
val csucc_le = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   627
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   628
goal CardinalArith.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   629
    "!!K. [| Ord(i); Card(K) |] ==> i < csucc(K) <-> |i| le K";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   630
by (resolve_tac [iffI] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   631
by (resolve_tac [Card_lt_imp_lt] 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   632
by (eresolve_tac [lt_trans1] 2);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   633
by (REPEAT (ares_tac [lt_csucc, Card_csucc, Card_is_Ord] 2));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   634
by (resolve_tac [notI RS not_lt_imp_le] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   635
by (resolve_tac [Card_cardinal RS csucc_le RS lt_trans1 RS lt_irrefl] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   636
by (assume_tac 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   637
by (resolve_tac [Ord_cardinal_le RS lt_trans1] 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   638
by (REPEAT (ares_tac [Ord_cardinal] 1
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   639
     ORELSE eresolve_tac [ltE, Card_is_Ord] 1));
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   640
val lt_csucc_iff = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   641
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   642
goal CardinalArith.thy
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   643
    "!!K' K. [| Card(K'); Card(K) |] ==> K' < csucc(K) <-> K' le K";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   644
by (asm_simp_tac 
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   645
    (ZF_ss addsimps [lt_csucc_iff, Card_cardinal_eq, Card_is_Ord]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   646
val Card_lt_csucc_iff = result();
488
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   647
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   648
goalw CardinalArith.thy [InfCard_def]
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   649
    "!!K. InfCard(K) ==> InfCard(csucc(K))";
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   650
by (asm_simp_tac (ZF_ss addsimps [Card_csucc, Card_is_Ord, 
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   651
				  lt_csucc RS leI RSN (2,le_trans)]) 1);
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
   652
val InfCard_csucc = result();