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