author | wenzelm |
Mon, 03 Feb 2025 20:22:51 +0100 | |
changeset 82073 | 879be333e939 |
parent 81125 | ec121999a9cb |
child 82248 | e8c96013ea8a |
permissions | -rw-r--r-- |
55056 | 1 |
(* Title: HOL/BNF_Cardinal_Order_Relation.thy |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Andrei Popescu, TU Muenchen |
75624 | 3 |
Author: Jan van Brügge, TU Muenchen |
4 |
Copyright 2012, 2022 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
|
55059 | 6 |
Cardinal-order relations as needed by bounded natural functors. |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
|
60758 | 9 |
section \<open>Cardinal-Order Relations as Needed by Bounded Natural Functors\<close> |
54473 | 10 |
|
55056 | 11 |
theory BNF_Cardinal_Order_Relation |
76951 | 12 |
imports Zorn BNF_Wellorder_Constructions |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
13 |
begin |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
|
60758 | 15 |
text\<open>In this section, we define cardinal-order relations to be minim well-orders |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
16 |
on their field. Then we define the cardinal of a set to be {\em some} cardinal-order |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
17 |
relation on that set, which will be unique up to order isomorphism. Then we study |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
18 |
the connection between cardinals and: |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
19 |
\begin{itemize} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
20 |
\item standard set-theoretic constructions: products, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
21 |
sums, unions, lists, powersets, set-of finite sets operator; |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
22 |
\item finiteness and infiniteness (in particular, with the numeric cardinal operator |
61799 | 23 |
for finite sets, \<open>card\<close>, from the theory \<open>Finite_Sets.thy\<close>). |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
24 |
\end{itemize} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
25 |
% |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
26 |
On the way, we define the canonical $\omega$ cardinal and finite cardinals. We also |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
27 |
define (again, up to order isomorphism) the successor of a cardinal, and show that |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
28 |
any cardinal admits a successor. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
29 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
30 |
Main results of this section are the existence of cardinal relations and the |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
31 |
facts that, in the presence of infiniteness, |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
32 |
most of the standard set-theoretic constructions (except for the powerset) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
33 |
{\em do not increase cardinality}. In particular, e.g., the set of words/lists over |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
34 |
any infinite set has the same cardinality (hence, is in bijection) with that set. |
60758 | 35 |
\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
36 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
37 |
|
60758 | 38 |
subsection \<open>Cardinal orders\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
39 |
|
60758 | 40 |
text\<open>A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the |
61799 | 41 |
order-embedding relation, \<open>\<le>o\<close> (which is the same as being {\em minimal} w.r.t. the |
42 |
strict order-embedding relation, \<open><o\<close>), among all the well-orders on its field.\<close> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
43 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
44 |
definition card_order_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool" |
76951 | 45 |
where |
46 |
"card_order_on A r \<equiv> well_order_on A r \<and> (\<forall>r'. well_order_on A r' \<longrightarrow> r \<le>o r')" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
47 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
48 |
abbreviation "Card_order r \<equiv> card_order_on (Field r) r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
49 |
abbreviation "card_order r \<equiv> card_order_on UNIV r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
50 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
51 |
lemma card_order_on_well_order_on: |
76951 | 52 |
assumes "card_order_on A r" |
53 |
shows "well_order_on A r" |
|
54 |
using assms unfolding card_order_on_def by simp |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
55 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
56 |
lemma card_order_on_Card_order: |
76951 | 57 |
"card_order_on A r \<Longrightarrow> A = Field r \<and> Card_order r" |
58 |
unfolding card_order_on_def using well_order_on_Field by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
59 |
|
60758 | 60 |
text\<open>The existence of a cardinal relation on any given set (which will mean |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
61 |
that any set has a cardinal) follows from two facts: |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
62 |
\begin{itemize} |
61799 | 63 |
\item Zermelo's theorem (proved in \<open>Zorn.thy\<close> as theorem \<open>well_order_on\<close>), |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
64 |
which states that on any given set there exists a well-order; |
61799 | 65 |
\item The well-founded-ness of \<open><o\<close>, ensuring that then there exists a minimal |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
66 |
such well-order, i.e., a cardinal order. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
67 |
\end{itemize} |
60758 | 68 |
\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
69 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
70 |
theorem card_order_on: "\<exists>r. card_order_on A r" |
76951 | 71 |
proof - |
72 |
define R where "R \<equiv> {r. well_order_on A r}" |
|
73 |
have "R \<noteq> {} \<and> (\<forall>r \<in> R. Well_order r)" |
|
74 |
using well_order_on[of A] R_def well_order_on_Well_order by blast |
|
75 |
with exists_minim_Well_order[of R] show ?thesis |
|
76 |
by (auto simp: R_def card_order_on_def) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
77 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
78 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
79 |
lemma card_order_on_ordIso: |
76951 | 80 |
assumes CO: "card_order_on A r" and CO': "card_order_on A r'" |
81 |
shows "r =o r'" |
|
82 |
using assms unfolding card_order_on_def |
|
83 |
using ordIso_iff_ordLeq by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
84 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
85 |
lemma Card_order_ordIso: |
76951 | 86 |
assumes CO: "Card_order r" and ISO: "r' =o r" |
87 |
shows "Card_order r'" |
|
88 |
using ISO unfolding ordIso_def |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
89 |
proof(unfold card_order_on_def, auto) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
90 |
fix p' assume "well_order_on (Field r') p'" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
91 |
hence 0: "Well_order p' \<and> Field p' = Field r'" |
76951 | 92 |
using well_order_on_Well_order by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
93 |
obtain f where 1: "iso r' r f" and 2: "Well_order r \<and> Well_order r'" |
76951 | 94 |
using ISO unfolding ordIso_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
95 |
hence 3: "inj_on f (Field r') \<and> f ` (Field r') = Field r" |
76951 | 96 |
by (auto simp add: iso_iff embed_inj_on) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
97 |
let ?p = "dir_image p' f" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
98 |
have 4: "p' =o ?p \<and> Well_order ?p" |
76951 | 99 |
using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image) |
100 |
moreover have "Field ?p = Field r" |
|
101 |
using 0 3 by (auto simp add: dir_image_Field) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
102 |
ultimately have "well_order_on (Field r) ?p" by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
103 |
hence "r \<le>o ?p" using CO unfolding card_order_on_def by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
104 |
thus "r' \<le>o p'" |
76951 | 105 |
using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
106 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
107 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
108 |
lemma Card_order_ordIso2: |
76951 | 109 |
assumes CO: "Card_order r" and ISO: "r =o r'" |
110 |
shows "Card_order r'" |
|
111 |
using assms Card_order_ordIso ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
112 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
113 |
|
60758 | 114 |
subsection \<open>Cardinal of a set\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
115 |
|
60758 | 116 |
text\<open>We define the cardinal of set to be {\em some} cardinal order on that set. |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
117 |
We shall prove that this notion is unique up to order isomorphism, meaning |
60758 | 118 |
that order isomorphism shall be the true identity of cardinals.\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
119 |
|
81125 | 120 |
definition card_of :: "'a set \<Rightarrow> 'a rel" (\<open>(\<open>open_block notation=\<open>mixfix card_of\<close>\<close>|_|)\<close>) |
76951 | 121 |
where "card_of A = (SOME r. card_order_on A r)" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
122 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
123 |
lemma card_of_card_order_on: "card_order_on A |A|" |
76951 | 124 |
unfolding card_of_def by (auto simp add: card_order_on someI_ex) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
125 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
126 |
lemma card_of_well_order_on: "well_order_on A |A|" |
76951 | 127 |
using card_of_card_order_on card_order_on_def by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
128 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
129 |
lemma Field_card_of: "Field |A| = A" |
76951 | 130 |
using card_of_card_order_on[of A] unfolding card_order_on_def |
131 |
using well_order_on_Field by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
132 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
133 |
lemma card_of_Card_order: "Card_order |A|" |
76951 | 134 |
by (simp only: card_of_card_order_on Field_card_of) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
135 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
136 |
corollary ordIso_card_of_imp_Card_order: |
76951 | 137 |
"r =o |A| \<Longrightarrow> Card_order r" |
138 |
using card_of_Card_order Card_order_ordIso by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
139 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
140 |
lemma card_of_Well_order: "Well_order |A|" |
76951 | 141 |
using card_of_Card_order unfolding card_order_on_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
142 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
143 |
lemma card_of_refl: "|A| =o |A|" |
76951 | 144 |
using card_of_Well_order ordIso_reflexive by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
145 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
146 |
lemma card_of_least: "well_order_on A r \<Longrightarrow> |A| \<le>o r" |
76951 | 147 |
using card_of_card_order_on unfolding card_order_on_def by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
148 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
149 |
lemma card_of_ordIso: |
76951 | 150 |
"(\<exists>f. bij_betw f A B) = ( |A| =o |B| )" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
151 |
proof(auto) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
152 |
fix f assume *: "bij_betw f A B" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
153 |
then obtain r where "well_order_on B r \<and> |A| =o r" |
76951 | 154 |
using Well_order_iso_copy card_of_well_order_on by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
155 |
hence "|B| \<le>o |A|" using card_of_least |
76951 | 156 |
ordLeq_ordIso_trans ordIso_symmetric by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
157 |
moreover |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
158 |
{let ?g = "inv_into A f" |
76951 | 159 |
have "bij_betw ?g B A" using * bij_betw_inv_into by blast |
160 |
then obtain r where "well_order_on A r \<and> |B| =o r" |
|
161 |
using Well_order_iso_copy card_of_well_order_on by blast |
|
162 |
hence "|A| \<le>o |B|" |
|
163 |
using card_of_least ordLeq_ordIso_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
164 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
165 |
ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
166 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
167 |
assume "|A| =o |B|" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
168 |
then obtain f where "iso ( |A| ) ( |B| ) f" |
76951 | 169 |
unfolding ordIso_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
170 |
hence "bij_betw f A B" unfolding iso_def Field_card_of by simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
171 |
thus "\<exists>f. bij_betw f A B" by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
172 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
173 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
174 |
lemma card_of_ordLeq: |
76951 | 175 |
"(\<exists>f. inj_on f A \<and> f ` A \<le> B) = ( |A| \<le>o |B| )" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
176 |
proof(auto) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
177 |
fix f assume *: "inj_on f A" and **: "f ` A \<le> B" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
178 |
{assume "|B| <o |A|" |
76951 | 179 |
hence "|B| \<le>o |A|" using ordLeq_iff_ordLess_or_ordIso by blast |
180 |
then obtain g where "embed ( |B| ) ( |A| ) g" |
|
181 |
unfolding ordLeq_def by auto |
|
182 |
hence 1: "inj_on g B \<and> g ` B \<le> A" using embed_inj_on[of "|B|" "|A|" "g"] |
|
183 |
card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"] |
|
184 |
embed_Field[of "|B|" "|A|" g] by auto |
|
185 |
obtain h where "bij_betw h A B" |
|
186 |
using * ** 1 Schroeder_Bernstein[of f] by fastforce |
|
187 |
hence "|A| \<le>o |B|" using card_of_ordIso ordIso_iff_ordLeq by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
188 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
189 |
thus "|A| \<le>o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"] |
76951 | 190 |
by (auto simp: card_of_Well_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
191 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
192 |
assume *: "|A| \<le>o |B|" |
76951 | 193 |
obtain f where "embed |A| |B| f" |
194 |
using * unfolding ordLeq_def by auto |
|
195 |
hence "inj_on f A \<and> f ` A \<le> B" |
|
196 |
using embed_inj_on[of "|A|" "|B|"] card_of_Well_order embed_Field[of "|A|" "|B|"] |
|
197 |
by (auto simp: Field_card_of) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
198 |
thus "\<exists>f. inj_on f A \<and> f ` A \<le> B" by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
199 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
200 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
201 |
lemma card_of_ordLeq2: |
76951 | 202 |
"A \<noteq> {} \<Longrightarrow> (\<exists>g. g ` B = A) = ( |A| \<le>o |B| )" |
203 |
using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
204 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
205 |
lemma card_of_ordLess: |
76951 | 206 |
"(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = ( |B| <o |A| )" |
207 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
208 |
have "(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = (\<not> |A| \<le>o |B| )" |
76951 | 209 |
using card_of_ordLeq by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
210 |
also have "\<dots> = ( |B| <o |A| )" |
76951 | 211 |
using not_ordLeq_iff_ordLess by (auto intro: card_of_Well_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
212 |
finally show ?thesis . |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
213 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
214 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
215 |
lemma card_of_ordLess2: |
76951 | 216 |
"B \<noteq> {} \<Longrightarrow> (\<not>(\<exists>f. f ` A = B)) = ( |A| <o |B| )" |
217 |
using card_of_ordLess[of B A] inj_on_iff_surj[of B A] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
218 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
219 |
lemma card_of_ordIsoI: |
76951 | 220 |
assumes "bij_betw f A B" |
221 |
shows "|A| =o |B|" |
|
222 |
using assms unfolding card_of_ordIso[symmetric] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
223 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
224 |
lemma card_of_ordLeqI: |
76951 | 225 |
assumes "inj_on f A" and "\<And> a. a \<in> A \<Longrightarrow> f a \<in> B" |
226 |
shows "|A| \<le>o |B|" |
|
227 |
using assms unfolding card_of_ordLeq[symmetric] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
228 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
229 |
lemma card_of_unique: |
76951 | 230 |
"card_order_on A r \<Longrightarrow> r =o |A|" |
231 |
by (simp only: card_order_on_ordIso card_of_card_order_on) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
232 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
233 |
lemma card_of_mono1: |
76951 | 234 |
"A \<le> B \<Longrightarrow> |A| \<le>o |B|" |
235 |
using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
236 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
237 |
lemma card_of_mono2: |
76951 | 238 |
assumes "r \<le>o r'" |
239 |
shows "|Field r| \<le>o |Field r'|" |
|
240 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
241 |
obtain f where |
76951 | 242 |
1: "well_order_on (Field r) r \<and> well_order_on (Field r) r \<and> embed r r' f" |
243 |
using assms unfolding ordLeq_def |
|
244 |
by (auto simp add: well_order_on_Well_order) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
245 |
hence "inj_on f (Field r) \<and> f ` (Field r) \<le> Field r'" |
76951 | 246 |
by (auto simp add: embed_inj_on embed_Field) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
247 |
thus "|Field r| \<le>o |Field r'|" using card_of_ordLeq by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
248 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
249 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
250 |
lemma card_of_cong: "r =o r' \<Longrightarrow> |Field r| =o |Field r'|" |
76951 | 251 |
by (simp add: ordIso_iff_ordLeq card_of_mono2) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
252 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
253 |
lemma card_of_Field_ordIso: |
76951 | 254 |
assumes "Card_order r" |
255 |
shows "|Field r| =o r" |
|
256 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
257 |
have "card_order_on (Field r) r" |
76951 | 258 |
using assms card_order_on_Card_order by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
259 |
moreover have "card_order_on (Field r) |Field r|" |
76951 | 260 |
using card_of_card_order_on by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
261 |
ultimately show ?thesis using card_order_on_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
262 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
263 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
264 |
lemma Card_order_iff_ordIso_card_of: |
76951 | 265 |
"Card_order r = (r =o |Field r| )" |
266 |
using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
267 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
268 |
lemma Card_order_iff_ordLeq_card_of: |
76951 | 269 |
"Card_order r = (r \<le>o |Field r| )" |
270 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
271 |
have "Card_order r = (r =o |Field r| )" |
76951 | 272 |
unfolding Card_order_iff_ordIso_card_of by simp |
273 |
also have "\<dots> = (r \<le>o |Field r| \<and> |Field r| \<le>o r)" |
|
274 |
unfolding ordIso_iff_ordLeq by simp |
|
275 |
also have "\<dots> = (r \<le>o |Field r| )" |
|
276 |
using card_of_least |
|
277 |
by (auto simp: card_of_least ordLeq_Well_order_simp) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
278 |
finally show ?thesis . |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
279 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
280 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
281 |
lemma Card_order_iff_Restr_underS: |
76951 | 282 |
assumes "Well_order r" |
283 |
shows "Card_order r = (\<forall>a \<in> Field r. Restr r (underS r a) <o |Field r| )" |
|
284 |
using assms ordLeq_iff_ordLess_Restr card_of_Well_order |
|
285 |
unfolding Card_order_iff_ordLeq_card_of by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
286 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
287 |
lemma card_of_underS: |
76951 | 288 |
assumes r: "Card_order r" and a: "a \<in> Field r" |
289 |
shows "|underS r a| <o r" |
|
290 |
proof - |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
291 |
let ?A = "underS r a" let ?r' = "Restr r ?A" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
292 |
have 1: "Well_order r" |
76951 | 293 |
using r unfolding card_order_on_def by simp |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
294 |
have "Well_order ?r'" using 1 Well_order_Restr by auto |
76951 | 295 |
with card_of_card_order_on have "|Field ?r'| \<le>o ?r'" |
296 |
unfolding card_order_on_def by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
297 |
moreover have "Field ?r' = ?A" |
76951 | 298 |
using 1 wo_rel.underS_ofilter Field_Restr_ofilter |
299 |
unfolding wo_rel_def by fastforce |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
300 |
ultimately have "|?A| \<le>o ?r'" by simp |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
301 |
also have "?r' <o |Field r|" |
76951 | 302 |
using 1 a r Card_order_iff_Restr_underS by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
303 |
also have "|Field r| =o r" |
76951 | 304 |
using r ordIso_symmetric unfolding Card_order_iff_ordIso_card_of by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
305 |
finally show ?thesis . |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
306 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
307 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
308 |
lemma ordLess_Field: |
76951 | 309 |
assumes "r <o r'" |
310 |
shows "|Field r| <o r'" |
|
311 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
312 |
have "well_order_on (Field r) r" using assms unfolding ordLess_def |
76951 | 313 |
by (auto simp add: well_order_on_Well_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
314 |
hence "|Field r| \<le>o r" using card_of_least by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
315 |
thus ?thesis using assms ordLeq_ordLess_trans by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
316 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
317 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
318 |
lemma internalize_card_of_ordLeq: |
76951 | 319 |
"( |A| \<le>o r) = (\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r)" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
320 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
321 |
assume "|A| \<le>o r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
322 |
then obtain p where 1: "Field p \<le> Field r \<and> |A| =o p \<and> p \<le>o r" |
76951 | 323 |
using internalize_ordLeq[of "|A|" r] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
324 |
hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
325 |
hence "|Field p| =o p" using card_of_Field_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
326 |
hence "|A| =o |Field p| \<and> |Field p| \<le>o r" |
76951 | 327 |
using 1 ordIso_equivalence ordIso_ordLeq_trans by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
328 |
thus "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r" using 1 by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
329 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
330 |
assume "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
331 |
thus "|A| \<le>o r" using ordIso_ordLeq_trans by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
332 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
333 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
334 |
lemma internalize_card_of_ordLeq2: |
76951 | 335 |
"( |A| \<le>o |C| ) = (\<exists>B \<le> C. |A| =o |B| \<and> |B| \<le>o |C| )" |
336 |
using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
337 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
338 |
|
60758 | 339 |
subsection \<open>Cardinals versus set operations on arbitrary sets\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
340 |
|
60758 | 341 |
text\<open>Here we embark in a long journey of simple results showing |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
342 |
that the standard set-theoretic operations are well-behaved w.r.t. the notion of |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
343 |
cardinal -- essentially, this means that they preserve the ``cardinal identity" |
61799 | 344 |
\<open>=o\<close> and are monotonic w.r.t. \<open>\<le>o\<close>. |
60758 | 345 |
\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
346 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
347 |
lemma card_of_empty: "|{}| \<le>o |A|" |
76951 | 348 |
using card_of_ordLeq inj_on_id by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
349 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
350 |
lemma card_of_empty1: |
76951 | 351 |
assumes "Well_order r \<or> Card_order r" |
352 |
shows "|{}| \<le>o r" |
|
353 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
354 |
have "Well_order r" using assms unfolding card_order_on_def by auto |
76951 | 355 |
hence "|Field r| \<le>o r" |
356 |
using assms card_of_least by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
357 |
moreover have "|{}| \<le>o |Field r|" by (simp add: card_of_empty) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
358 |
ultimately show ?thesis using ordLeq_transitive by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
359 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
360 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
361 |
corollary Card_order_empty: |
76951 | 362 |
"Card_order r \<Longrightarrow> |{}| \<le>o r" by (simp add: card_of_empty1) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
363 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
364 |
lemma card_of_empty2: |
76951 | 365 |
assumes "|A| =o |{}|" |
366 |
shows "A = {}" |
|
367 |
using assms card_of_ordIso[of A] bij_betw_empty2 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
368 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
369 |
lemma card_of_empty3: |
76951 | 370 |
assumes "|A| \<le>o |{}|" |
371 |
shows "A = {}" |
|
372 |
using assms |
|
373 |
by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2 |
|
374 |
ordLeq_Well_order_simp) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
375 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
376 |
lemma card_of_empty_ordIso: |
76951 | 377 |
"|{}::'a set| =o |{}::'b set|" |
378 |
using card_of_ordIso unfolding bij_betw_def inj_on_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
379 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
380 |
lemma card_of_image: |
76951 | 381 |
"|f ` A| \<le>o |A|" |
382 |
proof(cases "A = {}") |
|
383 |
case False |
|
67091 | 384 |
hence "f ` A \<noteq> {}" by auto |
76951 | 385 |
thus ?thesis |
386 |
using card_of_ordLeq2[of "f ` A" A] by auto |
|
387 |
qed (simp add: card_of_empty) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
388 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
389 |
lemma surj_imp_ordLeq: |
76951 | 390 |
assumes "B \<subseteq> f ` A" |
391 |
shows "|B| \<le>o |A|" |
|
392 |
proof - |
|
393 |
have "|B| \<le>o |f ` A|" using assms card_of_mono1 by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
394 |
thus ?thesis using card_of_image ordLeq_transitive by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
395 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
396 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
397 |
lemma card_of_singl_ordLeq: |
76951 | 398 |
assumes "A \<noteq> {}" |
399 |
shows "|{b}| \<le>o |A|" |
|
400 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
401 |
obtain a where *: "a \<in> A" using assms by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
402 |
let ?h = "\<lambda> b'::'b. if b' = b then a else undefined" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
403 |
have "inj_on ?h {b} \<and> ?h ` {b} \<le> A" |
76951 | 404 |
using * unfolding inj_on_def by auto |
55603
48596c45bf7f
less flex-flex pairs (thanks to Lars' statistics)
traytel
parents:
55206
diff
changeset
|
405 |
thus ?thesis unfolding card_of_ordLeq[symmetric] by (intro exI) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
406 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
407 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
408 |
corollary Card_order_singl_ordLeq: |
76951 | 409 |
"\<lbrakk>Card_order r; Field r \<noteq> {}\<rbrakk> \<Longrightarrow> |{b}| \<le>o r" |
410 |
using card_of_singl_ordLeq[of "Field r" b] |
|
411 |
card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
412 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
413 |
lemma card_of_Pow: "|A| <o |Pow A|" |
77140
9a60c1759543
Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents:
76951
diff
changeset
|
414 |
using card_of_ordLess2[of "Pow A" A] Cantors_theorem[of A] |
76951 | 415 |
Pow_not_empty[of A] by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
416 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
417 |
corollary Card_order_Pow: |
76951 | 418 |
"Card_order r \<Longrightarrow> r <o |Pow(Field r)|" |
419 |
using card_of_Pow card_of_Field_ordIso ordIso_ordLess_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
420 |
|
76951 | 421 |
lemma card_of_Plus1: "|A| \<le>o |A <+> B|" and card_of_Plus2: "|B| \<le>o |A <+> B|" |
422 |
using card_of_ordLeq by force+ |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
423 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
424 |
corollary Card_order_Plus1: |
76951 | 425 |
"Card_order r \<Longrightarrow> r \<le>o |(Field r) <+> B|" |
426 |
using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
427 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
428 |
corollary Card_order_Plus2: |
76951 | 429 |
"Card_order r \<Longrightarrow> r \<le>o |A <+> (Field r)|" |
430 |
using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
431 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
432 |
lemma card_of_Plus_empty1: "|A| =o |A <+> {}|" |
76951 | 433 |
proof - |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
434 |
have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
435 |
thus ?thesis using card_of_ordIso by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
436 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
437 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
438 |
lemma card_of_Plus_empty2: "|A| =o |{} <+> A|" |
76951 | 439 |
proof - |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
440 |
have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
441 |
thus ?thesis using card_of_ordIso by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
442 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
443 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
444 |
lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|" |
76951 | 445 |
proof - |
446 |
let ?f = "\<lambda>c. case c of Inl a \<Rightarrow> Inr a | Inr b \<Rightarrow> Inl b" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
447 |
have "bij_betw ?f (A <+> B) (B <+> A)" |
76951 | 448 |
unfolding bij_betw_def inj_on_def by force |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
449 |
thus ?thesis using card_of_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
450 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
451 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
452 |
lemma card_of_Plus_assoc: |
76951 | 453 |
fixes A :: "'a set" and B :: "'b set" and C :: "'c set" |
454 |
shows "|(A <+> B) <+> C| =o |A <+> B <+> C|" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
455 |
proof - |
63040 | 456 |
define f :: "('a + 'b) + 'c \<Rightarrow> 'a + 'b + 'c" |
457 |
where [abs_def]: "f k = |
|
458 |
(case k of |
|
459 |
Inl ab \<Rightarrow> |
|
460 |
(case ab of |
|
461 |
Inl a \<Rightarrow> Inl a |
|
462 |
| Inr b \<Rightarrow> Inr (Inl b)) |
|
463 |
| Inr c \<Rightarrow> Inr (Inr c))" |
|
464 |
for k |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
465 |
have "A <+> B <+> C \<subseteq> f ` ((A <+> B) <+> C)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
466 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
467 |
fix x assume x: "x \<in> A <+> B <+> C" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
468 |
show "x \<in> f ` ((A <+> B) <+> C)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
469 |
proof(cases x) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
470 |
case (Inl a) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
471 |
hence "a \<in> A" "x = f (Inl (Inl a))" |
76951 | 472 |
using x unfolding f_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
473 |
thus ?thesis by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
474 |
next |
76951 | 475 |
case (Inr bc) with x show ?thesis |
476 |
by (cases bc) (force simp: f_def)+ |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
477 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
478 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
479 |
hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)" |
63040 | 480 |
unfolding bij_betw_def inj_on_def f_def by fastforce |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
481 |
thus ?thesis using card_of_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
482 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
483 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
484 |
lemma card_of_Plus_mono1: |
76951 | 485 |
assumes "|A| \<le>o |B|" |
486 |
shows "|A <+> C| \<le>o |B <+> C|" |
|
487 |
proof - |
|
488 |
obtain f where f: "inj_on f A \<and> f ` A \<le> B" |
|
489 |
using assms card_of_ordLeq[of A] by fastforce |
|
490 |
define g where "g \<equiv> \<lambda>d. case d of Inl a \<Rightarrow> Inl(f a) | Inr (c::'c) \<Rightarrow> Inr c" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
491 |
have "inj_on g (A <+> C) \<and> g ` (A <+> C) \<le> (B <+> C)" |
76951 | 492 |
using f unfolding inj_on_def g_def by force |
55811 | 493 |
thus ?thesis using card_of_ordLeq by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
494 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
495 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
496 |
corollary ordLeq_Plus_mono1: |
76951 | 497 |
assumes "r \<le>o r'" |
498 |
shows "|(Field r) <+> C| \<le>o |(Field r') <+> C|" |
|
499 |
using assms card_of_mono2 card_of_Plus_mono1 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
500 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
501 |
lemma card_of_Plus_mono2: |
76951 | 502 |
assumes "|A| \<le>o |B|" |
503 |
shows "|C <+> A| \<le>o |C <+> B|" |
|
504 |
using card_of_Plus_mono1[OF assms] |
|
505 |
by (blast intro: card_of_Plus_commute ordIso_ordLeq_trans ordLeq_ordIso_trans) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
506 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
507 |
corollary ordLeq_Plus_mono2: |
76951 | 508 |
assumes "r \<le>o r'" |
509 |
shows "|A <+> (Field r)| \<le>o |A <+> (Field r')|" |
|
510 |
using assms card_of_mono2 card_of_Plus_mono2 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
511 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
512 |
lemma card_of_Plus_mono: |
76951 | 513 |
assumes "|A| \<le>o |B|" and "|C| \<le>o |D|" |
514 |
shows "|A <+> C| \<le>o |B <+> D|" |
|
515 |
using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B] |
|
516 |
ordLeq_transitive by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
517 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
518 |
corollary ordLeq_Plus_mono: |
76951 | 519 |
assumes "r \<le>o r'" and "p \<le>o p'" |
520 |
shows "|(Field r) <+> (Field p)| \<le>o |(Field r') <+> (Field p')|" |
|
521 |
using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
522 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
523 |
lemma card_of_Plus_cong1: |
76951 | 524 |
assumes "|A| =o |B|" |
525 |
shows "|A <+> C| =o |B <+> C|" |
|
526 |
using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
527 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
528 |
corollary ordIso_Plus_cong1: |
76951 | 529 |
assumes "r =o r'" |
530 |
shows "|(Field r) <+> C| =o |(Field r') <+> C|" |
|
531 |
using assms card_of_cong card_of_Plus_cong1 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
532 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
533 |
lemma card_of_Plus_cong2: |
76951 | 534 |
assumes "|A| =o |B|" |
535 |
shows "|C <+> A| =o |C <+> B|" |
|
536 |
using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
537 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
538 |
corollary ordIso_Plus_cong2: |
76951 | 539 |
assumes "r =o r'" |
540 |
shows "|A <+> (Field r)| =o |A <+> (Field r')|" |
|
541 |
using assms card_of_cong card_of_Plus_cong2 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
542 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
543 |
lemma card_of_Plus_cong: |
76951 | 544 |
assumes "|A| =o |B|" and "|C| =o |D|" |
545 |
shows "|A <+> C| =o |B <+> D|" |
|
546 |
using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
547 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
548 |
corollary ordIso_Plus_cong: |
76951 | 549 |
assumes "r =o r'" and "p =o p'" |
550 |
shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|" |
|
551 |
using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
552 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
553 |
lemma card_of_Un_Plus_ordLeq: |
76951 | 554 |
"|A \<union> B| \<le>o |A <+> B|" |
555 |
proof - |
|
556 |
let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c" |
|
557 |
have "inj_on ?f (A \<union> B) \<and> ?f ` (A \<union> B) \<le> A <+> B" |
|
558 |
unfolding inj_on_def by auto |
|
559 |
thus ?thesis using card_of_ordLeq by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
560 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
561 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
562 |
lemma card_of_Times1: |
76951 | 563 |
assumes "A \<noteq> {}" |
564 |
shows "|B| \<le>o |B \<times> A|" |
|
565 |
proof(cases "B = {}") |
|
566 |
case False |
|
56077 | 567 |
have "fst `(B \<times> A) = B" using assms by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
568 |
thus ?thesis using inj_on_iff_surj[of B "B \<times> A"] |
76951 | 569 |
card_of_ordLeq False by blast |
570 |
qed (simp add: card_of_empty) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
571 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
572 |
lemma card_of_Times_commute: "|A \<times> B| =o |B \<times> A|" |
76951 | 573 |
proof - |
574 |
have "bij_betw (\<lambda>(a,b). (b,a)) (A \<times> B) (B \<times> A)" |
|
575 |
unfolding bij_betw_def inj_on_def by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
576 |
thus ?thesis using card_of_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
577 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
578 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
579 |
lemma card_of_Times2: |
76951 | 580 |
assumes "A \<noteq> {}" shows "|B| \<le>o |A \<times> B|" |
581 |
using assms card_of_Times1[of A B] card_of_Times_commute[of B A] |
|
582 |
ordLeq_ordIso_trans by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
583 |
|
54475 | 584 |
corollary Card_order_Times1: |
76951 | 585 |
"\<lbrakk>Card_order r; B \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |(Field r) \<times> B|" |
586 |
using card_of_Times1[of B] card_of_Field_ordIso |
|
587 |
ordIso_ordLeq_trans ordIso_symmetric by blast |
|
54475 | 588 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
589 |
corollary Card_order_Times2: |
76951 | 590 |
"\<lbrakk>Card_order r; A \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |A \<times> (Field r)|" |
591 |
using card_of_Times2[of A] card_of_Field_ordIso |
|
592 |
ordIso_ordLeq_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
593 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
594 |
lemma card_of_Times3: "|A| \<le>o |A \<times> A|" |
76951 | 595 |
using card_of_Times1[of A] |
596 |
by(cases "A = {}", simp add: card_of_empty) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
597 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
598 |
lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \<times> (UNIV::bool set)|" |
76951 | 599 |
proof - |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
600 |
let ?f = "\<lambda>c::'a + 'a. case c of Inl a \<Rightarrow> (a,True) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
601 |
|Inr a \<Rightarrow> (a,False)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
602 |
have "bij_betw ?f (A <+> A) (A \<times> (UNIV::bool set))" |
76951 | 603 |
proof - |
604 |
have "\<And>c1 c2. ?f c1 = ?f c2 \<Longrightarrow> c1 = c2" |
|
605 |
by (force split: sum.split_asm) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
606 |
moreover |
76951 | 607 |
have "\<And>c. c \<in> A <+> A \<Longrightarrow> ?f c \<in> A \<times> (UNIV::bool set)" |
608 |
by (force split: sum.split_asm) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
609 |
moreover |
76951 | 610 |
{fix a bl assume "(a,bl) \<in> A \<times> (UNIV::bool set)" |
611 |
hence "(a,bl) \<in> ?f ` ( A <+> A)" |
|
612 |
by (cases bl) (force split: sum.split_asm)+ |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
613 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
614 |
ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
615 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
616 |
thus ?thesis using card_of_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
617 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
618 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
619 |
lemma card_of_Times_mono1: |
76951 | 620 |
assumes "|A| \<le>o |B|" |
621 |
shows "|A \<times> C| \<le>o |B \<times> C|" |
|
622 |
proof - |
|
623 |
obtain f where f: "inj_on f A \<and> f ` A \<le> B" |
|
624 |
using assms card_of_ordLeq[of A] by fastforce |
|
625 |
define g where "g \<equiv> (\<lambda>(a,c::'c). (f a,c))" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
626 |
have "inj_on g (A \<times> C) \<and> g ` (A \<times> C) \<le> (B \<times> C)" |
76951 | 627 |
using f unfolding inj_on_def using g_def by auto |
55811 | 628 |
thus ?thesis using card_of_ordLeq by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
629 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
630 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
631 |
corollary ordLeq_Times_mono1: |
76951 | 632 |
assumes "r \<le>o r'" |
633 |
shows "|(Field r) \<times> C| \<le>o |(Field r') \<times> C|" |
|
634 |
using assms card_of_mono2 card_of_Times_mono1 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
635 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
636 |
lemma card_of_Times_mono2: |
76951 | 637 |
assumes "|A| \<le>o |B|" |
638 |
shows "|C \<times> A| \<le>o |C \<times> B|" |
|
639 |
using assms card_of_Times_mono1[of A B C] |
|
640 |
by (blast intro: card_of_Times_commute ordIso_ordLeq_trans ordLeq_ordIso_trans) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
641 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
642 |
corollary ordLeq_Times_mono2: |
76951 | 643 |
assumes "r \<le>o r'" |
644 |
shows "|A \<times> (Field r)| \<le>o |A \<times> (Field r')|" |
|
645 |
using assms card_of_mono2 card_of_Times_mono2 by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
646 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
647 |
lemma card_of_Sigma_mono1: |
76951 | 648 |
assumes "\<forall>i \<in> I. |A i| \<le>o |B i|" |
649 |
shows "|SIGMA i : I. A i| \<le>o |SIGMA i : I. B i|" |
|
650 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
651 |
have "\<forall>i. i \<in> I \<longrightarrow> (\<exists>f. inj_on f (A i) \<and> f ` (A i) \<le> B i)" |
76951 | 652 |
using assms by (auto simp add: card_of_ordLeq) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
653 |
with choice[of "\<lambda> i f. i \<in> I \<longrightarrow> inj_on f (A i) \<and> f ` (A i) \<le> B i"] |
76951 | 654 |
obtain F where F: "\<forall>i \<in> I. inj_on (F i) (A i) \<and> (F i) ` (A i) \<le> B i" |
55811 | 655 |
by atomize_elim (auto intro: bchoice) |
76951 | 656 |
define g where "g \<equiv> (\<lambda>(i,a::'b). (i,F i a))" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
657 |
have "inj_on g (Sigma I A) \<and> g ` (Sigma I A) \<le> (Sigma I B)" |
76951 | 658 |
using F unfolding inj_on_def using g_def by force |
55811 | 659 |
thus ?thesis using card_of_ordLeq by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
660 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
661 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
662 |
lemma card_of_UNION_Sigma: |
76951 | 663 |
"|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|" |
664 |
using Ex_inj_on_UNION_Sigma [of A I] card_of_ordLeq by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
665 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
666 |
lemma card_of_bool: |
76951 | 667 |
assumes "a1 \<noteq> a2" |
668 |
shows "|UNIV::bool set| =o |{a1,a2}|" |
|
669 |
proof - |
|
670 |
let ?f = "\<lambda> bl. if bl then a1 else a2" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
671 |
have "bij_betw ?f UNIV {a1,a2}" |
76951 | 672 |
proof - |
673 |
have "\<And>bl1 bl2. ?f bl1 = ?f bl2 \<Longrightarrow> bl1 = bl2" |
|
674 |
using assms by (force split: if_split_asm) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
675 |
moreover |
76951 | 676 |
have "\<And>bl. ?f bl \<in> {a1,a2}" |
677 |
using assms by (force split: if_split_asm) |
|
678 |
ultimately show ?thesis unfolding bij_betw_def inj_on_def by force |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
679 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
680 |
thus ?thesis using card_of_ordIso by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
681 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
682 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
683 |
lemma card_of_Plus_Times_aux: |
76951 | 684 |
assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and |
685 |
LEQ: "|A| \<le>o |B|" |
|
686 |
shows "|A <+> B| \<le>o |A \<times> B|" |
|
687 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
688 |
have 1: "|UNIV::bool set| \<le>o |A|" |
76951 | 689 |
using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2] |
690 |
by (blast intro: ordIso_ordLeq_trans) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
691 |
have "|A <+> B| \<le>o |B <+> B|" |
76951 | 692 |
using LEQ card_of_Plus_mono1 by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
693 |
moreover have "|B <+> B| =o |B \<times> (UNIV::bool set)|" |
76951 | 694 |
using card_of_Plus_Times_bool by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
695 |
moreover have "|B \<times> (UNIV::bool set)| \<le>o |B \<times> A|" |
76951 | 696 |
using 1 by (simp add: card_of_Times_mono2) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
697 |
moreover have " |B \<times> A| =o |A \<times> B|" |
76951 | 698 |
using card_of_Times_commute by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
699 |
ultimately show "|A <+> B| \<le>o |A \<times> B|" |
76951 | 700 |
by (blast intro: ordLeq_transitive dest: ordLeq_ordIso_trans) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
701 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
702 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
703 |
lemma card_of_Plus_Times: |
76951 | 704 |
assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B" |
705 |
shows "|A <+> B| \<le>o |A \<times> B|" |
|
706 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
707 |
{assume "|A| \<le>o |B|" |
76951 | 708 |
hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
709 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
710 |
moreover |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
711 |
{assume "|B| \<le>o |A|" |
76951 | 712 |
hence "|B <+> A| \<le>o |B \<times> A|" |
713 |
using assms by (auto simp add: card_of_Plus_Times_aux) |
|
714 |
hence ?thesis |
|
715 |
using card_of_Plus_commute card_of_Times_commute |
|
716 |
ordIso_ordLeq_trans ordLeq_ordIso_trans by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
717 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
718 |
ultimately show ?thesis |
76951 | 719 |
using card_of_Well_order[of A] card_of_Well_order[of B] |
720 |
ordLeq_total[of "|A|"] by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
721 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
722 |
|
56191 | 723 |
lemma card_of_Times_Plus_distrib: |
61943 | 724 |
"|A \<times> (B <+> C)| =o |A \<times> B <+> A \<times> C|" (is "|?RHS| =o |?LHS|") |
56191 | 725 |
proof - |
726 |
let ?f = "\<lambda>(a, bc). case bc of Inl b \<Rightarrow> Inl (a, b) | Inr c \<Rightarrow> Inr (a, c)" |
|
727 |
have "bij_betw ?f ?RHS ?LHS" unfolding bij_betw_def inj_on_def by force |
|
728 |
thus ?thesis using card_of_ordIso by blast |
|
729 |
qed |
|
730 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
731 |
lemma card_of_ordLeq_finite: |
76951 | 732 |
assumes "|A| \<le>o |B|" and "finite B" |
733 |
shows "finite A" |
|
734 |
using assms unfolding ordLeq_def |
|
735 |
using embed_inj_on[of "|A|" "|B|"] embed_Field[of "|A|" "|B|"] |
|
736 |
Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
737 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
738 |
lemma card_of_ordLeq_infinite: |
76951 | 739 |
assumes "|A| \<le>o |B|" and "\<not> finite A" |
740 |
shows "\<not> finite B" |
|
741 |
using assms card_of_ordLeq_finite by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
742 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
743 |
lemma card_of_ordIso_finite: |
76951 | 744 |
assumes "|A| =o |B|" |
745 |
shows "finite A = finite B" |
|
746 |
using assms unfolding ordIso_def iso_def[abs_def] |
|
747 |
by (auto simp: bij_betw_finite Field_card_of) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
748 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
749 |
lemma card_of_ordIso_finite_Field: |
76951 | 750 |
assumes "Card_order r" and "r =o |A|" |
751 |
shows "finite(Field r) = finite A" |
|
752 |
using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
753 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
754 |
|
60758 | 755 |
subsection \<open>Cardinals versus set operations involving infinite sets\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
756 |
|
60758 | 757 |
text\<open>Here we show that, for infinite sets, most set-theoretic constructions |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
758 |
do not increase the cardinality. The cornerstone for this is |
61799 | 759 |
theorem \<open>Card_order_Times_same_infinite\<close>, which states that self-product |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
760 |
does not increase cardinality -- the proof of this fact adapts a standard |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
761 |
set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11 |
77172 | 762 |
at page 47 in \<^cite>\<open>"card-book"\<close>. Then everything else follows fairly easily.\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
763 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
764 |
lemma infinite_iff_card_of_nat: |
76951 | 765 |
"\<not> finite A \<longleftrightarrow> ( |UNIV::nat set| \<le>o |A| )" |
766 |
unfolding infinite_iff_countable_subset card_of_ordLeq .. |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
767 |
|
60758 | 768 |
text\<open>The next two results correspond to the ZF fact that all infinite cardinals are |
769 |
limit ordinals:\<close> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
770 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
771 |
lemma Card_order_infinite_not_under: |
76951 | 772 |
assumes CARD: "Card_order r" and INF: "\<not>finite (Field r)" |
773 |
shows "\<not> (\<exists>a. Field r = under r a)" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
774 |
proof(auto) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
775 |
have 0: "Well_order r \<and> wo_rel r \<and> Refl r" |
76951 | 776 |
using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
777 |
fix a assume *: "Field r = under r a" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
778 |
show False |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
779 |
proof(cases "a \<in> Field r") |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
780 |
assume Case1: "a \<notin> Field r" |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
781 |
hence "under r a = {}" unfolding Field_def under_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
782 |
thus False using INF * by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
783 |
next |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
784 |
let ?r' = "Restr r (underS r a)" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
785 |
assume Case2: "a \<in> Field r" |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
786 |
hence 1: "under r a = underS r a \<union> {a} \<and> a \<notin> underS r a" |
76951 | 787 |
using 0 Refl_under_underS[of r a] underS_notIn[of a r] by blast |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
788 |
have 2: "wo_rel.ofilter r (underS r a) \<and> underS r a < Field r" |
76951 | 789 |
using 0 wo_rel.underS_ofilter * 1 Case2 by fast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
790 |
hence "?r' <o r" using 0 using ofilter_ordLess by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
791 |
moreover |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
792 |
have "Field ?r' = underS r a \<and> Well_order ?r'" |
76951 | 793 |
using 2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
794 |
ultimately have "|underS r a| <o r" using ordLess_Field[of ?r'] by auto |
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
795 |
moreover have "|under r a| =o r" using * CARD card_of_Field_ordIso[of r] by auto |
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
796 |
ultimately have "|underS r a| <o |under r a|" |
76951 | 797 |
using ordIso_symmetric ordLess_ordIso_trans by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
798 |
moreover |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
799 |
{have "\<exists>f. bij_betw f (under r a) (underS r a)" |
76951 | 800 |
using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto |
801 |
hence "|under r a| =o |underS r a|" using card_of_ordIso by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
802 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
803 |
ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
804 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
805 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
806 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
807 |
lemma infinite_Card_order_limit: |
76951 | 808 |
assumes r: "Card_order r" and "\<not>finite (Field r)" |
809 |
and a: "a \<in> Field r" |
|
810 |
shows "\<exists>b \<in> Field r. a \<noteq> b \<and> (a,b) \<in> r" |
|
811 |
proof - |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
812 |
have "Field r \<noteq> under r a" |
76951 | 813 |
using assms Card_order_infinite_not_under by blast |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
814 |
moreover have "under r a \<le> Field r" |
76951 | 815 |
using under_Field . |
816 |
ultimately obtain b where b: "b \<in> Field r \<and> \<not> (b,a) \<in> r" |
|
817 |
unfolding under_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
818 |
moreover have ba: "b \<noteq> a" |
76951 | 819 |
using b r unfolding card_order_on_def well_order_on_def |
820 |
linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto |
|
67613 | 821 |
ultimately have "(a,b) \<in> r" |
76951 | 822 |
using a r unfolding card_order_on_def well_order_on_def linear_order_on_def |
823 |
total_on_def by blast |
|
824 |
thus ?thesis using b ba by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
825 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
826 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
827 |
theorem Card_order_Times_same_infinite: |
76951 | 828 |
assumes CO: "Card_order r" and INF: "\<not>finite(Field r)" |
829 |
shows "|Field r \<times> Field r| \<le>o r" |
|
830 |
proof - |
|
831 |
define phi where |
|
832 |
"phi \<equiv> \<lambda>r::'a rel. Card_order r \<and> \<not>finite(Field r) \<and> \<not> |Field r \<times> Field r| \<le>o r" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
833 |
have temp1: "\<forall>r. phi r \<longrightarrow> Well_order r" |
76951 | 834 |
unfolding phi_def card_order_on_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
835 |
have Ft: "\<not>(\<exists>r. phi r)" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
836 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
837 |
assume "\<exists>r. phi r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
838 |
hence "{r. phi r} \<noteq> {} \<and> {r. phi r} \<le> {r. Well_order r}" |
76951 | 839 |
using temp1 by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
840 |
then obtain r where 1: "phi r" and 2: "\<forall>r'. phi r' \<longrightarrow> r \<le>o r'" and |
76951 | 841 |
3: "Card_order r \<and> Well_order r" |
842 |
using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
843 |
let ?A = "Field r" let ?r' = "bsqr r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
844 |
have 4: "Well_order ?r' \<and> Field ?r' = ?A \<times> ?A \<and> |?A| =o r" |
76951 | 845 |
using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
846 |
have 5: "Card_order |?A \<times> ?A| \<and> Well_order |?A \<times> ?A|" |
76951 | 847 |
using card_of_Card_order card_of_Well_order by blast |
848 |
(* *) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
849 |
have "r <o |?A \<times> ?A|" |
76951 | 850 |
using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
851 |
moreover have "|?A \<times> ?A| \<le>o ?r'" |
76951 | 852 |
using card_of_least[of "?A \<times> ?A"] 4 by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
853 |
ultimately have "r <o ?r'" using ordLess_ordLeq_trans by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
854 |
then obtain f where 6: "embed r ?r' f" and 7: "\<not> bij_betw f ?A (?A \<times> ?A)" |
76951 | 855 |
unfolding ordLess_def embedS_def[abs_def] |
856 |
by (auto simp add: Field_bsqr) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
857 |
let ?B = "f ` ?A" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
858 |
have "|?A| =o |?B|" |
76951 | 859 |
using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
860 |
hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast |
76951 | 861 |
(* *) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
862 |
have "wo_rel.ofilter ?r' ?B" |
76951 | 863 |
using 6 embed_Field_ofilter 3 4 by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
864 |
hence "wo_rel.ofilter ?r' ?B \<and> ?B \<noteq> ?A \<times> ?A \<and> ?B \<noteq> Field ?r'" |
76951 | 865 |
using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
866 |
hence temp2: "wo_rel.ofilter ?r' ?B \<and> ?B < ?A \<times> ?A" |
76951 | 867 |
using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast |
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
868 |
have "\<not> (\<exists>a. Field r = under r a)" |
76951 | 869 |
using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
870 |
then obtain A1 where temp3: "wo_rel.ofilter r A1 \<and> A1 < ?A" and 9: "?B \<le> A1 \<times> A1" |
76951 | 871 |
using temp2 3 bsqr_ofilter[of r ?B] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
872 |
hence "|?B| \<le>o |A1 \<times> A1|" using card_of_mono1 by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
873 |
hence 10: "r \<le>o |A1 \<times> A1|" using 8 ordIso_ordLeq_trans by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
874 |
let ?r1 = "Restr r A1" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
875 |
have "?r1 <o r" using temp3 ofilter_ordLess 3 by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
876 |
moreover |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
877 |
{have "well_order_on A1 ?r1" using 3 temp3 well_order_on_Restr by blast |
76951 | 878 |
hence "|A1| \<le>o ?r1" using 3 Well_order_Restr card_of_least by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
879 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
880 |
ultimately have 11: "|A1| <o r" using ordLeq_ordLess_trans by blast |
76951 | 881 |
(* *) |
54578
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents:
54482
diff
changeset
|
882 |
have "\<not> finite (Field r)" using 1 unfolding phi_def by simp |
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents:
54482
diff
changeset
|
883 |
hence "\<not> finite ?B" using 8 3 card_of_ordIso_finite_Field[of r ?B] by blast |
55811 | 884 |
hence "\<not> finite A1" using 9 finite_cartesian_product finite_subset by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
885 |
moreover have temp4: "Field |A1| = A1 \<and> Well_order |A1| \<and> Card_order |A1|" |
76951 | 886 |
using card_of_Card_order[of A1] card_of_Well_order[of A1] |
887 |
by (simp add: Field_card_of) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
888 |
moreover have "\<not> r \<le>o | A1 |" |
76951 | 889 |
using temp4 11 3 using not_ordLeq_iff_ordLess by blast |
54578
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents:
54482
diff
changeset
|
890 |
ultimately have "\<not> finite(Field |A1| ) \<and> Card_order |A1| \<and> \<not> r \<le>o | A1 |" |
76951 | 891 |
by (simp add: card_of_card_order_on) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
892 |
hence "|Field |A1| \<times> Field |A1| | \<le>o |A1|" |
76951 | 893 |
using 2 unfolding phi_def by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
894 |
hence "|A1 \<times> A1 | \<le>o |A1|" using temp4 by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
895 |
hence "r \<le>o |A1|" using 10 ordLeq_transitive by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
896 |
thus False using 11 not_ordLess_ordLeq by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
897 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
898 |
thus ?thesis using assms unfolding phi_def by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
899 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
900 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
901 |
corollary card_of_Times_same_infinite: |
76951 | 902 |
assumes "\<not>finite A" |
903 |
shows "|A \<times> A| =o |A|" |
|
904 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
905 |
let ?r = "|A|" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
906 |
have "Field ?r = A \<and> Card_order ?r" |
76951 | 907 |
using Field_card_of card_of_Card_order[of A] by fastforce |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
908 |
hence "|A \<times> A| \<le>o |A|" |
76951 | 909 |
using Card_order_Times_same_infinite[of ?r] assms by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
910 |
thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
911 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
912 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
913 |
lemma card_of_Times_infinite: |
76951 | 914 |
assumes INF: "\<not>finite A" and NE: "B \<noteq> {}" and LEQ: "|B| \<le>o |A|" |
915 |
shows "|A \<times> B| =o |A| \<and> |B \<times> A| =o |A|" |
|
916 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
917 |
have "|A| \<le>o |A \<times> B| \<and> |A| \<le>o |B \<times> A|" |
76951 | 918 |
using assms by (simp add: card_of_Times1 card_of_Times2) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
919 |
moreover |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
920 |
{have "|A \<times> B| \<le>o |A \<times> A| \<and> |B \<times> A| \<le>o |A \<times> A|" |
76951 | 921 |
using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast |
922 |
moreover have "|A \<times> A| =o |A|" using INF card_of_Times_same_infinite by blast |
|
923 |
ultimately have "|A \<times> B| \<le>o |A| \<and> |B \<times> A| \<le>o |A|" |
|
924 |
using ordLeq_ordIso_trans[of "|A \<times> B|"] ordLeq_ordIso_trans[of "|B \<times> A|"] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
925 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
926 |
ultimately show ?thesis by (simp add: ordIso_iff_ordLeq) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
927 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
928 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
929 |
corollary Card_order_Times_infinite: |
76951 | 930 |
assumes INF: "\<not>finite(Field r)" and CARD: "Card_order r" and |
931 |
NE: "Field p \<noteq> {}" and LEQ: "p \<le>o r" |
|
932 |
shows "| (Field r) \<times> (Field p) | =o r \<and> | (Field p) \<times> (Field r) | =o r" |
|
933 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
934 |
have "|Field r \<times> Field p| =o |Field r| \<and> |Field p \<times> Field r| =o |Field r|" |
76951 | 935 |
using assms by (simp add: card_of_Times_infinite card_of_mono2) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
936 |
thus ?thesis |
76951 | 937 |
using assms card_of_Field_ordIso by (blast intro: ordIso_transitive) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
938 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
939 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
940 |
lemma card_of_Sigma_ordLeq_infinite: |
76951 | 941 |
assumes INF: "\<not>finite B" and |
942 |
LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|" |
|
943 |
shows "|SIGMA i : I. A i| \<le>o |B|" |
|
944 |
proof(cases "I = {}") |
|
945 |
case False |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
946 |
have "|SIGMA i : I. A i| \<le>o |I \<times> B|" |
76951 | 947 |
using card_of_Sigma_mono1[OF LEQ] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
948 |
moreover have "|I \<times> B| =o |B|" |
76951 | 949 |
using INF False LEQ_I by (auto simp add: card_of_Times_infinite) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
950 |
ultimately show ?thesis using ordLeq_ordIso_trans by blast |
76951 | 951 |
qed (simp add: card_of_empty) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
952 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
953 |
lemma card_of_Sigma_ordLeq_infinite_Field: |
76951 | 954 |
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and |
955 |
LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r" |
|
956 |
shows "|SIGMA i : I. A i| \<le>o r" |
|
957 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
958 |
let ?B = "Field r" |
76951 | 959 |
have 1: "r =o |?B| \<and> |?B| =o r" |
960 |
using r card_of_Field_ordIso ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
961 |
hence "|I| \<le>o |?B|" "\<forall>i \<in> I. |A i| \<le>o |?B|" |
76951 | 962 |
using LEQ_I LEQ ordLeq_ordIso_trans by blast+ |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
963 |
hence "|SIGMA i : I. A i| \<le>o |?B|" using INF LEQ |
76951 | 964 |
card_of_Sigma_ordLeq_infinite by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
965 |
thus ?thesis using 1 ordLeq_ordIso_trans by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
966 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
967 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
968 |
lemma card_of_Times_ordLeq_infinite_Field: |
76951 | 969 |
"\<lbrakk>\<not>finite (Field r); |A| \<le>o r; |B| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> |A \<times> B| \<le>o r" |
970 |
by(simp add: card_of_Sigma_ordLeq_infinite_Field) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
971 |
|
54475 | 972 |
lemma card_of_Times_infinite_simps: |
76951 | 973 |
"\<lbrakk>\<not>finite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A \<times> B| =o |A|" |
974 |
"\<lbrakk>\<not>finite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |A \<times> B|" |
|
975 |
"\<lbrakk>\<not>finite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |B \<times> A| =o |A|" |
|
976 |
"\<lbrakk>\<not>finite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |B \<times> A|" |
|
977 |
by (auto simp add: card_of_Times_infinite ordIso_symmetric) |
|
54475 | 978 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
979 |
lemma card_of_UNION_ordLeq_infinite: |
76951 | 980 |
assumes INF: "\<not>finite B" and LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|" |
981 |
shows "|\<Union>i \<in> I. A i| \<le>o |B|" |
|
982 |
proof(cases "I = {}") |
|
983 |
case False |
|
60585 | 984 |
have "|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|" |
76951 | 985 |
using card_of_UNION_Sigma by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
986 |
moreover have "|SIGMA i : I. A i| \<le>o |B|" |
76951 | 987 |
using assms card_of_Sigma_ordLeq_infinite by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
988 |
ultimately show ?thesis using ordLeq_transitive by blast |
76951 | 989 |
qed (simp add: card_of_empty) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
990 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
991 |
corollary card_of_UNION_ordLeq_infinite_Field: |
76951 | 992 |
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and |
993 |
LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r" |
|
994 |
shows "|\<Union>i \<in> I. A i| \<le>o r" |
|
995 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
996 |
let ?B = "Field r" |
76951 | 997 |
have 1: "r =o |?B| \<and> |?B| =o r" |
998 |
using r card_of_Field_ordIso ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
999 |
hence "|I| \<le>o |?B|" "\<forall>i \<in> I. |A i| \<le>o |?B|" |
76951 | 1000 |
using LEQ_I LEQ ordLeq_ordIso_trans by blast+ |
60585 | 1001 |
hence "|\<Union>i \<in> I. A i| \<le>o |?B|" using INF LEQ |
76951 | 1002 |
card_of_UNION_ordLeq_infinite by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1003 |
thus ?thesis using 1 ordLeq_ordIso_trans by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1004 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1005 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1006 |
lemma card_of_Plus_infinite1: |
76951 | 1007 |
assumes INF: "\<not>finite A" and LEQ: "|B| \<le>o |A|" |
1008 |
shows "|A <+> B| =o |A|" |
|
1009 |
proof(cases "B = {}") |
|
1010 |
case True |
|
1011 |
then show ?thesis |
|
1012 |
by (simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric) |
|
1013 |
next |
|
1014 |
case False |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1015 |
let ?Inl = "Inl::'a \<Rightarrow> 'a + 'b" let ?Inr = "Inr::'b \<Rightarrow> 'a + 'b" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1016 |
assume *: "B \<noteq> {}" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1017 |
then obtain b1 where 1: "b1 \<in> B" by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1018 |
show ?thesis |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1019 |
proof(cases "B = {b1}") |
76951 | 1020 |
case True |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1021 |
have 2: "bij_betw ?Inl A ((?Inl ` A))" |
76951 | 1022 |
unfolding bij_betw_def inj_on_def by auto |
54578
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents:
54482
diff
changeset
|
1023 |
hence 3: "\<not>finite (?Inl ` A)" |
76951 | 1024 |
using INF bij_betw_finite[of ?Inl A] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1025 |
let ?A' = "?Inl ` A \<union> {?Inr b1}" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1026 |
obtain g where "bij_betw g (?Inl ` A) ?A'" |
76951 | 1027 |
using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto |
1028 |
moreover have "?A' = A <+> B" using True by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1029 |
ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp |
67091 | 1030 |
hence "bij_betw (g \<circ> ?Inl) A (A <+> B)" |
76951 | 1031 |
using 2 by (auto simp add: bij_betw_trans) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1032 |
thus ?thesis using card_of_ordIso ordIso_symmetric by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1033 |
next |
76951 | 1034 |
case False |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1035 |
with * 1 obtain b2 where 3: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B" by fastforce |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1036 |
obtain f where "inj_on f B \<and> f ` B \<le> A" |
76951 | 1037 |
using LEQ card_of_ordLeq[of B] by fastforce |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1038 |
with 3 have "f b1 \<noteq> f b2 \<and> {f b1, f b2} \<le> A" |
76951 | 1039 |
unfolding inj_on_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1040 |
with 3 have "|A <+> B| \<le>o |A \<times> B|" |
76951 | 1041 |
by (auto simp add: card_of_Plus_Times) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1042 |
moreover have "|A \<times> B| =o |A|" |
76951 | 1043 |
using assms * by (simp add: card_of_Times_infinite_simps) |
55811 | 1044 |
ultimately have "|A <+> B| \<le>o |A|" using ordLeq_ordIso_trans by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1045 |
thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1046 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1047 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1048 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1049 |
lemma card_of_Plus_infinite2: |
76951 | 1050 |
assumes INF: "\<not>finite A" and LEQ: "|B| \<le>o |A|" |
1051 |
shows "|B <+> A| =o |A|" |
|
1052 |
using assms card_of_Plus_commute card_of_Plus_infinite1 |
|
1053 |
ordIso_equivalence by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1054 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1055 |
lemma card_of_Plus_infinite: |
76951 | 1056 |
assumes INF: "\<not>finite A" and LEQ: "|B| \<le>o |A|" |
1057 |
shows "|A <+> B| =o |A| \<and> |B <+> A| =o |A|" |
|
1058 |
using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1059 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1060 |
corollary Card_order_Plus_infinite: |
76951 | 1061 |
assumes INF: "\<not>finite(Field r)" and CARD: "Card_order r" and |
1062 |
LEQ: "p \<le>o r" |
|
1063 |
shows "| (Field r) <+> (Field p) | =o r \<and> | (Field p) <+> (Field r) | =o r" |
|
1064 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1065 |
have "| Field r <+> Field p | =o | Field r | \<and> |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1066 |
| Field p <+> Field r | =o | Field r |" |
76951 | 1067 |
using assms by (simp add: card_of_Plus_infinite card_of_mono2) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1068 |
thus ?thesis |
76951 | 1069 |
using assms card_of_Field_ordIso by (blast intro: ordIso_transitive) |
1070 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1071 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1072 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1073 |
|
60758 | 1074 |
subsection \<open>The cardinal $\omega$ and the finite cardinals\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1075 |
|
60758 | 1076 |
text\<open>The cardinal $\omega$, of natural numbers, shall be the standard non-strict |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1077 |
order relation on |
61799 | 1078 |
\<open>nat\<close>, that we abbreviate by \<open>natLeq\<close>. The finite cardinals |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1079 |
shall be the restrictions of these relations to the numbers smaller than |
61799 | 1080 |
fixed numbers \<open>n\<close>, that we abbreviate by \<open>natLeq_on n\<close>.\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1081 |
|
56011
39d5043ce8a3
made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents:
55936
diff
changeset
|
1082 |
definition "(natLeq::(nat * nat) set) \<equiv> {(x,y). x \<le> y}" |
39d5043ce8a3
made natLe{q,ss} constants (yields smaller terms in composition)
traytel
parents:
55936
diff
changeset
|
1083 |
definition "(natLess::(nat * nat) set) \<equiv> {(x,y). x < y}" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1084 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1085 |
abbreviation natLeq_on :: "nat \<Rightarrow> (nat * nat) set" |
76951 | 1086 |
where "natLeq_on n \<equiv> {(x,y). x < n \<and> y < n \<and> x \<le> y}" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1087 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1088 |
lemma infinite_cartesian_product: |
76951 | 1089 |
assumes "\<not>finite A" "\<not>finite B" |
1090 |
shows "\<not>finite (A \<times> B)" |
|
1091 |
using assms finite_cartesian_productD2 by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1092 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1093 |
|
60758 | 1094 |
subsubsection \<open>First as well-orders\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1095 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1096 |
lemma Field_natLeq: "Field natLeq = (UNIV::nat set)" |
76951 | 1097 |
by(unfold Field_def natLeq_def, auto) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1098 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1099 |
lemma natLeq_Refl: "Refl natLeq" |
76951 | 1100 |
unfolding refl_on_def Field_def natLeq_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1101 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1102 |
lemma natLeq_trans: "trans natLeq" |
76951 | 1103 |
unfolding trans_def natLeq_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1104 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1105 |
lemma natLeq_Preorder: "Preorder natLeq" |
76951 | 1106 |
unfolding preorder_on_def |
1107 |
by (auto simp add: natLeq_Refl natLeq_trans) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1108 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1109 |
lemma natLeq_antisym: "antisym natLeq" |
76951 | 1110 |
unfolding antisym_def natLeq_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1111 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1112 |
lemma natLeq_Partial_order: "Partial_order natLeq" |
76951 | 1113 |
unfolding partial_order_on_def |
1114 |
by (auto simp add: natLeq_Preorder natLeq_antisym) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1115 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1116 |
lemma natLeq_Total: "Total natLeq" |
76951 | 1117 |
unfolding total_on_def natLeq_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1118 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1119 |
lemma natLeq_Linear_order: "Linear_order natLeq" |
76951 | 1120 |
unfolding linear_order_on_def |
1121 |
by (auto simp add: natLeq_Partial_order natLeq_Total) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1122 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1123 |
lemma natLeq_natLess_Id: "natLess = natLeq - Id" |
76951 | 1124 |
unfolding natLeq_def natLess_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1125 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1126 |
lemma natLeq_Well_order: "Well_order natLeq" |
76951 | 1127 |
unfolding well_order_on_def |
1128 |
using natLeq_Linear_order wf_less natLeq_natLess_Id natLeq_def natLess_def by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1129 |
|
54581
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents:
54578
diff
changeset
|
1130 |
lemma Field_natLeq_on: "Field (natLeq_on n) = {x. x < n}" |
76951 | 1131 |
unfolding Field_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1132 |
|
55023
38db7814481d
get rid of 'rel' locale, to facilitate inclusion of 'Order_Relation_More_FP' into 'Order_Relation'
blanchet
parents:
54980
diff
changeset
|
1133 |
lemma natLeq_underS_less: "underS natLeq n = {x. x < n}" |
76951 | 1134 |
unfolding underS_def natLeq_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1135 |
|
54581
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents:
54578
diff
changeset
|
1136 |
lemma Restr_natLeq: "Restr natLeq {x. x < n} = natLeq_on n" |
76951 | 1137 |
unfolding natLeq_def by force |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1138 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1139 |
lemma Restr_natLeq2: |
76951 | 1140 |
"Restr natLeq (underS natLeq n) = natLeq_on n" |
1141 |
by (auto simp add: Restr_natLeq natLeq_underS_less) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1142 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1143 |
lemma natLeq_on_Well_order: "Well_order(natLeq_on n)" |
76951 | 1144 |
using Restr_natLeq[of n] natLeq_Well_order |
1145 |
Well_order_Restr[of natLeq "{x. x < n}"] by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1146 |
|
54581
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents:
54578
diff
changeset
|
1147 |
corollary natLeq_on_well_order_on: "well_order_on {x. x < n} (natLeq_on n)" |
76951 | 1148 |
using natLeq_on_Well_order Field_natLeq_on by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1149 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1150 |
lemma natLeq_on_wo_rel: "wo_rel(natLeq_on n)" |
76951 | 1151 |
unfolding wo_rel_def using natLeq_on_Well_order . |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1152 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1153 |
|
60758 | 1154 |
subsubsection \<open>Then as cardinals\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1155 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1156 |
lemma natLeq_Card_order: "Card_order natLeq" |
76951 | 1157 |
proof - |
1158 |
have "natLeq_on n <o |UNIV::nat set|" for n |
|
1159 |
proof - |
|
1160 |
have "finite(Field (natLeq_on n))" by (auto simp: Field_def) |
|
1161 |
moreover have "\<not>finite(UNIV::nat set)" by auto |
|
1162 |
ultimately show ?thesis |
|
1163 |
using finite_ordLess_infinite[of "natLeq_on n" "|UNIV::nat set|"] |
|
1164 |
card_of_Well_order[of "UNIV::nat set"] natLeq_on_Well_order |
|
1165 |
by (force simp add: Field_card_of) |
|
1166 |
qed |
|
1167 |
then show ?thesis |
|
1168 |
apply (simp add: natLeq_Well_order Card_order_iff_Restr_underS Restr_natLeq2) |
|
1169 |
apply (force simp add: Field_natLeq) |
|
1170 |
done |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1171 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1172 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1173 |
corollary card_of_Field_natLeq: |
76951 | 1174 |
"|Field natLeq| =o natLeq" |
1175 |
using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq] |
|
1176 |
ordIso_symmetric[of natLeq] by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1177 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1178 |
corollary card_of_nat: |
76951 | 1179 |
"|UNIV::nat set| =o natLeq" |
1180 |
using Field_natLeq card_of_Field_natLeq by auto |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1181 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1182 |
corollary infinite_iff_natLeq_ordLeq: |
76951 | 1183 |
"\<not>finite A = ( natLeq \<le>o |A| )" |
1184 |
using infinite_iff_card_of_nat[of A] card_of_nat |
|
1185 |
ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1186 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1187 |
corollary finite_iff_ordLess_natLeq: |
76951 | 1188 |
"finite A = ( |A| <o natLeq)" |
1189 |
using infinite_iff_natLeq_ordLeq not_ordLeq_iff_ordLess |
|
1190 |
card_of_Well_order natLeq_Well_order by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1191 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1192 |
|
60758 | 1193 |
subsection \<open>The successor of a cardinal\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1194 |
|
61799 | 1195 |
text\<open>First we define \<open>isCardSuc r r'\<close>, the notion of \<open>r'\<close> |
1196 |
being a successor cardinal of \<open>r\<close>. Although the definition does |
|
1197 |
not require \<open>r\<close> to be a cardinal, only this case will be meaningful.\<close> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1198 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1199 |
definition isCardSuc :: "'a rel \<Rightarrow> 'a set rel \<Rightarrow> bool" |
76951 | 1200 |
where |
1201 |
"isCardSuc r r' \<equiv> |
|
1202 |
Card_order r' \<and> r <o r' \<and> |
|
1203 |
(\<forall>(r''::'a set rel). Card_order r'' \<and> r <o r'' \<longrightarrow> r' \<le>o r'')" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1204 |
|
61799 | 1205 |
text\<open>Now we introduce the cardinal-successor operator \<open>cardSuc\<close>, |
1206 |
by picking {\em some} cardinal-order relation fulfilling \<open>isCardSuc\<close>. |
|
60758 | 1207 |
Again, the picked item shall be proved unique up to order-isomorphism.\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1208 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1209 |
definition cardSuc :: "'a rel \<Rightarrow> 'a set rel" |
76951 | 1210 |
where "cardSuc r \<equiv> SOME r'. isCardSuc r r'" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1211 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1212 |
lemma exists_minim_Card_order: |
76951 | 1213 |
"\<lbrakk>R \<noteq> {}; \<forall>r \<in> R. Card_order r\<rbrakk> \<Longrightarrow> \<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'" |
1214 |
unfolding card_order_on_def using exists_minim_Well_order by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1215 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1216 |
lemma exists_isCardSuc: |
76951 | 1217 |
assumes "Card_order r" |
1218 |
shows "\<exists>r'. isCardSuc r r'" |
|
1219 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1220 |
let ?R = "{(r'::'a set rel). Card_order r' \<and> r <o r'}" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1221 |
have "|Pow(Field r)| \<in> ?R \<and> (\<forall>r \<in> ?R. Card_order r)" using assms |
76951 | 1222 |
by (simp add: card_of_Card_order Card_order_Pow) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1223 |
then obtain r where "r \<in> ?R \<and> (\<forall>r' \<in> ?R. r \<le>o r')" |
76951 | 1224 |
using exists_minim_Card_order[of ?R] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1225 |
thus ?thesis unfolding isCardSuc_def by auto |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1226 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1227 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1228 |
lemma cardSuc_isCardSuc: |
76951 | 1229 |
assumes "Card_order r" |
1230 |
shows "isCardSuc r (cardSuc r)" |
|
1231 |
unfolding cardSuc_def using assms |
|
1232 |
by (simp add: exists_isCardSuc someI_ex) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1233 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1234 |
lemma cardSuc_Card_order: |
76951 | 1235 |
"Card_order r \<Longrightarrow> Card_order(cardSuc r)" |
1236 |
using cardSuc_isCardSuc unfolding isCardSuc_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1237 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1238 |
lemma cardSuc_greater: |
76951 | 1239 |
"Card_order r \<Longrightarrow> r <o cardSuc r" |
1240 |
using cardSuc_isCardSuc unfolding isCardSuc_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1241 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1242 |
lemma cardSuc_ordLeq: |
76951 | 1243 |
"Card_order r \<Longrightarrow> r \<le>o cardSuc r" |
1244 |
using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1245 |
|
61799 | 1246 |
text\<open>The minimality property of \<open>cardSuc\<close> originally present in its definition |
1247 |
is local to the type \<open>'a set rel\<close>, i.e., that of \<open>cardSuc r\<close>:\<close> |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1248 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1249 |
lemma cardSuc_least_aux: |
76951 | 1250 |
"\<lbrakk>Card_order (r::'a rel); Card_order (r'::'a set rel); r <o r'\<rbrakk> \<Longrightarrow> cardSuc r \<le>o r'" |
1251 |
using cardSuc_isCardSuc unfolding isCardSuc_def by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1252 |
|
60758 | 1253 |
text\<open>But from this we can infer general minimality:\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1254 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1255 |
lemma cardSuc_least: |
76951 | 1256 |
assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r <o r'" |
1257 |
shows "cardSuc r \<le>o r'" |
|
1258 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1259 |
let ?p = "cardSuc r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1260 |
have 0: "Well_order ?p \<and> Well_order r'" |
76951 | 1261 |
using assms cardSuc_Card_order unfolding card_order_on_def by blast |
1262 |
{ assume "r' <o ?p" |
|
1263 |
then obtain r'' where 1: "Field r'' < Field ?p" and 2: "r' =o r'' \<and> r'' <o ?p" |
|
1264 |
using internalize_ordLess[of r' ?p] by blast |
|
1265 |
(* *) |
|
1266 |
have "Card_order r''" using CARD' Card_order_ordIso2 2 by blast |
|
1267 |
moreover have "r <o r''" using LESS 2 ordLess_ordIso_trans by blast |
|
1268 |
ultimately have "?p \<le>o r''" using cardSuc_least_aux CARD by blast |
|
1269 |
hence False using 2 not_ordLess_ordLeq by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1270 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1271 |
thus ?thesis using 0 ordLess_or_ordLeq by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1272 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1273 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1274 |
lemma cardSuc_ordLess_ordLeq: |
76951 | 1275 |
assumes CARD: "Card_order r" and CARD': "Card_order r'" |
1276 |
shows "(r <o r') = (cardSuc r \<le>o r')" |
|
1277 |
proof |
|
1278 |
show "cardSuc r \<le>o r' \<Longrightarrow> r <o r'" |
|
1279 |
using assms cardSuc_greater ordLess_ordLeq_trans by blast |
|
1280 |
qed (auto simp add: assms cardSuc_least) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1281 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1282 |
lemma cardSuc_ordLeq_ordLess: |
76951 | 1283 |
assumes CARD: "Card_order r" and CARD': "Card_order r'" |
1284 |
shows "(r' <o cardSuc r) = (r' \<le>o r)" |
|
1285 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1286 |
have "Well_order r \<and> Well_order r'" |
76951 | 1287 |
using assms unfolding card_order_on_def by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1288 |
moreover have "Well_order(cardSuc r)" |
76951 | 1289 |
using assms cardSuc_Card_order card_order_on_def by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1290 |
ultimately show ?thesis |
76951 | 1291 |
using assms cardSuc_ordLess_ordLeq by (blast dest: not_ordLeq_iff_ordLess) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1292 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1293 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1294 |
lemma cardSuc_mono_ordLeq: |
76951 | 1295 |
assumes CARD: "Card_order r" and CARD': "Card_order r'" |
1296 |
shows "(cardSuc r \<le>o cardSuc r') = (r \<le>o r')" |
|
1297 |
using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1298 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1299 |
lemma cardSuc_invar_ordIso: |
76951 | 1300 |
assumes CARD: "Card_order r" and CARD': "Card_order r'" |
1301 |
shows "(cardSuc r =o cardSuc r') = (r =o r')" |
|
1302 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1303 |
have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')" |
76951 | 1304 |
using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1305 |
thus ?thesis |
76951 | 1306 |
using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq |
1307 |
using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1308 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1309 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1310 |
lemma card_of_cardSuc_finite: |
76951 | 1311 |
"finite(Field(cardSuc |A| )) = finite A" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1312 |
proof |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1313 |
assume *: "finite (Field (cardSuc |A| ))" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1314 |
have 0: "|Field(cardSuc |A| )| =o cardSuc |A|" |
76951 | 1315 |
using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1316 |
hence "|A| \<le>o |Field(cardSuc |A| )|" |
76951 | 1317 |
using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric |
1318 |
ordLeq_ordIso_trans by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1319 |
thus "finite A" using * card_of_ordLeq_finite by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1320 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1321 |
assume "finite A" |
54581
1502a1f707d9
eliminated dependence of Cardinals_FP on Set_Intervals, more precise imports
traytel
parents:
54578
diff
changeset
|
1322 |
then have "finite ( Field |Pow A| )" unfolding Field_card_of by simp |
76951 | 1323 |
moreover |
1324 |
have "cardSuc |A| \<le>o |Pow A|" |
|
1325 |
by (rule iffD1[OF cardSuc_ordLess_ordLeq card_of_Pow]) (simp_all add: card_of_Card_order) |
|
1326 |
ultimately show "finite (Field (cardSuc |A| ))" |
|
1327 |
by (blast intro: card_of_ordLeq_finite card_of_mono2) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1328 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1329 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1330 |
lemma cardSuc_finite: |
76951 | 1331 |
assumes "Card_order r" |
1332 |
shows "finite (Field (cardSuc r)) = finite (Field r)" |
|
1333 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1334 |
let ?A = "Field r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1335 |
have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1336 |
hence "cardSuc |?A| =o cardSuc r" using assms |
76951 | 1337 |
by (simp add: card_of_Card_order cardSuc_invar_ordIso) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1338 |
moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|" |
76951 | 1339 |
by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1340 |
moreover |
76951 | 1341 |
{ have "|Field (cardSuc r) | =o cardSuc r" |
1342 |
using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order) |
|
1343 |
hence "cardSuc r =o |Field (cardSuc r) |" |
|
1344 |
using ordIso_symmetric by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1345 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1346 |
ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |" |
76951 | 1347 |
using ordIso_transitive by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1348 |
hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))" |
76951 | 1349 |
using card_of_ordIso_finite by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1350 |
thus ?thesis by (simp only: card_of_cardSuc_finite) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1351 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1352 |
|
75624 | 1353 |
lemma Field_cardSuc_not_empty: |
76951 | 1354 |
assumes "Card_order r" |
1355 |
shows "Field (cardSuc r) \<noteq> {}" |
|
75624 | 1356 |
proof |
1357 |
assume "Field(cardSuc r) = {}" |
|
1358 |
then have "|Field(cardSuc r)| \<le>o r" using assms Card_order_empty[of r] by auto |
|
1359 |
then have "cardSuc r \<le>o r" using assms card_of_Field_ordIso |
|
76951 | 1360 |
cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast |
75624 | 1361 |
then show False using cardSuc_greater not_ordLess_ordLeq assms by blast |
1362 |
qed |
|
1363 |
||
1364 |
typedef 'a suc = "Field (cardSuc |UNIV :: 'a set| )" |
|
1365 |
using Field_cardSuc_not_empty card_of_Card_order by blast |
|
1366 |
||
1367 |
definition card_suc :: "'a rel \<Rightarrow> 'a suc rel" where |
|
1368 |
"card_suc \<equiv> \<lambda>_. map_prod Abs_suc Abs_suc ` cardSuc |UNIV :: 'a set|" |
|
1369 |
||
1370 |
lemma Field_card_suc: "Field (card_suc r) = UNIV" |
|
1371 |
proof - |
|
1372 |
let ?r = "cardSuc |UNIV|" |
|
1373 |
let ?ar = "\<lambda>x. Abs_suc (Rep_suc x)" |
|
1374 |
have 1: "\<And>P. (\<forall>x\<in>Field ?r. P x) = (\<forall>x. P (Rep_suc x))" using Rep_suc_induct Rep_suc by blast |
|
1375 |
have 2: "\<And>P. (\<exists>x\<in>Field ?r. P x) = (\<exists>x. P (Rep_suc x))" using Rep_suc_cases Rep_suc by blast |
|
1376 |
have 3: "\<And>A a b. (a, b) \<in> A \<Longrightarrow> (Abs_suc a, Abs_suc b) \<in> map_prod Abs_suc Abs_suc ` A" unfolding map_prod_def by auto |
|
1377 |
have "\<forall>x\<in>Field ?r. (\<exists>b\<in>Field ?r. (x, b) \<in> ?r) \<or> (\<exists>a\<in>Field ?r. (a, x) \<in> ?r)" |
|
1378 |
unfolding Field_def Range.simps Domain.simps Un_iff by blast |
|
1379 |
then have "\<forall>x. (\<exists>b. (Rep_suc x, Rep_suc b) \<in> ?r) \<or> (\<exists>a. (Rep_suc a, Rep_suc x) \<in> ?r)" unfolding 1 2 . |
|
1380 |
then have "\<forall>x. (\<exists>b. (?ar x, ?ar b) \<in> map_prod Abs_suc Abs_suc ` ?r) \<or> (\<exists>a. (?ar a, ?ar x) \<in> map_prod Abs_suc Abs_suc ` ?r)" using 3 by fast |
|
1381 |
then have "\<forall>x. (\<exists>b. (x, b) \<in> card_suc r) \<or> (\<exists>a. (a, x) \<in> card_suc r)" unfolding card_suc_def Rep_suc_inverse . |
|
1382 |
then show ?thesis unfolding Field_def Domain.simps Range.simps set_eq_iff Un_iff eqTrueI[OF UNIV_I] ex_simps simp_thms . |
|
1383 |
qed |
|
1384 |
||
1385 |
lemma card_suc_alt: "card_suc r = dir_image (cardSuc |UNIV :: 'a set| ) Abs_suc" |
|
1386 |
unfolding card_suc_def dir_image_def by auto |
|
1387 |
||
1388 |
lemma cardSuc_Well_order: "Card_order r \<Longrightarrow> Well_order(cardSuc r)" |
|
1389 |
using cardSuc_Card_order unfolding card_order_on_def by blast |
|
1390 |
||
1391 |
lemma cardSuc_ordIso_card_suc: |
|
1392 |
assumes "card_order r" |
|
1393 |
shows "cardSuc r =o card_suc (r :: 'a rel)" |
|
1394 |
proof - |
|
1395 |
have "cardSuc (r :: 'a rel) =o cardSuc |UNIV :: 'a set|" |
|
1396 |
using cardSuc_invar_ordIso[THEN iffD2, OF _ card_of_Card_order card_of_unique[OF assms]] assms |
|
1397 |
by (simp add: card_order_on_Card_order) |
|
1398 |
also have "cardSuc |UNIV :: 'a set| =o card_suc (r :: 'a rel)" |
|
1399 |
unfolding card_suc_alt |
|
1400 |
by (rule dir_image_ordIso) (simp_all add: inj_on_def Abs_suc_inject cardSuc_Well_order card_of_Card_order) |
|
1401 |
finally show ?thesis . |
|
1402 |
qed |
|
1403 |
||
1404 |
lemma Card_order_card_suc: "card_order r \<Longrightarrow> Card_order (card_suc r)" |
|
1405 |
using cardSuc_Card_order[THEN Card_order_ordIso2[OF _ cardSuc_ordIso_card_suc]] card_order_on_Card_order by blast |
|
1406 |
||
1407 |
lemma card_order_card_suc: "card_order r \<Longrightarrow> card_order (card_suc r)" |
|
1408 |
using Card_order_card_suc arg_cong2[OF Field_card_suc refl, of "card_order_on"] by blast |
|
1409 |
||
1410 |
lemma card_suc_greater: "card_order r \<Longrightarrow> r <o card_suc r" |
|
1411 |
using ordLess_ordIso_trans[OF cardSuc_greater cardSuc_ordIso_card_suc] card_order_on_Card_order by blast |
|
1412 |
||
54475 | 1413 |
lemma card_of_Plus_ordLess_infinite: |
76951 | 1414 |
assumes INF: "\<not>finite C" and LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|" |
1415 |
shows "|A <+> B| <o |C|" |
|
54475 | 1416 |
proof(cases "A = {} \<or> B = {}") |
76951 | 1417 |
case True |
54475 | 1418 |
hence "|A| =o |A <+> B| \<or> |B| =o |A <+> B|" |
76951 | 1419 |
using card_of_Plus_empty1 card_of_Plus_empty2 by blast |
54475 | 1420 |
hence "|A <+> B| =o |A| \<or> |A <+> B| =o |B|" |
76951 | 1421 |
using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast |
54475 | 1422 |
thus ?thesis using LESS1 LESS2 |
76951 | 1423 |
ordIso_ordLess_trans[of "|A <+> B|" "|A|"] |
1424 |
ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast |
|
54475 | 1425 |
next |
76951 | 1426 |
case False |
1427 |
have False if "|C| \<le>o |A <+> B|" |
|
1428 |
proof - |
|
1429 |
have \<section>: "\<not>finite A \<or> \<not>finite B" |
|
1430 |
using that INF card_of_ordLeq_finite finite_Plus by blast |
|
1431 |
consider "|A| \<le>o |B|" | "|B| \<le>o |A|" |
|
1432 |
using ordLeq_total [OF card_of_Well_order card_of_Well_order] by blast |
|
1433 |
then show False |
|
1434 |
proof cases |
|
1435 |
case 1 |
|
1436 |
hence "\<not>finite B" using \<section> card_of_ordLeq_finite by blast |
|
1437 |
hence "|A <+> B| =o |B|" using False 1 |
|
1438 |
by (auto simp add: card_of_Plus_infinite) |
|
1439 |
thus False using LESS2 not_ordLess_ordLeq that ordLeq_ordIso_trans by blast |
|
1440 |
next |
|
1441 |
case 2 |
|
1442 |
hence "\<not>finite A" using \<section> card_of_ordLeq_finite by blast |
|
1443 |
hence "|A <+> B| =o |A|" using False 2 |
|
1444 |
by (auto simp add: card_of_Plus_infinite) |
|
1445 |
thus False using LESS1 not_ordLess_ordLeq that ordLeq_ordIso_trans by blast |
|
1446 |
qed |
|
1447 |
qed |
|
1448 |
thus ?thesis |
|
1449 |
using ordLess_or_ordLeq[of "|A <+> B|" "|C|"] |
|
1450 |
card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto |
|
54475 | 1451 |
qed |
1452 |
||
1453 |
lemma card_of_Plus_ordLess_infinite_Field: |
|
76951 | 1454 |
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and |
1455 |
LESS1: "|A| <o r" and LESS2: "|B| <o r" |
|
1456 |
shows "|A <+> B| <o r" |
|
1457 |
proof - |
|
54475 | 1458 |
let ?C = "Field r" |
76951 | 1459 |
have 1: "r =o |?C| \<and> |?C| =o r" |
1460 |
using r card_of_Field_ordIso ordIso_symmetric by blast |
|
54475 | 1461 |
hence "|A| <o |?C|" "|B| <o |?C|" |
76951 | 1462 |
using LESS1 LESS2 ordLess_ordIso_trans by blast+ |
54475 | 1463 |
hence "|A <+> B| <o |?C|" using INF |
76951 | 1464 |
card_of_Plus_ordLess_infinite by blast |
54475 | 1465 |
thus ?thesis using 1 ordLess_ordIso_trans by blast |
1466 |
qed |
|
1467 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1468 |
lemma card_of_Plus_ordLeq_infinite_Field: |
76951 | 1469 |
assumes r: "\<not>finite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r" |
1470 |
and c: "Card_order r" |
|
1471 |
shows "|A <+> B| \<le>o r" |
|
1472 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1473 |
let ?r' = "cardSuc r" |
54578
9387251b6a46
eliminated dependence of BNF on Infinite_Set by moving 3 theorems from the latter to Main
traytel
parents:
54482
diff
changeset
|
1474 |
have "Card_order ?r' \<and> \<not>finite (Field ?r')" using assms |
76951 | 1475 |
by (simp add: cardSuc_Card_order cardSuc_finite) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1476 |
moreover have "|A| <o ?r'" and "|B| <o ?r'" using A B c |
76951 | 1477 |
by (auto simp: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1478 |
ultimately have "|A <+> B| <o ?r'" |
76951 | 1479 |
using card_of_Plus_ordLess_infinite_Field by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1480 |
thus ?thesis using c r |
76951 | 1481 |
by (simp add: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1482 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1483 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1484 |
lemma card_of_Un_ordLeq_infinite_Field: |
76951 | 1485 |
assumes C: "\<not>finite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r" |
1486 |
and "Card_order r" |
|
1487 |
shows "|A Un B| \<le>o r" |
|
1488 |
using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq |
|
1489 |
ordLeq_transitive by fast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1490 |
|
75624 | 1491 |
lemma card_of_Un_ordLess_infinite: |
76951 | 1492 |
assumes INF: "\<not>finite C" and |
1493 |
LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|" |
|
1494 |
shows "|A \<union> B| <o |C|" |
|
1495 |
using assms card_of_Plus_ordLess_infinite card_of_Un_Plus_ordLeq |
|
1496 |
ordLeq_ordLess_trans by blast |
|
75624 | 1497 |
|
1498 |
lemma card_of_Un_ordLess_infinite_Field: |
|
76951 | 1499 |
assumes INF: "\<not>finite (Field r)" and r: "Card_order r" and |
1500 |
LESS1: "|A| <o r" and LESS2: "|B| <o r" |
|
1501 |
shows "|A Un B| <o r" |
|
1502 |
proof - |
|
75624 | 1503 |
let ?C = "Field r" |
1504 |
have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso |
|
76951 | 1505 |
ordIso_symmetric by blast |
75624 | 1506 |
hence "|A| <o |?C|" "|B| <o |?C|" |
76951 | 1507 |
using LESS1 LESS2 ordLess_ordIso_trans by blast+ |
75624 | 1508 |
hence "|A Un B| <o |?C|" using INF |
76951 | 1509 |
card_of_Un_ordLess_infinite by blast |
75624 | 1510 |
thus ?thesis using 1 ordLess_ordIso_trans by blast |
1511 |
qed |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1512 |
|
60758 | 1513 |
subsection \<open>Regular cardinals\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1514 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1515 |
definition cofinal where |
76951 | 1516 |
"cofinal A r \<equiv> \<forall>a \<in> Field r. \<exists>b \<in> A. a \<noteq> b \<and> (a,b) \<in> r" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1517 |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55059
diff
changeset
|
1518 |
definition regularCard where |
76951 | 1519 |
"regularCard r \<equiv> \<forall>K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1520 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1521 |
definition relChain where |
76951 | 1522 |
"relChain r As \<equiv> \<forall>i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1523 |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55059
diff
changeset
|
1524 |
lemma regularCard_UNION: |
76951 | 1525 |
assumes r: "Card_order r" "regularCard r" |
1526 |
and As: "relChain r As" |
|
1527 |
and Bsub: "B \<le> (\<Union>i \<in> Field r. As i)" |
|
1528 |
and cardB: "|B| <o r" |
|
1529 |
shows "\<exists>i \<in> Field r. B \<le> As i" |
|
1530 |
proof - |
|
67091 | 1531 |
let ?phi = "\<lambda>b j. j \<in> Field r \<and> b \<in> As j" |
1532 |
have "\<forall>b\<in>B. \<exists>j. ?phi b j" using Bsub by blast |
|
67613 | 1533 |
then obtain f where f: "\<And>b. b \<in> B \<Longrightarrow> ?phi b (f b)" |
76951 | 1534 |
using bchoice[of B ?phi] by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1535 |
let ?K = "f ` B" |
67091 | 1536 |
{assume 1: "\<And>i. i \<in> Field r \<Longrightarrow> \<not> B \<le> As i" |
76951 | 1537 |
have 2: "cofinal ?K r" |
1538 |
unfolding cofinal_def |
|
1539 |
proof (intro strip) |
|
1540 |
fix i assume i: "i \<in> Field r" |
|
1541 |
with 1 obtain b where b: "b \<in> B \<and> b \<notin> As i" by blast |
|
1542 |
hence "i \<noteq> f b \<and> \<not> (f b,i) \<in> r" |
|
1543 |
using As f unfolding relChain_def by auto |
|
1544 |
hence "i \<noteq> f b \<and> (i, f b) \<in> r" using r |
|
1545 |
unfolding card_order_on_def well_order_on_def linear_order_on_def |
|
1546 |
total_on_def using i f b by auto |
|
1547 |
with b show "\<exists>b \<in> f`B. i \<noteq> b \<and> (i,b) \<in> r" by blast |
|
1548 |
qed |
|
1549 |
moreover have "?K \<le> Field r" using f by blast |
|
1550 |
ultimately have "|?K| =o r" using 2 r unfolding regularCard_def by blast |
|
1551 |
moreover |
|
1552 |
have "|?K| <o r" using cardB ordLeq_ordLess_trans card_of_image by blast |
|
1553 |
ultimately have False using not_ordLess_ordIso by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1554 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1555 |
thus ?thesis by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1556 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1557 |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55059
diff
changeset
|
1558 |
lemma infinite_cardSuc_regularCard: |
76951 | 1559 |
assumes r_inf: "\<not>finite (Field r)" and r_card: "Card_order r" |
1560 |
shows "regularCard (cardSuc r)" |
|
1561 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1562 |
let ?r' = "cardSuc r" |
76951 | 1563 |
have r': "Card_order ?r'" "\<And>p. Card_order p \<longrightarrow> (p \<le>o r) = (p <o ?r')" |
1564 |
using r_card by (auto simp: cardSuc_Card_order cardSuc_ordLeq_ordLess) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1565 |
show ?thesis |
76951 | 1566 |
unfolding regularCard_def proof auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1567 |
fix K assume 1: "K \<le> Field ?r'" and 2: "cofinal K ?r'" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1568 |
hence "|K| \<le>o |Field ?r'|" by (simp only: card_of_mono1) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1569 |
also have 22: "|Field ?r'| =o ?r'" |
76951 | 1570 |
using r' by (simp add: card_of_Field_ordIso[of ?r']) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1571 |
finally have "|K| \<le>o ?r'" . |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1572 |
moreover |
76951 | 1573 |
{ let ?L = "\<Union> j \<in> K. underS ?r' j" |
1574 |
let ?J = "Field r" |
|
1575 |
have rJ: "r =o |?J|" |
|
1576 |
using r_card card_of_Field_ordIso ordIso_symmetric by blast |
|
1577 |
assume "|K| <o ?r'" |
|
1578 |
hence "|K| \<le>o r" using r' card_of_Card_order[of K] by blast |
|
1579 |
hence "|K| \<le>o |?J|" using rJ ordLeq_ordIso_trans by blast |
|
1580 |
moreover |
|
1581 |
{have "\<forall>j\<in>K. |underS ?r' j| <o ?r'" |
|
1582 |
using r' 1 by (auto simp: card_of_underS) |
|
1583 |
hence "\<forall>j\<in>K. |underS ?r' j| \<le>o r" |
|
1584 |
using r' card_of_Card_order by blast |
|
1585 |
hence "\<forall>j\<in>K. |underS ?r' j| \<le>o |?J|" |
|
1586 |
using rJ ordLeq_ordIso_trans by blast |
|
1587 |
} |
|
1588 |
ultimately have "|?L| \<le>o |?J|" |
|
1589 |
using r_inf card_of_UNION_ordLeq_infinite by blast |
|
1590 |
hence "|?L| \<le>o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast |
|
1591 |
hence "|?L| <o ?r'" using r' card_of_Card_order by blast |
|
1592 |
moreover |
|
1593 |
{ |
|
1594 |
have "Field ?r' \<le> ?L" |
|
1595 |
using 2 unfolding underS_def cofinal_def by auto |
|
1596 |
hence "|Field ?r'| \<le>o |?L|" by (simp add: card_of_mono1) |
|
1597 |
hence "?r' \<le>o |?L|" |
|
1598 |
using 22 ordIso_ordLeq_trans ordIso_symmetric by blast |
|
1599 |
} |
|
1600 |
ultimately have "|?L| <o |?L|" using ordLess_ordLeq_trans by blast |
|
1601 |
hence False using ordLess_irreflexive by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1602 |
} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1603 |
ultimately show "|K| =o ?r'" |
76951 | 1604 |
unfolding ordLeq_iff_ordLess_or_ordIso by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1605 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1606 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1607 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1608 |
lemma cardSuc_UNION: |
76951 | 1609 |
assumes r: "Card_order r" and "\<not>finite (Field r)" |
1610 |
and As: "relChain (cardSuc r) As" |
|
1611 |
and Bsub: "B \<le> (\<Union> i \<in> Field (cardSuc r). As i)" |
|
1612 |
and cardB: "|B| \<le>o r" |
|
1613 |
shows "\<exists>i \<in> Field (cardSuc r). B \<le> As i" |
|
1614 |
proof - |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1615 |
let ?r' = "cardSuc r" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1616 |
have "Card_order ?r' \<and> |B| <o ?r'" |
76951 | 1617 |
using r cardB cardSuc_ordLeq_ordLess cardSuc_Card_order |
1618 |
card_of_Card_order by blast |
|
55087
252c7fec4119
renamed 'regular' to 'regularCard' to avoid clashes (e.g. in Meson_Test)
blanchet
parents:
55059
diff
changeset
|
1619 |
moreover have "regularCard ?r'" |
76951 | 1620 |
using assms by(simp add: infinite_cardSuc_regularCard) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1621 |
ultimately show ?thesis |
76951 | 1622 |
using As Bsub cardB regularCard_UNION by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1623 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1624 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1625 |
|
60758 | 1626 |
subsection \<open>Others\<close> |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1627 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1628 |
lemma card_of_Func_Times: |
76951 | 1629 |
"|Func (A \<times> B) C| =o |Func A (Func B C)|" |
1630 |
unfolding card_of_ordIso[symmetric] |
|
1631 |
using bij_betw_curr by blast |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1632 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1633 |
lemma card_of_Pow_Func: |
76951 | 1634 |
"|Pow A| =o |Func A (UNIV::bool set)|" |
1635 |
proof - |
|
1636 |
define F where [abs_def]: "F A' a \<equiv> |
|
63040 | 1637 |
(if a \<in> A then (if a \<in> A' then True else False) else undefined)" for A' a |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1638 |
have "bij_betw F (Pow A) (Func A (UNIV::bool set))" |
76951 | 1639 |
unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI) |
52545
d2ad6eae514f
Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents:
52544
diff
changeset
|
1640 |
fix A1 A2 assume "A1 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2" |
62390 | 1641 |
thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: if_split_asm) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1642 |
next |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1643 |
show "F ` Pow A = Func A UNIV" |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1644 |
proof safe |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1645 |
fix f assume f: "f \<in> Func A (UNIV::bool set)" |
76951 | 1646 |
show "f \<in> F ` Pow A" |
1647 |
unfolding image_iff |
|
1648 |
proof |
|
1649 |
show "f = F {a \<in> A. f a = True}" |
|
1650 |
using f unfolding Func_def F_def by force |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1651 |
qed auto |
76951 | 1652 |
qed(unfold Func_def F_def, auto) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1653 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1654 |
thus ?thesis unfolding card_of_ordIso[symmetric] by blast |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1655 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1656 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1657 |
lemma card_of_Func_UNIV: |
76951 | 1658 |
"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \<Rightarrow> 'b. range f \<subseteq> B}|" |
1659 |
proof - |
|
52545
d2ad6eae514f
Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents:
52544
diff
changeset
|
1660 |
let ?F = "\<lambda> f (a::'a). ((f a)::'b)" |
76951 | 1661 |
have "bij_betw ?F {f. range f \<subseteq> B} (Func UNIV B)" |
1662 |
unfolding bij_betw_def inj_on_def |
|
1663 |
proof safe |
|
52545
d2ad6eae514f
Func -> Func_option, Ffunc -> Func (avoids dependence of codatatypes on the option type)
traytel
parents:
52544
diff
changeset
|
1664 |
fix h :: "'a \<Rightarrow> 'b" assume h: "h \<in> Func UNIV B" |
55811 | 1665 |
then obtain f where f: "\<forall> a. h a = f a" by blast |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1666 |
hence "range f \<subseteq> B" using h unfolding Func_def by auto |
56077 | 1667 |
thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f by auto |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1668 |
qed(unfold Func_def fun_eq_iff, auto) |
76951 | 1669 |
then show ?thesis |
1670 |
by (blast intro: ordIso_symmetric card_of_ordIsoI) |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1671 |
qed |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1672 |
|
56191 | 1673 |
lemma Func_Times_Range: |
61943 | 1674 |
"|Func A (B \<times> C)| =o |Func A B \<times> Func A C|" (is "|?LHS| =o |?RHS|") |
56191 | 1675 |
proof - |
1676 |
let ?F = "\<lambda>fg. (\<lambda>x. if x \<in> A then fst (fg x) else undefined, |
|
1677 |
\<lambda>x. if x \<in> A then snd (fg x) else undefined)" |
|
1678 |
let ?G = "\<lambda>(f, g) x. if x \<in> A then (f x, g x) else undefined" |
|
1679 |
have "bij_betw ?F ?LHS ?RHS" unfolding bij_betw_def inj_on_def |
|
1680 |
proof (intro conjI impI ballI equalityI subsetI) |
|
1681 |
fix f g assume *: "f \<in> Func A (B \<times> C)" "g \<in> Func A (B \<times> C)" "?F f = ?F g" |
|
1682 |
show "f = g" |
|
1683 |
proof |
|
1684 |
fix x from * have "fst (f x) = fst (g x) \<and> snd (f x) = snd (g x)" |
|
69850 | 1685 |
by (cases "x \<in> A") (auto simp: Func_def fun_eq_iff split: if_splits) |
56191 | 1686 |
then show "f x = g x" by (subst (1 2) surjective_pairing) simp |
1687 |
qed |
|
1688 |
next |
|
1689 |
fix fg assume "fg \<in> Func A B \<times> Func A C" |
|
1690 |
thus "fg \<in> ?F ` Func A (B \<times> C)" |
|
1691 |
by (intro image_eqI[of _ _ "?G fg"]) (auto simp: Func_def) |
|
1692 |
qed (auto simp: Func_def fun_eq_iff) |
|
1693 |
thus ?thesis using card_of_ordIso by blast |
|
1694 |
qed |
|
1695 |
||
75624 | 1696 |
subsection \<open>Regular vs. stable cardinals\<close> |
1697 |
||
1698 |
definition stable :: "'a rel \<Rightarrow> bool" |
|
76951 | 1699 |
where |
1700 |
"stable r \<equiv> \<forall>(A::'a set) (F :: 'a \<Rightarrow> 'a set). |
|
75624 | 1701 |
|A| <o r \<and> (\<forall>a \<in> A. |F a| <o r) |
1702 |
\<longrightarrow> |SIGMA a : A. F a| <o r" |
|
1703 |
||
1704 |
lemma regularCard_stable: |
|
1705 |
assumes cr: "Card_order r" and ir: "\<not>finite (Field r)" and reg: "regularCard r" |
|
1706 |
shows "stable r" |
|
1707 |
unfolding stable_def proof safe |
|
1708 |
fix A :: "'a set" and F :: "'a \<Rightarrow> 'a set" assume A: "|A| <o r" and F: "\<forall>a\<in>A. |F a| <o r" |
|
1709 |
{assume "r \<le>o |Sigma A F|" |
|
1710 |
hence "|Field r| \<le>o |Sigma A F|" using card_of_Field_ordIso[OF cr] ordIso_ordLeq_trans by blast |
|
1711 |
moreover have Fi: "Field r \<noteq> {}" using ir by auto |
|
1712 |
ultimately have "\<exists>f. f ` Sigma A F = Field r" using card_of_ordLeq2[OF Fi] by blast |
|
1713 |
then obtain f where f: "f ` Sigma A F = Field r" by blast |
|
1714 |
have r: "wo_rel r" using cr unfolding card_order_on_def wo_rel_def by auto |
|
1715 |
{fix a assume a: "a \<in> A" |
|
1716 |
define L where "L = {(a,u) | u. u \<in> F a}" |
|
1717 |
have fL: "f ` L \<subseteq> Field r" using f a unfolding L_def by auto |
|
76951 | 1718 |
have "bij_betw snd {(a, u) |u. u \<in> F a} (F a)" |
1719 |
unfolding bij_betw_def inj_on_def by (auto simp: image_def) |
|
1720 |
then have "|L| =o |F a|" unfolding L_def card_of_ordIso[symmetric] by blast |
|
75624 | 1721 |
hence "|L| <o r" using F a ordIso_ordLess_trans[of "|L|" "|F a|"] unfolding L_def by auto |
1722 |
hence "|f ` L| <o r" using ordLeq_ordLess_trans[OF card_of_image, of "L"] unfolding L_def by auto |
|
1723 |
hence "\<not> cofinal (f ` L) r" using reg fL unfolding regularCard_def |
|
76951 | 1724 |
by (force simp add: dest: not_ordLess_ordIso) |
75624 | 1725 |
then obtain k where k: "k \<in> Field r" and "\<forall> l \<in> L. \<not> (f l \<noteq> k \<and> (k, f l) \<in> r)" |
1726 |
unfolding cofinal_def image_def by auto |
|
1727 |
hence "\<exists> k \<in> Field r. \<forall> l \<in> L. (f l, k) \<in> r" |
|
1728 |
using wo_rel.in_notinI[OF r _ _ \<open>k \<in> Field r\<close>] fL unfolding image_subset_iff by fast |
|
1729 |
hence "\<exists> k \<in> Field r. \<forall> u \<in> F a. (f (a,u), k) \<in> r" unfolding L_def by auto |
|
1730 |
} |
|
1731 |
then have x: "\<And>a. a\<in>A \<Longrightarrow> \<exists>k. k \<in> Field r \<and> (\<forall>u\<in>F a. (f (a, u), k) \<in> r)" by blast |
|
1732 |
obtain gg where "\<And>a. a\<in>A \<Longrightarrow> gg a = (SOME k. k \<in> Field r \<and> (\<forall>u\<in>F a. (f (a, u), k) \<in> r))" by simp |
|
1733 |
then have gg: "\<forall>a\<in>A. \<forall>u\<in>F a. (f (a, u), gg a) \<in> r" using someI_ex[OF x] by auto |
|
1734 |
obtain j0 where j0: "j0 \<in> Field r" using Fi by auto |
|
1735 |
define g where [abs_def]: "g a = (if F a \<noteq> {} then gg a else j0)" for a |
|
1736 |
have g: "\<forall> a \<in> A. \<forall> u \<in> F a. (f (a,u),g a) \<in> r" using gg unfolding g_def by auto |
|
1737 |
hence 1: "Field r \<subseteq> (\<Union>a \<in> A. under r (g a))" |
|
1738 |
using f[symmetric] unfolding under_def image_def by auto |
|
1739 |
have gA: "g ` A \<subseteq> Field r" using gg j0 unfolding Field_def g_def by auto |
|
76951 | 1740 |
moreover have "cofinal (g ` A) r" unfolding cofinal_def |
1741 |
proof safe |
|
75624 | 1742 |
fix i assume "i \<in> Field r" |
1743 |
then obtain j where ij: "(i,j) \<in> r" "i \<noteq> j" using cr ir infinite_Card_order_limit by fast |
|
1744 |
hence "j \<in> Field r" using card_order_on_def cr well_order_on_domain by fast |
|
76951 | 1745 |
then obtain a where a: "a \<in> A" and j: "(j, g a) \<in> r" |
1746 |
using 1 unfolding under_def by auto |
|
75624 | 1747 |
hence "(i, g a) \<in> r" using ij wo_rel.TRANS[OF r] unfolding trans_def by blast |
1748 |
moreover have "i \<noteq> g a" |
|
1749 |
using ij j r unfolding wo_rel_def unfolding well_order_on_def linear_order_on_def |
|
1750 |
partial_order_on_def antisym_def by auto |
|
1751 |
ultimately show "\<exists>j\<in>g ` A. i \<noteq> j \<and> (i, j) \<in> r" using a by auto |
|
1752 |
qed |
|
1753 |
ultimately have "|g ` A| =o r" using reg unfolding regularCard_def by auto |
|
1754 |
moreover have "|g ` A| \<le>o |A|" using card_of_image by blast |
|
1755 |
ultimately have False using A using not_ordLess_ordIso ordLeq_ordLess_trans by blast |
|
1756 |
} |
|
1757 |
thus "|Sigma A F| <o r" |
|
1758 |
using cr not_ordLess_iff_ordLeq using card_of_Well_order card_order_on_well_order_on by blast |
|
1759 |
qed |
|
1760 |
||
1761 |
lemma stable_regularCard: |
|
76951 | 1762 |
assumes cr: "Card_order r" and ir: "\<not>finite (Field r)" and st: "stable r" |
1763 |
shows "regularCard r" |
|
1764 |
unfolding regularCard_def proof safe |
|
75624 | 1765 |
fix K assume K: "K \<subseteq> Field r" and cof: "cofinal K r" |
1766 |
have "|K| \<le>o r" using K card_of_Field_ordIso card_of_mono1 cr ordLeq_ordIso_trans by blast |
|
1767 |
moreover |
|
1768 |
{assume Kr: "|K| <o r" |
|
76951 | 1769 |
have x: "\<And>a. a \<in> Field r \<Longrightarrow> \<exists>b. b \<in> K \<and> a \<noteq> b \<and> (a, b) \<in> r" using cof unfolding cofinal_def by blast |
1770 |
then obtain f where "\<And>a. a \<in> Field r \<Longrightarrow> f a = (SOME b. b \<in> K \<and> a \<noteq> b \<and> (a, b) \<in> r)" by simp |
|
1771 |
then have "\<forall>a\<in>Field r. f a \<in> K \<and> a \<noteq> f a \<and> (a, f a) \<in> r" using someI_ex[OF x] by simp |
|
1772 |
hence "Field r \<subseteq> (\<Union>a \<in> K. underS r a)" unfolding underS_def by auto |
|
1773 |
hence "r \<le>o |\<Union>a \<in> K. underS r a|" |
|
1774 |
using cr Card_order_iff_ordLeq_card_of card_of_mono1 ordLeq_transitive by blast |
|
1775 |
also have "|\<Union>a \<in> K. underS r a| \<le>o |SIGMA a: K. underS r a|" by (rule card_of_UNION_Sigma) |
|
1776 |
also |
|
1777 |
{have "\<forall> a \<in> K. |underS r a| <o r" using K card_of_underS[OF cr] subsetD by auto |
|
1778 |
hence "|SIGMA a: K. underS r a| <o r" using st Kr unfolding stable_def by auto |
|
1779 |
} |
|
1780 |
finally have "r <o r" . |
|
1781 |
hence False using ordLess_irreflexive by blast |
|
75624 | 1782 |
} |
1783 |
ultimately show "|K| =o r" using ordLeq_iff_ordLess_or_ordIso by blast |
|
1784 |
qed |
|
1785 |
||
1786 |
lemma internalize_card_of_ordLess: |
|
76951 | 1787 |
"( |A| <o r) = (\<exists>B < Field r. |A| =o |B| \<and> |B| <o r)" |
75624 | 1788 |
proof |
1789 |
assume "|A| <o r" |
|
1790 |
then obtain p where 1: "Field p < Field r \<and> |A| =o p \<and> p <o r" |
|
76951 | 1791 |
using internalize_ordLess[of "|A|" r] by blast |
75624 | 1792 |
hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast |
1793 |
hence "|Field p| =o p" using card_of_Field_ordIso by blast |
|
1794 |
hence "|A| =o |Field p| \<and> |Field p| <o r" |
|
76951 | 1795 |
using 1 ordIso_equivalence ordIso_ordLess_trans by blast |
75624 | 1796 |
thus "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r" using 1 by blast |
1797 |
next |
|
1798 |
assume "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r" |
|
1799 |
thus "|A| <o r" using ordIso_ordLess_trans by blast |
|
1800 |
qed |
|
1801 |
||
1802 |
lemma card_of_Sigma_cong1: |
|
76951 | 1803 |
assumes "\<forall>i \<in> I. |A i| =o |B i|" |
1804 |
shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|" |
|
1805 |
using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq) |
|
75624 | 1806 |
|
1807 |
lemma card_of_Sigma_cong2: |
|
76951 | 1808 |
assumes "bij_betw f (I::'i set) (J::'j set)" |
1809 |
shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| =o |SIGMA j : J. A j|" |
|
1810 |
proof - |
|
75624 | 1811 |
let ?LEFT = "SIGMA i : I. A (f i)" |
1812 |
let ?RIGHT = "SIGMA j : J. A j" |
|
76951 | 1813 |
define u where "u \<equiv> \<lambda>(i::'i,a::'a). (f i,a)" |
75624 | 1814 |
have "bij_betw u ?LEFT ?RIGHT" |
76951 | 1815 |
using assms unfolding u_def bij_betw_def inj_on_def by auto |
75624 | 1816 |
thus ?thesis using card_of_ordIso by blast |
1817 |
qed |
|
1818 |
||
1819 |
lemma card_of_Sigma_cong: |
|
76951 | 1820 |
assumes BIJ: "bij_betw f I J" and ISO: "\<forall>j \<in> J. |A j| =o |B j|" |
1821 |
shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|" |
|
1822 |
proof - |
|
75624 | 1823 |
have "\<forall>i \<in> I. |A(f i)| =o |B(f i)|" |
76951 | 1824 |
using ISO BIJ unfolding bij_betw_def by blast |
75624 | 1825 |
hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|" by (rule card_of_Sigma_cong1) |
1826 |
moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|" |
|
76951 | 1827 |
using BIJ card_of_Sigma_cong2 by blast |
75624 | 1828 |
ultimately show ?thesis using ordIso_transitive by blast |
1829 |
qed |
|
1830 |
||
1831 |
(* Note that below the types of A and F are now unconstrained: *) |
|
1832 |
lemma stable_elim: |
|
76951 | 1833 |
assumes ST: "stable r" and A_LESS: "|A| <o r" and |
1834 |
F_LESS: "\<And> a. a \<in> A \<Longrightarrow> |F a| <o r" |
|
1835 |
shows "|SIGMA a : A. F a| <o r" |
|
1836 |
proof - |
|
75624 | 1837 |
obtain A' where 1: "A' < Field r \<and> |A'| <o r" and 2: " |A| =o |A'|" |
76951 | 1838 |
using internalize_card_of_ordLess[of A r] A_LESS by blast |
75624 | 1839 |
then obtain G where 3: "bij_betw G A' A" |
76951 | 1840 |
using card_of_ordIso ordIso_symmetric by blast |
1841 |
(* *) |
|
75624 | 1842 |
{fix a assume "a \<in> A" |
76951 | 1843 |
hence "\<exists>B'. B' \<le> Field r \<and> |F a| =o |B'| \<and> |B'| <o r" |
1844 |
using internalize_card_of_ordLess[of "F a" r] F_LESS by blast |
|
75624 | 1845 |
} |
1846 |
then obtain F' where |
|
76951 | 1847 |
temp: "\<forall>a \<in> A. F' a \<le> Field r \<and> |F a| =o |F' a| \<and> |F' a| <o r" |
1848 |
using bchoice[of A "\<lambda> a B'. B' \<le> Field r \<and> |F a| =o |B'| \<and> |B'| <o r"] by blast |
|
75624 | 1849 |
hence 4: "\<forall>a \<in> A. F' a \<le> Field r \<and> |F' a| <o r" by auto |
1850 |
have 5: "\<forall>a \<in> A. |F' a| =o |F a|" using temp ordIso_symmetric by auto |
|
76951 | 1851 |
(* *) |
75624 | 1852 |
have "\<forall>a' \<in> A'. F'(G a') \<le> Field r \<and> |F'(G a')| <o r" |
76951 | 1853 |
using 3 4 bij_betw_ball[of G A' A] by auto |
75624 | 1854 |
hence "|SIGMA a' : A'. F'(G a')| <o r" |
76951 | 1855 |
using ST 1 unfolding stable_def by auto |
75624 | 1856 |
moreover have "|SIGMA a' : A'. F'(G a')| =o |SIGMA a : A. F a|" |
76951 | 1857 |
using card_of_Sigma_cong[of G A' A F' F] 5 3 by blast |
75624 | 1858 |
ultimately show ?thesis using ordIso_symmetric ordIso_ordLess_trans by blast |
1859 |
qed |
|
1860 |
||
1861 |
lemma stable_natLeq: "stable natLeq" |
|
1862 |
proof(unfold stable_def, safe) |
|
1863 |
fix A :: "'a set" and F :: "'a \<Rightarrow> 'a set" |
|
1864 |
assume "|A| <o natLeq" and "\<forall>a\<in>A. |F a| <o natLeq" |
|
1865 |
hence "finite A \<and> (\<forall>a \<in> A. finite(F a))" |
|
76951 | 1866 |
by (auto simp add: finite_iff_ordLess_natLeq) |
75624 | 1867 |
hence "finite(Sigma A F)" by (simp only: finite_SigmaI[of A F]) |
1868 |
thus "|Sigma A F | <o natLeq" |
|
76951 | 1869 |
by (auto simp add: finite_iff_ordLess_natLeq) |
75624 | 1870 |
qed |
1871 |
||
1872 |
corollary regularCard_natLeq: "regularCard natLeq" |
|
76951 | 1873 |
using stable_regularCard[OF natLeq_Card_order _ stable_natLeq] Field_natLeq by simp |
75624 | 1874 |
|
1875 |
lemma stable_ordIso1: |
|
76951 | 1876 |
assumes ST: "stable r" and ISO: "r' =o r" |
1877 |
shows "stable r'" |
|
75624 | 1878 |
proof(unfold stable_def, auto) |
1879 |
fix A::"'b set" and F::"'b \<Rightarrow> 'b set" |
|
1880 |
assume "|A| <o r'" and "\<forall>a \<in> A. |F a| <o r'" |
|
1881 |
hence "( |A| <o r) \<and> (\<forall>a \<in> A. |F a| <o r)" |
|
76951 | 1882 |
using ISO ordLess_ordIso_trans by blast |
75624 | 1883 |
hence "|SIGMA a : A. F a| <o r" using assms stable_elim by blast |
1884 |
thus "|SIGMA a : A. F a| <o r'" |
|
76951 | 1885 |
using ISO ordIso_symmetric ordLess_ordIso_trans by blast |
75624 | 1886 |
qed |
1887 |
||
1888 |
lemma stable_UNION: |
|
76951 | 1889 |
assumes "stable r" and "|A| <o r" and "\<And> a. a \<in> A \<Longrightarrow> |F a| <o r" |
1890 |
shows "|\<Union>a \<in> A. F a| <o r" |
|
1891 |
using assms card_of_UNION_Sigma stable_elim ordLeq_ordLess_trans by blast |
|
75624 | 1892 |
|
1893 |
corollary card_of_UNION_ordLess_infinite: |
|
76951 | 1894 |
assumes "stable |B|" and "|I| <o |B|" and "\<forall>i \<in> I. |A i| <o |B|" |
1895 |
shows "|\<Union>i \<in> I. A i| <o |B|" |
|
75624 | 1896 |
using assms stable_UNION by blast |
1897 |
||
1898 |
corollary card_of_UNION_ordLess_infinite_Field: |
|
76951 | 1899 |
assumes ST: "stable r" and r: "Card_order r" and |
1900 |
LEQ_I: "|I| <o r" and LEQ: "\<forall>i \<in> I. |A i| <o r" |
|
1901 |
shows "|\<Union>i \<in> I. A i| <o r" |
|
1902 |
proof - |
|
75624 | 1903 |
let ?B = "Field r" |
1904 |
have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso |
|
76951 | 1905 |
ordIso_symmetric by blast |
75624 | 1906 |
hence "|I| <o |?B|" "\<forall>i \<in> I. |A i| <o |?B|" |
1907 |
using LEQ_I LEQ ordLess_ordIso_trans by blast+ |
|
1908 |
moreover have "stable |?B|" using stable_ordIso1 ST 1 by blast |
|
1909 |
ultimately have "|\<Union>i \<in> I. A i| <o |?B|" using LEQ |
|
76951 | 1910 |
card_of_UNION_ordLess_infinite by blast |
75624 | 1911 |
thus ?thesis using 1 ordLess_ordIso_trans by blast |
1912 |
qed |
|
1913 |
||
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
1914 |
end |