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