src/ZF/Cardinal.ML
author wenzelm
Sat, 01 Jul 2000 19:55:22 +0200
changeset 9230 17ae63f82ad8
parent 9211 6236c5285bd8
child 9683 f87c8c449018
permissions -rw-r--r--
GPLed;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
     1
(*  Title:      ZF/Cardinal.ML
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     5
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     6
Cardinals in Zermelo-Fraenkel Set Theory 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     7
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     8
This theory does NOT assume the Axiom of Choice
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     9
*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    10
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    11
(*** The Schroeder-Bernstein Theorem -- see Davey & Priestly, page 106 ***)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    12
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    13
(** Lemma: Banach's Decomposition Theorem **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    14
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    15
Goal "bnd_mono(X, %W. X - g``(Y - f``W))";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    16
by (rtac bnd_monoI 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    17
by (REPEAT (ares_tac [Diff_subset, subset_refl, Diff_mono, image_mono] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    18
qed "decomp_bnd_mono";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    19
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    20
val [gfun] = goal Cardinal.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
    21
    "g: Y->X ==>                                        \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
    22
\    g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) =       \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    23
\    X - lfp(X, %W. X - g``(Y - f``W)) ";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    24
by (res_inst_tac [("P", "%u. ?v = X-u")] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    25
     (decomp_bnd_mono RS lfp_Tarski RS ssubst) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
    26
by (simp_tac (simpset() addsimps [subset_refl, double_complement,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
    27
                             gfun RS fun_is_rel RS image_subset]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    28
qed "Banach_last_equation";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    29
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
    30
Goal "[| f: X->Y;  g: Y->X |] ==>   \
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
    31
\     EX XA XB YA YB. (XA Int XB = 0) & (XA Un XB = X) &    \
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
    32
\                     (YA Int YB = 0) & (YA Un YB = Y) &    \
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
    33
\                     f``XA=YA & g``YB=XB";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    34
by (REPEAT 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    35
    (FIRSTGOAL
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    36
     (resolve_tac [refl, exI, conjI, Diff_disjoint, Diff_partition])));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    37
by (rtac Banach_last_equation 3);
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
    38
by (REPEAT (ares_tac [fun_is_rel, image_subset, lfp_subset] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    39
qed "decomposition";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    40
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    41
val prems = goal Cardinal.thy
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    42
    "[| f: inj(X,Y);  g: inj(Y,X) |] ==> EX h. h: bij(X,Y)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    43
by (cut_facts_tac prems 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    44
by (cut_facts_tac [(prems RL [inj_is_fun]) MRS decomposition] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
    45
by (blast_tac (claset() addSIs [restrict_bij,bij_disjoint_Un]
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    46
                    addIs [bij_converse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    47
(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))"
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    48
   is forced by the context!! *)
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    49
qed "schroeder_bernstein";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    50
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    51
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    52
(** Equipollence is an equivalence relation **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    53
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    54
Goalw [eqpoll_def] "f: bij(A,B) ==> A eqpoll B";
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
    55
by (etac exI 1);
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
    56
qed "bij_imp_eqpoll";
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
    57
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
    58
(*A eqpoll A*)
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
    59
bind_thm ("eqpoll_refl", id_bij RS bij_imp_eqpoll);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    60
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    61
Goalw [eqpoll_def] "X eqpoll Y ==> Y eqpoll X";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
    62
by (blast_tac (claset() addIs [bij_converse_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    63
qed "eqpoll_sym";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    64
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    65
Goalw [eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    66
    "[| X eqpoll Y;  Y eqpoll Z |] ==> X eqpoll Z";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
    67
by (blast_tac (claset() addIs [comp_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    68
qed "eqpoll_trans";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    69
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    70
(** Le-pollence is a partial ordering **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    71
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    72
Goalw [lepoll_def] "X<=Y ==> X lepoll Y";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    73
by (rtac exI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    74
by (etac id_subset_inj 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    75
qed "subset_imp_lepoll";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    76
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
    77
bind_thm ("lepoll_refl", subset_refl RS subset_imp_lepoll);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
    78
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
    79
bind_thm ("le_imp_lepoll", le_imp_subset RS subset_imp_lepoll);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    80
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    81
Goalw [eqpoll_def, bij_def, lepoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    82
    "X eqpoll Y ==> X lepoll Y";
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
    83
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    84
qed "eqpoll_imp_lepoll";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    85
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    86
Goalw [lepoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    87
    "[| X lepoll Y;  Y lepoll Z |] ==> X lepoll Z";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
    88
by (blast_tac (claset() addIs [comp_inj]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    89
qed "lepoll_trans";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    90
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    91
(*Asymmetry law*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
    92
Goalw [lepoll_def,eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
    93
    "[| X lepoll Y;  Y lepoll X |] ==> X eqpoll Y";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    94
by (REPEAT (etac exE 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    95
by (rtac schroeder_bernstein 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    96
by (REPEAT (assume_tac 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
    97
qed "eqpollI";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    98
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
    99
val [major,minor] = Goal
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   100
    "[| X eqpoll Y; [| X lepoll Y; Y lepoll X |] ==> P |] ==> P";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   101
by (rtac minor 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   102
by (REPEAT (resolve_tac [major, eqpoll_imp_lepoll, eqpoll_sym] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   103
qed "eqpollE";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   104
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   105
Goal "X eqpoll Y <-> X lepoll Y & Y lepoll X";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   106
by (blast_tac (claset() addIs [eqpollI] addSEs [eqpollE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   107
qed "eqpoll_iff";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   108
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   109
Goalw [lepoll_def, inj_def] "A lepoll 0 ==> A = 0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   110
by (blast_tac (claset() addDs [apply_type]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   111
qed "lepoll_0_is_0";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   112
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   113
(*0 lepoll Y*)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   114
bind_thm ("empty_lepollI", empty_subsetI RS subset_imp_lepoll);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   115
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   116
Goal "A lepoll 0 <-> A=0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   117
by (blast_tac (claset() addIs [lepoll_0_is_0, lepoll_refl]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   118
qed "lepoll_0_iff";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   119
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   120
Goalw [lepoll_def] 
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   121
    "[| A lepoll B; C lepoll D; B Int D = 0 |] ==> A Un C lepoll B Un D";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   122
by (blast_tac (claset() addIs [inj_disjoint_Un]) 1);
1709
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1623
diff changeset
   123
qed "Un_lepoll_Un";
220dd588bfc9 New lemma inspired by KG
paulson
parents: 1623
diff changeset
   124
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   125
(*A eqpoll 0 ==> A=0*)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   126
bind_thm ("eqpoll_0_is_0",  eqpoll_imp_lepoll RS lepoll_0_is_0);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   127
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   128
Goal "A eqpoll 0 <-> A=0";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   129
by (blast_tac (claset() addIs [eqpoll_0_is_0, eqpoll_refl]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   130
qed "eqpoll_0_iff";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   131
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   132
Goalw [eqpoll_def] 
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   133
    "[| A eqpoll B;  C eqpoll D;  A Int C = 0;  B Int D = 0 |]  \
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   134
\    ==> A Un C eqpoll B Un D";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   135
by (blast_tac (claset() addIs [bij_disjoint_Un]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   136
qed "eqpoll_disjoint_Un";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   137
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   138
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   139
(*** lesspoll: contributions by Krzysztof Grabczewski ***)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   140
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   141
Goalw [lesspoll_def] "A lesspoll B ==> A lepoll B";
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
   142
by (Blast_tac 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   143
qed "lesspoll_imp_lepoll";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   144
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   145
Goalw [lepoll_def] "[| A lepoll B; well_ord(B,r) |] ==> EX s. well_ord(A,s)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   146
by (blast_tac (claset() addIs [well_ord_rvimage]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   147
qed "lepoll_well_ord";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   148
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   149
Goalw [lesspoll_def] "A lepoll B <-> A lesspoll B | A eqpoll B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   150
by (blast_tac (claset() addSIs [eqpollI] addSEs [eqpollE]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   151
qed "lepoll_iff_leqpoll";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   152
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   153
Goalw [inj_def, surj_def] 
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   154
  "[| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   155
by (safe_tac (claset_of ZF.thy));
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   156
by (swap_res_tac [exI] 1);
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   157
by (res_inst_tac [("a", "lam z:A. if f`z=m then y else f`z")] CollectI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   158
by (best_tac (claset() addSIs [if_type RS lam_type]
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   159
                       addEs [apply_funtype RS succE]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   160
(*Proving it's injective*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   161
by (Asm_simp_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   162
by (blast_tac (claset() delrules [equalityI]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   163
qed "inj_not_surj_succ";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   164
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   165
(** Variations on transitivity **)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   166
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   167
Goalw [lesspoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   168
      "[| X lesspoll Y; Y lesspoll Z |] ==> X lesspoll Z";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   169
by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   170
qed "lesspoll_trans";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   171
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   172
Goalw [lesspoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   173
      "[| X lesspoll Y; Y lepoll Z |] ==> X lesspoll Z";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   174
by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   175
qed "lesspoll_lepoll_lesspoll";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   176
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   177
Goalw [lesspoll_def] 
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   178
      "[| X lesspoll Y; Z lepoll X |] ==> Z lesspoll Y";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   179
by (blast_tac (claset() addSEs [eqpollE] addIs [eqpollI, lepoll_trans]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   180
qed "lepoll_lesspoll_lesspoll";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   181
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   182
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   183
(** LEAST -- the least number operator [from HOL/Univ.ML] **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   184
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   185
val [premP,premOrd,premNot] = Goalw [Least_def]
3840
e0baea4d485a fixed dots;
wenzelm
parents: 3736
diff changeset
   186
    "[| P(i);  Ord(i);  !!x. x<i ==> ~P(x) |] ==> (LEAST x. P(x)) = i";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   187
by (rtac the_equality 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   188
by (blast_tac (claset() addSIs [premP,premOrd,premNot]) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   189
by (REPEAT (etac conjE 1));
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   190
by (etac (premOrd RS Ord_linear_lt) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   191
by (ALLGOALS (blast_tac (claset() addSIs [premP] addSDs [premNot])));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   192
qed "Least_equality";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   193
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   194
Goal "[| P(i);  Ord(i) |] ==> P(LEAST x. P(x))";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   195
by (etac rev_mp 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   196
by (trans_ind_tac "i" [] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   197
by (rtac impI 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   198
by (rtac classical 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   199
by (EVERY1 [stac Least_equality, assume_tac, assume_tac]);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   200
by (assume_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   201
by (blast_tac (claset() addSEs [ltE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   202
qed "LeastI";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   203
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   204
(*Proof is almost identical to the one above!*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   205
Goal "[| P(i);  Ord(i) |] ==> (LEAST x. P(x)) le i";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   206
by (etac rev_mp 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   207
by (trans_ind_tac "i" [] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   208
by (rtac impI 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   209
by (rtac classical 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   210
by (EVERY1 [stac Least_equality, assume_tac, assume_tac]);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   211
by (etac le_refl 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   212
by (blast_tac (claset() addEs [ltE] addIs [leI, ltI, lt_trans1]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   213
qed "Least_le";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   214
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   215
(*LEAST really is the smallest*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   216
Goal "[| P(i);  i < (LEAST x. P(x)) |] ==> Q";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   217
by (rtac (Least_le RSN (2,lt_trans2) RS lt_irrefl) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   218
by (REPEAT (eresolve_tac [asm_rl, ltE] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   219
qed "less_LeastE";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   220
1031
a53cbb4b06c5 Proved lesspoll_succ_iff.
lcp
parents: 845
diff changeset
   221
(*Easier to apply than LeastI: conclusion has only one occurrence of P*)
9211
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   222
val prems = goal Cardinal.thy
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   223
    "[| P(i);  Ord(i);  !!j. P(j) ==> Q(j) |] ==> Q(LEAST j. P(j))";
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   224
by (resolve_tac prems 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   225
by (rtac LeastI 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   226
by (resolve_tac prems 1);
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   227
by (resolve_tac prems 1) ;
6236c5285bd8 removal of batch-style proofs
paulson
parents: 9173
diff changeset
   228
qed "LeastI2";
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   229
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   230
(*If there is no such P then LEAST is vacuously 0*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   231
Goalw [Least_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   232
    "[| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   233
by (rtac the_0 1);
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
   234
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   235
qed "Least_0";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   236
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   237
Goal "Ord(LEAST x. P(x))";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   238
by (excluded_middle_tac "EX i. Ord(i) & P(i)" 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   239
by Safe_tac;
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   240
by (rtac (Least_le RS ltE) 2);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   241
by (REPEAT_SOME assume_tac);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   242
by (etac (Least_0 RS ssubst) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   243
by (rtac Ord_0 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   244
qed "Ord_Least";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   245
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   246
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   247
(** Basic properties of cardinals **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   248
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   249
(*Not needed for simplification, but helpful below*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   250
val prems = Goal "(!!y. P(y) <-> Q(y)) ==> (LEAST x. P(x)) = (LEAST x. Q(x))";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   251
by (simp_tac (simpset() addsimps prems) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   252
qed "Least_cong";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   253
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   254
(*Need AC to get X lepoll Y ==> |X| le |Y|;  see well_ord_lepoll_imp_Card_le
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   255
  Converse also requires AC, but see well_ord_cardinal_eqE*)
5143
b94cd208f073 Removal of leading "\!\!..." from most Goal commands
paulson
parents: 5137
diff changeset
   256
Goalw [eqpoll_def,cardinal_def] "X eqpoll Y ==> |X| = |Y|";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   257
by (rtac Least_cong 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   258
by (blast_tac (claset() addIs [comp_bij, bij_converse_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   259
qed "cardinal_cong";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   260
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   261
(*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   262
Goalw [cardinal_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   263
    "well_ord(A,r) ==> |A| eqpoll A";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   264
by (rtac LeastI 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   265
by (etac Ord_ordertype 2);
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   266
by (etac (ordermap_bij RS bij_converse_bij RS bij_imp_eqpoll) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   267
qed "well_ord_cardinal_eqpoll";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   268
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   269
(* Ord(A) ==> |A| eqpoll A *)
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 792
diff changeset
   270
bind_thm ("Ord_cardinal_eqpoll", well_ord_Memrel RS well_ord_cardinal_eqpoll);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   271
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   272
Goal "[| well_ord(X,r);  well_ord(Y,s);  |X| = |Y| |] ==> X eqpoll Y";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   273
by (rtac (eqpoll_sym RS eqpoll_trans) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   274
by (etac well_ord_cardinal_eqpoll 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   275
by (asm_simp_tac (simpset() addsimps [well_ord_cardinal_eqpoll]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   276
qed "well_ord_cardinal_eqE";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   277
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   278
Goal "[| well_ord(X,r);  well_ord(Y,s) |] ==> |X| = |Y| <-> X eqpoll Y";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   279
by (blast_tac (claset() addIs [cardinal_cong, well_ord_cardinal_eqE]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   280
qed "well_ord_cardinal_eqpoll_iff";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   281
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   282
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   283
(** Observations from Kunen, page 28 **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   284
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   285
Goalw [cardinal_def] "Ord(i) ==> |i| le i";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   286
by (etac (eqpoll_refl RS Least_le) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   287
qed "Ord_cardinal_le";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   288
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   289
Goalw [Card_def] "Card(K) ==> |K| = K";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   290
by (etac sym 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   291
qed "Card_cardinal_eq";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   292
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   293
(* Could replace the  ~(j eqpoll i)  by  ~(i lepoll j) *)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   294
val prems = Goalw [Card_def,cardinal_def]
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   295
    "[| Ord(i);  !!j. j<i ==> ~(j eqpoll i) |] ==> Card(i)";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   296
by (stac Least_equality 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   297
by (REPEAT (ares_tac ([refl,eqpoll_refl]@prems) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   298
qed "CardI";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   299
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5143
diff changeset
   300
Goalw [Card_def, cardinal_def] "Card(i) ==> Ord(i)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   301
by (etac ssubst 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   302
by (rtac Ord_Least 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   303
qed "Card_is_Ord";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   304
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   305
Goal "Card(K) ==> K le |K|";
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   306
by (asm_simp_tac (simpset() addsimps [Card_is_Ord, Card_cardinal_eq]) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 765
diff changeset
   307
qed "Card_cardinal_le";
765
06a484afc603 Card_cardinal_le: new
lcp
parents: 760
diff changeset
   308
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   309
Goalw [cardinal_def] "Ord(|A|)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   310
by (rtac Ord_Least 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   311
qed "Ord_cardinal";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   312
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   313
Addsimps [Ord_cardinal];
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   314
AddSIs [Ord_cardinal];
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   315
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   316
(*The cardinals are the initial ordinals*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   317
Goal "Card(K) <-> Ord(K) & (ALL j. j<K --> ~ j eqpoll K)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   318
by (safe_tac (claset() addSIs [CardI, Card_is_Ord]));
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
   319
by (Blast_tac 2);
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   320
by (rewrite_goals_tac [Card_def, cardinal_def]);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   321
by (rtac less_LeastE 1);
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   322
by (etac subst 2);
845
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   323
by (ALLGOALS assume_tac);
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   324
qed "Card_iff_initial";
825e96b87ef7 Added Krzysztof's theorem LeastI2. Proof of sum_eqpoll_cong
lcp
parents: 833
diff changeset
   325
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   326
Goalw [lesspoll_def] "[| Card(a); i<a |] ==> i lesspoll a";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   327
by (dresolve_tac [Card_iff_initial RS iffD1] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   328
by (blast_tac (claset() addSIs [leI RS le_imp_lepoll]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   329
qed "lt_Card_imp_lesspoll";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   330
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   331
Goal "Card(0)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   332
by (rtac (Ord_0 RS CardI) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   333
by (blast_tac (claset() addSEs [ltE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   334
qed "Card_0";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   335
522
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   336
val [premK,premL] = goal Cardinal.thy
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   337
    "[| Card(K);  Card(L) |] ==> Card(K Un L)";
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   338
by (rtac ([premK RS Card_is_Ord, premL RS Card_is_Ord] MRS Ord_linear_le) 1);
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   339
by (asm_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   340
    (simpset() addsimps [premL, le_imp_subset, subset_Un_iff RS iffD1]) 1);
522
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   341
by (asm_simp_tac
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   342
    (simpset() addsimps [premK, le_imp_subset, subset_Un_iff2 RS iffD1]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   343
qed "Card_Un";
522
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   344
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   345
(*Infinite unions of cardinals?  See Devlin, Lemma 6.7, page 98*)
e1de521e012a ZF/Cardinal/Card_Un: new
lcp
parents: 484
diff changeset
   346
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   347
Goalw [cardinal_def] "Card(|A|)";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   348
by (excluded_middle_tac "EX i. Ord(i) & i eqpoll A" 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   349
by (etac (Least_0 RS ssubst) 1 THEN rtac Card_0 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   350
by (rtac (Ord_Least RS CardI) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   351
by Safe_tac;
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   352
by (rtac less_LeastE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   353
by (assume_tac 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   354
by (etac eqpoll_trans 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   355
by (REPEAT (ares_tac [LeastI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   356
qed "Card_cardinal";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   357
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   358
(*Kunen's Lemma 10.5*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   359
Goal "[| |i| le j;  j le i |] ==> |j| = |i|";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   360
by (rtac (eqpollI RS cardinal_cong) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   361
by (etac le_imp_lepoll 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   362
by (rtac lepoll_trans 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   363
by (etac le_imp_lepoll 2);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   364
by (rtac (eqpoll_sym RS eqpoll_imp_lepoll) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   365
by (rtac Ord_cardinal_eqpoll 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   366
by (REPEAT (eresolve_tac [ltE, Ord_succD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   367
qed "cardinal_eq_lemma";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   368
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   369
Goal "i le j ==> |i| le |j|";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   370
by (res_inst_tac [("i","|i|"),("j","|j|")] Ord_linear_le 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   371
by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI]));
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   372
by (rtac cardinal_eq_lemma 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   373
by (assume_tac 2);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   374
by (etac le_trans 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   375
by (etac ltE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   376
by (etac Ord_cardinal_le 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   377
qed "cardinal_mono";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   378
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   379
(*Since we have |succ(nat)| le |nat|, the converse of cardinal_mono fails!*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   380
Goal "[| |i| < |j|;  Ord(i);  Ord(j) |] ==> i < j";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   381
by (rtac Ord_linear2 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   382
by (REPEAT_SOME assume_tac);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   383
by (etac (lt_trans2 RS lt_irrefl) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   384
by (etac cardinal_mono 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   385
qed "cardinal_lt_imp_lt";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   386
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   387
Goal "[| |i| < K;  Ord(i);  Card(K) |] ==> i < K";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   388
by (asm_simp_tac (simpset() addsimps 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   389
                  [cardinal_lt_imp_lt, Card_is_Ord, Card_cardinal_eq]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   390
qed "Card_lt_imp_lt";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   391
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   392
Goal "[| Ord(i);  Card(K) |] ==> (|i| < K) <-> (i < K)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   393
by (blast_tac (claset() addIs [Card_lt_imp_lt, Ord_cardinal_le RS lt_trans1]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   394
qed "Card_lt_iff";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   395
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   396
Goal "[| Ord(i);  Card(K) |] ==> (K le |i|) <-> (K le i)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   397
by (asm_simp_tac (simpset() addsimps 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   398
                  [Card_lt_iff, Card_is_Ord, Ord_cardinal, 
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   399
                   not_lt_iff_le RS iff_sym]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   400
qed "Card_le_iff";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   401
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   402
(*Can use AC or finiteness to discharge first premise*)
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   403
Goal "[| well_ord(B,r);  A lepoll B |] ==> |A| le |B|";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   404
by (res_inst_tac [("i","|A|"),("j","|B|")] Ord_linear_le 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   405
by (REPEAT_FIRST (ares_tac [Ord_cardinal, le_eqI]));
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   406
by (rtac (eqpollI RS cardinal_cong) 1 THEN assume_tac 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   407
by (rtac lepoll_trans 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   408
by (rtac (well_ord_cardinal_eqpoll RS eqpoll_sym RS eqpoll_imp_lepoll) 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   409
by (assume_tac 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   410
by (etac (le_imp_lepoll RS lepoll_trans) 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   411
by (rtac eqpoll_imp_lepoll 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   412
by (rewtac lepoll_def);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   413
by (etac exE 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   414
by (rtac well_ord_cardinal_eqpoll 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   415
by (etac well_ord_rvimage 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   416
by (assume_tac 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   417
qed "well_ord_lepoll_imp_Card_le";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   418
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   419
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   420
Goal "[| A lepoll i; Ord(i) |] ==> |A| le i";
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   421
by (rtac le_trans 1);
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   422
by (etac (well_ord_Memrel RS well_ord_lepoll_imp_Card_le) 1);
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   423
by (assume_tac 1);
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   424
by (etac Ord_cardinal_le 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   425
qed "lepoll_cardinal_le";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   426
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   427
Goal "[| A lepoll i; Ord(i) |] ==> |A| eqpoll A";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   428
by (blast_tac (claset() addIs [lepoll_cardinal_le, well_ord_Memrel,
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   429
                                well_ord_cardinal_eqpoll]
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   430
                        addSDs [lepoll_well_ord]) 1);
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   431
qed "lepoll_Ord_imp_eqpoll";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   432
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   433
Goalw [lesspoll_def]
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   434
     "[| A lesspoll i; Ord(i) |] ==> |A| eqpoll A";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   435
by (blast_tac (claset() addIs [lepoll_Ord_imp_eqpoll]) 1);
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   436
qed "lesspoll_imp_eqpoll";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   437
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   438
Goalw [lesspoll_def]
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   439
     "[| A lesspoll i; Ord(i) |] ==> |A| eqpoll A";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   440
by (blast_tac (claset() addIs [lepoll_Ord_imp_eqpoll]) 1);
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   441
qed "lesspoll_imp_eqpoll";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   442
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   443
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   444
(*** The finite cardinals ***)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   445
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   446
Goalw [lepoll_def, inj_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   447
 "[| cons(u,A) lepoll cons(v,B);  u~:A;  v~:B |] ==> A lepoll B";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   448
by Safe_tac;
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   449
by (res_inst_tac [("x", "lam x:A. if f`x=v then f`u else f`x")] exI 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   450
by (rtac CollectI 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   451
(*Proving it's in the function space A->B*)
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   452
by (rtac (if_type RS lam_type) 1);
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 9099
diff changeset
   453
by (blast_tac (claset() addDs [apply_funtype]) 1);
422968aeed49 fixed some weak elim rules
paulson
parents: 9099
diff changeset
   454
by (blast_tac (claset() addSEs [mem_irrefl] addDs [apply_funtype]) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   455
(*Proving it's injective*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   456
by (Asm_simp_tac 1);
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
   457
by (Blast_tac 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   458
qed "cons_lepoll_consD";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   459
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   460
Goal "[| cons(u,A) eqpoll cons(v,B);  u~:A;  v~:B |] ==> A eqpoll B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   461
by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   462
by (blast_tac (claset() addIs [cons_lepoll_consD]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   463
qed "cons_eqpoll_consD";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   464
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   465
(*Lemma suggested by Mike Fourman*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   466
Goalw [succ_def] "succ(m) lepoll succ(n) ==> m lepoll n";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   467
by (etac cons_lepoll_consD 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   468
by (REPEAT (rtac mem_not_refl 1));
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   469
qed "succ_lepoll_succD";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   470
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   471
Goal "m:nat ==> ALL n: nat. m lepoll n --> m le n";
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   472
by (nat_ind_tac "m" [] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   473
by (blast_tac (claset() addSIs [nat_0_le]) 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   474
by (rtac ballI 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   475
by (eres_inst_tac [("n","n")] natE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   476
by (asm_simp_tac (simpset() addsimps [lepoll_def, inj_def, 
6112
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
   477
				      succI1 RS Pi_empty2]) 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   478
by (blast_tac (claset() addSIs [succ_leI] addSDs [succ_lepoll_succD]) 1);
6112
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
   479
qed_spec_mp "nat_lepoll_imp_le";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   480
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   481
Goal "[| m:nat; n: nat |] ==> m eqpoll n <-> m = n";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   482
by (rtac iffI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   483
by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   484
by (blast_tac (claset() addIs [nat_lepoll_imp_le, le_anti_sym] 
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 9099
diff changeset
   485
                        addSEs [eqpollE]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   486
qed "nat_eqpoll_iff";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   487
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   488
(*The object of all this work: every natural number is a (finite) cardinal*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   489
Goalw [Card_def,cardinal_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   490
    "n: nat ==> Card(n)";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   491
by (stac Least_equality 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   492
by (REPEAT_FIRST (ares_tac [eqpoll_refl, nat_into_Ord, refl]));
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   493
by (asm_simp_tac (simpset() addsimps [lt_nat_in_nat RS nat_eqpoll_iff]) 1);
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   494
by (blast_tac (claset() addSEs [lt_irrefl]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   495
qed "nat_into_Card";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   496
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   497
(*Part of Kunen's Lemma 10.6*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   498
Goal "[| succ(n) lepoll n;  n:nat |] ==> P";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   499
by (rtac (nat_lepoll_imp_le RS lt_irrefl) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   500
by (REPEAT (ares_tac [nat_succI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   501
qed "succ_lepoll_natE";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   502
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   503
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   504
(** lepoll, lesspoll and natural numbers **)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   505
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   506
Goalw [lesspoll_def]
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   507
     "[| A lepoll m; m:nat |] ==> A lesspoll succ(m)";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   508
by (rtac conjI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   509
by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2,lepoll_trans)]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   510
by (rtac notI 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   511
by (dresolve_tac [eqpoll_sym RS eqpoll_imp_lepoll] 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   512
by (dtac lepoll_trans 1 THEN assume_tac 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   513
by (etac succ_lepoll_natE 1 THEN assume_tac 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   514
qed "lepoll_imp_lesspoll_succ";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   515
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   516
Goalw [lesspoll_def, lepoll_def, eqpoll_def, bij_def]
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   517
     "[| A lesspoll succ(m); m:nat |] ==> A lepoll m";
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3016
diff changeset
   518
by (Clarify_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   519
by (blast_tac (claset() addSIs [inj_not_surj_succ]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   520
qed "lesspoll_succ_imp_lepoll";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   521
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   522
Goal "m:nat ==> A lesspoll succ(m) <-> A lepoll m";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   523
by (blast_tac (claset() addSIs [lepoll_imp_lesspoll_succ, 
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   524
				lesspoll_succ_imp_lepoll]) 1);
1031
a53cbb4b06c5 Proved lesspoll_succ_iff.
lcp
parents: 845
diff changeset
   525
qed "lesspoll_succ_iff";
a53cbb4b06c5 Proved lesspoll_succ_iff.
lcp
parents: 845
diff changeset
   526
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   527
Goal "[| A lepoll succ(m);  m:nat |] ==> A lepoll m | A eqpoll succ(m)";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   528
by (rtac disjCI 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   529
by (rtac lesspoll_succ_imp_lepoll 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   530
by (assume_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   531
by (asm_simp_tac (simpset() addsimps [lesspoll_def]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   532
qed "lepoll_succ_disj";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   533
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   534
Goalw [lesspoll_def] "[| A lesspoll i; Ord(i) |] ==> |A| < i";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   535
by (Clarify_tac 1);
8183
344888de76c4 expandshort
paulson
parents: 8127
diff changeset
   536
by (ftac lepoll_cardinal_le 1);
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   537
by (assume_tac 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   538
by (blast_tac (claset() addIs [well_ord_Memrel,
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   539
			       well_ord_cardinal_eqpoll RS eqpoll_sym]
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   540
                        addDs [lepoll_well_ord] 
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   541
                        addSEs [leE]) 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   542
qed "lesspoll_cardinal_lt";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   543
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   544
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   545
(*** The first infinite cardinal: Omega, or nat ***)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   546
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   547
(*This implies Kunen's Lemma 10.6*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   548
Goal "[| n<i;  n:nat |] ==> ~ i lepoll n";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   549
by (rtac notI 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   550
by (rtac succ_lepoll_natE 1 THEN assume_tac 2);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   551
by (rtac lepoll_trans 1 THEN assume_tac 2);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   552
by (etac ltE 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   553
by (REPEAT (ares_tac [Ord_succ_subsetI RS subset_imp_lepoll] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   554
qed "lt_not_lepoll";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   555
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   556
Goal "[| Ord(i);  n:nat |] ==> i eqpoll n <-> i=n";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   557
by (rtac iffI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   558
by (asm_simp_tac (simpset() addsimps [eqpoll_refl]) 2);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   559
by (rtac Ord_linear_lt 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   560
by (REPEAT_SOME (eresolve_tac [asm_rl, nat_into_Ord]));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   561
by (etac (lt_nat_in_nat RS nat_eqpoll_iff RS iffD1) 1 THEN
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   562
    REPEAT (assume_tac 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   563
by (rtac (lt_not_lepoll RS notE) 1 THEN (REPEAT (assume_tac 1)));
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   564
by (etac eqpoll_imp_lepoll 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   565
qed "Ord_nat_eqpoll_iff";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   566
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   567
Goalw [Card_def,cardinal_def] "Card(nat)";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1791
diff changeset
   568
by (stac Least_equality 1);
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   569
by (REPEAT_FIRST (ares_tac [eqpoll_refl, Ord_nat, refl]));
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   570
by (etac ltE 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   571
by (asm_simp_tac (simpset() addsimps [eqpoll_iff, lt_not_lepoll, ltI]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   572
qed "Card_nat";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   573
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   574
(*Allows showing that |i| is a limit cardinal*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   575
Goal  "nat le i ==> nat le |i|";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   576
by (rtac (Card_nat RS Card_cardinal_eq RS subst) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   577
by (etac cardinal_mono 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   578
qed "nat_le_cardinal";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   579
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   580
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   581
(*** Towards Cardinal Arithmetic ***)
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   582
(** Congruence laws for successor, cardinal addition and multiplication **)
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   583
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   584
(*Congruence law for  cons  under equipollence*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   585
Goalw [lepoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   586
    "[| A lepoll B;  b ~: B |] ==> cons(a,A) lepoll cons(b,B)";
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   587
by Safe_tac;
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   588
by (res_inst_tac [("x", "lam y: cons(a,A). if y=a then b else f`y")] exI 1);
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   589
by (res_inst_tac [("d","%z. if z:B then converse(f)`z else a")] 
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   590
    lam_injective 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   591
by (asm_simp_tac (simpset() addsimps [inj_is_fun RS apply_type, cons_iff]
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   592
                        setloop etac consE') 1);
6176
707b6f9859d2 tidied, with left_inverse & right_inverse as default simprules
paulson
parents: 6112
diff changeset
   593
by (asm_simp_tac (simpset() addsimps [inj_is_fun RS apply_type]
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   594
                        setloop etac consE') 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   595
qed "cons_lepoll_cong";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   596
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   597
Goal "[| A eqpoll B;  a ~: A;  b ~: B |] ==> cons(a,A) eqpoll cons(b,B)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   598
by (asm_full_simp_tac (simpset() addsimps [eqpoll_iff, cons_lepoll_cong]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   599
qed "cons_eqpoll_cong";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   600
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   601
Goal "[| a ~: A;  b ~: B |] ==> \
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   602
\           cons(a,A) lepoll cons(b,B)  <->  A lepoll B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   603
by (blast_tac (claset() addIs [cons_lepoll_cong, cons_lepoll_consD]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   604
qed "cons_lepoll_cons_iff";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   605
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   606
Goal "[| a ~: A;  b ~: B |] ==> \
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   607
\           cons(a,A) eqpoll cons(b,B)  <->  A eqpoll B";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   608
by (blast_tac (claset() addIs [cons_eqpoll_cong, cons_eqpoll_consD]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   609
qed "cons_eqpoll_cons_iff";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   610
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   611
Goalw [succ_def] "{a} eqpoll 1";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   612
by (blast_tac (claset() addSIs [eqpoll_refl RS cons_eqpoll_cong]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   613
qed "singleton_eqpoll_1";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   614
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   615
Goal "|{a}| = 1";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   616
by (resolve_tac [singleton_eqpoll_1 RS cardinal_cong RS trans] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   617
by (simp_tac (simpset() addsimps [nat_into_Card RS Card_cardinal_eq]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   618
qed "cardinal_singleton";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   619
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   620
(*Congruence law for  succ  under equipollence*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   621
Goalw [succ_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   622
    "A eqpoll B ==> succ(A) eqpoll succ(B)";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   623
by (REPEAT (ares_tac [cons_eqpoll_cong, mem_not_refl] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   624
qed "succ_eqpoll_cong";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   625
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   626
(*Congruence law for + under equipollence*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   627
Goalw [eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   628
    "[| A eqpoll C;  B eqpoll D |] ==> A+B eqpoll C+D";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   629
by (blast_tac (claset() addSIs [sum_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   630
qed "sum_eqpoll_cong";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   631
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   632
(*Congruence law for * under equipollence*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   633
Goalw [eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   634
    "[| A eqpoll C;  B eqpoll D |] ==> A*B eqpoll C*D";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   635
by (blast_tac (claset() addSIs [prod_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   636
qed "prod_eqpoll_cong";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   637
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   638
Goalw [eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   639
    "[| f: inj(A,B);  A Int B = 0 |] ==> A Un (B - range(f)) eqpoll B";
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   640
by (rtac exI 1);
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   641
by (res_inst_tac [("c", "%x. if x:A then f`x else x"),
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   642
                  ("d", "%y. if y: range(f) then converse(f)`y else y")] 
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   643
    lam_bijective 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   644
by (blast_tac (claset() addSIs [if_type, inj_is_fun RS apply_type]) 1);
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   645
by (asm_simp_tac 
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   646
    (simpset() addsimps [inj_converse_fun RS apply_funtype]) 1);
6176
707b6f9859d2 tidied, with left_inverse & right_inverse as default simprules
paulson
parents: 6112
diff changeset
   647
by (asm_simp_tac (simpset() addsimps [inj_is_fun RS apply_rangeI]
571
0b03ce5b62f7 ZF/Cardinal: some results moved here from CardinalArith
lcp
parents: 522
diff changeset
   648
                        setloop etac UnE') 1);
6176
707b6f9859d2 tidied, with left_inverse & right_inverse as default simprules
paulson
parents: 6112
diff changeset
   649
by (asm_simp_tac (simpset() addsimps [inj_converse_fun RS apply_funtype]) 1);
5265
9d1d4c43c76d Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents: 5242
diff changeset
   650
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 571
diff changeset
   651
qed "inj_disjoint_eqpoll";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   652
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   653
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   654
(*** Lemmas by Krzysztof Grabczewski.  New proofs using cons_lepoll_cons.
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   655
     Could easily generalise from succ to cons. ***)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   656
1055
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   657
(*If A has at most n+1 elements and a:A then A-{a} has at most n.*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   658
Goalw [succ_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   659
      "[| a:A;  A lepoll succ(n) |] ==> A - {a} lepoll n";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   660
by (rtac cons_lepoll_consD 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   661
by (rtac mem_not_refl 3);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   662
by (eresolve_tac [cons_Diff RS ssubst] 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   663
by Safe_tac;
1055
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   664
qed "Diff_sing_lepoll";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   665
1055
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   666
(*If A has at least n+1 elements then A-{a} has at least n.*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   667
Goalw [succ_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   668
      "[| succ(n) lepoll A |] ==> n lepoll A - {a}";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   669
by (rtac cons_lepoll_consD 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   670
by (rtac mem_not_refl 2);
2875
6e3ccb94836c Mostly converted to blast_tac
paulson
parents: 2802
diff changeset
   671
by (Blast_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   672
by (blast_tac (claset() addIs [subset_imp_lepoll RSN (2, lepoll_trans)]) 1);
1055
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   673
qed "lepoll_Diff_sing";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   674
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   675
Goal "[| a:A; A eqpoll succ(n) |] ==> A - {a} eqpoll n";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   676
by (blast_tac (claset() addSIs [eqpollI] addSEs [eqpollE] 
1055
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   677
                    addIs [Diff_sing_lepoll,lepoll_Diff_sing]) 1);
67f5344605b7 Renamed diff_sing_lepoll, lepoll_diff_sing and diff_sing_eqpoll
lcp
parents: 1031
diff changeset
   678
qed "Diff_sing_eqpoll";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   679
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   680
Goal "[| A lepoll 1; a:A |] ==> A = {a}";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 6176
diff changeset
   681
by (ftac Diff_sing_lepoll 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   682
by (assume_tac 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   683
by (dtac lepoll_0_is_0 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   684
by (blast_tac (claset() addEs [equalityE]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   685
qed "lepoll_1_is_sing";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   686
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   687
Goalw [lepoll_def] "A Un B lepoll A+B";
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   688
by (res_inst_tac [("x",
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   689
		   "lam x: A Un B. if x:A then Inl(x) else Inr(x)")] exI 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   690
by (res_inst_tac [("d","%z. snd(z)")] lam_injective 1);
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   691
by (asm_full_simp_tac (simpset() addsimps [Inl_def, Inr_def]) 2);
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   692
by Auto_tac;
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   693
qed "Un_lepoll_sum";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   694
5284
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   695
Goal "[| well_ord(X,R); well_ord(Y,S) |] ==> EX T. well_ord(X Un Y, T)";
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   696
by (eresolve_tac [well_ord_radd RS (Un_lepoll_sum RS lepoll_well_ord)] 1);
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   697
by (assume_tac 1);
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   698
qed "well_ord_Un";
c77e9dd9bafc Tidying of AC, especially of AC16_WO4 using a locale
paulson
parents: 5268
diff changeset
   699
5242
3087dafb70ec Renamed equals0D to equals0E
paulson
parents: 5147
diff changeset
   700
(*Krzysztof Grabczewski*)
3087dafb70ec Renamed equals0D to equals0E
paulson
parents: 5147
diff changeset
   701
Goalw [eqpoll_def] "A Int B = 0 ==> A Un B eqpoll A + B";
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 5325
diff changeset
   702
by (res_inst_tac [("x","lam a:A Un B. if a:A then Inl(a) else Inr(a)")] exI 1);
5242
3087dafb70ec Renamed equals0D to equals0E
paulson
parents: 5147
diff changeset
   703
by (res_inst_tac [("d","%z. case(%x. x, %x. x, z)")] lam_bijective 1);
5265
9d1d4c43c76d Disjointness reasoning by AddEs [equals0E, sym RS equals0E]
paulson
parents: 5242
diff changeset
   704
by Auto_tac;
5242
3087dafb70ec Renamed equals0D to equals0E
paulson
parents: 5147
diff changeset
   705
qed "disj_Un_eqpoll_sum";
3087dafb70ec Renamed equals0D to equals0E
paulson
parents: 5147
diff changeset
   706
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   707
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   708
(*** Finite and infinite sets ***)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   709
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   710
Goalw [Finite_def] "Finite(0)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   711
by (blast_tac (claset() addSIs [eqpoll_refl, nat_0I]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   712
qed "Finite_0";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   713
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   714
Goalw [Finite_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   715
    "[| A lepoll n;  n:nat |] ==> Finite(A)";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   716
by (etac rev_mp 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   717
by (etac nat_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   718
by (blast_tac (claset() addSDs [lepoll_0_is_0] addSIs [eqpoll_refl,nat_0I]) 1);
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   719
by (blast_tac (claset() addSDs [lepoll_succ_disj]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   720
qed "lepoll_nat_imp_Finite";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   721
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   722
Goalw [Finite_def]
9099
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   723
     "A lesspoll nat ==> Finite(A)";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   724
by (blast_tac (claset() addDs [ltD, lesspoll_cardinal_lt,
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   725
	       lesspoll_imp_eqpoll RS eqpoll_sym]) 1);;
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   726
qed "lesspoll_nat_is_Finite";
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   727
f713ef362ad0 new theorems lepoll_Ord_imp_eqpoll, lesspoll_imp_eqpoll, lesspoll_nat_is_Finite
paulson
parents: 8183
diff changeset
   728
Goalw [Finite_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   729
     "[| Y lepoll X;  Finite(X) |] ==> Finite(Y)";
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   730
by (blast_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   731
    (claset() addSEs [eqpollE] 
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   732
             addIs [lepoll_trans RS 
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   733
		    rewrite_rule [Finite_def] lepoll_nat_imp_Finite]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   734
qed "lepoll_Finite";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   735
1623
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   736
bind_thm ("subset_Finite", subset_imp_lepoll RS lepoll_Finite);
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   737
2b8573c1b1c1 Ran expandshort
paulson
parents: 1609
diff changeset
   738
bind_thm ("Finite_Diff", Diff_subset RS subset_Finite);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   739
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   740
Goalw [Finite_def] "Finite(x) ==> Finite(cons(y,x))";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   741
by (excluded_middle_tac "y:x" 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   742
by (asm_simp_tac (simpset() addsimps [cons_absorb]) 2);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   743
by (etac bexE 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   744
by (rtac bexI 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   745
by (etac nat_succI 2);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   746
by (asm_simp_tac 
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   747
    (simpset() addsimps [succ_def, cons_eqpoll_cong, mem_not_refl]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   748
qed "Finite_cons";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   749
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   750
Goalw [succ_def] "Finite(x) ==> Finite(succ(x))";
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   751
by (etac Finite_cons 1);
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   752
qed "Finite_succ";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   753
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   754
Goalw [Finite_def] 
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   755
      "[| Ord(i);  ~ Finite(i) |] ==> nat le i";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   756
by (eresolve_tac [Ord_nat RSN (2,Ord_linear2)] 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   757
by (assume_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   758
by (blast_tac (claset() addSIs [eqpoll_refl] addSEs [ltE]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   759
qed "nat_le_infinite_Ord";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   760
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   761
Goalw [Finite_def, eqpoll_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   762
    "Finite(A) ==> EX r. well_ord(A,r)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   763
by (blast_tac (claset() addIs [well_ord_rvimage, bij_is_inj, well_ord_Memrel, 
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   764
			      nat_into_Ord]) 1);
1609
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   765
qed "Finite_imp_well_ord";
5324067d993f New lemmas for Mutilated Checkerboard
paulson
parents: 1461
diff changeset
   766
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   767
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   768
(*Krzysztof Grabczewski's proof that the converse of a finite, well-ordered
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   769
  set is well-ordered.  Proofs simplified by lcp. *)
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   770
5325
f7a5e06adea1 Yet more removal of "goal" commands, especially "goal ZF.thy", so ZF.thy
paulson
parents: 5284
diff changeset
   771
Goal "n:nat ==> wf[n](converse(Memrel(n)))";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   772
by (etac nat_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   773
by (blast_tac (claset() addIs [wf_onI]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   774
by (rtac wf_onI 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   775
by (asm_full_simp_tac (simpset() addsimps [wf_on_def, wf_def, Memrel_iff]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   776
by (excluded_middle_tac "x:Z" 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   777
by (dres_inst_tac [("x", "x")] bspec 2 THEN assume_tac 2);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   778
by (blast_tac (claset() addEs [mem_irrefl, mem_asym]) 2);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   779
by (dres_inst_tac [("x", "Z")] spec 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   780
by (Blast.depth_tac (claset()) 4 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   781
qed "nat_wf_on_converse_Memrel";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   782
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   783
Goal "n:nat ==> well_ord(n,converse(Memrel(n)))";
3894
wenzelm
parents: 3891
diff changeset
   784
by (forward_tac [transfer thy Ord_nat RS Ord_in_Ord RS well_ord_Memrel] 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   785
by (rewtac well_ord_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   786
by (blast_tac (claset() addSIs [tot_ord_converse, 
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   787
			       nat_wf_on_converse_Memrel]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   788
qed "nat_well_ord_converse_Memrel";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   789
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   790
Goal "[| well_ord(A,r);     \
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   791
\            well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) \
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   792
\         |] ==> well_ord(A,converse(r))";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   793
by (resolve_tac [well_ord_Int_iff RS iffD1] 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   794
by (forward_tac [ordermap_bij RS bij_is_inj RS well_ord_rvimage] 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   795
by (assume_tac 1);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   796
by (asm_full_simp_tac
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   797
    (simpset() addsimps [rvimage_converse, converse_Int, converse_prod, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1055
diff changeset
   798
                     ordertype_ord_iso RS ord_iso_rvimage_eq]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   799
qed "well_ord_converse";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   800
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5265
diff changeset
   801
Goal "[| well_ord(A,r);  A eqpoll n;  n:nat |] ==> ordertype(A,r)=n";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   802
by (rtac (Ord_ordertype RS Ord_nat_eqpoll_iff RS iffD1) 1 THEN 
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   803
    REPEAT (assume_tac 1));
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   804
by (rtac eqpoll_trans 1 THEN assume_tac 2);
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   805
by (rewtac eqpoll_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   806
by (blast_tac (claset() addSIs [ordermap_bij RS bij_converse_bij]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   807
qed "ordertype_eq_n";
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   808
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4152
diff changeset
   809
Goalw [Finite_def]
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5116
diff changeset
   810
    "[| Finite(A);  well_ord(A,r) |] ==> well_ord(A,converse(r))";
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   811
by (rtac well_ord_converse 1 THEN assume_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3894
diff changeset
   812
by (blast_tac (claset() addDs [ordertype_eq_n] 
3016
15763781afb0 Conversion to use blast_tac
paulson
parents: 2896
diff changeset
   813
                       addSIs [nat_well_ord_converse_Memrel]) 1);
833
ba386650df2c Proved cons_lepoll_consD, succ_lepoll_succD, cons_eqpoll_consD,
lcp
parents: 803
diff changeset
   814
qed "Finite_well_ord_converse";
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   815
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   816
Goalw [Finite_def] "n:nat ==> Finite(n)";
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   817
by (fast_tac (claset() addSIs [eqpoll_refl]) 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 7499
diff changeset
   818
qed "nat_into_Finite";