src/ZF/Cardinal.thy
author nipkow
Sun, 22 Feb 2009 17:25:28 +0100
changeset 30056 0a35bee25c20
parent 26056 6a0801279f4c
child 32960 69916a850301
permissions -rw-r--r--
added lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/Cardinal.thy
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
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
13328
703de709a64b better document preparation
paulson
parents: 13269
diff changeset
     6
*)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
     7
13328
703de709a64b better document preparation
paulson
parents: 13269
diff changeset
     8
header{*Cardinal Numbers Without the Axiom of Choice*}
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     9
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents: 24893
diff changeset
    10
theory Cardinal imports OrderType Finite Nat_ZF Sum begin
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    11
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    12
definition
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    13
  (*least ordinal operator*)
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    14
   Least    :: "(i=>o) => i"    (binder "LEAST " 10)  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    15
     "Least(P) == THE i. Ord(i) & P(i) & (ALL j. j<i --> ~P(j))"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    16
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    17
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    18
  eqpoll   :: "[i,i] => o"     (infixl "eqpoll" 50)  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    19
    "A eqpoll B == EX f. f: bij(A,B)"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    20
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    21
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    22
  lepoll   :: "[i,i] => o"     (infixl "lepoll" 50)  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    23
    "A lepoll B == EX f. f: inj(A,B)"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    24
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    25
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    26
  lesspoll :: "[i,i] => o"     (infixl "lesspoll" 50)  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    27
    "A lesspoll B == A lepoll B & ~(A eqpoll B)"
832
ad50a9e74eaf Added Krzysztof's constants lesspoll and Finite
lcp
parents: 753
diff changeset
    28
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    29
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    30
  cardinal :: "i=>i"           ("|_|")  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    31
    "|A| == LEAST i. i eqpoll A"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    32
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    33
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    34
  Finite   :: "i=>o"  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    35
    "Finite(A) == EX n:nat. A eqpoll n"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    36
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    37
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 21524
diff changeset
    38
  Card     :: "i=>o"  where
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    39
    "Card(i) == (i = |i|)"
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    40
21524
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    41
notation (xsymbols)
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    42
  eqpoll    (infixl "\<approx>" 50) and
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    43
  lepoll    (infixl "\<lesssim>" 50) and
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    44
  lesspoll  (infixl "\<prec>" 50) and
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    45
  Least     (binder "\<mu>" 10)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    46
21524
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    47
notation (HTML output)
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    48
  eqpoll    (infixl "\<approx>" 50) and
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    49
  Least     (binder "\<mu>" 10)
7843e2fd14a9 updated (binder) syntax/notation;
wenzelm
parents: 16417
diff changeset
    50
14565
c6dc17aab88a use more symbols in HTML output
kleing
parents: 14153
diff changeset
    51
13357
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    52
subsection{*The Schroeder-Bernstein Theorem*}
6f54e992777e Removal of mono.thy
paulson
parents: 13356
diff changeset
    53
text{*See Davey and Priestly, page 106*}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    54
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    55
(** Lemma: Banach's Decomposition Theorem **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    56
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    57
lemma decomp_bnd_mono: "bnd_mono(X, %W. X - g``(Y - f``W))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    58
by (rule bnd_monoI, blast+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    59
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    60
lemma Banach_last_equation:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    61
    "g: Y->X
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    62
     ==> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) =        
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    63
	 X - lfp(X, %W. X - g``(Y - f``W))" 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    64
apply (rule_tac P = "%u. ?v = X-u" 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    65
       in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    66
apply (simp add: double_complement  fun_is_rel [THEN image_subset])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    67
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    68
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    69
lemma decomposition:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    70
     "[| f: X->Y;  g: Y->X |] ==>    
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    71
      EX XA XB YA YB. (XA Int XB = 0) & (XA Un XB = X) &     
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    72
                      (YA Int YB = 0) & (YA Un YB = Y) &     
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    73
                      f``XA=YA & g``YB=XB"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    74
apply (intro exI conjI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    75
apply (rule_tac [6] Banach_last_equation)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    76
apply (rule_tac [5] refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    77
apply (assumption | 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    78
       rule  Diff_disjoint Diff_partition fun_is_rel image_subset lfp_subset)+
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    79
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    80
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    81
lemma schroeder_bernstein:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    82
    "[| f: inj(X,Y);  g: inj(Y,X) |] ==> EX h. h: bij(X,Y)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    83
apply (insert decomposition [of f X Y g]) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    84
apply (simp add: inj_is_fun)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    85
apply (blast intro!: restrict_bij bij_disjoint_Un intro: bij_converse_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    86
(* The instantiation of exI to "restrict(f,XA) Un converse(restrict(g,YB))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    87
   is forced by the context!! *)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    88
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    89
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    90
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    91
(** Equipollence is an equivalence relation **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    92
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    93
lemma bij_imp_eqpoll: "f: bij(A,B) ==> A \<approx> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    94
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    95
apply (erule exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    96
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    97
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    98
(*A eqpoll A*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
    99
lemmas eqpoll_refl = id_bij [THEN bij_imp_eqpoll, standard, simp]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   100
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   101
lemma eqpoll_sym: "X \<approx> Y ==> Y \<approx> X"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   102
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   103
apply (blast intro: bij_converse_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   104
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   105
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   106
lemma eqpoll_trans: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   107
    "[| X \<approx> Y;  Y \<approx> Z |] ==> X \<approx> Z"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   108
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   109
apply (blast intro: comp_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   110
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   111
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   112
(** Le-pollence is a partial ordering **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   113
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   114
lemma subset_imp_lepoll: "X<=Y ==> X \<lesssim> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   115
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   116
apply (rule exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   117
apply (erule id_subset_inj)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   118
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   119
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   120
lemmas lepoll_refl = subset_refl [THEN subset_imp_lepoll, standard, simp]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   121
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   122
lemmas le_imp_lepoll = le_imp_subset [THEN subset_imp_lepoll, standard]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   123
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   124
lemma eqpoll_imp_lepoll: "X \<approx> Y ==> X \<lesssim> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   125
by (unfold eqpoll_def bij_def lepoll_def, blast)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   126
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   127
lemma lepoll_trans: "[| X \<lesssim> Y;  Y \<lesssim> Z |] ==> X \<lesssim> Z"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   128
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   129
apply (blast intro: comp_inj)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   130
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   131
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   132
(*Asymmetry law*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   133
lemma eqpollI: "[| X \<lesssim> Y;  Y \<lesssim> X |] ==> X \<approx> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   134
apply (unfold lepoll_def eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   135
apply (elim exE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   136
apply (rule schroeder_bernstein, assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   137
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   138
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   139
lemma eqpollE:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   140
    "[| X \<approx> Y; [| X \<lesssim> Y; Y \<lesssim> X |] ==> P |] ==> P"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   141
by (blast intro: eqpoll_imp_lepoll eqpoll_sym) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   142
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   143
lemma eqpoll_iff: "X \<approx> Y <-> X \<lesssim> Y & Y \<lesssim> X"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   144
by (blast intro: eqpollI elim!: eqpollE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   145
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   146
lemma lepoll_0_is_0: "A \<lesssim> 0 ==> A = 0"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   147
apply (unfold lepoll_def inj_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   148
apply (blast dest: apply_type)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   149
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   150
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   151
(*0 \<lesssim> Y*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   152
lemmas empty_lepollI = empty_subsetI [THEN subset_imp_lepoll, standard]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   153
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   154
lemma lepoll_0_iff: "A \<lesssim> 0 <-> A=0"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   155
by (blast intro: lepoll_0_is_0 lepoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   156
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   157
lemma Un_lepoll_Un: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   158
    "[| A \<lesssim> B; C \<lesssim> D; B Int D = 0 |] ==> A Un C \<lesssim> B Un D"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   159
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   160
apply (blast intro: inj_disjoint_Un)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   161
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   162
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   163
(*A eqpoll 0 ==> A=0*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   164
lemmas eqpoll_0_is_0 = eqpoll_imp_lepoll [THEN lepoll_0_is_0, standard]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   165
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   166
lemma eqpoll_0_iff: "A \<approx> 0 <-> A=0"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   167
by (blast intro: eqpoll_0_is_0 eqpoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   168
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   169
lemma eqpoll_disjoint_Un: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   170
    "[| A \<approx> B;  C \<approx> D;  A Int C = 0;  B Int D = 0 |]   
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   171
     ==> A Un C \<approx> B Un D"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   172
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   173
apply (blast intro: bij_disjoint_Un)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   174
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   175
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   176
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   177
subsection{*lesspoll: contributions by Krzysztof Grabczewski *}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   178
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   179
lemma lesspoll_not_refl: "~ (i \<prec> i)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   180
by (simp add: lesspoll_def) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   181
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   182
lemma lesspoll_irrefl [elim!]: "i \<prec> i ==> P"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   183
by (simp add: lesspoll_def) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   184
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   185
lemma lesspoll_imp_lepoll: "A \<prec> B ==> A \<lesssim> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   186
by (unfold lesspoll_def, blast)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   187
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   188
lemma lepoll_well_ord: "[| A \<lesssim> B; well_ord(B,r) |] ==> EX s. well_ord(A,s)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   189
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   190
apply (blast intro: well_ord_rvimage)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   191
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   192
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   193
lemma lepoll_iff_leqpoll: "A \<lesssim> B <-> A \<prec> B | A \<approx> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   194
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   195
apply (blast intro!: eqpollI elim!: eqpollE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   196
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   197
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   198
lemma inj_not_surj_succ: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   199
  "[| f : inj(A, succ(m)); f ~: surj(A, succ(m)) |] ==> EX f. f:inj(A,m)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   200
apply (unfold inj_def surj_def) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   201
apply (safe del: succE) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   202
apply (erule swap, rule exI) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   203
apply (rule_tac a = "lam z:A. if f`z=m then y else f`z" in CollectI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   204
txt{*the typing condition*}
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   205
 apply (best intro!: if_type [THEN lam_type] elim: apply_funtype [THEN succE])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   206
txt{*Proving it's injective*}
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   207
apply simp
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   208
apply blast 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   209
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   210
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   211
(** Variations on transitivity **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   212
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   213
lemma lesspoll_trans: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   214
      "[| X \<prec> Y; Y \<prec> Z |] ==> X \<prec> Z"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   215
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   216
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   217
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   218
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   219
lemma lesspoll_trans1: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   220
      "[| X \<lesssim> Y; Y \<prec> Z |] ==> X \<prec> Z"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   221
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   222
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   223
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   224
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   225
lemma lesspoll_trans2: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   226
      "[| X \<prec> Y; Y \<lesssim> Z |] ==> X \<prec> Z"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   227
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   228
apply (blast elim!: eqpollE intro: eqpollI lepoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   229
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   230
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   231
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   232
(** LEAST -- the least number operator [from HOL/Univ.ML] **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   233
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   234
lemma Least_equality: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   235
    "[| P(i);  Ord(i);  !!x. x<i ==> ~P(x) |] ==> (LEAST x. P(x)) = i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   236
apply (unfold Least_def) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   237
apply (rule the_equality, blast)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   238
apply (elim conjE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   239
apply (erule Ord_linear_lt, assumption, blast+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   240
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   241
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   242
lemma LeastI: "[| P(i);  Ord(i) |] ==> P(LEAST x. P(x))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   243
apply (erule rev_mp)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   244
apply (erule_tac i=i in trans_induct) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   245
apply (rule impI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   246
apply (rule classical)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   247
apply (blast intro: Least_equality [THEN ssubst]  elim!: ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   248
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   249
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   250
(*Proof is almost identical to the one above!*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   251
lemma Least_le: "[| P(i);  Ord(i) |] ==> (LEAST x. P(x)) le i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   252
apply (erule rev_mp)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   253
apply (erule_tac i=i in trans_induct) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   254
apply (rule impI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   255
apply (rule classical)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   256
apply (subst Least_equality, assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   257
apply (erule_tac [2] le_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   258
apply (blast elim: ltE intro: leI ltI lt_trans1)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   259
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   260
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   261
(*LEAST really is the smallest*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   262
lemma less_LeastE: "[| P(i);  i < (LEAST x. P(x)) |] ==> Q"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   263
apply (rule Least_le [THEN [2] lt_trans2, THEN lt_irrefl], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   264
apply (simp add: lt_Ord) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   265
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   266
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   267
(*Easier to apply than LeastI: conclusion has only one occurrence of P*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   268
lemma LeastI2:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   269
    "[| P(i);  Ord(i);  !!j. P(j) ==> Q(j) |] ==> Q(LEAST j. P(j))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   270
by (blast intro: LeastI ) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   271
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   272
(*If there is no such P then LEAST is vacuously 0*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   273
lemma Least_0: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   274
    "[| ~ (EX i. Ord(i) & P(i)) |] ==> (LEAST x. P(x)) = 0"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   275
apply (unfold Least_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   276
apply (rule the_0, blast)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   277
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   278
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   279
lemma Ord_Least [intro,simp,TC]: "Ord(LEAST x. P(x))"
14153
76a6ba67bd15 new case_tac
paulson
parents: 14076
diff changeset
   280
apply (case_tac "\<exists>i. Ord(i) & P(i)")  
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   281
apply safe
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   282
apply (rule Least_le [THEN ltE])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   283
prefer 3 apply assumption+
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   284
apply (erule Least_0 [THEN ssubst])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   285
apply (rule Ord_0)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   286
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   287
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   288
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   289
(** Basic properties of cardinals **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   290
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   291
(*Not needed for simplification, but helpful below*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   292
lemma Least_cong:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   293
     "(!!y. P(y) <-> Q(y)) ==> (LEAST x. P(x)) = (LEAST x. Q(x))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   294
by simp
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   295
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   296
(*Need AC to get X \<lesssim> Y ==> |X| le |Y|;  see well_ord_lepoll_imp_Card_le
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   297
  Converse also requires AC, but see well_ord_cardinal_eqE*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   298
lemma cardinal_cong: "X \<approx> Y ==> |X| = |Y|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   299
apply (unfold eqpoll_def cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   300
apply (rule Least_cong)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   301
apply (blast intro: comp_bij bij_converse_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   302
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   303
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   304
(*Under AC, the premise becomes trivial; one consequence is ||A|| = |A|*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   305
lemma well_ord_cardinal_eqpoll: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   306
    "well_ord(A,r) ==> |A| \<approx> A"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   307
apply (unfold cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   308
apply (rule LeastI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   309
apply (erule_tac [2] Ord_ordertype)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   310
apply (erule ordermap_bij [THEN bij_converse_bij, THEN bij_imp_eqpoll])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   311
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   312
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   313
(* Ord(A) ==> |A| \<approx> A *)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   314
lemmas Ord_cardinal_eqpoll = well_ord_Memrel [THEN well_ord_cardinal_eqpoll]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   315
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   316
lemma well_ord_cardinal_eqE:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   317
     "[| well_ord(X,r);  well_ord(Y,s);  |X| = |Y| |] ==> X \<approx> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   318
apply (rule eqpoll_sym [THEN eqpoll_trans])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   319
apply (erule well_ord_cardinal_eqpoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   320
apply (simp (no_asm_simp) add: well_ord_cardinal_eqpoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   321
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   322
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   323
lemma well_ord_cardinal_eqpoll_iff:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   324
     "[| well_ord(X,r);  well_ord(Y,s) |] ==> |X| = |Y| <-> X \<approx> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   325
by (blast intro: cardinal_cong well_ord_cardinal_eqE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   326
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   327
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   328
(** Observations from Kunen, page 28 **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   329
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   330
lemma Ord_cardinal_le: "Ord(i) ==> |i| le i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   331
apply (unfold cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   332
apply (erule eqpoll_refl [THEN Least_le])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   333
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   334
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   335
lemma Card_cardinal_eq: "Card(K) ==> |K| = K"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   336
apply (unfold Card_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   337
apply (erule sym)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   338
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   339
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   340
(* Could replace the  ~(j \<approx> i)  by  ~(i \<lesssim> j) *)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   341
lemma CardI: "[| Ord(i);  !!j. j<i ==> ~(j \<approx> i) |] ==> Card(i)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   342
apply (unfold Card_def cardinal_def) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   343
apply (subst Least_equality)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   344
apply (blast intro: eqpoll_refl )+
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   345
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   346
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   347
lemma Card_is_Ord: "Card(i) ==> Ord(i)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   348
apply (unfold Card_def cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   349
apply (erule ssubst)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   350
apply (rule Ord_Least)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   351
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   352
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   353
lemma Card_cardinal_le: "Card(K) ==> K le |K|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   354
apply (simp (no_asm_simp) add: Card_is_Ord Card_cardinal_eq)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   355
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   356
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   357
lemma Ord_cardinal [simp,intro!]: "Ord(|A|)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   358
apply (unfold cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   359
apply (rule Ord_Least)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   360
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   361
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   362
(*The cardinals are the initial ordinals*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   363
lemma Card_iff_initial: "Card(K) <-> Ord(K) & (ALL j. j<K --> ~ j \<approx> K)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   364
apply (safe intro!: CardI Card_is_Ord)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   365
 prefer 2 apply blast
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   366
apply (unfold Card_def cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   367
apply (rule less_LeastE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   368
apply (erule_tac [2] subst, assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   369
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   370
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   371
lemma lt_Card_imp_lesspoll: "[| Card(a); i<a |] ==> i \<prec> a"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   372
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   373
apply (drule Card_iff_initial [THEN iffD1])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   374
apply (blast intro!: leI [THEN le_imp_lepoll])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   375
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   376
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   377
lemma Card_0: "Card(0)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   378
apply (rule Ord_0 [THEN CardI])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   379
apply (blast elim!: ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   380
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   381
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   382
lemma Card_Un: "[| Card(K);  Card(L) |] ==> Card(K Un L)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   383
apply (rule Ord_linear_le [of K L])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   384
apply (simp_all add: subset_Un_iff [THEN iffD1]  Card_is_Ord le_imp_subset
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   385
                     subset_Un_iff2 [THEN iffD1])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   386
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   387
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   388
(*Infinite unions of cardinals?  See Devlin, Lemma 6.7, page 98*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   389
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   390
lemma Card_cardinal: "Card(|A|)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   391
apply (unfold cardinal_def)
14153
76a6ba67bd15 new case_tac
paulson
parents: 14076
diff changeset
   392
apply (case_tac "EX i. Ord (i) & i \<approx> A")
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   393
 txt{*degenerate case*}
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   394
 prefer 2 apply (erule Least_0 [THEN ssubst], rule Card_0)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   395
txt{*real case: A is isomorphic to some ordinal*}
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   396
apply (rule Ord_Least [THEN CardI], safe)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   397
apply (rule less_LeastE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   398
prefer 2 apply assumption
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   399
apply (erule eqpoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   400
apply (best intro: LeastI ) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   401
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   402
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   403
(*Kunen's Lemma 10.5*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   404
lemma cardinal_eq_lemma: "[| |i| le j;  j le i |] ==> |j| = |i|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   405
apply (rule eqpollI [THEN cardinal_cong])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   406
apply (erule le_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   407
apply (rule lepoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   408
apply (erule_tac [2] le_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   409
apply (rule eqpoll_sym [THEN eqpoll_imp_lepoll])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   410
apply (rule Ord_cardinal_eqpoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   411
apply (elim ltE Ord_succD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   412
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   413
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   414
lemma cardinal_mono: "i le j ==> |i| le |j|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   415
apply (rule_tac i = "|i|" and j = "|j|" in Ord_linear_le)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   416
apply (safe intro!: Ord_cardinal le_eqI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   417
apply (rule cardinal_eq_lemma)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   418
prefer 2 apply assumption
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   419
apply (erule le_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   420
apply (erule ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   421
apply (erule Ord_cardinal_le)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   422
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   423
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   424
(*Since we have |succ(nat)| le |nat|, the converse of cardinal_mono fails!*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   425
lemma cardinal_lt_imp_lt: "[| |i| < |j|;  Ord(i);  Ord(j) |] ==> i < j"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   426
apply (rule Ord_linear2 [of i j], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   427
apply (erule lt_trans2 [THEN lt_irrefl])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   428
apply (erule cardinal_mono)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   429
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   430
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   431
lemma Card_lt_imp_lt: "[| |i| < K;  Ord(i);  Card(K) |] ==> i < K"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   432
apply (simp (no_asm_simp) add: cardinal_lt_imp_lt Card_is_Ord Card_cardinal_eq)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   433
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   434
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   435
lemma Card_lt_iff: "[| Ord(i);  Card(K) |] ==> (|i| < K) <-> (i < K)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   436
by (blast intro: Card_lt_imp_lt Ord_cardinal_le [THEN lt_trans1])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   437
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   438
lemma Card_le_iff: "[| Ord(i);  Card(K) |] ==> (K le |i|) <-> (K le i)"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13244
diff changeset
   439
by (simp add: Card_lt_iff Card_is_Ord Ord_cardinal not_lt_iff_le [THEN iff_sym])
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   440
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   441
(*Can use AC or finiteness to discharge first premise*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   442
lemma well_ord_lepoll_imp_Card_le:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   443
     "[| well_ord(B,r);  A \<lesssim> B |] ==> |A| le |B|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   444
apply (rule_tac i = "|A|" and j = "|B|" in Ord_linear_le)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   445
apply (safe intro!: Ord_cardinal le_eqI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   446
apply (rule eqpollI [THEN cardinal_cong], assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   447
apply (rule lepoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   448
apply (rule well_ord_cardinal_eqpoll [THEN eqpoll_sym, THEN eqpoll_imp_lepoll], assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   449
apply (erule le_imp_lepoll [THEN lepoll_trans])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   450
apply (rule eqpoll_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   451
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   452
apply (erule exE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   453
apply (rule well_ord_cardinal_eqpoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   454
apply (erule well_ord_rvimage, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   455
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   456
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   457
lemma lepoll_cardinal_le: "[| A \<lesssim> i; Ord(i) |] ==> |A| le i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   458
apply (rule le_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   459
apply (erule well_ord_Memrel [THEN well_ord_lepoll_imp_Card_le], assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   460
apply (erule Ord_cardinal_le)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   461
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   462
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   463
lemma lepoll_Ord_imp_eqpoll: "[| A \<lesssim> i; Ord(i) |] ==> |A| \<approx> A"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   464
by (blast intro: lepoll_cardinal_le well_ord_Memrel well_ord_cardinal_eqpoll dest!: lepoll_well_ord)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   465
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   466
lemma lesspoll_imp_eqpoll: "[| A \<prec> i; Ord(i) |] ==> |A| \<approx> A"
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   467
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   468
apply (blast intro: lepoll_Ord_imp_eqpoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   469
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   470
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   471
lemma cardinal_subset_Ord: "[|A<=i; Ord(i)|] ==> |A| <= i"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   472
apply (drule subset_imp_lepoll [THEN lepoll_cardinal_le])
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   473
apply (auto simp add: lt_def)
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   474
apply (blast intro: Ord_trans)
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
   475
done
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   476
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   477
subsection{*The finite cardinals *}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   478
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   479
lemma cons_lepoll_consD: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   480
 "[| cons(u,A) \<lesssim> cons(v,B);  u~:A;  v~:B |] ==> A \<lesssim> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   481
apply (unfold lepoll_def inj_def, safe)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   482
apply (rule_tac x = "lam x:A. if f`x=v then f`u else f`x" in exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   483
apply (rule CollectI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   484
(*Proving it's in the function space A->B*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   485
apply (rule if_type [THEN lam_type])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   486
apply (blast dest: apply_funtype)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   487
apply (blast elim!: mem_irrefl dest: apply_funtype)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   488
(*Proving it's injective*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   489
apply (simp (no_asm_simp))
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   490
apply blast
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   491
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   492
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   493
lemma cons_eqpoll_consD: "[| cons(u,A) \<approx> cons(v,B);  u~:A;  v~:B |] ==> A \<approx> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   494
apply (simp add: eqpoll_iff)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   495
apply (blast intro: cons_lepoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   496
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   497
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   498
(*Lemma suggested by Mike Fourman*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   499
lemma succ_lepoll_succD: "succ(m) \<lesssim> succ(n) ==> m \<lesssim> n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   500
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   501
apply (erule cons_lepoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   502
apply (rule mem_not_refl)+
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   503
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   504
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   505
lemma nat_lepoll_imp_le [rule_format]:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   506
     "m:nat ==> ALL n: nat. m \<lesssim> n --> m le n"
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   507
apply (induct_tac m)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   508
apply (blast intro!: nat_0_le)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   509
apply (rule ballI)
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
   510
apply (erule_tac n = n in natE)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   511
apply (simp (no_asm_simp) add: lepoll_def inj_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   512
apply (blast intro!: succ_leI dest!: succ_lepoll_succD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   513
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   514
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   515
lemma nat_eqpoll_iff: "[| m:nat; n: nat |] ==> m \<approx> n <-> m = n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   516
apply (rule iffI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   517
apply (blast intro: nat_lepoll_imp_le le_anti_sym elim!: eqpollE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   518
apply (simp add: eqpoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   519
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   520
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   521
(*The object of all this work: every natural number is a (finite) cardinal*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   522
lemma nat_into_Card: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   523
    "n: nat ==> Card(n)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   524
apply (unfold Card_def cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   525
apply (subst Least_equality)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   526
apply (rule eqpoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   527
apply (erule nat_into_Ord) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   528
apply (simp (no_asm_simp) add: lt_nat_in_nat [THEN nat_eqpoll_iff])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   529
apply (blast elim!: lt_irrefl)+
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   530
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   531
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   532
lemmas cardinal_0 = nat_0I [THEN nat_into_Card, THEN Card_cardinal_eq, iff]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   533
lemmas cardinal_1 = nat_1I [THEN nat_into_Card, THEN Card_cardinal_eq, iff]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   534
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   535
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   536
(*Part of Kunen's Lemma 10.6*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   537
lemma succ_lepoll_natE: "[| succ(n) \<lesssim> n;  n:nat |] ==> P"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   538
by (rule nat_lepoll_imp_le [THEN lt_irrefl], auto)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   539
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   540
lemma n_lesspoll_nat: "n \<in> nat ==> n \<prec> nat"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   541
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   542
apply (fast elim!: Ord_nat [THEN [2] ltI [THEN leI, THEN le_imp_lepoll]]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   543
                   eqpoll_sym [THEN eqpoll_imp_lepoll] 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   544
    intro: Ord_nat [THEN [2] nat_succI [THEN ltI], THEN leI, 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   545
                 THEN le_imp_lepoll, THEN lepoll_trans, THEN succ_lepoll_natE])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   546
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   547
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   548
lemma nat_lepoll_imp_ex_eqpoll_n: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   549
     "[| n \<in> nat;  nat \<lesssim> X |] ==> \<exists>Y. Y \<subseteq> X & n \<approx> Y"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   550
apply (unfold lepoll_def eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   551
apply (fast del: subsetI subsetCE
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   552
            intro!: subset_SIs
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   553
            dest!: Ord_nat [THEN [2] OrdmemD, THEN [2] restrict_inj]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   554
            elim!: restrict_bij 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   555
                   inj_is_fun [THEN fun_is_rel, THEN image_subset])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   556
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   557
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   558
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   559
(** lepoll, \<prec> and natural numbers **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   560
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   561
lemma lepoll_imp_lesspoll_succ: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   562
     "[| A \<lesssim> m; m:nat |] ==> A \<prec> succ(m)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   563
apply (unfold lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   564
apply (rule conjI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   565
apply (blast intro: subset_imp_lepoll [THEN [2] lepoll_trans])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   566
apply (rule notI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   567
apply (drule eqpoll_sym [THEN eqpoll_imp_lepoll])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   568
apply (drule lepoll_trans, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   569
apply (erule succ_lepoll_natE, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   570
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   571
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   572
lemma lesspoll_succ_imp_lepoll: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   573
     "[| A \<prec> succ(m); m:nat |] ==> A \<lesssim> m"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   574
apply (unfold lesspoll_def lepoll_def eqpoll_def bij_def, clarify)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   575
apply (blast intro!: inj_not_surj_succ)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   576
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   577
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   578
lemma lesspoll_succ_iff: "m:nat ==> A \<prec> succ(m) <-> A \<lesssim> m"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   579
by (blast intro!: lepoll_imp_lesspoll_succ lesspoll_succ_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   580
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   581
lemma lepoll_succ_disj: "[| A \<lesssim> succ(m);  m:nat |] ==> A \<lesssim> m | A \<approx> succ(m)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   582
apply (rule disjCI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   583
apply (rule lesspoll_succ_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   584
prefer 2 apply assumption
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   585
apply (simp (no_asm_simp) add: lesspoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   586
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   587
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   588
lemma lesspoll_cardinal_lt: "[| A \<prec> i; Ord(i) |] ==> |A| < i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   589
apply (unfold lesspoll_def, clarify)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   590
apply (frule lepoll_cardinal_le, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   591
apply (blast intro: well_ord_Memrel well_ord_cardinal_eqpoll [THEN eqpoll_sym]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   592
             dest: lepoll_well_ord  elim!: leE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   593
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   594
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   595
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   596
subsection{*The first infinite cardinal: Omega, or nat *}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   597
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   598
(*This implies Kunen's Lemma 10.6*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   599
lemma lt_not_lepoll: "[| n<i;  n:nat |] ==> ~ i \<lesssim> n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   600
apply (rule notI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   601
apply (rule succ_lepoll_natE [of n])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   602
apply (rule lepoll_trans [of _ i])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   603
apply (erule ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   604
apply (rule Ord_succ_subsetI [THEN subset_imp_lepoll], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   605
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   606
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   607
lemma Ord_nat_eqpoll_iff: "[| Ord(i);  n:nat |] ==> i \<approx> n <-> i=n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   608
apply (rule iffI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   609
 prefer 2 apply (simp add: eqpoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   610
apply (rule Ord_linear_lt [of i n])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   611
apply (simp_all add: nat_into_Ord)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   612
apply (erule lt_nat_in_nat [THEN nat_eqpoll_iff, THEN iffD1], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   613
apply (rule lt_not_lepoll [THEN notE], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   614
apply (erule eqpoll_imp_lepoll)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   615
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   616
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   617
lemma Card_nat: "Card(nat)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   618
apply (unfold Card_def cardinal_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   619
apply (subst Least_equality)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   620
apply (rule eqpoll_refl) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   621
apply (rule Ord_nat) 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   622
apply (erule ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   623
apply (simp_all add: eqpoll_iff lt_not_lepoll ltI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   624
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   625
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   626
(*Allows showing that |i| is a limit cardinal*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   627
lemma nat_le_cardinal: "nat le i ==> nat le |i|"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   628
apply (rule Card_nat [THEN Card_cardinal_eq, THEN subst])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   629
apply (erule cardinal_mono)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   630
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   631
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   632
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   633
subsection{*Towards Cardinal Arithmetic *}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   634
(** Congruence laws for successor, cardinal addition and multiplication **)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   635
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   636
(*Congruence law for  cons  under equipollence*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   637
lemma cons_lepoll_cong: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   638
    "[| A \<lesssim> B;  b ~: B |] ==> cons(a,A) \<lesssim> cons(b,B)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   639
apply (unfold lepoll_def, safe)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   640
apply (rule_tac x = "lam y: cons (a,A) . if y=a then b else f`y" in exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   641
apply (rule_tac d = "%z. if z:B then converse (f) `z else a" in lam_injective)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   642
apply (safe elim!: consE') 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   643
   apply simp_all
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   644
apply (blast intro: inj_is_fun [THEN apply_type])+ 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   645
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   646
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   647
lemma cons_eqpoll_cong:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   648
     "[| A \<approx> B;  a ~: A;  b ~: B |] ==> cons(a,A) \<approx> cons(b,B)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   649
by (simp add: eqpoll_iff cons_lepoll_cong)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   650
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   651
lemma cons_lepoll_cons_iff:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   652
     "[| a ~: A;  b ~: B |] ==> cons(a,A) \<lesssim> cons(b,B)  <->  A \<lesssim> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   653
by (blast intro: cons_lepoll_cong cons_lepoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   654
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   655
lemma cons_eqpoll_cons_iff:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   656
     "[| a ~: A;  b ~: B |] ==> cons(a,A) \<approx> cons(b,B)  <->  A \<approx> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   657
by (blast intro: cons_eqpoll_cong cons_eqpoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   658
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   659
lemma singleton_eqpoll_1: "{a} \<approx> 1"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   660
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   661
apply (blast intro!: eqpoll_refl [THEN cons_eqpoll_cong])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   662
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   663
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   664
lemma cardinal_singleton: "|{a}| = 1"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   665
apply (rule singleton_eqpoll_1 [THEN cardinal_cong, THEN trans])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   666
apply (simp (no_asm) add: nat_into_Card [THEN Card_cardinal_eq])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   667
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   668
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   669
lemma not_0_is_lepoll_1: "A ~= 0 ==> 1 \<lesssim> A"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   670
apply (erule not_emptyE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   671
apply (rule_tac a = "cons (x, A-{x}) " in subst)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   672
apply (rule_tac [2] a = "cons(0,0)" and P= "%y. y \<lesssim> cons (x, A-{x})" in subst)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   673
prefer 3 apply (blast intro: cons_lepoll_cong subset_imp_lepoll, auto)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   674
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   675
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   676
(*Congruence law for  succ  under equipollence*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   677
lemma succ_eqpoll_cong: "A \<approx> B ==> succ(A) \<approx> succ(B)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   678
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   679
apply (simp add: cons_eqpoll_cong mem_not_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   680
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   681
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   682
(*Congruence law for + under equipollence*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   683
lemma sum_eqpoll_cong: "[| A \<approx> C;  B \<approx> D |] ==> A+B \<approx> C+D"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   684
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   685
apply (blast intro!: sum_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   686
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   687
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   688
(*Congruence law for * under equipollence*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   689
lemma prod_eqpoll_cong: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   690
    "[| A \<approx> C;  B \<approx> D |] ==> A*B \<approx> C*D"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   691
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   692
apply (blast intro!: prod_bij)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   693
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   694
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   695
lemma inj_disjoint_eqpoll: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   696
    "[| f: inj(A,B);  A Int B = 0 |] ==> A Un (B - range(f)) \<approx> B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   697
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   698
apply (rule exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   699
apply (rule_tac c = "%x. if x:A then f`x else x" 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   700
            and d = "%y. if y: range (f) then converse (f) `y else y" 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   701
       in lam_bijective)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   702
apply (blast intro!: if_type inj_is_fun [THEN apply_type])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   703
apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   704
apply (safe elim!: UnE') 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   705
   apply (simp_all add: inj_is_fun [THEN apply_rangeI])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   706
apply (blast intro: inj_converse_fun [THEN apply_type])+ 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   707
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   708
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   709
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   710
subsection{*Lemmas by Krzysztof Grabczewski*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   711
c9cfe1638bf2 improved presentation markup
paulson
parents: 13328
diff changeset
   712
(*New proofs using cons_lepoll_cons. Could generalise from succ to cons.*)
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   713
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   714
(*If A has at most n+1 elements and a:A then A-{a} has at most n.*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   715
lemma Diff_sing_lepoll: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   716
      "[| a:A;  A \<lesssim> succ(n) |] ==> A - {a} \<lesssim> n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   717
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   718
apply (rule cons_lepoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   719
apply (rule_tac [3] mem_not_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   720
apply (erule cons_Diff [THEN ssubst], safe)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   721
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   722
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   723
(*If A has at least n+1 elements then A-{a} has at least n.*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   724
lemma lepoll_Diff_sing: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   725
      "[| succ(n) \<lesssim> A |] ==> n \<lesssim> A - {a}"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   726
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   727
apply (rule cons_lepoll_consD)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   728
apply (rule_tac [2] mem_not_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   729
prefer 2 apply blast
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   730
apply (blast intro: subset_imp_lepoll [THEN [2] lepoll_trans])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   731
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   732
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   733
lemma Diff_sing_eqpoll: "[| a:A; A \<approx> succ(n) |] ==> A - {a} \<approx> n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   734
by (blast intro!: eqpollI 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   735
          elim!: eqpollE 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   736
          intro: Diff_sing_lepoll lepoll_Diff_sing)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   737
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   738
lemma lepoll_1_is_sing: "[| A \<lesssim> 1; a:A |] ==> A = {a}"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   739
apply (frule Diff_sing_lepoll, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   740
apply (drule lepoll_0_is_0)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   741
apply (blast elim: equalityE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   742
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   743
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   744
lemma Un_lepoll_sum: "A Un B \<lesssim> A+B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   745
apply (unfold lepoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   746
apply (rule_tac x = "lam x: A Un B. if x:A then Inl (x) else Inr (x) " in exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   747
apply (rule_tac d = "%z. snd (z) " in lam_injective)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   748
apply force 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   749
apply (simp add: Inl_def Inr_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   750
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   751
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   752
lemma well_ord_Un:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   753
     "[| well_ord(X,R); well_ord(Y,S) |] ==> EX T. well_ord(X Un Y, T)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   754
by (erule well_ord_radd [THEN Un_lepoll_sum [THEN lepoll_well_ord]], 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   755
    assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   756
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   757
(*Krzysztof Grabczewski*)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   758
lemma disj_Un_eqpoll_sum: "A Int B = 0 ==> A Un B \<approx> A + B"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   759
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   760
apply (rule_tac x = "lam a:A Un B. if a:A then Inl (a) else Inr (a) " in exI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   761
apply (rule_tac d = "%z. case (%x. x, %x. x, z) " in lam_bijective)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   762
apply auto
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   763
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   764
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   765
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   766
subsection {*Finite and infinite sets*}
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   767
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   768
lemma Finite_0 [simp]: "Finite(0)"
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   769
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   770
apply (blast intro!: eqpoll_refl nat_0I)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   771
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   772
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   773
lemma lepoll_nat_imp_Finite: "[| A \<lesssim> n;  n:nat |] ==> Finite(A)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   774
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   775
apply (erule rev_mp)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   776
apply (erule nat_induct)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   777
apply (blast dest!: lepoll_0_is_0 intro!: eqpoll_refl nat_0I)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   778
apply (blast dest!: lepoll_succ_disj)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   779
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   780
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   781
lemma lesspoll_nat_is_Finite: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   782
     "A \<prec> nat ==> Finite(A)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   783
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   784
apply (blast dest: ltD lesspoll_cardinal_lt 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   785
                   lesspoll_imp_eqpoll [THEN eqpoll_sym])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   786
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   787
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   788
lemma lepoll_Finite: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   789
     "[| Y \<lesssim> X;  Finite(X) |] ==> Finite(Y)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   790
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   791
apply (blast elim!: eqpollE
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   792
             intro: lepoll_trans [THEN lepoll_nat_imp_Finite
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   793
                                       [unfolded Finite_def]])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   794
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   795
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   796
lemmas subset_Finite = subset_imp_lepoll [THEN lepoll_Finite, standard]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   797
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   798
lemma Finite_Int: "Finite(A) | Finite(B) ==> Finite(A Int B)"
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   799
by (blast intro: subset_Finite) 
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   800
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   801
lemmas Finite_Diff = Diff_subset [THEN subset_Finite, standard]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   802
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   803
lemma Finite_cons: "Finite(x) ==> Finite(cons(y,x))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   804
apply (unfold Finite_def)
14153
76a6ba67bd15 new case_tac
paulson
parents: 14076
diff changeset
   805
apply (case_tac "y:x")
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   806
apply (simp add: cons_absorb)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   807
apply (erule bexE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   808
apply (rule bexI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   809
apply (erule_tac [2] nat_succI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   810
apply (simp (no_asm_simp) add: succ_def cons_eqpoll_cong mem_not_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   811
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   812
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   813
lemma Finite_succ: "Finite(x) ==> Finite(succ(x))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   814
apply (unfold succ_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   815
apply (erule Finite_cons)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   816
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   817
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13244
diff changeset
   818
lemma Finite_cons_iff [iff]: "Finite(cons(y,x)) <-> Finite(x)"
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   819
by (blast intro: Finite_cons subset_Finite)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   820
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13244
diff changeset
   821
lemma Finite_succ_iff [iff]: "Finite(succ(x)) <-> Finite(x)"
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   822
by (simp add: succ_def)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   823
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   824
lemma nat_le_infinite_Ord: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   825
      "[| Ord(i);  ~ Finite(i) |] ==> nat le i"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   826
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   827
apply (erule Ord_nat [THEN [2] Ord_linear2])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   828
prefer 2 apply assumption
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   829
apply (blast intro!: eqpoll_refl elim!: ltE)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   830
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   831
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   832
lemma Finite_imp_well_ord: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   833
    "Finite(A) ==> EX r. well_ord(A,r)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   834
apply (unfold Finite_def eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   835
apply (blast intro: well_ord_rvimage bij_is_inj well_ord_Memrel nat_into_Ord)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   836
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   837
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   838
lemma succ_lepoll_imp_not_empty: "succ(x) \<lesssim> y ==> y \<noteq> 0"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   839
by (fast dest!: lepoll_0_is_0)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   840
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   841
lemma eqpoll_succ_imp_not_empty: "x \<approx> succ(n) ==> x \<noteq> 0"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   842
by (fast elim!: eqpoll_sym [THEN eqpoll_0_is_0, THEN succ_neq_0])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   843
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   844
lemma Finite_Fin_lemma [rule_format]:
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   845
     "n \<in> nat ==> \<forall>A. (A\<approx>n & A \<subseteq> X) --> A \<in> Fin(X)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   846
apply (induct_tac n)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   847
apply (rule allI)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   848
apply (fast intro!: Fin.emptyI dest!: eqpoll_imp_lepoll [THEN lepoll_0_is_0])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   849
apply (rule allI)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   850
apply (rule impI)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   851
apply (erule conjE)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   852
apply (rule eqpoll_succ_imp_not_empty [THEN not_emptyE], assumption)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   853
apply (frule Diff_sing_eqpoll, assumption)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   854
apply (erule allE)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   855
apply (erule impE, fast)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   856
apply (drule subsetD, assumption)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   857
apply (drule Fin.consI, assumption)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   858
apply (simp add: cons_Diff)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   859
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   860
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   861
lemma Finite_Fin: "[| Finite(A); A \<subseteq> X |] ==> A \<in> Fin(X)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   862
by (unfold Finite_def, blast intro: Finite_Fin_lemma) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   863
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   864
lemma eqpoll_imp_Finite_iff: "A \<approx> B ==> Finite(A) <-> Finite(B)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   865
apply (unfold Finite_def) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   866
apply (blast intro: eqpoll_trans eqpoll_sym) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   867
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   868
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   869
lemma Fin_lemma [rule_format]: "n: nat ==> ALL A. A \<approx> n --> A : Fin(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   870
apply (induct_tac n)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   871
apply (simp add: eqpoll_0_iff, clarify)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   872
apply (subgoal_tac "EX u. u:A")
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   873
apply (erule exE)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   874
apply (rule Diff_sing_eqpoll [THEN revcut_rl])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   875
prefer 2 apply assumption
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   876
apply assumption
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
   877
apply (rule_tac b = A in cons_Diff [THEN subst], assumption)
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   878
apply (rule Fin.consI, blast)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   879
apply (blast intro: subset_consI [THEN Fin_mono, THEN subsetD])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   880
(*Now for the lemma assumed above*)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   881
apply (unfold eqpoll_def)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   882
apply (blast intro: bij_converse_bij [THEN bij_is_fun, THEN apply_type])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   883
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   884
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   885
lemma Finite_into_Fin: "Finite(A) ==> A : Fin(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   886
apply (unfold Finite_def)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   887
apply (blast intro: Fin_lemma)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   888
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   889
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   890
lemma Fin_into_Finite: "A : Fin(U) ==> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   891
by (fast intro!: Finite_0 Finite_cons elim: Fin_induct)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   892
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   893
lemma Finite_Fin_iff: "Finite(A) <-> A : Fin(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   894
by (blast intro: Finite_into_Fin Fin_into_Finite)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   895
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   896
lemma Finite_Un: "[| Finite(A); Finite(B) |] ==> Finite(A Un B)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   897
by (blast intro!: Fin_into_Finite Fin_UnI 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   898
          dest!: Finite_into_Fin
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   899
          intro: Un_upper1 [THEN Fin_mono, THEN subsetD] 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   900
                 Un_upper2 [THEN Fin_mono, THEN subsetD])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   901
14883
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   902
lemma Finite_Un_iff [simp]: "Finite(A Un B) <-> (Finite(A) & Finite(B))"
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   903
by (blast intro: subset_Finite Finite_Un) 
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   904
ca000a495448 Groups, Rings and supporting lemmas
paulson
parents: 14565
diff changeset
   905
text{*The converse must hold too.*}
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   906
lemma Finite_Union: "[| ALL y:X. Finite(y);  Finite(X) |] ==> Finite(Union(X))"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   907
apply (simp add: Finite_Fin_iff)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   908
apply (rule Fin_UnionI)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   909
apply (erule Fin_induct, simp)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   910
apply (blast intro: Fin.consI Fin_mono [THEN [2] rev_subsetD])
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   911
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   912
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   913
(* Induction principle for Finite(A), by Sidi Ehmety *)
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 13357
diff changeset
   914
lemma Finite_induct [case_names 0 cons, induct set: Finite]:
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   915
"[| Finite(A); P(0);
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   916
    !! x B.   [| Finite(B); x ~: B; P(B) |] ==> P(cons(x, B)) |]
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   917
 ==> P(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   918
apply (erule Finite_into_Fin [THEN Fin_induct]) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   919
apply (blast intro: Fin_into_Finite)+
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   920
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   921
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   922
(*Sidi Ehmety.  The contrapositive says ~Finite(A) ==> ~Finite(A-{a}) *)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   923
lemma Diff_sing_Finite: "Finite(A - {a}) ==> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   924
apply (unfold Finite_def)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   925
apply (case_tac "a:A")
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   926
apply (subgoal_tac [2] "A-{a}=A", auto)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   927
apply (rule_tac x = "succ (n) " in bexI)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   928
apply (subgoal_tac "cons (a, A - {a}) = A & cons (n, n) = succ (n) ")
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
   929
apply (drule_tac a = a and b = n in cons_eqpoll_cong)
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   930
apply (auto dest: mem_irrefl)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   931
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   932
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   933
(*Sidi Ehmety.  And the contrapositive of this says
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   934
   [| ~Finite(A); Finite(B) |] ==> ~Finite(A-B) *)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   935
lemma Diff_Finite [rule_format]: "Finite(B) ==> Finite(A-B) --> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   936
apply (erule Finite_induct, auto)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   937
apply (case_tac "x:A")
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   938
 apply (subgoal_tac [2] "A-cons (x, B) = A - B")
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13524
diff changeset
   939
apply (subgoal_tac "A - cons (x, B) = (A - B) - {x}", simp)
13244
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   940
apply (drule Diff_sing_Finite, auto)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   941
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   942
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   943
lemma Finite_RepFun: "Finite(A) ==> Finite(RepFun(A,f))"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   944
by (erule Finite_induct, simp_all)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   945
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   946
lemma Finite_RepFun_iff_lemma [rule_format]:
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   947
     "[|Finite(x); !!x y. f(x)=f(y) ==> x=y|] 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   948
      ==> \<forall>A. x = RepFun(A,f) --> Finite(A)" 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   949
apply (erule Finite_induct)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   950
 apply clarify 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   951
 apply (case_tac "A=0", simp)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   952
 apply (blast del: allE, clarify) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   953
apply (subgoal_tac "\<exists>z\<in>A. x = f(z)") 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   954
 prefer 2 apply (blast del: allE elim: equalityE, clarify) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   955
apply (subgoal_tac "B = {f(u) . u \<in> A - {z}}")
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   956
 apply (blast intro: Diff_sing_Finite) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   957
apply (thin_tac "\<forall>A. ?P(A) --> Finite(A)") 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   958
apply (rule equalityI) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   959
 apply (blast intro: elim: equalityE) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   960
apply (blast intro: elim: equalityCE) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   961
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   962
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   963
text{*I don't know why, but if the premise is expressed using meta-connectives
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   964
then  the simplifier cannot prove it automatically in conditional rewriting.*}
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   965
lemma Finite_RepFun_iff:
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   966
     "(\<forall>x y. f(x)=f(y) --> x=y) ==> Finite(RepFun(A,f)) <-> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   967
by (blast intro: Finite_RepFun Finite_RepFun_iff_lemma [of _ f]) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   968
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   969
lemma Finite_Pow: "Finite(A) ==> Finite(Pow(A))"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   970
apply (erule Finite_induct) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   971
apply (simp_all add: Pow_insert Finite_Un Finite_RepFun) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   972
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   973
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   974
lemma Finite_Pow_imp_Finite: "Finite(Pow(A)) ==> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   975
apply (subgoal_tac "Finite({{x} . x \<in> A})")
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   976
 apply (simp add: Finite_RepFun_iff ) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   977
apply (blast intro: subset_Finite) 
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   978
done
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   979
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   980
lemma Finite_Pow_iff [iff]: "Finite(Pow(A)) <-> Finite(A)"
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   981
by (blast intro: Finite_Pow Finite_Pow_imp_Finite)
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   982
7b37e218f298 moving some results around
paulson
parents: 13221
diff changeset
   983
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   984
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   985
(*Krzysztof Grabczewski's proof that the converse of a finite, well-ordered
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   986
  set is well-ordered.  Proofs simplified by lcp. *)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   987
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   988
lemma nat_wf_on_converse_Memrel: "n:nat ==> wf[n](converse(Memrel(n)))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   989
apply (erule nat_induct)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   990
apply (blast intro: wf_onI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   991
apply (rule wf_onI)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   992
apply (simp add: wf_on_def wf_def)
14153
76a6ba67bd15 new case_tac
paulson
parents: 14076
diff changeset
   993
apply (case_tac "x:Z")
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   994
 txt{*x:Z case*}
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   995
 apply (drule_tac x = x in bspec, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   996
 apply (blast elim: mem_irrefl mem_asym)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   997
txt{*other case*} 
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13615
diff changeset
   998
apply (drule_tac x = Z in spec, blast) 
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
   999
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1000
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1001
lemma nat_well_ord_converse_Memrel: "n:nat ==> well_ord(n,converse(Memrel(n)))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1002
apply (frule Ord_nat [THEN Ord_in_Ord, THEN well_ord_Memrel])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1003
apply (unfold well_ord_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1004
apply (blast intro!: tot_ord_converse nat_wf_on_converse_Memrel)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1005
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1006
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1007
lemma well_ord_converse:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1008
     "[|well_ord(A,r);      
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1009
        well_ord(ordertype(A,r), converse(Memrel(ordertype(A, r)))) |]
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1010
      ==> well_ord(A,converse(r))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1011
apply (rule well_ord_Int_iff [THEN iffD1])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1012
apply (frule ordermap_bij [THEN bij_is_inj, THEN well_ord_rvimage], assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1013
apply (simp add: rvimage_converse converse_Int converse_prod
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1014
                 ordertype_ord_iso [THEN ord_iso_rvimage_eq])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1015
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1016
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1017
lemma ordertype_eq_n:
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1018
     "[| well_ord(A,r);  A \<approx> n;  n:nat |] ==> ordertype(A,r)=n"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1019
apply (rule Ord_ordertype [THEN Ord_nat_eqpoll_iff, THEN iffD1], assumption+)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1020
apply (rule eqpoll_trans)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1021
 prefer 2 apply assumption
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1022
apply (unfold eqpoll_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1023
apply (blast intro!: ordermap_bij [THEN bij_converse_bij])
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1024
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1025
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1026
lemma Finite_well_ord_converse: 
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1027
    "[| Finite(A);  well_ord(A,r) |] ==> well_ord(A,converse(r))"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1028
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1029
apply (rule well_ord_converse, assumption)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1030
apply (blast dest: ordertype_eq_n intro!: nat_well_ord_converse_Memrel)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1031
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1032
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1033
lemma nat_into_Finite: "n:nat ==> Finite(n)"
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1034
apply (unfold Finite_def)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1035
apply (fast intro!: eqpoll_refl)
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1036
done
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1037
14076
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1038
lemma nat_not_Finite: "~Finite(nat)"
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1039
apply (unfold Finite_def, clarify) 
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1040
apply (drule eqpoll_imp_lepoll [THEN lepoll_cardinal_le], simp) 
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1041
apply (insert Card_nat) 
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1042
apply (simp add: Card_def)
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1043
apply (drule le_imp_subset)
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1044
apply (blast elim: mem_irrefl)
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1045
done
5cfc8b9fb880 Conversion of AllocBase to new-style
paulson
parents: 14046
diff changeset
  1046
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1047
ML
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1048
{*
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1049
val Least_def = thm "Least_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1050
val eqpoll_def = thm "eqpoll_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1051
val lepoll_def = thm "lepoll_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1052
val lesspoll_def = thm "lesspoll_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1053
val cardinal_def = thm "cardinal_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1054
val Finite_def = thm "Finite_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1055
val Card_def = thm "Card_def";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1056
val eq_imp_not_mem = thm "eq_imp_not_mem";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1057
val decomp_bnd_mono = thm "decomp_bnd_mono";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1058
val Banach_last_equation = thm "Banach_last_equation";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1059
val decomposition = thm "decomposition";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1060
val schroeder_bernstein = thm "schroeder_bernstein";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1061
val bij_imp_eqpoll = thm "bij_imp_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1062
val eqpoll_refl = thm "eqpoll_refl";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1063
val eqpoll_sym = thm "eqpoll_sym";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1064
val eqpoll_trans = thm "eqpoll_trans";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1065
val subset_imp_lepoll = thm "subset_imp_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1066
val lepoll_refl = thm "lepoll_refl";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1067
val le_imp_lepoll = thm "le_imp_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1068
val eqpoll_imp_lepoll = thm "eqpoll_imp_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1069
val lepoll_trans = thm "lepoll_trans";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1070
val eqpollI = thm "eqpollI";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1071
val eqpollE = thm "eqpollE";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1072
val eqpoll_iff = thm "eqpoll_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1073
val lepoll_0_is_0 = thm "lepoll_0_is_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1074
val empty_lepollI = thm "empty_lepollI";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1075
val lepoll_0_iff = thm "lepoll_0_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1076
val Un_lepoll_Un = thm "Un_lepoll_Un";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1077
val eqpoll_0_is_0 = thm "eqpoll_0_is_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1078
val eqpoll_0_iff = thm "eqpoll_0_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1079
val eqpoll_disjoint_Un = thm "eqpoll_disjoint_Un";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1080
val lesspoll_not_refl = thm "lesspoll_not_refl";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1081
val lesspoll_irrefl = thm "lesspoll_irrefl";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1082
val lesspoll_imp_lepoll = thm "lesspoll_imp_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1083
val lepoll_well_ord = thm "lepoll_well_ord";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1084
val lepoll_iff_leqpoll = thm "lepoll_iff_leqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1085
val inj_not_surj_succ = thm "inj_not_surj_succ";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1086
val lesspoll_trans = thm "lesspoll_trans";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1087
val lesspoll_trans1 = thm "lesspoll_trans1";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1088
val lesspoll_trans2 = thm "lesspoll_trans2";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1089
val Least_equality = thm "Least_equality";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1090
val LeastI = thm "LeastI";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1091
val Least_le = thm "Least_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1092
val less_LeastE = thm "less_LeastE";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1093
val LeastI2 = thm "LeastI2";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1094
val Least_0 = thm "Least_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1095
val Ord_Least = thm "Ord_Least";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1096
val Least_cong = thm "Least_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1097
val cardinal_cong = thm "cardinal_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1098
val well_ord_cardinal_eqpoll = thm "well_ord_cardinal_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1099
val Ord_cardinal_eqpoll = thm "Ord_cardinal_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1100
val well_ord_cardinal_eqE = thm "well_ord_cardinal_eqE";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1101
val well_ord_cardinal_eqpoll_iff = thm "well_ord_cardinal_eqpoll_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1102
val Ord_cardinal_le = thm "Ord_cardinal_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1103
val Card_cardinal_eq = thm "Card_cardinal_eq";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1104
val CardI = thm "CardI";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1105
val Card_is_Ord = thm "Card_is_Ord";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1106
val Card_cardinal_le = thm "Card_cardinal_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1107
val Ord_cardinal = thm "Ord_cardinal";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1108
val Card_iff_initial = thm "Card_iff_initial";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1109
val lt_Card_imp_lesspoll = thm "lt_Card_imp_lesspoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1110
val Card_0 = thm "Card_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1111
val Card_Un = thm "Card_Un";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1112
val Card_cardinal = thm "Card_cardinal";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1113
val cardinal_mono = thm "cardinal_mono";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1114
val cardinal_lt_imp_lt = thm "cardinal_lt_imp_lt";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1115
val Card_lt_imp_lt = thm "Card_lt_imp_lt";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1116
val Card_lt_iff = thm "Card_lt_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1117
val Card_le_iff = thm "Card_le_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1118
val well_ord_lepoll_imp_Card_le = thm "well_ord_lepoll_imp_Card_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1119
val lepoll_cardinal_le = thm "lepoll_cardinal_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1120
val lepoll_Ord_imp_eqpoll = thm "lepoll_Ord_imp_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1121
val lesspoll_imp_eqpoll = thm "lesspoll_imp_eqpoll";
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 13784
diff changeset
  1122
val cardinal_subset_Ord = thm "cardinal_subset_Ord";
13221
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1123
val cons_lepoll_consD = thm "cons_lepoll_consD";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1124
val cons_eqpoll_consD = thm "cons_eqpoll_consD";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1125
val succ_lepoll_succD = thm "succ_lepoll_succD";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1126
val nat_lepoll_imp_le = thm "nat_lepoll_imp_le";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1127
val nat_eqpoll_iff = thm "nat_eqpoll_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1128
val nat_into_Card = thm "nat_into_Card";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1129
val cardinal_0 = thm "cardinal_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1130
val cardinal_1 = thm "cardinal_1";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1131
val succ_lepoll_natE = thm "succ_lepoll_natE";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1132
val n_lesspoll_nat = thm "n_lesspoll_nat";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1133
val nat_lepoll_imp_ex_eqpoll_n = thm "nat_lepoll_imp_ex_eqpoll_n";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1134
val lepoll_imp_lesspoll_succ = thm "lepoll_imp_lesspoll_succ";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1135
val lesspoll_succ_imp_lepoll = thm "lesspoll_succ_imp_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1136
val lesspoll_succ_iff = thm "lesspoll_succ_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1137
val lepoll_succ_disj = thm "lepoll_succ_disj";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1138
val lesspoll_cardinal_lt = thm "lesspoll_cardinal_lt";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1139
val lt_not_lepoll = thm "lt_not_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1140
val Ord_nat_eqpoll_iff = thm "Ord_nat_eqpoll_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1141
val Card_nat = thm "Card_nat";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1142
val nat_le_cardinal = thm "nat_le_cardinal";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1143
val cons_lepoll_cong = thm "cons_lepoll_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1144
val cons_eqpoll_cong = thm "cons_eqpoll_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1145
val cons_lepoll_cons_iff = thm "cons_lepoll_cons_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1146
val cons_eqpoll_cons_iff = thm "cons_eqpoll_cons_iff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1147
val singleton_eqpoll_1 = thm "singleton_eqpoll_1";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1148
val cardinal_singleton = thm "cardinal_singleton";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1149
val not_0_is_lepoll_1 = thm "not_0_is_lepoll_1";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1150
val succ_eqpoll_cong = thm "succ_eqpoll_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1151
val sum_eqpoll_cong = thm "sum_eqpoll_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1152
val prod_eqpoll_cong = thm "prod_eqpoll_cong";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1153
val inj_disjoint_eqpoll = thm "inj_disjoint_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1154
val Diff_sing_lepoll = thm "Diff_sing_lepoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1155
val lepoll_Diff_sing = thm "lepoll_Diff_sing";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1156
val Diff_sing_eqpoll = thm "Diff_sing_eqpoll";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1157
val lepoll_1_is_sing = thm "lepoll_1_is_sing";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1158
val Un_lepoll_sum = thm "Un_lepoll_sum";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1159
val well_ord_Un = thm "well_ord_Un";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1160
val disj_Un_eqpoll_sum = thm "disj_Un_eqpoll_sum";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1161
val Finite_0 = thm "Finite_0";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1162
val lepoll_nat_imp_Finite = thm "lepoll_nat_imp_Finite";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1163
val lesspoll_nat_is_Finite = thm "lesspoll_nat_is_Finite";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1164
val lepoll_Finite = thm "lepoll_Finite";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1165
val subset_Finite = thm "subset_Finite";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1166
val Finite_Diff = thm "Finite_Diff";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1167
val Finite_cons = thm "Finite_cons";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1168
val Finite_succ = thm "Finite_succ";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1169
val nat_le_infinite_Ord = thm "nat_le_infinite_Ord";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1170
val Finite_imp_well_ord = thm "Finite_imp_well_ord";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1171
val nat_wf_on_converse_Memrel = thm "nat_wf_on_converse_Memrel";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1172
val nat_well_ord_converse_Memrel = thm "nat_well_ord_converse_Memrel";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1173
val well_ord_converse = thm "well_ord_converse";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1174
val ordertype_eq_n = thm "ordertype_eq_n";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1175
val Finite_well_ord_converse = thm "Finite_well_ord_converse";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1176
val nat_into_Finite = thm "nat_into_Finite";
e29378f347e4 conversion of Cardinal, CardinalArith
paulson
parents: 12861
diff changeset
  1177
*}
9683
f87c8c449018 added some xsymbols, and tidied
paulson
parents: 1478
diff changeset
  1178
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
  1179
end