| author | Fabian Huch <huch@in.tum.de> | 
| Wed, 10 Jul 2024 17:42:48 +0200 | |
| changeset 80547 | 819402eeac34 | 
| parent 79589 | 9dee3b4fdb06 | 
| child 80790 | 07c51801c2ea | 
| permissions | -rw-r--r-- | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | section \<open>Equipollence and Other Relations Connected with Cardinality\<close> | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | theory "Equipollence" | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 4 | imports FuncSet Countable_Set | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | begin | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | subsection\<open>Eqpoll\<close> | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | definition eqpoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl "\<approx>" 50) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | where "eqpoll A B \<equiv> \<exists>f. bij_betw f A B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | definition lepoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl "\<lesssim>" 50) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | where "lepoll A B \<equiv> \<exists>f. inj_on f A \<and> f ` A \<subseteq> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | definition lesspoll :: "'a set \<Rightarrow> 'b set \<Rightarrow> bool" (infixl \<open>\<prec>\<close> 50) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | where "A \<prec> B == A \<lesssim> B \<and> ~(A \<approx> B)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | |
| 78274 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 18 | lemma lepoll_def': "lepoll A B \<equiv> \<exists>f. inj_on f A \<and> f \<in> A \<rightarrow> B" | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 19 | by (simp add: Pi_iff image_subset_iff lepoll_def) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 20 | |
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 21 | lemma eqpoll_empty_iff_empty [simp]: "A \<approx> {} \<longleftrightarrow> A={}"
 | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 22 | by (simp add: bij_betw_iff_bijections eqpoll_def) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 23 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | lemma lepoll_empty_iff_empty [simp]: "A \<lesssim> {} \<longleftrightarrow> A = {}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | by (auto simp: lepoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | |
| 78274 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 27 | lemma not_lesspoll_empty: "\<not> A \<prec> {}"
 | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 28 | by (simp add: lesspoll_def) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 29 | |
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 30 | (*The HOL Light CARD_LE_RELATIONAL_FULL*) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 31 | lemma lepoll_relational_full: | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 32 | assumes "\<And>y. y \<in> B \<Longrightarrow> \<exists>x. x \<in> A \<and> R x y" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 33 | and "\<And>x y y'. \<lbrakk>x \<in> A; y \<in> B; y' \<in> B; R x y; R x y'\<rbrakk> \<Longrightarrow> y = y'" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 34 | shows "B \<lesssim> A" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 35 | proof - | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 36 | obtain f where f: "\<And>y. y \<in> B \<Longrightarrow> f y \<in> A \<and> R (f y) y" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 37 | using assms by metis | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 38 | with assms have "inj_on f B" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 39 | by (metis inj_onI) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 40 | with f show ?thesis | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 41 | unfolding lepoll_def by blast | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 42 | qed | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 43 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | lemma eqpoll_iff_card_of_ordIso: "A \<approx> B \<longleftrightarrow> ordIso2 (card_of A) (card_of B)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | by (simp add: card_of_ordIso eqpoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 47 | lemma eqpoll_refl [iff]: "A \<approx> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 48 | by (simp add: card_of_refl eqpoll_iff_card_of_ordIso) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 49 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | lemma eqpoll_finite_iff: "A \<approx> B \<Longrightarrow> finite A \<longleftrightarrow> finite B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | by (meson bij_betw_finite eqpoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | lemma eqpoll_iff_card: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | assumes "finite A" "finite B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | shows "A \<approx> B \<longleftrightarrow> card A = card B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | using assms by (auto simp: bij_betw_iff_card eqpoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | |
| 78274 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 58 | lemma eqpoll_singleton_iff: "A \<approx> {x} \<longleftrightarrow> (\<exists>u. A = {u})"
 | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 59 | by (metis card.infinite card_1_singleton_iff eqpoll_finite_iff eqpoll_iff_card not_less_eq_eq) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 60 | |
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 61 | lemma eqpoll_doubleton_iff: "A \<approx> {x,y} \<longleftrightarrow> (\<exists>u v. A = {u,v} \<and> (u=v \<longleftrightarrow> x=y))"
 | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 62 | proof (cases "x=y") | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 63 | case True | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 64 | then show ?thesis | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 65 | by (simp add: eqpoll_singleton_iff) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 66 | next | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 67 | case False | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 68 | then show ?thesis | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 69 | by (smt (verit, ccfv_threshold) card_1_singleton_iff card_Suc_eq_finite eqpoll_finite_iff | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 70 | eqpoll_iff_card finite.insertI singleton_iff) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 71 | qed | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 72 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | lemma lepoll_antisym: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | assumes "A \<lesssim> B" "B \<lesssim> A" shows "A \<approx> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | using assms unfolding eqpoll_def lepoll_def by (metis Schroeder_Bernstein) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | |
| 78274 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 77 | lemma lepoll_trans [trans]: | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 78 | assumes "A \<lesssim> B" " B \<lesssim> C" shows "A \<lesssim> C" | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 79 | proof - | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 80 | obtain f g where fg: "inj_on f A" "inj_on g B" and "f : A \<rightarrow> B" "g \<in> B \<rightarrow> C" | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 81 | by (metis assms lepoll_def') | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 82 | then have "g \<circ> f \<in> A \<rightarrow> C" | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 83 | by auto | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 84 | with fg show ?thesis | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 85 | unfolding lepoll_def | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 86 | by (metis \<open>f \<in> A \<rightarrow> B\<close> comp_inj_on image_subset_iff_funcset inj_on_subset) | 
| 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 87 | qed | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | lemma lepoll_trans1 [trans]: "\<lbrakk>A \<approx> B; B \<lesssim> C\<rbrakk> \<Longrightarrow> A \<lesssim> C" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | by (meson card_of_ordLeq eqpoll_iff_card_of_ordIso lepoll_def lepoll_trans ordIso_iff_ordLeq) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | lemma lepoll_trans2 [trans]: "\<lbrakk>A \<lesssim> B; B \<approx> C\<rbrakk> \<Longrightarrow> A \<lesssim> C" | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 93 | by (metis bij_betw_def eqpoll_def lepoll_def lepoll_trans order_refl) | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | lemma eqpoll_sym: "A \<approx> B \<Longrightarrow> B \<approx> A" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | unfolding eqpoll_def | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | using bij_betw_the_inv_into by auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | lemma eqpoll_trans [trans]: "\<lbrakk>A \<approx> B; B \<approx> C\<rbrakk> \<Longrightarrow> A \<approx> C" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 | unfolding eqpoll_def using bij_betw_trans by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | lemma eqpoll_imp_lepoll: "A \<approx> B \<Longrightarrow> A \<lesssim> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | unfolding eqpoll_def lepoll_def by (metis bij_betw_def order_refl) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | lemma subset_imp_lepoll: "A \<subseteq> B \<Longrightarrow> A \<lesssim> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | by (force simp: lepoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 108 | lemma lepoll_refl [iff]: "A \<lesssim> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 109 | by (simp add: subset_imp_lepoll) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 110 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | lemma lepoll_iff: "A \<lesssim> B \<longleftrightarrow> (\<exists>g. A \<subseteq> g ` B)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | unfolding lepoll_def | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | proof safe | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | fix g assume "A \<subseteq> g ` B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | then show "\<exists>f. inj_on f A \<and> f ` A \<subseteq> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | by (rule_tac x="inv_into B g" in exI) (auto simp: inv_into_into inj_on_inv_into) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | qed (metis image_mono the_inv_into_onto) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 119 | lemma empty_lepoll [iff]: "{} \<lesssim> A"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 120 | by (simp add: lepoll_iff) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 121 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | lemma subset_image_lepoll: "B \<subseteq> f ` A \<Longrightarrow> B \<lesssim> A" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | by (auto simp: lepoll_iff) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | lemma image_lepoll: "f ` A \<lesssim> A" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | by (auto simp: lepoll_iff) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | lemma infinite_le_lepoll: "infinite A \<longleftrightarrow> (UNIV::nat set) \<lesssim> A" | 
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 129 | by (simp add: infinite_iff_countable_subset lepoll_def) | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 131 | lemma lepoll_Pow_self: "A \<lesssim> Pow A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 132 | unfolding lepoll_def inj_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 133 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 134 |     show "inj_on (\<lambda>x. {x}) A"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 135 | by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 136 | qed auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 137 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | lemma eqpoll_iff_bijections: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | "A \<approx> B \<longleftrightarrow> (\<exists>f g. (\<forall>x \<in> A. f x \<in> B \<and> g(f x) = x) \<and> (\<forall>y \<in> B. g y \<in> A \<and> f(g y) = y))" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | by (auto simp: eqpoll_def bij_betw_iff_bijections) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | lemma lepoll_restricted_funspace: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 |    "{f. f ` A \<subseteq> B \<and> {x. f x \<noteq> k x} \<subseteq> A \<and> finite {x. f x \<noteq> k x}} \<lesssim> Fpow (A \<times> B)"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | have *: "\<exists>U \<in> Fpow (A \<times> B). f = (\<lambda>x. if \<exists>y. (x, y) \<in> U then SOME y. (x,y) \<in> U else k x)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 |     if "f ` A \<subseteq> B" "{x. f x \<noteq> k x} \<subseteq> A" "finite {x. f x \<noteq> k x}" for f
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 |     apply (rule_tac x="(\<lambda>x. (x, f x)) ` {x. f x \<noteq> k x}" in bexI)
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | using that by (auto simp: image_def Fpow_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | apply (rule subset_image_lepoll [where f = "\<lambda>U x. if \<exists>y. (x,y) \<in> U then @y. (x,y) \<in> U else k x"]) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | using * by (auto simp: image_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | |
| 69874 | 154 | lemma singleton_lepoll: "{x} \<lesssim> insert y A"
 | 
| 155 | by (force simp: lepoll_def) | |
| 156 | ||
| 157 | lemma singleton_eqpoll: "{x} \<approx> {y}"
 | |
| 158 | by (blast intro: lepoll_antisym singleton_lepoll) | |
| 159 | ||
| 160 | lemma subset_singleton_iff_lepoll: "(\<exists>x. S \<subseteq> {x}) \<longleftrightarrow> S \<lesssim> {()}"
 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 161 | using lepoll_iff by fastforce | 
| 69874 | 162 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 163 | lemma infinite_insert_lepoll: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 164 | assumes "infinite A" shows "insert a A \<lesssim> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 165 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 166 | obtain f :: "nat \<Rightarrow> 'a" where "inj f" and f: "range f \<subseteq> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 167 | using assms infinite_countable_subset by blast | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 168 | let ?g = "(\<lambda>z. if z=a then f 0 else if z \<in> range f then f (Suc (inv f z)) else z)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 169 | show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 170 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 171 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 172 | show "inj_on ?g (insert a A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 173 | using inj_on_eq_iff [OF \<open>inj f\<close>] | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 174 | by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 175 | show "?g ` insert a A \<subseteq> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 176 | using f by auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 177 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 178 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 179 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 180 | lemma infinite_insert_eqpoll: "infinite A \<Longrightarrow> insert a A \<approx> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 181 | by (simp add: lepoll_antisym infinite_insert_lepoll subset_imp_lepoll subset_insertI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 182 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 183 | lemma finite_lepoll_infinite: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 184 | assumes "infinite A" "finite B" shows "B \<lesssim> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 185 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 186 | have "B \<lesssim> (UNIV::nat set)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 187 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 188 | using finite_imp_inj_to_nat_seg [OF \<open>finite B\<close>] by blast | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 189 | then show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 190 | using \<open>infinite A\<close> infinite_le_lepoll lepoll_trans by auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 191 | qed | 
| 69874 | 192 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 193 | lemma countable_lepoll: "\<lbrakk>countable A; B \<lesssim> A\<rbrakk> \<Longrightarrow> countable B" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 194 | by (meson countable_image countable_subset lepoll_iff) | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 195 | |
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 196 | lemma countable_eqpoll: "\<lbrakk>countable A; B \<approx> A\<rbrakk> \<Longrightarrow> countable B" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 197 | using countable_lepoll eqpoll_imp_lepoll by blast | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 198 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | subsection\<open>The strict relation\<close> | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | lemma lesspoll_not_refl [iff]: "~ (i \<prec> i)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | by (simp add: lepoll_antisym lesspoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | lemma lesspoll_imp_lepoll: "A \<prec> B ==> A \<lesssim> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | by (unfold lesspoll_def, blast) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | lemma lepoll_iff_leqpoll: "A \<lesssim> B \<longleftrightarrow> A \<prec> B | A \<approx> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | using eqpoll_imp_lepoll lesspoll_def by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | lemma lesspoll_trans [trans]: "\<lbrakk>X \<prec> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | by (meson eqpoll_sym lepoll_antisym lepoll_trans lepoll_trans1 lesspoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | lemma lesspoll_trans1 [trans]: "\<lbrakk>X \<lesssim> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | by (meson eqpoll_sym lepoll_antisym lepoll_trans lepoll_trans1 lesspoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | lemma lesspoll_trans2 [trans]: "\<lbrakk>X \<prec> Y; Y \<lesssim> Z\<rbrakk> \<Longrightarrow> X \<prec> Z" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | by (meson eqpoll_imp_lepoll eqpoll_sym lepoll_antisym lepoll_trans lesspoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | lemma eq_lesspoll_trans [trans]: "\<lbrakk>X \<approx> Y; Y \<prec> Z\<rbrakk> \<Longrightarrow> X \<prec> Z" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | using eqpoll_imp_lepoll lesspoll_trans1 by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | lemma lesspoll_eq_trans [trans]: "\<lbrakk>X \<prec> Y; Y \<approx> Z\<rbrakk> \<Longrightarrow> X \<prec> Z" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | using eqpoll_imp_lepoll lesspoll_trans2 by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 225 | lemma lesspoll_Pow_self: "A \<prec> Pow A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 226 | unfolding lesspoll_def bij_betw_def eqpoll_def | 
| 77140 
9a60c1759543
Lots more new material thanks to Manuel Eberl
 paulson <lp15@cam.ac.uk> parents: 
75331diff
changeset | 227 | by (meson lepoll_Pow_self Cantors_theorem) | 
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 228 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 229 | lemma finite_lesspoll_infinite: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 230 | assumes "infinite A" "finite B" shows "B \<prec> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 231 | by (meson assms eqpoll_finite_iff finite_lepoll_infinite lesspoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 232 | |
| 78200 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 233 | lemma countable_lesspoll: "\<lbrakk>countable A; B \<prec> A\<rbrakk> \<Longrightarrow> countable B" | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 234 | using countable_lepoll lesspoll_def by blast | 
| 
264f2b69d09c
New and generalised analysis lemmas
 paulson <lp15@cam.ac.uk> parents: 
77140diff
changeset | 235 | |
| 78250 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 236 | lemma lepoll_iff_card_le: "\<lbrakk>finite A; finite B\<rbrakk> \<Longrightarrow> A \<lesssim> B \<longleftrightarrow> card A \<le> card B" | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 237 | by (simp add: inj_on_iff_card_le lepoll_def) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 238 | |
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 239 | lemma lepoll_iff_finite_card: "A \<lesssim> {..<n::nat} \<longleftrightarrow> finite A \<and> card A \<le> n"
 | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 240 | by (metis card_lessThan finite_lessThan finite_surj lepoll_iff lepoll_iff_card_le) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 241 | |
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 242 | lemma eqpoll_iff_finite_card: "A \<approx> {..<n::nat} \<longleftrightarrow> finite A \<and> card A = n"
 | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 243 | by (metis card_lessThan eqpoll_finite_iff eqpoll_iff_card finite_lessThan) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 244 | |
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 245 | lemma lesspoll_iff_finite_card: "A \<prec> {..<n::nat} \<longleftrightarrow> finite A \<and> card A < n"
 | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 246 | by (metis eqpoll_iff_finite_card lepoll_iff_finite_card lesspoll_def order_less_le) | 
| 
400aecdfd71f
Another tranche of HOL Light material on metric and topological spaces
 paulson <lp15@cam.ac.uk> parents: 
78200diff
changeset | 247 | |
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 248 | subsection\<open>Mapping by an injection\<close> | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 249 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 250 | lemma inj_on_image_eqpoll_self: "inj_on f A \<Longrightarrow> f ` A \<approx> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 251 | by (meson bij_betw_def eqpoll_def eqpoll_sym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 252 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 253 | lemma inj_on_image_lepoll_1 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 254 | assumes "inj_on f A" shows "f ` A \<lesssim> B \<longleftrightarrow> A \<lesssim> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 255 | by (meson assms image_lepoll lepoll_def lepoll_trans order_refl) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 256 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 257 | lemma inj_on_image_lepoll_2 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 258 | assumes "inj_on f B" shows "A \<lesssim> f ` B \<longleftrightarrow> A \<lesssim> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 259 | by (meson assms eq_iff image_lepoll lepoll_def lepoll_trans) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 260 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 261 | lemma inj_on_image_lesspoll_1 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 262 | assumes "inj_on f A" shows "f ` A \<prec> B \<longleftrightarrow> A \<prec> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 263 | by (meson assms image_lepoll le_less lepoll_def lesspoll_trans1) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 264 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 265 | lemma inj_on_image_lesspoll_2 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 266 | assumes "inj_on f B" shows "A \<prec> f ` B \<longleftrightarrow> A \<prec> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 267 | by (meson assms eqpoll_sym inj_on_image_eqpoll_self lesspoll_eq_trans) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 268 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 269 | lemma inj_on_image_eqpoll_1 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 270 | assumes "inj_on f A" shows "f ` A \<approx> B \<longleftrightarrow> A \<approx> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 271 | by (metis assms eqpoll_trans inj_on_image_eqpoll_self eqpoll_sym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 272 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 273 | lemma inj_on_image_eqpoll_2 [simp]: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 274 | assumes "inj_on f B" shows "A \<approx> f ` B \<longleftrightarrow> A \<approx> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 275 | by (metis assms inj_on_image_eqpoll_1 eqpoll_sym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 276 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 277 | subsection \<open>Inserting elements into sets\<close> | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 278 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 279 | lemma insert_lepoll_insertD: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 280 | assumes "insert u A \<lesssim> insert v B" "u \<notin> A" "v \<notin> B" shows "A \<lesssim> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 281 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 282 | obtain f where inj: "inj_on f (insert u A)" and fim: "f ` (insert u A) \<subseteq> insert v B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 283 | by (meson assms lepoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 284 | show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 285 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 286 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 287 | let ?g = "\<lambda>x\<in>A. if f x = v then f u else f x" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 288 | show "inj_on ?g A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 289 | using inj \<open>u \<notin> A\<close> by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 290 | show "?g ` A \<subseteq> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 291 | using fim \<open>u \<notin> A\<close> image_subset_iff inj inj_on_image_mem_iff by fastforce | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 292 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 293 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 294 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 295 | lemma insert_eqpoll_insertD: "\<lbrakk>insert u A \<approx> insert v B; u \<notin> A; v \<notin> B\<rbrakk> \<Longrightarrow> A \<approx> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 296 | by (meson insert_lepoll_insertD eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 297 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 298 | lemma insert_lepoll_cong: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 299 | assumes "A \<lesssim> B" "b \<notin> B" shows "insert a A \<lesssim> insert b B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 300 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 301 | obtain f where f: "inj_on f A" "f ` A \<subseteq> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 302 | by (meson assms lepoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 303 | let ?f = "\<lambda>u \<in> insert a A. if u=a then b else f u" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 304 | show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 305 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 306 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 307 | show "inj_on ?f (insert a A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 308 | using f \<open>b \<notin> B\<close> by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 309 | show "?f ` insert a A \<subseteq> insert b B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 310 | using f \<open>b \<notin> B\<close> by auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 311 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 312 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 313 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 314 | lemma insert_eqpoll_cong: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 315 | "\<lbrakk>A \<approx> B; a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> insert a A \<approx> insert b B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 316 | apply (rule lepoll_antisym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 317 | apply (simp add: eqpoll_imp_lepoll insert_lepoll_cong)+ | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 318 | by (meson eqpoll_imp_lepoll eqpoll_sym insert_lepoll_cong) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 319 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 320 | lemma insert_eqpoll_insert_iff: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 321 | "\<lbrakk>a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> insert a A \<approx> insert b B \<longleftrightarrow> A \<approx> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 322 | by (meson insert_eqpoll_insertD insert_eqpoll_cong) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 323 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 324 | lemma insert_lepoll_insert_iff: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 325 | " \<lbrakk>a \<notin> A; b \<notin> B\<rbrakk> \<Longrightarrow> (insert a A \<lesssim> insert b B) \<longleftrightarrow> (A \<lesssim> B)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 326 | by (meson insert_lepoll_insertD insert_lepoll_cong) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 327 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 328 | lemma less_imp_insert_lepoll: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 329 | assumes "A \<prec> B" shows "insert a A \<lesssim> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 330 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 331 | obtain f where "inj_on f A" "f ` A \<subset> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 332 | using assms by (metis bij_betw_def eqpoll_def lepoll_def lesspoll_def psubset_eq) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 333 | then obtain b where b: "b \<in> B" "b \<notin> f ` A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 334 | by auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 335 | show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 336 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 337 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 338 | show "inj_on (f(a:=b)) (insert a A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 339 | using b \<open>inj_on f A\<close> by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 340 | show "(f(a:=b)) ` insert a A \<subseteq> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 341 | using \<open>f ` A \<subset> B\<close> by (auto simp: b) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 342 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 343 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 344 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 345 | lemma finite_insert_lepoll: "finite A \<Longrightarrow> (insert a A \<lesssim> A) \<longleftrightarrow> (a \<in> A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 346 | proof (induction A rule: finite_induct) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 347 | case (insert x A) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 348 | then show ?case | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 349 | apply (auto simp: insert_absorb) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 350 | by (metis insert_commute insert_iff insert_lepoll_insertD) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 351 | qed auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 352 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 353 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 354 | subsection\<open>Binary sums and unions\<close> | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 355 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 356 | lemma Un_lepoll_mono: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 357 | assumes "A \<lesssim> C" "B \<lesssim> D" "disjnt C D" shows "A \<union> B \<lesssim> C \<union> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 358 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 359 | obtain f g where inj: "inj_on f A" "inj_on g B" and fg: "f ` A \<subseteq> C" "g ` B \<subseteq> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 360 | by (meson assms lepoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 361 | have "inj_on (\<lambda>x. if x \<in> A then f x else g x) (A \<union> B)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 362 | using inj \<open>disjnt C D\<close> fg unfolding disjnt_iff | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 363 | by (fastforce intro: inj_onI dest: inj_on_contraD split: if_split_asm) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 364 | with fg show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 365 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 366 | by (rule_tac x="\<lambda>x. if x \<in> A then f x else g x" in exI) auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 367 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 368 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 369 | lemma Un_eqpoll_cong: "\<lbrakk>A \<approx> C; B \<approx> D; disjnt A B; disjnt C D\<rbrakk> \<Longrightarrow> A \<union> B \<approx> C \<union> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 370 | by (meson Un_lepoll_mono eqpoll_imp_lepoll eqpoll_sym lepoll_antisym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 371 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 372 | lemma sum_lepoll_mono: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 373 | assumes "A \<lesssim> C" "B \<lesssim> D" shows "A <+> B \<lesssim> C <+> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 374 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 375 | obtain f g where "inj_on f A" "f ` A \<subseteq> C" "inj_on g B" "g ` B \<subseteq> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 376 | by (meson assms lepoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 377 | then show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 378 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 379 | by (rule_tac x="case_sum (Inl \<circ> f) (Inr \<circ> g)" in exI) (force simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 380 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 381 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 382 | lemma sum_eqpoll_cong: "\<lbrakk>A \<approx> C; B \<approx> D\<rbrakk> \<Longrightarrow> A <+> B \<approx> C <+> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 383 | by (meson eqpoll_imp_lepoll eqpoll_sym lepoll_antisym sum_lepoll_mono) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 384 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 385 | subsection\<open>Binary Cartesian products\<close> | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 386 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 387 | lemma times_square_lepoll: "A \<lesssim> A \<times> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 388 | unfolding lepoll_def inj_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 389 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 390 | show "inj_on (\<lambda>x. (x,x)) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 391 | by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 392 | qed auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 393 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 394 | lemma times_commute_eqpoll: "A \<times> B \<approx> B \<times> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 395 | unfolding eqpoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 396 | by (force intro: bij_betw_byWitness [where f = "\<lambda>(x,y). (y,x)" and f' = "\<lambda>(x,y). (y,x)"]) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 397 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 398 | lemma times_assoc_eqpoll: "(A \<times> B) \<times> C \<approx> A \<times> (B \<times> C)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 399 | unfolding eqpoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 400 | by (force intro: bij_betw_byWitness [where f = "\<lambda>((x,y),z). (x,(y,z))" and f' = "\<lambda>(x,(y,z)). ((x,y),z)"]) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 401 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 402 | lemma times_singleton_eqpoll: "{a} \<times> A \<approx> A"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 403 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 404 |   have "{a} \<times> A = (\<lambda>x. (a,x)) ` A"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 405 | by auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 406 | also have "\<dots> \<approx> A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 407 | proof (rule inj_on_image_eqpoll_self) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 408 | show "inj_on (Pair a) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 409 | by (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 410 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 411 | finally show ?thesis . | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 412 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 413 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 414 | lemma times_lepoll_mono: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 415 | assumes "A \<lesssim> C" "B \<lesssim> D" shows "A \<times> B \<lesssim> C \<times> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 416 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 417 | obtain f g where "inj_on f A" "f ` A \<subseteq> C" "inj_on g B" "g ` B \<subseteq> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 418 | by (meson assms lepoll_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 419 | then show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 420 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 421 | by (rule_tac x="\<lambda>(x,y). (f x, g y)" in exI) (auto simp: inj_on_def) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 422 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 423 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 424 | lemma times_eqpoll_cong: "\<lbrakk>A \<approx> C; B \<approx> D\<rbrakk> \<Longrightarrow> A \<times> B \<approx> C \<times> D" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 425 | by (metis eqpoll_imp_lepoll eqpoll_sym lepoll_antisym times_lepoll_mono) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 426 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 427 | lemma | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 428 |   assumes "B \<noteq> {}" shows lepoll_times1: "A \<lesssim> A \<times> B" and lepoll_times2:  "A \<lesssim> B \<times> A"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 429 | using assms lepoll_iff by fastforce+ | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 430 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 431 | lemma times_0_eqpoll: "{} \<times> A \<approx> {}"
 | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 432 | by (simp add: eqpoll_iff_bijections) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 433 | |
| 75331 | 434 | lemma Sigma_inj_lepoll_mono: | 
| 435 | assumes h: "inj_on h A" "h ` A \<subseteq> C" and "\<And>x. x \<in> A \<Longrightarrow> B x \<lesssim> D (h x)" | |
| 436 | shows "Sigma A B \<lesssim> Sigma C D" | |
| 437 | proof - | |
| 438 | have "\<And>x. x \<in> A \<Longrightarrow> \<exists>f. inj_on f (B x) \<and> f ` (B x) \<subseteq> D (h x)" | |
| 439 | by (meson assms lepoll_def) | |
| 440 | then obtain f where "\<And>x. x \<in> A \<Longrightarrow> inj_on (f x) (B x) \<and> f x ` B x \<subseteq> D (h x)" | |
| 441 | by metis | |
| 442 | with h show ?thesis | |
| 443 | unfolding lepoll_def inj_on_def | |
| 444 | by (rule_tac x="\<lambda>(x,y). (h x, f x y)" in exI) force | |
| 445 | qed | |
| 446 | ||
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 447 | lemma Sigma_lepoll_mono: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 448 | assumes "A \<subseteq> C" "\<And>x. x \<in> A \<Longrightarrow> B x \<lesssim> D x" shows "Sigma A B \<lesssim> Sigma C D" | 
| 75331 | 449 | using Sigma_inj_lepoll_mono [of id] assms by auto | 
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 450 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 451 | lemma sum_times_distrib_eqpoll: "(A <+> B) \<times> C \<approx> (A \<times> C) <+> (B \<times> C)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 452 | unfolding eqpoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 453 | proof | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 454 | show "bij_betw (\<lambda>(x,z). case_sum(\<lambda>y. Inl(y,z)) (\<lambda>y. Inr(y,z)) x) ((A <+> B) \<times> C) (A \<times> C <+> B \<times> C)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 455 | by (rule bij_betw_byWitness [where f' = "case_sum (\<lambda>(x,z). (Inl x, z)) (\<lambda>(y,z). (Inr y, z))"]) auto | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 456 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 457 | |
| 75331 | 458 | lemma Sigma_eqpoll_cong: | 
| 459 | assumes h: "bij_betw h A C" and BD: "\<And>x. x \<in> A \<Longrightarrow> B x \<approx> D (h x)" | |
| 460 | shows "Sigma A B \<approx> Sigma C D" | |
| 461 | proof (intro lepoll_antisym) | |
| 462 | show "Sigma A B \<lesssim> Sigma C D" | |
| 463 | by (metis Sigma_inj_lepoll_mono bij_betw_def eqpoll_imp_lepoll subset_refl assms) | |
| 464 | have "inj_on (inv_into A h) C \<and> inv_into A h ` C \<subseteq> A" | |
| 465 | by (metis bij_betw_def bij_betw_inv_into h set_eq_subset) | |
| 466 | then show "Sigma C D \<lesssim> Sigma A B" | |
| 467 | by (smt (verit, best) BD Sigma_inj_lepoll_mono bij_betw_inv_into_right eqpoll_sym h image_subset_iff lepoll_refl lepoll_trans2) | |
| 468 | qed | |
| 469 | ||
| 71096 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 470 | lemma prod_insert_eqpoll: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 471 | assumes "a \<notin> A" shows "insert a A \<times> B \<approx> B <+> A \<times> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 472 | unfolding eqpoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 473 | proof | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 474 | show "bij_betw (\<lambda>(x,y). if x=a then Inl y else Inr (x,y)) (insert a A \<times> B) (B <+> A \<times> B)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 475 | by (rule bij_betw_byWitness [where f' = "case_sum (\<lambda>y. (a,y)) id"]) (auto simp: assms) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 476 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 477 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 478 | subsection\<open>General Unions\<close> | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 479 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 480 | lemma Union_eqpoll_Times: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 481 | assumes B: "\<And>x. x \<in> A \<Longrightarrow> F x \<approx> B" and disj: "pairwise (\<lambda>x y. disjnt (F x) (F y)) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 482 | shows "(\<Union>x\<in>A. F x) \<approx> A \<times> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 483 | proof (rule lepoll_antisym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 484 | obtain b where b: "\<And>x. x \<in> A \<Longrightarrow> bij_betw (b x) (F x) B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 485 | using B unfolding eqpoll_def by metis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 486 | show "\<Union>(F ` A) \<lesssim> A \<times> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 487 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 488 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 489 | define \<chi> where "\<chi> \<equiv> \<lambda>z. THE x. x \<in> A \<and> z \<in> F x" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 490 | have \<chi>: "\<chi> z = x" if "x \<in> A" "z \<in> F x" for x z | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 491 | unfolding \<chi>_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 492 | apply (rule the_equality) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 493 | apply (simp add: that) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 494 | by (metis disj disjnt_iff pairwiseD that) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 495 | let ?f = "\<lambda>z. (\<chi> z, b (\<chi> z) z)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 496 | show "inj_on ?f (\<Union>(F ` A))" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 497 | unfolding inj_on_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 498 | by clarify (metis \<chi> b bij_betw_inv_into_left) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 499 | show "?f ` \<Union>(F ` A) \<subseteq> A \<times> B" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 500 | using \<chi> b bij_betwE by blast | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 501 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 502 | show "A \<times> B \<lesssim> \<Union>(F ` A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 503 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 504 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 505 | let ?f = "\<lambda>(x,y). inv_into (F x) (b x) y" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 506 | have *: "inv_into (F x) (b x) y \<in> F x" if "x \<in> A" "y \<in> B" for x y | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 507 | by (metis b bij_betw_imp_surj_on inv_into_into that) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 508 | then show "inj_on ?f (A \<times> B)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 509 | unfolding inj_on_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 510 | by clarsimp (metis (mono_tags, lifting) b bij_betw_inv_into_right disj disjnt_iff pairwiseD) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 511 | show "?f ` (A \<times> B) \<subseteq> \<Union> (F ` A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 512 | by clarsimp (metis b bij_betw_imp_surj_on inv_into_into) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 513 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 514 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 515 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 516 | lemma UN_lepoll_UN: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 517 | assumes A: "\<And>x. x \<in> A \<Longrightarrow> B x \<lesssim> C x" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 518 | and disj: "pairwise (\<lambda>x y. disjnt (C x) (C y)) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 519 | shows "\<Union> (B`A) \<lesssim> \<Union> (C`A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 520 | proof - | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 521 | obtain f where f: "\<And>x. x \<in> A \<Longrightarrow> inj_on (f x) (B x) \<and> f x ` (B x) \<subseteq> (C x)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 522 | using A unfolding lepoll_def by metis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 523 | show ?thesis | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 524 | unfolding lepoll_def | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 525 | proof (intro exI conjI) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 526 | define \<chi> where "\<chi> \<equiv> \<lambda>z. @x. x \<in> A \<and> z \<in> B x" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 527 | have \<chi>: "\<chi> z \<in> A \<and> z \<in> B (\<chi> z)" if "x \<in> A" "z \<in> B x" for x z | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 528 | unfolding \<chi>_def by (metis (mono_tags, lifting) someI_ex that) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 529 | let ?f = "\<lambda>z. (f (\<chi> z) z)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 530 | show "inj_on ?f (\<Union>(B ` A))" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 531 | using disj f unfolding inj_on_def disjnt_iff pairwise_def image_subset_iff | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 532 | by (metis UN_iff \<chi>) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 533 | show "?f ` \<Union> (B ` A) \<subseteq> \<Union> (C ` A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 534 | using \<chi> f unfolding image_subset_iff by blast | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 535 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 536 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 537 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 538 | lemma UN_eqpoll_UN: | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 539 | assumes A: "\<And>x. x \<in> A \<Longrightarrow> B x \<approx> C x" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 540 | and B: "pairwise (\<lambda>x y. disjnt (B x) (B y)) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 541 | and C: "pairwise (\<lambda>x y. disjnt (C x) (C y)) A" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 542 | shows "(\<Union>x\<in>A. B x) \<approx> (\<Union>x\<in>A. C x)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 543 | proof (rule lepoll_antisym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 544 | show "\<Union> (B ` A) \<lesssim> \<Union> (C ` A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 545 | by (meson A C UN_lepoll_UN eqpoll_imp_lepoll) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 546 | show "\<Union> (C ` A) \<lesssim> \<Union> (B ` A)" | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 547 | by (simp add: A B UN_lepoll_UN eqpoll_imp_lepoll eqpoll_sym) | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 548 | qed | 
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 549 | |
| 
ec7cc76e88e5
New library material from the AFP entry ZFC_in_HOL
 paulson <lp15@cam.ac.uk> parents: 
69874diff
changeset | 550 | subsection\<open>General Cartesian products (Pi)\<close> | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 | lemma PiE_sing_eqpoll_self: "({a} \<rightarrow>\<^sub>E B) \<approx> B"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 | have 1: "x = y" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 |     if "x \<in> {a} \<rightarrow>\<^sub>E B" "y \<in> {a} \<rightarrow>\<^sub>E B" "x a = y a" for x y
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 | by (metis IntD2 PiE_def extensionalityI singletonD that) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 |   have 2: "x \<in> (\<lambda>h. h a) ` ({a} \<rightarrow>\<^sub>E B)" if "x \<in> B" for x
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 |     using that by (rule_tac x="\<lambda>z\<in>{a}. x" in image_eqI) auto
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | unfolding eqpoll_def bij_betw_def inj_on_def | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | by (force intro: 1 2) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | lemma lepoll_funcset_right: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | "B \<lesssim> B' \<Longrightarrow> A \<rightarrow>\<^sub>E B \<lesssim> A \<rightarrow>\<^sub>E B'" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | apply (auto simp: lepoll_def inj_on_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 | apply (rule_tac x = "\<lambda>g. \<lambda>z \<in> A. f(g z)" in exI) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | apply (auto simp: fun_eq_iff) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 | apply (metis PiE_E) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | lemma lepoll_funcset_left: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 |   assumes "B \<noteq> {}" "A \<lesssim> A'"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | shows "A \<rightarrow>\<^sub>E B \<lesssim> A' \<rightarrow>\<^sub>E B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | obtain b where "b \<in> B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 | using assms by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | obtain f where "inj_on f A" and fim: "f ` A \<subseteq> A'" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | using assms by (auto simp: lepoll_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | then obtain h where h: "\<And>x. x \<in> A \<Longrightarrow> h (f x) = x" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 | using the_inv_into_f_f by fastforce | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 | let ?F = "\<lambda>g. \<lambda>u \<in> A'. if h u \<in> A then g(h u) else b" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | unfolding lepoll_def inj_on_def | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | proof (intro exI conjI ballI impI ext) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 | fix k l x | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | assume k: "k \<in> A \<rightarrow>\<^sub>E B" and l: "l \<in> A \<rightarrow>\<^sub>E B" and "?F k = ?F l" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | then have "?F k (f x) = ?F l (f x)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | by simp | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | then show "k x = l x" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | apply (auto simp: h split: if_split_asm) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | apply (metis PiE_arb h k l) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 | apply (metis (full_types) PiE_E h k l) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | using fim k l by fastforce | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | next | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | show "?F ` (A \<rightarrow>\<^sub>E B) \<subseteq> A' \<rightarrow>\<^sub>E B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | using \<open>b \<in> B\<close> by force | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | lemma lepoll_funcset: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 |    "\<lbrakk>B \<noteq> {}; A \<lesssim> A'; B \<lesssim> B'\<rbrakk> \<Longrightarrow> A \<rightarrow>\<^sub>E B \<lesssim> A' \<rightarrow>\<^sub>E B'"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | by (rule lepoll_trans [OF lepoll_funcset_right lepoll_funcset_left]) auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | lemma lepoll_PiE: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | assumes "\<And>i. i \<in> A \<Longrightarrow> B i \<lesssim> C i" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | shows "PiE A B \<lesssim> PiE A C" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | obtain f where f: "\<And>i. i \<in> A \<Longrightarrow> inj_on (f i) (B i) \<and> (f i) ` B i \<subseteq> C i" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | using assms unfolding lepoll_def by metis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | then show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | unfolding lepoll_def | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | apply (rule_tac x = "\<lambda>g. \<lambda>i \<in> A. f i (g i)" in exI) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | apply (auto simp: inj_on_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | apply (rule PiE_ext, auto) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 | apply (metis (full_types) PiE_mem restrict_apply') | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | lemma card_le_PiE_subindex: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 |   assumes "A \<subseteq> A'" "Pi\<^sub>E A' B \<noteq> {}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | shows "PiE A B \<lesssim> PiE A' B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | have "\<And>x. x \<in> A' \<Longrightarrow> \<exists>y. y \<in> B x" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | using assms by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | then obtain g where g: "\<And>x. x \<in> A' \<Longrightarrow> g x \<in> B x" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | by metis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | let ?F = "\<lambda>f x. if x \<in> A then f x else if x \<in> A' then g x else undefined" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | have "Pi\<^sub>E A B \<subseteq> (\<lambda>f. restrict f A) ` Pi\<^sub>E A' B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 | proof | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | show "f \<in> Pi\<^sub>E A B \<Longrightarrow> f \<in> (\<lambda>f. restrict f A) ` Pi\<^sub>E A' B" for f | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | using \<open>A \<subseteq> A'\<close> | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | by (rule_tac x="?F f" in image_eqI) (auto simp: g fun_eq_iff) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | then have "Pi\<^sub>E A B \<lesssim> (\<lambda>f. \<lambda>i \<in> A. f i) ` Pi\<^sub>E A' B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 | by (simp add: subset_imp_lepoll) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | also have "\<dots> \<lesssim> PiE A' B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | by (rule image_lepoll) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | finally show ?thesis . | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | lemma finite_restricted_funspace: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | assumes "finite A" "finite B" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 |   shows "finite {f. f ` A \<subseteq> B \<and> {x. f x \<noteq> k x} \<subseteq> A}" (is "finite ?F")
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 | proof (rule finite_subset) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | show "finite ((\<lambda>U x. if \<exists>y. (x,y) \<in> U then @y. (x,y) \<in> U else k x) ` Pow(A \<times> B))" (is "finite ?G") | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | using assms by auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | show "?F \<subseteq> ?G" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | proof | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 | fix f | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | assume "f \<in> ?F" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | then show "f \<in> ?G" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 |       by (rule_tac x="(\<lambda>x. (x,f x)) ` {x. f x \<noteq> k x}" in image_eqI) (auto simp: fun_eq_iff image_def)
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | proposition finite_PiE_iff: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 |    "finite(PiE I S) \<longleftrightarrow> PiE I S = {} \<or> finite {i \<in> I. ~(\<exists>a. S i \<subseteq> {a})} \<and> (\<forall>i \<in> I. finite(S i))"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | (is "?lhs = ?rhs") | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | proof (cases "PiE I S = {}")
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | case False | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 |   define J where "J \<equiv> {i \<in> I. \<nexists>a. S i \<subseteq> {a}}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 | show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | proof | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | assume L: ?lhs | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 669 | have "infinite (Pi\<^sub>E I S)" if "infinite J" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | have "(UNIV::nat set) \<lesssim> (UNIV::(nat\<Rightarrow>bool) set)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | have "\<forall>N::nat set. inj_on (=) N" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | by (simp add: inj_on_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | then show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | by (meson infinite_iff_countable_subset infinite_le_lepoll top.extremum) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | also have "\<dots> = (UNIV::nat set) \<rightarrow>\<^sub>E (UNIV::bool set)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | by auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | also have "\<dots> \<lesssim> J \<rightarrow>\<^sub>E (UNIV::bool set)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | apply (rule lepoll_funcset_left) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | using infinite_le_lepoll that by auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | also have "\<dots> \<lesssim> Pi\<^sub>E J S" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 |         have *: "(UNIV::bool set) \<lesssim> S i" if "i \<in> I" and "\<forall>a. \<not> S i \<subseteq> {a}" for i
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | proof - | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 |           obtain a b where "{a,b} \<subseteq> S i" "a \<noteq> b"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 |             by (metis \<open>\<forall>a. \<not> S i \<subseteq> {a}\<close> all_not_in_conv empty_subsetI insertCI insert_subset set_eq_subset subsetI)
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | then show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | apply (clarsimp simp: lepoll_def inj_on_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | apply (rule_tac x="\<lambda>x. if x then a else b" in exI, auto) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 | done | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | by (auto simp: * J_def intro: lepoll_PiE) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | also have "\<dots> \<lesssim> Pi\<^sub>E I S" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | using False by (auto simp: J_def intro: card_le_PiE_subindex) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | finally have "(UNIV::nat set) \<lesssim> Pi\<^sub>E I S" . | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | then show ?thesis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | by (simp add: infinite_le_lepoll) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 | moreover have "finite (S i)" if "i \<in> I" for i | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | proof (rule finite_subset) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 | obtain f where f: "f \<in> PiE I S" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | using False by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | show "S i \<subseteq> (\<lambda>f. f i) ` Pi\<^sub>E I S" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | proof | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | show "s \<in> (\<lambda>f. f i) ` Pi\<^sub>E I S" if "s \<in> S i" for s | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 | using that f \<open>i \<in> I\<close> | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | by (rule_tac x="\<lambda>j. if j = i then s else f j" in image_eqI) auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | next | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 | show "finite ((\<lambda>x. x i) ` Pi\<^sub>E I S)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | using L by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 | ultimately show ?rhs | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | using L | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | by (auto simp: J_def False) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 | next | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | assume R: ?rhs | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 |     have "\<forall>i \<in> I - J. \<exists>a. S i = {a}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | using False J_def by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 |     then obtain a where a: "\<forall>i \<in> I - J. S i = {a i}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | by metis | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 |     let ?F = "{f. f ` J \<subseteq> (\<Union>i \<in> J. S i) \<and> {i. f i \<noteq> (if i \<in> I then a i else undefined)} \<subseteq> J}"
 | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 | have *: "finite (Pi\<^sub>E I S)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 | if "finite J" and "\<forall>i\<in>I. finite (S i)" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | proof (rule finite_subset) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | show "Pi\<^sub>E I S \<subseteq> ?F" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | apply safe | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | using J_def apply blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | by (metis DiffI PiE_E a singletonD) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | show "finite ?F" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | proof (rule finite_restricted_funspace [OF \<open>finite J\<close>]) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | show "finite (\<Union> (S ` J))" | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 | using that J_def by blast | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | show ?lhs | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 | using R by (auto simp: * J_def) | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | qed | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 | qed auto | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | |
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | corollary finite_funcset_iff: | 
| 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 |   "finite(I \<rightarrow>\<^sub>E S) \<longleftrightarrow> (\<exists>a. S \<subseteq> {a}) \<or> I = {} \<or> finite I \<and> finite S"
 | 
| 78274 
f44aec9a6894
Last of the HOL Light metric space imports, and some supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
78250diff
changeset | 748 | by (fastforce simp: finite_PiE_iff PiE_eq_empty_iff dest: subset_singletonD) | 
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | |
| 79589 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 750 | subsection \<open>Misc other resultd\<close> | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 751 | |
| 71226 | 752 | lemma lists_lepoll_mono: | 
| 753 | assumes "A \<lesssim> B" shows "lists A \<lesssim> lists B" | |
| 754 | proof - | |
| 755 | obtain f where f: "inj_on f A" "f ` A \<subseteq> B" | |
| 756 | by (meson assms lepoll_def) | |
| 757 | moreover have "inj_on (map f) (lists A)" | |
| 758 | using f unfolding inj_on_def | |
| 759 | by clarsimp (metis list.inj_map_strong) | |
| 760 | ultimately show ?thesis | |
| 761 | unfolding lepoll_def by force | |
| 762 | qed | |
| 763 | ||
| 764 | lemma lepoll_lists: "A \<lesssim> lists A" | |
| 765 | unfolding lepoll_def inj_on_def by(rule_tac x="\<lambda>x. [x]" in exI) auto | |
| 766 | ||
| 79589 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 767 | text \<open>Dedekind's definition of infinite set\<close> | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 768 | |
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 769 | lemma infinite_iff_psubset: "infinite A \<longleftrightarrow> (\<exists>B. B \<subset> A \<and> A\<approx>B)" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 770 | proof | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 771 | assume "infinite A" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 772 | then obtain f :: "nat \<Rightarrow> 'a" where "inj f" and f: "range f \<subseteq> A" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 773 | by (meson infinite_countable_subset) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 774 | define C where "C \<equiv> A - range f" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 775 |   have C: "A = range f \<union> C" "range f \<inter> C = {}"
 | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 776 | using f by (auto simp: C_def) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 777 | have *: "range (f \<circ> Suc) \<subset> range f" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 778 | using inj_eq [OF \<open>inj f\<close>] by (fastforce simp: set_eq_iff) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 779 | have "range f \<union> C \<approx> range (f \<circ> Suc) \<union> C" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 780 | proof (intro Un_eqpoll_cong) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 781 | show "range f \<approx> range (f \<circ> Suc)" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 782 | by (meson \<open>inj f\<close> eqpoll_refl inj_Suc inj_compose inj_on_image_eqpoll_2) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 783 | show "disjnt (range f) C" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 784 | by (simp add: C disjnt_def) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 785 | then show "disjnt (range (f \<circ> Suc)) C" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 786 | using "*" disjnt_subset1 by blast | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 787 | qed auto | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 788 | moreover have "range (f \<circ> Suc) \<union> C \<subset> A" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 789 | using "*" f C_def by blast | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 790 | ultimately show "\<exists>B\<subset>A. A \<approx> B" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 791 | by (metis C(1)) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 792 | next | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 793 | assume "\<exists>B\<subset>A. A \<approx> B" then show "infinite A" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 794 | by (metis card_subset_eq eqpoll_finite_iff eqpoll_iff_card psubsetE) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 795 | qed | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 796 | |
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 797 | lemma infinite_iff_psubset_le: "infinite A \<longleftrightarrow> (\<exists>B. B \<subset> A \<and> A \<lesssim> B)" | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 798 | by (meson eqpoll_imp_lepoll infinite_iff_psubset lepoll_antisym psubsetE subset_imp_lepoll) | 
| 
9dee3b4fdb06
new lemmas involving Ramsey numbers, infinite sets
 paulson <lp15@cam.ac.uk> parents: 
78274diff
changeset | 799 | |
| 69735 
8230dca028eb
the theory of Equipollence, and moving Fpow from Cardinals into Main
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | end |