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