started three-way split of 'HOL-Cardinals'
authorblanchet
Mon, 18 Nov 2013 18:04:44 +0100
changeset 54473 8bee5ca99e63
parent 54472 073f041d83ae
child 54474 6d5941722fae
started three-way split of 'HOL-Cardinals'
src/HOL/Cardinals/Cardinal_Arithmetic.thy
src/HOL/Cardinals/Cardinal_Order_Relation.thy
src/HOL/Cardinals/Cardinal_Order_Relation_Base.thy
src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy
src/HOL/Cardinals/Constructions_on_Wellorders.thy
src/HOL/Cardinals/Constructions_on_Wellorders_Base.thy
src/HOL/Cardinals/Constructions_on_Wellorders_LFP.thy
src/HOL/Cardinals/Fun_More.thy
src/HOL/Cardinals/Fun_More_Base.thy
src/HOL/Cardinals/Fun_More_LFP.thy
src/HOL/Cardinals/Order_Relation_More.thy
src/HOL/Cardinals/Order_Relation_More_Base.thy
src/HOL/Cardinals/Order_Relation_More_LFP.thy
src/HOL/Cardinals/README.txt
src/HOL/Cardinals/Wellfounded_More.thy
src/HOL/Cardinals/Wellfounded_More_Base.thy
src/HOL/Cardinals/Wellfounded_More_LFP.thy
src/HOL/Cardinals/Wellorder_Embedding.thy
src/HOL/Cardinals/Wellorder_Embedding_Base.thy
src/HOL/Cardinals/Wellorder_Embedding_LFP.thy
src/HOL/Cardinals/Wellorder_Relation.thy
src/HOL/Cardinals/Wellorder_Relation_Base.thy
src/HOL/Cardinals/Wellorder_Relation_LFP.thy
src/HOL/Library/Order_Union.thy
src/HOL/ROOT
--- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy	Mon Nov 18 18:04:44 2013 +0100
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -8,7 +8,7 @@
 header {* Cardinal Arithmetic  *}
 
 theory Cardinal_Arithmetic
-imports Cardinal_Order_Relation_Base
+imports Cardinal_Order_Relation_LFP
 begin
 
 text {*
--- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy	Mon Nov 18 18:04:44 2013 +0100
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -8,7 +8,7 @@
 header {* Cardinal-Order Relations *}
 
 theory Cardinal_Order_Relation
-imports Cardinal_Order_Relation_Base Constructions_on_Wellorders
+imports Cardinal_Order_Relation_LFP Constructions_on_Wellorders
 begin
 
 declare
--- a/src/HOL/Cardinals/Cardinal_Order_Relation_Base.thy	Mon Nov 18 18:04:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2438 +0,0 @@
-(*  Title:      HOL/Cardinals/Cardinal_Order_Relation_Base.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Cardinal-order relations (base).
-*)
-
-header {* Cardinal-Order Relations (Base)  *}
-
-theory Cardinal_Order_Relation_Base
-imports Constructions_on_Wellorders_Base
-begin
-
-
-text{* In this section, we define cardinal-order relations to be minim well-orders
-on their field.  Then we define the cardinal of a set to be {\em some} cardinal-order
-relation on that set, which will be unique up to order isomorphism.  Then we study
-the connection between cardinals and:
-\begin{itemize}
-\item standard set-theoretic constructions: products,
-sums, unions, lists, powersets, set-of finite sets operator;
-\item finiteness and infiniteness (in particular, with the numeric cardinal operator
-for finite sets, @{text "card"}, from the theory @{text "Finite_Sets.thy"}).
-\end{itemize}
-%
-On the way, we define the canonical $\omega$ cardinal and finite cardinals.  We also
-define (again, up to order isomorphism) the successor of a cardinal, and show that
-any cardinal admits a successor.
-
-Main results of this section are the existence of cardinal relations and the
-facts that, in the presence of infiniteness,
-most of the standard set-theoretic constructions (except for the powerset)
-{\em do not increase cardinality}.  In particular, e.g., the set of words/lists over
-any infinite set has the same cardinality (hence, is in bijection) with that set.
-*}
-
-
-subsection {* Cardinal orders *}
-
-
-text{* A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the
-order-embedding relation, @{text "\<le>o"} (which is the same as being {\em minimal} w.r.t. the
-strict order-embedding relation, @{text "<o"}), among all the well-orders on its field.  *}
-
-definition card_order_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
-where
-"card_order_on A r \<equiv> well_order_on A r \<and> (\<forall>r'. well_order_on A r' \<longrightarrow> r \<le>o r')"
-
-
-abbreviation "Card_order r \<equiv> card_order_on (Field r) r"
-abbreviation "card_order r \<equiv> card_order_on UNIV r"
-
-
-lemma card_order_on_well_order_on:
-assumes "card_order_on A r"
-shows "well_order_on A r"
-using assms unfolding card_order_on_def by simp
-
-
-lemma card_order_on_Card_order:
-"card_order_on A r \<Longrightarrow> A = Field r \<and> Card_order r"
-unfolding card_order_on_def using rel.well_order_on_Field by blast
-
-
-text{* The existence of a cardinal relation on any given set (which will mean
-that any set has a cardinal) follows from two facts:
-\begin{itemize}
-\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}),
-which states that on any given set there exists a well-order;
-\item The well-founded-ness of @{text "<o"}, ensuring that then there exists a minimal
-such well-order, i.e., a cardinal order.
-\end{itemize}
-*}
-
-
-theorem card_order_on: "\<exists>r. card_order_on A r"
-proof-
-  obtain R where R_def: "R = {r. well_order_on A r}" by blast
-  have 1: "R \<noteq> {} \<and> (\<forall>r \<in> R. Well_order r)"
-  using well_order_on[of A] R_def rel.well_order_on_Well_order by blast
-  hence "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
-  using  exists_minim_Well_order[of R] by auto
-  thus ?thesis using R_def unfolding card_order_on_def by auto
-qed
-
-
-lemma card_order_on_ordIso:
-assumes CO: "card_order_on A r" and CO': "card_order_on A r'"
-shows "r =o r'"
-using assms unfolding card_order_on_def
-using ordIso_iff_ordLeq by blast
-
-
-lemma Card_order_ordIso:
-assumes CO: "Card_order r" and ISO: "r' =o r"
-shows "Card_order r'"
-using ISO unfolding ordIso_def
-proof(unfold card_order_on_def, auto)
-  fix p' assume "well_order_on (Field r') p'"
-  hence 0: "Well_order p' \<and> Field p' = Field r'"
-  using rel.well_order_on_Well_order by blast
-  obtain f where 1: "iso r' r f" and 2: "Well_order r \<and> Well_order r'"
-  using ISO unfolding ordIso_def by auto
-  hence 3: "inj_on f (Field r') \<and> f ` (Field r') = Field r"
-  by (auto simp add: iso_iff embed_inj_on)
-  let ?p = "dir_image p' f"
-  have 4: "p' =o ?p \<and> Well_order ?p"
-  using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image)
-  moreover have "Field ?p =  Field r"
-  using 0 3 by (auto simp add: dir_image_Field2 order_on_defs)
-  ultimately have "well_order_on (Field r) ?p" by auto
-  hence "r \<le>o ?p" using CO unfolding card_order_on_def by auto
-  thus "r' \<le>o p'"
-  using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast
-qed
-
-
-lemma Card_order_ordIso2:
-assumes CO: "Card_order r" and ISO: "r =o r'"
-shows "Card_order r'"
-using assms Card_order_ordIso ordIso_symmetric by blast
-
-
-subsection {* Cardinal of a set *}
-
-
-text{* We define the cardinal of set to be {\em some} cardinal order on that set.
-We shall prove that this notion is unique up to order isomorphism, meaning
-that order isomorphism shall be the true identity of cardinals.  *}
-
-
-definition card_of :: "'a set \<Rightarrow> 'a rel" ("|_|" )
-where "card_of A = (SOME r. card_order_on A r)"
-
-
-lemma card_of_card_order_on: "card_order_on A |A|"
-unfolding card_of_def by (auto simp add: card_order_on someI_ex)
-
-
-lemma card_of_well_order_on: "well_order_on A |A|"
-using card_of_card_order_on card_order_on_def by blast
-
-
-lemma Field_card_of: "Field |A| = A"
-using card_of_card_order_on[of A] unfolding card_order_on_def
-using rel.well_order_on_Field by blast
-
-
-lemma card_of_Card_order: "Card_order |A|"
-by (simp only: card_of_card_order_on Field_card_of)
-
-
-corollary ordIso_card_of_imp_Card_order:
-"r =o |A| \<Longrightarrow> Card_order r"
-using card_of_Card_order Card_order_ordIso by blast
-
-
-lemma card_of_Well_order: "Well_order |A|"
-using card_of_Card_order unfolding  card_order_on_def by auto
-
-
-lemma card_of_refl: "|A| =o |A|"
-using card_of_Well_order ordIso_reflexive by blast
-
-
-lemma card_of_least: "well_order_on A r \<Longrightarrow> |A| \<le>o r"
-using card_of_card_order_on unfolding card_order_on_def by blast
-
-
-lemma card_of_ordIso:
-"(\<exists>f. bij_betw f A B) = ( |A| =o |B| )"
-proof(auto)
-  fix f assume *: "bij_betw f A B"
-  then obtain r where "well_order_on B r \<and> |A| =o r"
-  using Well_order_iso_copy card_of_well_order_on by blast
-  hence "|B| \<le>o |A|" using card_of_least
-  ordLeq_ordIso_trans ordIso_symmetric by blast
-  moreover
-  {let ?g = "inv_into A f"
-   have "bij_betw ?g B A" using * bij_betw_inv_into by blast
-   then obtain r where "well_order_on A r \<and> |B| =o r"
-   using Well_order_iso_copy card_of_well_order_on by blast
-   hence "|A| \<le>o |B|" using card_of_least
-   ordLeq_ordIso_trans ordIso_symmetric by blast
-  }
-  ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast
-next
-  assume "|A| =o |B|"
-  then obtain f where "iso ( |A| ) ( |B| ) f"
-  unfolding ordIso_def by auto
-  hence "bij_betw f A B" unfolding iso_def Field_card_of by simp
-  thus "\<exists>f. bij_betw f A B" by auto
-qed
-
-
-lemma card_of_ordLeq:
-"(\<exists>f. inj_on f A \<and> f ` A \<le> B) = ( |A| \<le>o |B| )"
-proof(auto)
-  fix f assume *: "inj_on f A" and **: "f ` A \<le> B"
-  {assume "|B| <o |A|"
-   hence "|B| \<le>o |A|" using ordLeq_iff_ordLess_or_ordIso by blast
-   then obtain g where "embed ( |B| ) ( |A| ) g"
-   unfolding ordLeq_def by auto
-   hence 1: "inj_on g B \<and> g ` B \<le> A" using embed_inj_on[of "|B|" "|A|" "g"]
-   card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"]
-   embed_Field[of "|B|" "|A|" g] by auto
-   obtain h where "bij_betw h A B"
-   using * ** 1 Cantor_Bernstein[of f] by fastforce
-   hence "|A| =o |B|" using card_of_ordIso by blast
-   hence "|A| \<le>o |B|" using ordIso_iff_ordLeq by auto
-  }
-  thus "|A| \<le>o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"]
-  by (auto simp: card_of_Well_order)
-next
-  assume *: "|A| \<le>o |B|"
-  obtain f where "embed ( |A| ) ( |B| ) f"
-  using * unfolding ordLeq_def by auto
-  hence "inj_on f A \<and> f ` A \<le> B" using embed_inj_on[of "|A|" "|B|" f]
-  card_of_Well_order[of "A"] Field_card_of[of "A"] Field_card_of[of "B"]
-  embed_Field[of "|A|" "|B|" f] by auto
-  thus "\<exists>f. inj_on f A \<and> f ` A \<le> B" by auto
-qed
-
-
-lemma card_of_ordLeq2:
-"A \<noteq> {} \<Longrightarrow> (\<exists>g. g ` B = A) = ( |A| \<le>o |B| )"
-using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto
-
-
-lemma card_of_ordLess:
-"(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = ( |B| <o |A| )"
-proof-
-  have "(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = (\<not> |A| \<le>o |B| )"
-  using card_of_ordLeq by blast
-  also have "\<dots> = ( |B| <o |A| )"
-  using card_of_Well_order[of A] card_of_Well_order[of B]
-        not_ordLeq_iff_ordLess by blast
-  finally show ?thesis .
-qed
-
-
-lemma card_of_ordLess2:
-"B \<noteq> {} \<Longrightarrow> (\<not>(\<exists>f. f ` A = B)) = ( |A| <o |B| )"
-using card_of_ordLess[of B A] inj_on_iff_surj[of B A] by auto
-
-
-lemma card_of_ordIsoI:
-assumes "bij_betw f A B"
-shows "|A| =o |B|"
-using assms unfolding card_of_ordIso[symmetric] by auto
-
-
-lemma card_of_ordLeqI:
-assumes "inj_on f A" and "\<And> a. a \<in> A \<Longrightarrow> f a \<in> B"
-shows "|A| \<le>o |B|"
-using assms unfolding card_of_ordLeq[symmetric] by auto
-
-
-lemma card_of_unique:
-"card_order_on A r \<Longrightarrow> r =o |A|"
-by (simp only: card_order_on_ordIso card_of_card_order_on)
-
-
-lemma card_of_mono1:
-"A \<le> B \<Longrightarrow> |A| \<le>o |B|"
-using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce
-
-
-lemma card_of_mono2:
-assumes "r \<le>o r'"
-shows "|Field r| \<le>o |Field r'|"
-proof-
-  obtain f where
-  1: "well_order_on (Field r) r \<and> well_order_on (Field r) r \<and> embed r r' f"
-  using assms unfolding ordLeq_def
-  by (auto simp add: rel.well_order_on_Well_order)
-  hence "inj_on f (Field r) \<and> f ` (Field r) \<le> Field r'"
-  by (auto simp add: embed_inj_on embed_Field)
-  thus "|Field r| \<le>o |Field r'|" using card_of_ordLeq by blast
-qed
-
-
-lemma card_of_cong: "r =o r' \<Longrightarrow> |Field r| =o |Field r'|"
-by (simp add: ordIso_iff_ordLeq card_of_mono2)
-
-
-lemma card_of_Field_ordLess: "Well_order r \<Longrightarrow> |Field r| \<le>o r"
-using card_of_least card_of_well_order_on rel.well_order_on_Well_order by blast
-
-
-lemma card_of_Field_ordIso:
-assumes "Card_order r"
-shows "|Field r| =o r"
-proof-
-  have "card_order_on (Field r) r"
-  using assms card_order_on_Card_order by blast
-  moreover have "card_order_on (Field r) |Field r|"
-  using card_of_card_order_on by blast
-  ultimately show ?thesis using card_order_on_ordIso by blast
-qed
-
-
-lemma Card_order_iff_ordIso_card_of:
-"Card_order r = (r =o |Field r| )"
-using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast
-
-
-lemma Card_order_iff_ordLeq_card_of:
-"Card_order r = (r \<le>o |Field r| )"
-proof-
-  have "Card_order r = (r =o |Field r| )"
-  unfolding Card_order_iff_ordIso_card_of by simp
-  also have "... = (r \<le>o |Field r| \<and> |Field r| \<le>o r)"
-  unfolding ordIso_iff_ordLeq by simp
-  also have "... = (r \<le>o |Field r| )"
-  using card_of_Field_ordLess
-  by (auto simp: card_of_Field_ordLess ordLeq_Well_order_simp)
-  finally show ?thesis .
-qed
-
-
-lemma Card_order_iff_Restr_underS:
-assumes "Well_order r"
-shows "Card_order r = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o |Field r| )"
-using assms unfolding Card_order_iff_ordLeq_card_of
-using ordLeq_iff_ordLess_Restr card_of_Well_order by blast
-
-
-lemma card_of_underS:
-assumes r: "Card_order r" and a: "a : Field r"
-shows "|rel.underS r a| <o r"
-proof-
-  let ?A = "rel.underS r a"  let ?r' = "Restr r ?A"
-  have 1: "Well_order r"
-  using r unfolding card_order_on_def by simp
-  have "Well_order ?r'" using 1 Well_order_Restr by auto
-  moreover have "card_order_on (Field ?r') |Field ?r'|"
-  using card_of_card_order_on .
-  ultimately have "|Field ?r'| \<le>o ?r'"
-  unfolding card_order_on_def by simp
-  moreover have "Field ?r' = ?A"
-  using 1 wo_rel.underS_ofilter Field_Restr_ofilter
-  unfolding wo_rel_def by fastforce
-  ultimately have "|?A| \<le>o ?r'" by simp
-  also have "?r' <o |Field r|"
-  using 1 a r Card_order_iff_Restr_underS by blast
-  also have "|Field r| =o r"
-  using r ordIso_symmetric unfolding Card_order_iff_ordIso_card_of by auto
-  finally show ?thesis .
-qed
-
-
-lemma ordLess_Field:
-assumes "r <o r'"
-shows "|Field r| <o r'"
-proof-
-  have "well_order_on (Field r) r" using assms unfolding ordLess_def
-  by (auto simp add: rel.well_order_on_Well_order)
-  hence "|Field r| \<le>o r" using card_of_least by blast
-  thus ?thesis using assms ordLeq_ordLess_trans by blast
-qed
-
-
-lemma internalize_card_of_ordLeq:
-"( |A| \<le>o r) = (\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r)"
-proof
-  assume "|A| \<le>o r"
-  then obtain p where 1: "Field p \<le> Field r \<and> |A| =o p \<and> p \<le>o r"
-  using internalize_ordLeq[of "|A|" r] by blast
-  hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast
-  hence "|Field p| =o p" using card_of_Field_ordIso by blast
-  hence "|A| =o |Field p| \<and> |Field p| \<le>o r"
-  using 1 ordIso_equivalence ordIso_ordLeq_trans by blast
-  thus "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r" using 1 by blast
-next
-  assume "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r"
-  thus "|A| \<le>o r" using ordIso_ordLeq_trans by blast
-qed
-
-
-lemma internalize_card_of_ordLeq2:
-"( |A| \<le>o |C| ) = (\<exists>B \<le> C. |A| =o |B| \<and> |B| \<le>o |C| )"
-using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto
-
-
-
-subsection {* Cardinals versus set operations on arbitrary sets *}
-
-
-text{* Here we embark in a long journey of simple results showing
-that the standard set-theoretic operations are well-behaved w.r.t. the notion of
-cardinal -- essentially, this means that they preserve the ``cardinal identity"
-@{text "=o"} and are monotonic w.r.t. @{text "\<le>o"}.
-*}
-
-
-lemma card_of_empty: "|{}| \<le>o |A|"
-using card_of_ordLeq inj_on_id by blast
-
-
-lemma card_of_empty1:
-assumes "Well_order r \<or> Card_order r"
-shows "|{}| \<le>o r"
-proof-
-  have "Well_order r" using assms unfolding card_order_on_def by auto
-  hence "|Field r| <=o r"
-  using assms card_of_Field_ordLess by blast
-  moreover have "|{}| \<le>o |Field r|" by (simp add: card_of_empty)
-  ultimately show ?thesis using ordLeq_transitive by blast
-qed
-
-
-corollary Card_order_empty:
-"Card_order r \<Longrightarrow> |{}| \<le>o r" by (simp add: card_of_empty1)
-
-
-lemma card_of_empty2:
-assumes LEQ: "|A| =o |{}|"
-shows "A = {}"
-using assms card_of_ordIso[of A] bij_betw_empty2 by blast
-
-
-lemma card_of_empty3:
-assumes LEQ: "|A| \<le>o |{}|"
-shows "A = {}"
-using assms
-by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2
-              ordLeq_Well_order_simp)
-
-
-lemma card_of_empty_ordIso:
-"|{}::'a set| =o |{}::'b set|"
-using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
-
-
-lemma card_of_image:
-"|f ` A| <=o |A|"
-proof(cases "A = {}", simp add: card_of_empty)
-  assume "A ~= {}"
-  hence "f ` A ~= {}" by auto
-  thus "|f ` A| \<le>o |A|"
-  using card_of_ordLeq2[of "f ` A" A] by auto
-qed
-
-
-lemma surj_imp_ordLeq:
-assumes "B <= f ` A"
-shows "|B| <=o |A|"
-proof-
-  have "|B| <=o |f ` A|" using assms card_of_mono1 by auto
-  thus ?thesis using card_of_image ordLeq_transitive by blast
-qed
-
-
-lemma card_of_ordLeqI2:
-assumes "B \<subseteq> f ` A"
-shows "|B| \<le>o |A|"
-using assms by (metis surj_imp_ordLeq)
-
-
-lemma card_of_singl_ordLeq:
-assumes "A \<noteq> {}"
-shows "|{b}| \<le>o |A|"
-proof-
-  obtain a where *: "a \<in> A" using assms by auto
-  let ?h = "\<lambda> b'::'b. if b' = b then a else undefined"
-  have "inj_on ?h {b} \<and> ?h ` {b} \<le> A"
-  using * unfolding inj_on_def by auto
-  thus ?thesis using card_of_ordLeq by blast
-qed
-
-
-corollary Card_order_singl_ordLeq:
-"\<lbrakk>Card_order r; Field r \<noteq> {}\<rbrakk> \<Longrightarrow> |{b}| \<le>o r"
-using card_of_singl_ordLeq[of "Field r" b]
-      card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast
-
-
-lemma card_of_Pow: "|A| <o |Pow A|"
-using card_of_ordLess2[of "Pow A" A]  Cantors_paradox[of A]
-      Pow_not_empty[of A] by auto
-
-
-lemma infinite_Pow:
-assumes "infinite A"
-shows "infinite (Pow A)"
-proof-
-  have "|A| \<le>o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq)
-  thus ?thesis by (metis assms finite_Pow_iff)
-qed
-
-
-corollary Card_order_Pow:
-"Card_order r \<Longrightarrow> r <o |Pow(Field r)|"
-using card_of_Pow card_of_Field_ordIso ordIso_ordLess_trans ordIso_symmetric by blast
-
-
-corollary card_of_set_type: "|UNIV::'a set| <o |UNIV::'a set set|"
-using card_of_Pow[of "UNIV::'a set"] by simp
-
-
-lemma card_of_Plus1: "|A| \<le>o |A <+> B|"
-proof-
-  have "Inl ` A \<le> A <+> B" by auto
-  thus ?thesis using inj_Inl[of A] card_of_ordLeq by blast
-qed
-
-
-corollary Card_order_Plus1:
-"Card_order r \<Longrightarrow> r \<le>o |(Field r) <+> B|"
-using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
-
-
-lemma card_of_Plus2: "|B| \<le>o |A <+> B|"
-proof-
-  have "Inr ` B \<le> A <+> B" by auto
-  thus ?thesis using inj_Inr[of B] card_of_ordLeq by blast
-qed
-
-
-corollary Card_order_Plus2:
-"Card_order r \<Longrightarrow> r \<le>o |A <+> (Field r)|"
-using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
-
-
-lemma card_of_Plus_empty1: "|A| =o |A <+> {}|"
-proof-
-  have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto
-  thus ?thesis using card_of_ordIso by auto
-qed
-
-
-lemma card_of_Plus_empty2: "|A| =o |{} <+> A|"
-proof-
-  have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto
-  thus ?thesis using card_of_ordIso by auto
-qed
-
-
-lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|"
-proof-
-  let ?f = "\<lambda>(c::'a + 'b). case c of Inl a \<Rightarrow> Inr a
-                                   | Inr b \<Rightarrow> Inl b"
-  have "bij_betw ?f (A <+> B) (B <+> A)"
-  unfolding bij_betw_def inj_on_def by force
-  thus ?thesis using card_of_ordIso by blast
-qed
-
-
-lemma card_of_Plus_assoc:
-fixes A :: "'a set" and B :: "'b set" and C :: "'c set"
-shows "|(A <+> B) <+> C| =o |A <+> B <+> C|"
-proof -
-  def f \<equiv> "\<lambda>(k::('a + 'b) + 'c).
-  case k of Inl ab \<Rightarrow> (case ab of Inl a \<Rightarrow> Inl a
-                                 |Inr b \<Rightarrow> Inr (Inl b))
-           |Inr c \<Rightarrow> Inr (Inr c)"
-  have "A <+> B <+> C \<subseteq> f ` ((A <+> B) <+> C)"
-  proof
-    fix x assume x: "x \<in> A <+> B <+> C"
-    show "x \<in> f ` ((A <+> B) <+> C)"
-    proof(cases x)
-      case (Inl a)
-      hence "a \<in> A" "x = f (Inl (Inl a))"
-      using x unfolding f_def by auto
-      thus ?thesis by auto
-    next
-      case (Inr bc) note 1 = Inr show ?thesis
-      proof(cases bc)
-        case (Inl b)
-        hence "b \<in> B" "x = f (Inl (Inr b))"
-        using x 1 unfolding f_def by auto
-        thus ?thesis by auto
-      next
-        case (Inr c)
-        hence "c \<in> C" "x = f (Inr c)"
-        using x 1 unfolding f_def by auto
-        thus ?thesis by auto
-      qed
-    qed
-  qed
-  hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)"
-  unfolding bij_betw_def inj_on_def f_def by force
-  thus ?thesis using card_of_ordIso by blast
-qed
-
-
-lemma card_of_Plus_mono1:
-assumes "|A| \<le>o |B|"
-shows "|A <+> C| \<le>o |B <+> C|"
-proof-
-  obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
-  using assms card_of_ordLeq[of A] by fastforce
-  obtain g where g_def:
-  "g = (\<lambda>d. case d of Inl a \<Rightarrow> Inl(f a) | Inr (c::'c) \<Rightarrow> Inr c)" by blast
-  have "inj_on g (A <+> C) \<and> g ` (A <+> C) \<le> (B <+> C)"
-  proof-
-    {fix d1 and d2 assume "d1 \<in> A <+> C \<and> d2 \<in> A <+> C" and
-                          "g d1 = g d2"
-     hence "d1 = d2" using 1 unfolding inj_on_def
-     by(case_tac d1, case_tac d2, auto simp add: g_def)
-    }
-    moreover
-    {fix d assume "d \<in> A <+> C"
-     hence "g d \<in> B <+> C"  using 1
-     by(case_tac d, auto simp add: g_def)
-    }
-    ultimately show ?thesis unfolding inj_on_def by auto
-  qed
-  thus ?thesis using card_of_ordLeq by metis
-qed
-
-
-corollary ordLeq_Plus_mono1:
-assumes "r \<le>o r'"
-shows "|(Field r) <+> C| \<le>o |(Field r') <+> C|"
-using assms card_of_mono2 card_of_Plus_mono1 by blast
-
-
-lemma card_of_Plus_mono2:
-assumes "|A| \<le>o |B|"
-shows "|C <+> A| \<le>o |C <+> B|"
-using assms card_of_Plus_mono1[of A B C]
-      card_of_Plus_commute[of C A]  card_of_Plus_commute[of B C]
-      ordIso_ordLeq_trans[of "|C <+> A|"] ordLeq_ordIso_trans[of "|C <+> A|"]
-by blast
-
-
-corollary ordLeq_Plus_mono2:
-assumes "r \<le>o r'"
-shows "|A <+> (Field r)| \<le>o |A <+> (Field r')|"
-using assms card_of_mono2 card_of_Plus_mono2 by blast
-
-
-lemma card_of_Plus_mono:
-assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
-shows "|A <+> C| \<le>o |B <+> D|"
-using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B]
-      ordLeq_transitive[of "|A <+> C|"] by blast
-
-
-corollary ordLeq_Plus_mono:
-assumes "r \<le>o r'" and "p \<le>o p'"
-shows "|(Field r) <+> (Field p)| \<le>o |(Field r') <+> (Field p')|"
-using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast
-
-
-lemma card_of_Plus_cong1:
-assumes "|A| =o |B|"
-shows "|A <+> C| =o |B <+> C|"
-using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1)
-
-
-corollary ordIso_Plus_cong1:
-assumes "r =o r'"
-shows "|(Field r) <+> C| =o |(Field r') <+> C|"
-using assms card_of_cong card_of_Plus_cong1 by blast
-
-
-lemma card_of_Plus_cong2:
-assumes "|A| =o |B|"
-shows "|C <+> A| =o |C <+> B|"
-using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2)
-
-
-corollary ordIso_Plus_cong2:
-assumes "r =o r'"
-shows "|A <+> (Field r)| =o |A <+> (Field r')|"
-using assms card_of_cong card_of_Plus_cong2 by blast
-
-
-lemma card_of_Plus_cong:
-assumes "|A| =o |B|" and "|C| =o |D|"
-shows "|A <+> C| =o |B <+> D|"
-using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono)
-
-
-corollary ordIso_Plus_cong:
-assumes "r =o r'" and "p =o p'"
-shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|"
-using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast
-
-
-lemma card_of_Un1:
-shows "|A| \<le>o |A \<union> B| "
-using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
-
-
-lemma card_of_diff:
-shows "|A - B| \<le>o |A|"
-using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce
-
-
-lemma card_of_Un_Plus_ordLeq:
-"|A \<union> B| \<le>o |A <+> B|"
-proof-
-   let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
-   have "inj_on ?f (A \<union> B) \<and> ?f ` (A \<union> B) \<le> A <+> B"
-   unfolding inj_on_def by auto
-   thus ?thesis using card_of_ordLeq by blast
-qed
-
-
-lemma card_of_Times1:
-assumes "A \<noteq> {}"
-shows "|B| \<le>o |B \<times> A|"
-proof(cases "B = {}", simp add: card_of_empty)
-  assume *: "B \<noteq> {}"
-  have "fst `(B \<times> A) = B" unfolding image_def using assms by auto
-  thus ?thesis using inj_on_iff_surj[of B "B \<times> A"]
-                     card_of_ordLeq[of B "B \<times> A"] * by blast
-qed
-
-
-corollary Card_order_Times1:
-"\<lbrakk>Card_order r; B \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |(Field r) \<times> B|"
-using card_of_Times1[of B] card_of_Field_ordIso
-      ordIso_ordLeq_trans ordIso_symmetric by blast
-
-
-lemma card_of_Times_commute: "|A \<times> B| =o |B \<times> A|"
-proof-
-  let ?f = "\<lambda>(a::'a,b::'b). (b,a)"
-  have "bij_betw ?f (A \<times> B) (B \<times> A)"
-  unfolding bij_betw_def inj_on_def by auto
-  thus ?thesis using card_of_ordIso by blast
-qed
-
-
-lemma card_of_Times2:
-assumes "A \<noteq> {}"   shows "|B| \<le>o |A \<times> B|"
-using assms card_of_Times1[of A B] card_of_Times_commute[of B A]
-      ordLeq_ordIso_trans by blast
-
-
-corollary Card_order_Times2:
-"\<lbrakk>Card_order r; A \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |A \<times> (Field r)|"
-using card_of_Times2[of A] card_of_Field_ordIso
-      ordIso_ordLeq_trans ordIso_symmetric by blast
-
-
-lemma card_of_Times3: "|A| \<le>o |A \<times> A|"
-using card_of_Times1[of A]
-by(cases "A = {}", simp add: card_of_empty, blast)
-
-
-lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \<times> (UNIV::bool set)|"
-proof-
-  let ?f = "\<lambda>c::'a + 'a. case c of Inl a \<Rightarrow> (a,True)
-                                  |Inr a \<Rightarrow> (a,False)"
-  have "bij_betw ?f (A <+> A) (A \<times> (UNIV::bool set))"
-  proof-
-    {fix  c1 and c2 assume "?f c1 = ?f c2"
-     hence "c1 = c2"
-     by(case_tac "c1", case_tac "c2", auto, case_tac "c2", auto)
-    }
-    moreover
-    {fix c assume "c \<in> A <+> A"
-     hence "?f c \<in> A \<times> (UNIV::bool set)"
-     by(case_tac c, auto)
-    }
-    moreover
-    {fix a bl assume *: "(a,bl) \<in> A \<times> (UNIV::bool set)"
-     have "(a,bl) \<in> ?f ` ( A <+> A)"
-     proof(cases bl)
-       assume bl hence "?f(Inl a) = (a,bl)" by auto
-       thus ?thesis using * by force
-     next
-       assume "\<not> bl" hence "?f(Inr a) = (a,bl)" by auto
-       thus ?thesis using * by force
-     qed
-    }
-    ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto
-  qed
-  thus ?thesis using card_of_ordIso by blast
-qed
-
-
-lemma card_of_Times_mono1:
-assumes "|A| \<le>o |B|"
-shows "|A \<times> C| \<le>o |B \<times> C|"
-proof-
-  obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
-  using assms card_of_ordLeq[of A] by fastforce
-  obtain g where g_def:
-  "g = (\<lambda>(a,c::'c). (f a,c))" by blast
-  have "inj_on g (A \<times> C) \<and> g ` (A \<times> C) \<le> (B \<times> C)"
-  using 1 unfolding inj_on_def using g_def by auto
-  thus ?thesis using card_of_ordLeq by metis
-qed
-
-
-corollary ordLeq_Times_mono1:
-assumes "r \<le>o r'"
-shows "|(Field r) \<times> C| \<le>o |(Field r') \<times> C|"
-using assms card_of_mono2 card_of_Times_mono1 by blast
-
-
-lemma card_of_Times_mono2:
-assumes "|A| \<le>o |B|"
-shows "|C \<times> A| \<le>o |C \<times> B|"
-using assms card_of_Times_mono1[of A B C]
-      card_of_Times_commute[of C A]  card_of_Times_commute[of B C]
-      ordIso_ordLeq_trans[of "|C \<times> A|"] ordLeq_ordIso_trans[of "|C \<times> A|"]
-by blast
-
-
-corollary ordLeq_Times_mono2:
-assumes "r \<le>o r'"
-shows "|A \<times> (Field r)| \<le>o |A \<times> (Field r')|"
-using assms card_of_mono2 card_of_Times_mono2 by blast
-
-
-lemma card_of_Times_cong1:
-assumes "|A| =o |B|"
-shows "|A \<times> C| =o |B \<times> C|"
-using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono1)
-
-
-lemma card_of_Times_cong2:
-assumes "|A| =o |B|"
-shows "|C \<times> A| =o |C \<times> B|"
-using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2)
-
-
-corollary ordIso_Times_cong2:
-assumes "r =o r'"
-shows "|A \<times> (Field r)| =o |A \<times> (Field r')|"
-using assms card_of_cong card_of_Times_cong2 by blast
-
-
-lemma card_of_Sigma_mono1:
-assumes "\<forall>i \<in> I. |A i| \<le>o |B i|"
-shows "|SIGMA i : I. A i| \<le>o |SIGMA i : I. B i|"
-proof-
-  have "\<forall>i. i \<in> I \<longrightarrow> (\<exists>f. inj_on f (A i) \<and> f ` (A i) \<le> B i)"
-  using assms by (auto simp add: card_of_ordLeq)
-  with choice[of "\<lambda> i f. i \<in> I \<longrightarrow> inj_on f (A i) \<and> f ` (A i) \<le> B i"]
-  obtain F where 1: "\<forall>i \<in> I. inj_on (F i) (A i) \<and> (F i) ` (A i) \<le> B i" by metis
-  obtain g where g_def: "g = (\<lambda>(i,a::'b). (i,F i a))" by blast
-  have "inj_on g (Sigma I A) \<and> g ` (Sigma I A) \<le> (Sigma I B)"
-  using 1 unfolding inj_on_def using g_def by force
-  thus ?thesis using card_of_ordLeq by metis
-qed
-
-
-corollary card_of_Sigma_Times:
-"\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> |SIGMA i : I. A i| \<le>o |I \<times> B|"
-using card_of_Sigma_mono1[of I A "\<lambda>i. B"] .
-
-
-lemma card_of_UNION_Sigma:
-"|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
-using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by metis
-
-
-lemma card_of_bool:
-assumes "a1 \<noteq> a2"
-shows "|UNIV::bool set| =o |{a1,a2}|"
-proof-
-  let ?f = "\<lambda> bl. case bl of True \<Rightarrow> a1 | False \<Rightarrow> a2"
-  have "bij_betw ?f UNIV {a1,a2}"
-  proof-
-    {fix bl1 and bl2 assume "?f  bl1 = ?f bl2"
-     hence "bl1 = bl2" using assms by (case_tac bl1, case_tac bl2, auto)
-    }
-    moreover
-    {fix bl have "?f bl \<in> {a1,a2}" by (case_tac bl, auto)
-    }
-    moreover
-    {fix a assume *: "a \<in> {a1,a2}"
-     have "a \<in> ?f ` UNIV"
-     proof(cases "a = a1")
-       assume "a = a1"
-       hence "?f True = a" by auto  thus ?thesis by blast
-     next
-       assume "a \<noteq> a1" hence "a = a2" using * by auto
-       hence "?f False = a" by auto  thus ?thesis by blast
-     qed
-    }
-    ultimately show ?thesis unfolding bij_betw_def inj_on_def
-    by (metis image_subsetI order_eq_iff subsetI)
-  qed
-  thus ?thesis using card_of_ordIso by blast
-qed
-
-
-lemma card_of_Plus_Times_aux:
-assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
-        LEQ: "|A| \<le>o |B|"
-shows "|A <+> B| \<le>o |A \<times> B|"
-proof-
-  have 1: "|UNIV::bool set| \<le>o |A|"
-  using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2]
-        ordIso_ordLeq_trans[of "|UNIV::bool set|"] by metis
-  (*  *)
-  have "|A <+> B| \<le>o |B <+> B|"
-  using LEQ card_of_Plus_mono1 by blast
-  moreover have "|B <+> B| =o |B \<times> (UNIV::bool set)|"
-  using card_of_Plus_Times_bool by blast
-  moreover have "|B \<times> (UNIV::bool set)| \<le>o |B \<times> A|"
-  using 1 by (simp add: card_of_Times_mono2)
-  moreover have " |B \<times> A| =o |A \<times> B|"
-  using card_of_Times_commute by blast
-  ultimately show "|A <+> B| \<le>o |A \<times> B|"
-  using ordLeq_ordIso_trans[of "|A <+> B|" "|B <+> B|" "|B \<times> (UNIV::bool set)|"]
-        ordLeq_transitive[of "|A <+> B|" "|B \<times> (UNIV::bool set)|" "|B \<times> A|"]
-        ordLeq_ordIso_trans[of "|A <+> B|" "|B \<times> A|" "|A \<times> B|"]
-  by blast
-qed
-
-
-lemma card_of_Plus_Times:
-assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
-        B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
-shows "|A <+> B| \<le>o |A \<times> B|"
-proof-
-  {assume "|A| \<le>o |B|"
-   hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux)
-  }
-  moreover
-  {assume "|B| \<le>o |A|"
-   hence "|B <+> A| \<le>o |B \<times> A|"
-   using assms by (auto simp add: card_of_Plus_Times_aux)
-   hence ?thesis
-   using card_of_Plus_commute card_of_Times_commute
-         ordIso_ordLeq_trans ordLeq_ordIso_trans by metis
-  }
-  ultimately show ?thesis
-  using card_of_Well_order[of A] card_of_Well_order[of B]
-        ordLeq_total[of "|A|"] by metis
-qed
-
-
-lemma card_of_ordLeq_finite:
-assumes "|A| \<le>o |B|" and "finite B"
-shows "finite A"
-using assms unfolding ordLeq_def
-using embed_inj_on[of "|A|" "|B|"]  embed_Field[of "|A|" "|B|"]
-      Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce
-
-
-lemma card_of_ordLeq_infinite:
-assumes "|A| \<le>o |B|" and "infinite A"
-shows "infinite B"
-using assms card_of_ordLeq_finite by auto
-
-
-lemma card_of_ordIso_finite:
-assumes "|A| =o |B|"
-shows "finite A = finite B"
-using assms unfolding ordIso_def iso_def[abs_def]
-by (auto simp: bij_betw_finite Field_card_of)
-
-
-lemma card_of_ordIso_finite_Field:
-assumes "Card_order r" and "r =o |A|"
-shows "finite(Field r) = finite A"
-using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast
-
-
-subsection {* Cardinals versus set operations involving infinite sets *}
-
-
-text{* Here we show that, for infinite sets, most set-theoretic constructions
-do not increase the cardinality.  The cornerstone for this is
-theorem @{text "Card_order_Times_same_infinite"}, which states that self-product
-does not increase cardinality -- the proof of this fact adapts a standard
-set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11
-at page 47 in \cite{card-book}. Then everything else follows fairly easily.  *}
-
-
-lemma infinite_iff_card_of_nat:
-"infinite A = ( |UNIV::nat set| \<le>o |A| )"
-by (auto simp add: infinite_iff_countable_subset card_of_ordLeq)
-
-
-lemma finite_iff_cardOf_nat:
-"finite A = ( |A| <o |UNIV :: nat set| )"
-using infinite_iff_card_of_nat[of A]
-not_ordLeq_iff_ordLess[of "|A|" "|UNIV :: nat set|"]
-by (fastforce simp: card_of_Well_order)
-
-lemma finite_ordLess_infinite2:
-assumes "finite A" and "infinite B"
-shows "|A| <o |B|"
-using assms
-finite_ordLess_infinite[of "|A|" "|B|"]
-card_of_Well_order[of A] card_of_Well_order[of B]
-Field_card_of[of A] Field_card_of[of B] by auto
-
-
-text{* The next two results correspond to the ZF fact that all infinite cardinals are
-limit ordinals: *}
-
-lemma Card_order_infinite_not_under:
-assumes CARD: "Card_order r" and INF: "infinite (Field r)"
-shows "\<not> (\<exists>a. Field r = rel.under r a)"
-proof(auto)
-  have 0: "Well_order r \<and> wo_rel r \<and> Refl r"
-  using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto
-  fix a assume *: "Field r = rel.under r a"
-  show False
-  proof(cases "a \<in> Field r")
-    assume Case1: "a \<notin> Field r"
-    hence "rel.under r a = {}" unfolding Field_def rel.under_def by auto
-    thus False using INF *  by auto
-  next
-    let ?r' = "Restr r (rel.underS r a)"
-    assume Case2: "a \<in> Field r"
-    hence 1: "rel.under r a = rel.underS r a \<union> {a} \<and> a \<notin> rel.underS r a"
-    using 0 rel.Refl_under_underS rel.underS_notIn by fastforce
-    have 2: "wo_rel.ofilter r (rel.underS r a) \<and> rel.underS r a < Field r"
-    using 0 wo_rel.underS_ofilter * 1 Case2 by auto
-    hence "?r' <o r" using 0 using ofilter_ordLess by blast
-    moreover
-    have "Field ?r' = rel.underS r a \<and> Well_order ?r'"
-    using  2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast
-    ultimately have "|rel.underS r a| <o r" using ordLess_Field[of ?r'] by auto
-    moreover have "|rel.under r a| =o r" using * CARD card_of_Field_ordIso[of r] by auto
-    ultimately have "|rel.underS r a| <o |rel.under r a|"
-    using ordIso_symmetric ordLess_ordIso_trans by blast
-    moreover
-    {have "\<exists>f. bij_betw f (rel.under r a) (rel.underS r a)"
-     using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto
-     hence "|rel.under r a| =o |rel.underS r a|" using card_of_ordIso by blast
-    }
-    ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast
-  qed
-qed
-
-
-lemma infinite_Card_order_limit:
-assumes r: "Card_order r" and "infinite (Field r)"
-and a: "a : Field r"
-shows "EX b : Field r. a \<noteq> b \<and> (a,b) : r"
-proof-
-  have "Field r \<noteq> rel.under r a"
-  using assms Card_order_infinite_not_under by blast
-  moreover have "rel.under r a \<le> Field r"
-  using rel.under_Field .
-  ultimately have "rel.under r a < Field r" by blast
-  then obtain b where 1: "b : Field r \<and> ~ (b,a) : r"
-  unfolding rel.under_def by blast
-  moreover have ba: "b \<noteq> a"
-  using 1 r unfolding card_order_on_def well_order_on_def
-  linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto
-  ultimately have "(a,b) : r"
-  using a r unfolding card_order_on_def well_order_on_def linear_order_on_def
-  total_on_def by blast
-  thus ?thesis using 1 ba by auto
-qed
-
-
-theorem Card_order_Times_same_infinite:
-assumes CO: "Card_order r" and INF: "infinite(Field r)"
-shows "|Field r \<times> Field r| \<le>o r"
-proof-
-  obtain phi where phi_def:
-  "phi = (\<lambda>r::'a rel. Card_order r \<and> infinite(Field r) \<and>
-                      \<not> |Field r \<times> Field r| \<le>o r )" by blast
-  have temp1: "\<forall>r. phi r \<longrightarrow> Well_order r"
-  unfolding phi_def card_order_on_def by auto
-  have Ft: "\<not>(\<exists>r. phi r)"
-  proof
-    assume "\<exists>r. phi r"
-    hence "{r. phi r} \<noteq> {} \<and> {r. phi r} \<le> {r. Well_order r}"
-    using temp1 by auto
-    then obtain r where 1: "phi r" and 2: "\<forall>r'. phi r' \<longrightarrow> r \<le>o r'" and
-                   3: "Card_order r \<and> Well_order r"
-    using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast
-    let ?A = "Field r"  let ?r' = "bsqr r"
-    have 4: "Well_order ?r' \<and> Field ?r' = ?A \<times> ?A \<and> |?A| =o r"
-    using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast
-    have 5: "Card_order |?A \<times> ?A| \<and> Well_order |?A \<times> ?A|"
-    using card_of_Card_order card_of_Well_order by blast
-    (*  *)
-    have "r <o |?A \<times> ?A|"
-    using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast
-    moreover have "|?A \<times> ?A| \<le>o ?r'"
-    using card_of_least[of "?A \<times> ?A"] 4 by auto
-    ultimately have "r <o ?r'" using ordLess_ordLeq_trans by auto
-    then obtain f where 6: "embed r ?r' f" and 7: "\<not> bij_betw f ?A (?A \<times> ?A)"
-    unfolding ordLess_def embedS_def[abs_def]
-    by (auto simp add: Field_bsqr)
-    let ?B = "f ` ?A"
-    have "|?A| =o |?B|"
-    using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast
-    hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast
-    (*  *)
-    have "wo_rel.ofilter ?r' ?B"
-    using 6 embed_Field_ofilter 3 4 by blast
-    hence "wo_rel.ofilter ?r' ?B \<and> ?B \<noteq> ?A \<times> ?A \<and> ?B \<noteq> Field ?r'"
-    using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto
-    hence temp2: "wo_rel.ofilter ?r' ?B \<and> ?B < ?A \<times> ?A"
-    using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast
-    have "\<not> (\<exists>a. Field r = rel.under r a)"
-    using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto
-    then obtain A1 where temp3: "wo_rel.ofilter r A1 \<and> A1 < ?A" and 9: "?B \<le> A1 \<times> A1"
-    using temp2 3 bsqr_ofilter[of r ?B] by blast
-    hence "|?B| \<le>o |A1 \<times> A1|" using card_of_mono1 by blast
-    hence 10: "r \<le>o |A1 \<times> A1|" using 8 ordIso_ordLeq_trans by blast
-    let ?r1 = "Restr r A1"
-    have "?r1 <o r" using temp3 ofilter_ordLess 3 by blast
-    moreover
-    {have "well_order_on A1 ?r1" using 3 temp3 well_order_on_Restr by blast
-     hence "|A1| \<le>o ?r1" using 3 Well_order_Restr card_of_least by blast
-    }
-    ultimately have 11: "|A1| <o r" using ordLeq_ordLess_trans by blast
-    (*  *)
-    have "infinite (Field r)" using 1 unfolding phi_def by simp
-    hence "infinite ?B" using 8 3 card_of_ordIso_finite_Field[of r ?B] by blast
-    hence "infinite A1" using 9 infinite_super finite_cartesian_product by blast
-    moreover have temp4: "Field |A1| = A1 \<and> Well_order |A1| \<and> Card_order |A1|"
-    using card_of_Card_order[of A1] card_of_Well_order[of A1]
-    by (simp add: Field_card_of)
-    moreover have "\<not> r \<le>o | A1 |"
-    using temp4 11 3 using not_ordLeq_iff_ordLess by blast
-    ultimately have "infinite(Field |A1| ) \<and> Card_order |A1| \<and> \<not> r \<le>o | A1 |"
-    by (simp add: card_of_card_order_on)
-    hence "|Field |A1| \<times> Field |A1| | \<le>o |A1|"
-    using 2 unfolding phi_def by blast
-    hence "|A1 \<times> A1 | \<le>o |A1|" using temp4 by auto
-    hence "r \<le>o |A1|" using 10 ordLeq_transitive by blast
-    thus False using 11 not_ordLess_ordLeq by auto
-  qed
-  thus ?thesis using assms unfolding phi_def by blast
-qed
-
-
-corollary card_of_Times_same_infinite:
-assumes "infinite A"
-shows "|A \<times> A| =o |A|"
-proof-
-  let ?r = "|A|"
-  have "Field ?r = A \<and> Card_order ?r"
-  using Field_card_of card_of_Card_order[of A] by fastforce
-  hence "|A \<times> A| \<le>o |A|"
-  using Card_order_Times_same_infinite[of ?r] assms by auto
-  thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast
-qed
-
-
-lemma card_of_Times_infinite:
-assumes INF: "infinite A" and NE: "B \<noteq> {}" and LEQ: "|B| \<le>o |A|"
-shows "|A \<times> B| =o |A| \<and> |B \<times> A| =o |A|"
-proof-
-  have "|A| \<le>o |A \<times> B| \<and> |A| \<le>o |B \<times> A|"
-  using assms by (simp add: card_of_Times1 card_of_Times2)
-  moreover
-  {have "|A \<times> B| \<le>o |A \<times> A| \<and> |B \<times> A| \<le>o |A \<times> A|"
-   using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast
-   moreover have "|A \<times> A| =o |A|" using INF card_of_Times_same_infinite by blast
-   ultimately have "|A \<times> B| \<le>o |A| \<and> |B \<times> A| \<le>o |A|"
-   using ordLeq_ordIso_trans[of "|A \<times> B|"] ordLeq_ordIso_trans[of "|B \<times> A|"] by auto
-  }
-  ultimately show ?thesis by (simp add: ordIso_iff_ordLeq)
-qed
-
-
-corollary card_of_Times_infinite_simps:
-"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A \<times> B| =o |A|"
-"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |A \<times> B|"
-"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |B \<times> A| =o |A|"
-"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |B \<times> A|"
-by (auto simp add: card_of_Times_infinite ordIso_symmetric)
-
-
-corollary Card_order_Times_infinite:
-assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
-        NE: "Field p \<noteq> {}" and LEQ: "p \<le>o r"
-shows "| (Field r) \<times> (Field p) | =o r \<and> | (Field p) \<times> (Field r) | =o r"
-proof-
-  have "|Field r \<times> Field p| =o |Field r| \<and> |Field p \<times> Field r| =o |Field r|"
-  using assms by (simp add: card_of_Times_infinite card_of_mono2)
-  thus ?thesis
-  using assms card_of_Field_ordIso[of r]
-        ordIso_transitive[of "|Field r \<times> Field p|"]
-        ordIso_transitive[of _ "|Field r|"] by blast
-qed
-
-
-lemma card_of_Sigma_ordLeq_infinite:
-assumes INF: "infinite B" and
-        LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
-shows "|SIGMA i : I. A i| \<le>o |B|"
-proof(cases "I = {}", simp add: card_of_empty)
-  assume *: "I \<noteq> {}"
-  have "|SIGMA i : I. A i| \<le>o |I \<times> B|"
-  using LEQ card_of_Sigma_Times by blast
-  moreover have "|I \<times> B| =o |B|"
-  using INF * LEQ_I by (auto simp add: card_of_Times_infinite)
-  ultimately show ?thesis using ordLeq_ordIso_trans by blast
-qed
-
-
-lemma card_of_Sigma_ordLeq_infinite_Field:
-assumes INF: "infinite (Field r)" and r: "Card_order r" and
-        LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
-shows "|SIGMA i : I. A i| \<le>o r"
-proof-
-  let ?B  = "Field r"
-  have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
-  ordIso_symmetric by blast
-  hence "|I| \<le>o |?B|"  "\<forall>i \<in> I. |A i| \<le>o |?B|"
-  using LEQ_I LEQ ordLeq_ordIso_trans by blast+
-  hence  "|SIGMA i : I. A i| \<le>o |?B|" using INF LEQ
-  card_of_Sigma_ordLeq_infinite by blast
-  thus ?thesis using 1 ordLeq_ordIso_trans by blast
-qed
-
-
-lemma card_of_Times_ordLeq_infinite_Field:
-"\<lbrakk>infinite (Field r); |A| \<le>o r; |B| \<le>o r; Card_order r\<rbrakk>
- \<Longrightarrow> |A <*> B| \<le>o r"
-by(simp add: card_of_Sigma_ordLeq_infinite_Field)
-
-
-lemma card_of_UNION_ordLeq_infinite:
-assumes INF: "infinite B" and
-        LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
-shows "|\<Union> i \<in> I. A i| \<le>o |B|"
-proof(cases "I = {}", simp add: card_of_empty)
-  assume *: "I \<noteq> {}"
-  have "|\<Union> i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
-  using card_of_UNION_Sigma by blast
-  moreover have "|SIGMA i : I. A i| \<le>o |B|"
-  using assms card_of_Sigma_ordLeq_infinite by blast
-  ultimately show ?thesis using ordLeq_transitive by blast
-qed
-
-
-corollary card_of_UNION_ordLeq_infinite_Field:
-assumes INF: "infinite (Field r)" and r: "Card_order r" and
-        LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
-shows "|\<Union> i \<in> I. A i| \<le>o r"
-proof-
-  let ?B  = "Field r"
-  have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
-  ordIso_symmetric by blast
-  hence "|I| \<le>o |?B|"  "\<forall>i \<in> I. |A i| \<le>o |?B|"
-  using LEQ_I LEQ ordLeq_ordIso_trans by blast+
-  hence  "|\<Union> i \<in> I. A i| \<le>o |?B|" using INF LEQ
-  card_of_UNION_ordLeq_infinite by blast
-  thus ?thesis using 1 ordLeq_ordIso_trans by blast
-qed
-
-
-lemma card_of_Plus_infinite1:
-assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
-shows "|A <+> B| =o |A|"
-proof(cases "B = {}", simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric)
-  let ?Inl = "Inl::'a \<Rightarrow> 'a + 'b"  let ?Inr = "Inr::'b \<Rightarrow> 'a + 'b"
-  assume *: "B \<noteq> {}"
-  then obtain b1 where 1: "b1 \<in> B" by blast
-  show ?thesis
-  proof(cases "B = {b1}")
-    assume Case1: "B = {b1}"
-    have 2: "bij_betw ?Inl A ((?Inl ` A))"
-    unfolding bij_betw_def inj_on_def by auto
-    hence 3: "infinite (?Inl ` A)"
-    using INF bij_betw_finite[of ?Inl A] by blast
-    let ?A' = "?Inl ` A \<union> {?Inr b1}"
-    obtain g where "bij_betw g (?Inl ` A) ?A'"
-    using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto
-    moreover have "?A' = A <+> B" using Case1 by blast
-    ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp
-    hence "bij_betw (g o ?Inl) A (A <+> B)"
-    using 2 by (auto simp add: bij_betw_trans)
-    thus ?thesis using card_of_ordIso ordIso_symmetric by blast
-  next
-    assume Case2: "B \<noteq> {b1}"
-    with * 1 obtain b2 where 3: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B" by fastforce
-    obtain f where "inj_on f B \<and> f ` B \<le> A"
-    using LEQ card_of_ordLeq[of B] by fastforce
-    with 3 have "f b1 \<noteq> f b2 \<and> {f b1, f b2} \<le> A"
-    unfolding inj_on_def by auto
-    with 3 have "|A <+> B| \<le>o |A \<times> B|"
-    by (auto simp add: card_of_Plus_Times)
-    moreover have "|A \<times> B| =o |A|"
-    using assms * by (simp add: card_of_Times_infinite_simps)
-    ultimately have "|A <+> B| \<le>o |A|" using ordLeq_ordIso_trans by metis
-    thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast
-  qed
-qed
-
-
-lemma card_of_Plus_infinite2:
-assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
-shows "|B <+> A| =o |A|"
-using assms card_of_Plus_commute card_of_Plus_infinite1
-ordIso_equivalence by blast
-
-
-lemma card_of_Plus_infinite:
-assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
-shows "|A <+> B| =o |A| \<and> |B <+> A| =o |A|"
-using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2)
-
-
-corollary Card_order_Plus_infinite:
-assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
-        LEQ: "p \<le>o r"
-shows "| (Field r) <+> (Field p) | =o r \<and> | (Field p) <+> (Field r) | =o r"
-proof-
-  have "| Field r <+> Field p | =o | Field r | \<and>
-        | Field p <+> Field r | =o | Field r |"
-  using assms by (simp add: card_of_Plus_infinite card_of_mono2)
-  thus ?thesis
-  using assms card_of_Field_ordIso[of r]
-        ordIso_transitive[of "|Field r <+> Field p|"]
-        ordIso_transitive[of _ "|Field r|"] by blast
-qed
-
-
-lemma card_of_Un_infinite:
-assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
-shows "|A \<union> B| =o |A| \<and> |B \<union> A| =o |A|"
-proof-
-  have "|A \<union> B| \<le>o |A <+> B|" by (rule card_of_Un_Plus_ordLeq)
-  moreover have "|A <+> B| =o |A|"
-  using assms by (metis card_of_Plus_infinite)
-  ultimately have "|A \<union> B| \<le>o |A|" using ordLeq_ordIso_trans by blast
-  hence "|A \<union> B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast
-  thus ?thesis using Un_commute[of B A] by auto
-qed
-
-
-lemma card_of_Un_diff_infinite:
-assumes INF: "infinite A" and LESS: "|B| <o |A|"
-shows "|A - B| =o |A|"
-proof-
-  obtain C where C_def: "C = A - B" by blast
-  have "|A \<union> B| =o |A|"
-  using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast
-  moreover have "C \<union> B = A \<union> B" unfolding C_def by auto
-  ultimately have 1: "|C \<union> B| =o |A|" by auto
-  (*  *)
-  {assume *: "|C| \<le>o |B|"
-   moreover
-   {assume **: "finite B"
-    hence "finite C"
-    using card_of_ordLeq_finite * by blast
-    hence False using ** INF card_of_ordIso_finite 1 by blast
-   }
-   hence "infinite B" by auto
-   ultimately have False
-   using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis
-  }
-  hence 2: "|B| \<le>o |C|" using card_of_Well_order ordLeq_total by blast
-  {assume *: "finite C"
-    hence "finite B" using card_of_ordLeq_finite 2 by blast
-    hence False using * INF card_of_ordIso_finite 1 by blast
-  }
-  hence "infinite C" by auto
-  hence "|C| =o |A|"
-  using  card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis
-  thus ?thesis unfolding C_def .
-qed
-
-
-lemma card_of_Plus_ordLess_infinite:
-assumes INF: "infinite C" and
-        LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
-shows "|A <+> B| <o |C|"
-proof(cases "A = {} \<or> B = {}")
-  assume Case1: "A = {} \<or> B = {}"
-  hence "|A| =o |A <+> B| \<or> |B| =o |A <+> B|"
-  using card_of_Plus_empty1 card_of_Plus_empty2 by blast
-  hence "|A <+> B| =o |A| \<or> |A <+> B| =o |B|"
-  using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast
-  thus ?thesis using LESS1 LESS2
-       ordIso_ordLess_trans[of "|A <+> B|" "|A|"]
-       ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast
-next
-  assume Case2: "\<not>(A = {} \<or> B = {})"
-  {assume *: "|C| \<le>o |A <+> B|"
-   hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast
-   hence 1: "infinite A \<or> infinite B" using finite_Plus by blast
-   {assume Case21: "|A| \<le>o |B|"
-    hence "infinite B" using 1 card_of_ordLeq_finite by blast
-    hence "|A <+> B| =o |B|" using Case2 Case21
-    by (auto simp add: card_of_Plus_infinite)
-    hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
-   }
-   moreover
-   {assume Case22: "|B| \<le>o |A|"
-    hence "infinite A" using 1 card_of_ordLeq_finite by blast
-    hence "|A <+> B| =o |A|" using Case2 Case22
-    by (auto simp add: card_of_Plus_infinite)
-    hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
-   }
-   ultimately have False using ordLeq_total card_of_Well_order[of A]
-   card_of_Well_order[of B] by blast
-  }
-  thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"]
-  card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto
-qed
-
-
-lemma card_of_Plus_ordLess_infinite_Field:
-assumes INF: "infinite (Field r)" and r: "Card_order r" and
-        LESS1: "|A| <o r" and LESS2: "|B| <o r"
-shows "|A <+> B| <o r"
-proof-
-  let ?C  = "Field r"
-  have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
-  ordIso_symmetric by blast
-  hence "|A| <o |?C|"  "|B| <o |?C|"
-  using LESS1 LESS2 ordLess_ordIso_trans by blast+
-  hence  "|A <+> B| <o |?C|" using INF
-  card_of_Plus_ordLess_infinite by blast
-  thus ?thesis using 1 ordLess_ordIso_trans by blast
-qed
-
-
-lemma infinite_card_of_insert:
-assumes "infinite A"
-shows "|insert a A| =o |A|"
-proof-
-  have iA: "insert a A = A \<union> {a}" by simp
-  show ?thesis
-  using infinite_imp_bij_betw2[OF assms] unfolding iA
-  by (metis bij_betw_inv card_of_ordIso)
-qed
-
-
-subsection {* Cardinals versus lists  *}
-
-
-text{* The next is an auxiliary operator, which shall be used for inductive
-proofs of facts concerning the cardinality of @{text "List"} : *}
-
-definition nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set"
-where "nlists A n \<equiv> {l. set l \<le> A \<and> length l = n}"
-
-
-lemma lists_def2: "lists A = {l. set l \<le> A}"
-using in_listsI by blast
-
-
-lemma lists_UNION_nlists: "lists A = (\<Union> n. nlists A n)"
-unfolding lists_def2 nlists_def by blast
-
-
-lemma card_of_lists: "|A| \<le>o |lists A|"
-proof-
-  let ?h = "\<lambda> a. [a]"
-  have "inj_on ?h A \<and> ?h ` A \<le> lists A"
-  unfolding inj_on_def lists_def2 by auto
-  thus ?thesis by (metis card_of_ordLeq)
-qed
-
-
-lemma nlists_0: "nlists A 0 = {[]}"
-unfolding nlists_def by auto
-
-
-lemma nlists_not_empty:
-assumes "A \<noteq> {}"
-shows "nlists A n \<noteq> {}"
-proof(induct n, simp add: nlists_0)
-  fix n assume "nlists A n \<noteq> {}"
-  then obtain a and l where "a \<in> A \<and> l \<in> nlists A n" using assms by auto
-  hence "a # l \<in> nlists A (Suc n)" unfolding nlists_def by auto
-  thus "nlists A (Suc n) \<noteq> {}" by auto
-qed
-
-
-lemma Nil_in_lists: "[] \<in> lists A"
-unfolding lists_def2 by auto
-
-
-lemma lists_not_empty: "lists A \<noteq> {}"
-using Nil_in_lists by blast
-
-
-lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
-proof-
-  let ?B = "A \<times> (nlists A n)"   let ?h = "\<lambda>(a,l). a # l"
-  have "inj_on ?h ?B \<and> ?h ` ?B \<le> nlists A (Suc n)"
-  unfolding inj_on_def nlists_def by auto
-  moreover have "nlists A (Suc n) \<le> ?h ` ?B"
-  proof(auto)
-    fix l assume "l \<in> nlists A (Suc n)"
-    hence 1: "length l = Suc n \<and> set l \<le> A" unfolding nlists_def by auto
-    then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv)
-    hence "a \<in> A \<and> set l' \<le> A \<and> length l' = n" using 1 by auto
-    thus "l \<in> ?h ` ?B"  using 2 unfolding nlists_def by auto
-  qed
-  ultimately have "bij_betw ?h ?B (nlists A (Suc n))"
-  unfolding bij_betw_def by auto
-  thus ?thesis using card_of_ordIso ordIso_symmetric by blast
-qed
-
-
-lemma card_of_nlists_infinite:
-assumes "infinite A"
-shows "|nlists A n| \<le>o |A|"
-proof(induct n)
-  have "A \<noteq> {}" using assms by auto
-  thus "|nlists A 0| \<le>o |A|" by (simp add: nlists_0 card_of_singl_ordLeq)
-next
-  fix n assume IH: "|nlists A n| \<le>o |A|"
-  have "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
-  using card_of_nlists_Succ by blast
-  moreover
-  {have "nlists A n \<noteq> {}" using assms nlists_not_empty[of A] by blast
-   hence "|A \<times> (nlists A n)| =o |A|"
-   using assms IH by (auto simp add: card_of_Times_infinite)
-  }
-  ultimately show "|nlists A (Suc n)| \<le>o |A|"
-  using ordIso_transitive ordIso_iff_ordLeq by blast
-qed
-
-
-lemma card_of_lists_infinite:
-assumes "infinite A"
-shows "|lists A| =o |A|"
-proof-
-  have "|lists A| \<le>o |A|"
-  using assms
-  by (auto simp add: lists_UNION_nlists card_of_UNION_ordLeq_infinite
-                     infinite_iff_card_of_nat card_of_nlists_infinite)
-  thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast
-qed
-
-
-lemma Card_order_lists_infinite:
-assumes "Card_order r" and "infinite(Field r)"
-shows "|lists(Field r)| =o r"
-using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast
-
-
-
-subsection {* The cardinal $\omega$ and the finite cardinals  *}
-
-
-text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict
-order relation on
-@{text "nat"}, that we abbreviate by @{text "natLeq"}.  The finite cardinals
-shall be the restrictions of these relations to the numbers smaller than
-fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}.  *}
-
-abbreviation "(natLeq::(nat * nat) set) \<equiv> {(x,y). x \<le> y}"
-abbreviation "(natLess::(nat * nat) set) \<equiv> {(x,y). x < y}"
-
-abbreviation natLeq_on :: "nat \<Rightarrow> (nat * nat) set"
-where "natLeq_on n \<equiv> {(x,y). x < n \<and> y < n \<and> x \<le> y}"
-
-lemma infinite_cartesian_product:
-assumes "infinite A" "infinite B"
-shows "infinite (A \<times> B)"
-proof
-  assume "finite (A \<times> B)"
-  from assms(1) have "A \<noteq> {}" by auto
-  with `finite (A \<times> B)` have "finite B" using finite_cartesian_productD2 by auto
-  with assms(2) show False by simp
-qed
-
-
-
-subsubsection {* First as well-orders *}
-
-
-lemma Field_natLeq: "Field natLeq = (UNIV::nat set)"
-by(unfold Field_def, auto)
-
-
-lemma natLeq_Refl: "Refl natLeq"
-unfolding refl_on_def Field_def by auto
-
-
-lemma natLeq_trans: "trans natLeq"
-unfolding trans_def by auto
-
-
-lemma natLeq_Preorder: "Preorder natLeq"
-unfolding preorder_on_def
-by (auto simp add: natLeq_Refl natLeq_trans)
-
-
-lemma natLeq_antisym: "antisym natLeq"
-unfolding antisym_def by auto
-
-
-lemma natLeq_Partial_order: "Partial_order natLeq"
-unfolding partial_order_on_def
-by (auto simp add: natLeq_Preorder natLeq_antisym)
-
-
-lemma natLeq_Total: "Total natLeq"
-unfolding total_on_def by auto
-
-
-lemma natLeq_Linear_order: "Linear_order natLeq"
-unfolding linear_order_on_def
-by (auto simp add: natLeq_Partial_order natLeq_Total)
-
-
-lemma natLeq_natLess_Id: "natLess = natLeq - Id"
-by auto
-
-
-lemma natLeq_Well_order: "Well_order natLeq"
-unfolding well_order_on_def
-using natLeq_Linear_order wf_less natLeq_natLess_Id by auto
-
-
-corollary natLeq_well_order_on: "well_order_on UNIV natLeq"
-using natLeq_Well_order Field_natLeq by auto
-
-
-lemma natLeq_wo_rel: "wo_rel natLeq"
-unfolding wo_rel_def using natLeq_Well_order .
-
-
-lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV"
-using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto
-
-
-lemma closed_nat_set_iff:
-assumes "\<forall>(m::nat) n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
-shows "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
-proof-
-  {assume "A \<noteq> UNIV" hence "\<exists>n. n \<notin> A" by blast
-   moreover obtain n where n_def: "n = (LEAST n. n \<notin> A)" by blast
-   ultimately have 1: "n \<notin> A \<and> (\<forall>m. m < n \<longrightarrow> m \<in> A)"
-   using LeastI_ex[of "\<lambda> n. n \<notin> A"] n_def Least_le[of "\<lambda> n. n \<notin> A"] by fastforce
-   have "A = {0 ..< n}"
-   proof(auto simp add: 1)
-     fix m assume *: "m \<in> A"
-     {assume "n \<le> m" with assms * have "n \<in> A" by blast
-      hence False using 1 by auto
-     }
-     thus "m < n" by fastforce
-   qed
-   hence "\<exists>n. A = {0 ..< n}" by blast
-  }
-  thus ?thesis by blast
-qed
-
-
-lemma Field_natLeq_on: "Field (natLeq_on n) = {0 ..< n}"
-unfolding Field_def by auto
-
-
-lemma natLeq_underS_less: "rel.underS natLeq n = {0 ..< n}"
-unfolding rel.underS_def by auto
-
-
-lemma Restr_natLeq: "Restr natLeq {0 ..< n} = natLeq_on n"
-by auto
-
-
-lemma Restr_natLeq2:
-"Restr natLeq (rel.underS natLeq n) = natLeq_on n"
-by (auto simp add: Restr_natLeq natLeq_underS_less)
-
-
-lemma natLeq_on_Well_order: "Well_order(natLeq_on n)"
-using Restr_natLeq[of n] natLeq_Well_order
-      Well_order_Restr[of natLeq "{0..<n}"] by auto
-
-
-corollary natLeq_on_well_order_on: "well_order_on {0 ..< n} (natLeq_on n)"
-using natLeq_on_Well_order Field_natLeq_on by auto
-
-
-lemma natLeq_on_wo_rel: "wo_rel(natLeq_on n)"
-unfolding wo_rel_def using natLeq_on_Well_order .
-
-
-lemma natLeq_on_ofilter_less_eq:
-"n \<le> m \<Longrightarrow> wo_rel.ofilter (natLeq_on m) {0 ..< n}"
-by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def,
-    simp add: Field_natLeq_on, unfold rel.under_def, auto)
-
-
-lemma natLeq_on_ofilter_iff:
-"wo_rel.ofilter (natLeq_on m) A = (\<exists>n \<le> m. A = {0 ..< n})"
-proof(rule iffI)
-  assume *: "wo_rel.ofilter (natLeq_on m) A"
-  hence 1: "A \<le> {0..<m}"
-  by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def Field_natLeq_on)
-  hence "\<forall>n1 n2. n2 \<in> A \<and> n1 \<le> n2 \<longrightarrow> n1 \<in> A"
-  using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def)
-  hence "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
-  thus "\<exists>n \<le> m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast
-next
-  assume "(\<exists>n\<le>m. A = {0 ..< n})"
-  thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq)
-qed
-
-
-
-subsubsection {* Then as cardinals *}
-
-
-lemma natLeq_Card_order: "Card_order natLeq"
-proof(auto simp add: natLeq_Well_order
-      Card_order_iff_Restr_underS Restr_natLeq2, simp add:  Field_natLeq)
-  fix n have "finite(Field (natLeq_on n))"
-  unfolding Field_natLeq_on by auto
-  moreover have "infinite(UNIV::nat set)" by auto
-  ultimately show "natLeq_on n <o |UNIV::nat set|"
-  using finite_ordLess_infinite[of "natLeq_on n" "|UNIV::nat set|"]
-        Field_card_of[of "UNIV::nat set"]
-        card_of_Well_order[of "UNIV::nat set"] natLeq_on_Well_order[of n] by auto
-qed
-
-
-corollary card_of_Field_natLeq:
-"|Field natLeq| =o natLeq"
-using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq]
-      ordIso_symmetric[of natLeq] by blast
-
-
-corollary card_of_nat:
-"|UNIV::nat set| =o natLeq"
-using Field_natLeq card_of_Field_natLeq by auto
-
-
-corollary infinite_iff_natLeq_ordLeq:
-"infinite A = ( natLeq \<le>o |A| )"
-using infinite_iff_card_of_nat[of A] card_of_nat
-      ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
-
-
-corollary finite_iff_ordLess_natLeq:
-"finite A = ( |A| <o natLeq)"
-using infinite_iff_natLeq_ordLeq not_ordLeq_iff_ordLess
-      card_of_Well_order natLeq_Well_order by blast
-
-
-lemma ordIso_natLeq_on_imp_finite:
-"|A| =o natLeq_on n \<Longrightarrow> finite A"
-unfolding ordIso_def iso_def[abs_def]
-by (auto simp: Field_natLeq_on bij_betw_finite Field_card_of)
-
-
-lemma natLeq_on_Card_order: "Card_order (natLeq_on n)"
-proof(unfold card_order_on_def,
-      auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on)
-  fix r assume "well_order_on {0..<n} r"
-  thus "natLeq_on n \<le>o r"
-  using finite_atLeastLessThan natLeq_on_well_order_on
-        finite_well_order_on_ordIso ordIso_iff_ordLeq by blast
-qed
-
-
-corollary card_of_Field_natLeq_on:
-"|Field (natLeq_on n)| =o natLeq_on n"
-using Field_natLeq_on natLeq_on_Card_order
-      Card_order_iff_ordIso_card_of[of "natLeq_on n"]
-      ordIso_symmetric[of "natLeq_on n"] by blast
-
-
-corollary card_of_less:
-"|{0 ..< n}| =o natLeq_on n"
-using Field_natLeq_on card_of_Field_natLeq_on by auto
-
-
-lemma natLeq_on_ordLeq_less_eq:
-"((natLeq_on m) \<le>o (natLeq_on n)) = (m \<le> n)"
-proof
-  assume "natLeq_on m \<le>o natLeq_on n"
-  then obtain f where "inj_on f {0..<m} \<and> f ` {0..<m} \<le> {0..<n}"
-  unfolding ordLeq_def using
-    embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on]
-     embed_Field[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] by blast
-  thus "m \<le> n" using atLeastLessThan_less_eq2 by blast
-next
-  assume "m \<le> n"
-  hence "inj_on id {0..<m} \<and> id ` {0..<m} \<le> {0..<n}" unfolding inj_on_def by auto
-  hence "|{0..<m}| \<le>o |{0..<n}|" using card_of_ordLeq by blast
-  thus "natLeq_on m \<le>o natLeq_on n"
-  using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
-qed
-
-
-lemma natLeq_on_ordLeq_less:
-"((natLeq_on m) <o (natLeq_on n)) = (m < n)"
-using not_ordLeq_iff_ordLess[of "natLeq_on m" "natLeq_on n"]
-natLeq_on_Well_order natLeq_on_ordLeq_less_eq by auto
-
-
-
-subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
-
-
-lemma finite_card_of_iff_card2:
-assumes FIN: "finite A" and FIN': "finite B"
-shows "( |A| \<le>o |B| ) = (card A \<le> card B)"
-using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast
-
-
-lemma finite_imp_card_of_natLeq_on:
-assumes "finite A"
-shows "|A| =o natLeq_on (card A)"
-proof-
-  obtain h where "bij_betw h A {0 ..< card A}"
-  using assms ex_bij_betw_finite_nat by blast
-  thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast
-qed
-
-
-lemma finite_iff_card_of_natLeq_on:
-"finite A = (\<exists>n. |A| =o natLeq_on n)"
-using finite_imp_card_of_natLeq_on[of A]
-by(auto simp add: ordIso_natLeq_on_imp_finite)
-
-
-
-subsection {* The successor of a cardinal *}
-
-
-text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"}
-being a successor cardinal of @{text "r"}. Although the definition does
-not require @{text "r"} to be a cardinal, only this case will be meaningful.  *}
-
-
-definition isCardSuc :: "'a rel \<Rightarrow> 'a set rel \<Rightarrow> bool"
-where
-"isCardSuc r r' \<equiv>
- Card_order r' \<and> r <o r' \<and>
- (\<forall>(r''::'a set rel). Card_order r'' \<and> r <o r'' \<longrightarrow> r' \<le>o r'')"
-
-
-text{* Now we introduce the cardinal-successor operator @{text "cardSuc"},
-by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}.
-Again, the picked item shall be proved unique up to order-isomorphism. *}
-
-
-definition cardSuc :: "'a rel \<Rightarrow> 'a set rel"
-where
-"cardSuc r \<equiv> SOME r'. isCardSuc r r'"
-
-
-lemma exists_minim_Card_order:
-"\<lbrakk>R \<noteq> {}; \<forall>r \<in> R. Card_order r\<rbrakk> \<Longrightarrow> \<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
-unfolding card_order_on_def using exists_minim_Well_order by blast
-
-
-lemma exists_isCardSuc:
-assumes "Card_order r"
-shows "\<exists>r'. isCardSuc r r'"
-proof-
-  let ?R = "{(r'::'a set rel). Card_order r' \<and> r <o r'}"
-  have "|Pow(Field r)| \<in> ?R \<and> (\<forall>r \<in> ?R. Card_order r)" using assms
-  by (simp add: card_of_Card_order Card_order_Pow)
-  then obtain r where "r \<in> ?R \<and> (\<forall>r' \<in> ?R. r \<le>o r')"
-  using exists_minim_Card_order[of ?R] by blast
-  thus ?thesis unfolding isCardSuc_def by auto
-qed
-
-
-lemma cardSuc_isCardSuc:
-assumes "Card_order r"
-shows "isCardSuc r (cardSuc r)"
-unfolding cardSuc_def using assms
-by (simp add: exists_isCardSuc someI_ex)
-
-
-lemma cardSuc_Card_order:
-"Card_order r \<Longrightarrow> Card_order(cardSuc r)"
-using cardSuc_isCardSuc unfolding isCardSuc_def by blast
-
-
-lemma cardSuc_greater:
-"Card_order r \<Longrightarrow> r <o cardSuc r"
-using cardSuc_isCardSuc unfolding isCardSuc_def by blast
-
-
-lemma cardSuc_ordLeq:
-"Card_order r \<Longrightarrow> r \<le>o cardSuc r"
-using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast
-
-
-text{* The minimality property of @{text "cardSuc"} originally present in its definition
-is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}:  *}
-
-lemma cardSuc_least_aux:
-"\<lbrakk>Card_order (r::'a rel); Card_order (r'::'a set rel); r <o r'\<rbrakk> \<Longrightarrow> cardSuc r \<le>o r'"
-using cardSuc_isCardSuc unfolding isCardSuc_def by blast
-
-
-text{* But from this we can infer general minimality: *}
-
-lemma cardSuc_least:
-assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r <o r'"
-shows "cardSuc r \<le>o r'"
-proof-
-  let ?p = "cardSuc r"
-  have 0: "Well_order ?p \<and> Well_order r'"
-  using assms cardSuc_Card_order unfolding card_order_on_def by blast
-  {assume "r' <o ?p"
-   then obtain r'' where 1: "Field r'' < Field ?p" and 2: "r' =o r'' \<and> r'' <o ?p"
-   using internalize_ordLess[of r' ?p] by blast
-   (*  *)
-   have "Card_order r''" using CARD' Card_order_ordIso2 2 by blast
-   moreover have "r <o r''" using LESS 2 ordLess_ordIso_trans by blast
-   ultimately have "?p \<le>o r''" using cardSuc_least_aux CARD by blast
-   hence False using 2 not_ordLess_ordLeq by blast
-  }
-  thus ?thesis using 0 ordLess_or_ordLeq by blast
-qed
-
-
-lemma cardSuc_ordLess_ordLeq:
-assumes CARD: "Card_order r" and CARD': "Card_order r'"
-shows "(r <o r') = (cardSuc r \<le>o r')"
-proof(auto simp add: assms cardSuc_least)
-  assume "cardSuc r \<le>o r'"
-  thus "r <o r'" using assms cardSuc_greater ordLess_ordLeq_trans by blast
-qed
-
-
-lemma cardSuc_ordLeq_ordLess:
-assumes CARD: "Card_order r" and CARD': "Card_order r'"
-shows "(r' <o cardSuc r) = (r' \<le>o r)"
-proof-
-  have "Well_order r \<and> Well_order r'"
-  using assms unfolding card_order_on_def by auto
-  moreover have "Well_order(cardSuc r)"
-  using assms cardSuc_Card_order card_order_on_def by blast
-  ultimately show ?thesis
-  using assms cardSuc_ordLess_ordLeq[of r r']
-  not_ordLeq_iff_ordLess[of r r'] not_ordLeq_iff_ordLess[of r' "cardSuc r"] by blast
-qed
-
-
-lemma cardSuc_mono_ordLeq:
-assumes CARD: "Card_order r" and CARD': "Card_order r'"
-shows "(cardSuc r \<le>o cardSuc r') = (r \<le>o r')"
-using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast
-
-
-lemma cardSuc_invar_ordIso:
-assumes CARD: "Card_order r" and CARD': "Card_order r'"
-shows "(cardSuc r =o cardSuc r') = (r =o r')"
-proof-
-  have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')"
-  using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order)
-  thus ?thesis
-  using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq
-  using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast
-qed
-
-
-lemma cardSuc_natLeq_on_Suc:
-"cardSuc(natLeq_on n) =o natLeq_on(Suc n)"
-proof-
-  obtain r r' p where r_def: "r = natLeq_on n" and
-                      r'_def: "r' = cardSuc(natLeq_on n)"  and
-                      p_def: "p = natLeq_on(Suc n)" by blast
-  (* Preliminary facts:  *)
-  have CARD: "Card_order r \<and> Card_order r' \<and> Card_order p" unfolding r_def r'_def p_def
-  using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast
-  hence WELL: "Well_order r \<and> Well_order r' \<and>  Well_order p"
-  unfolding card_order_on_def by force
-  have FIELD: "Field r = {0..<n} \<and> Field p = {0..<(Suc n)}"
-  unfolding r_def p_def Field_natLeq_on by simp
-  hence FIN: "finite (Field r)" by force
-  have "r <o r'" using CARD unfolding r_def r'_def using cardSuc_greater by blast
-  hence "|Field r| <o r'" using CARD card_of_Field_ordIso ordIso_ordLess_trans by blast
-  hence LESS: "|Field r| <o |Field r'|"
-  using CARD card_of_Field_ordIso ordLess_ordIso_trans ordIso_symmetric by blast
-  (* Main proof: *)
-  have "r' \<le>o p" using CARD unfolding r_def r'_def p_def
-  using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast
-  moreover have "p \<le>o r'"
-  proof-
-    {assume "r' <o p"
-     then obtain f where 0: "embedS r' p f" unfolding ordLess_def by force
-     let ?q = "Restr p (f ` Field r')"
-     have 1: "embed r' p f" using 0 unfolding embedS_def by force
-     hence 2: "f ` Field r' < {0..<(Suc n)}"
-     using WELL FIELD 0 by (auto simp add: embedS_iff)
-     have "wo_rel.ofilter p (f ` Field r')" using embed_Field_ofilter 1 WELL by blast
-     then obtain m where "m \<le> Suc n" and 3: "f ` (Field r') = {0..<m}"
-     unfolding p_def by (auto simp add: natLeq_on_ofilter_iff)
-     hence 4: "m \<le> n" using 2 by force
-     (*  *)
-     have "bij_betw f (Field r') (f ` (Field r'))"
-     using 1 WELL embed_inj_on unfolding bij_betw_def by force
-     moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force
-     ultimately have 5: "finite (Field r') \<and> card(Field r') = card (f ` (Field r'))"
-     using bij_betw_same_card bij_betw_finite by metis
-     hence "card(Field r') \<le> card(Field r)" using 3 4 FIELD by force
-     hence "|Field r'| \<le>o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast
-     hence False using LESS not_ordLess_ordLeq by auto
-    }
-    thus ?thesis using WELL CARD by (fastforce simp: not_ordLess_iff_ordLeq)
-  qed
-  ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast
-qed
-
-
-lemma card_of_cardSuc_finite:
-"finite(Field(cardSuc |A| )) = finite A"
-proof
-  assume *: "finite (Field (cardSuc |A| ))"
-  have 0: "|Field(cardSuc |A| )| =o cardSuc |A|"
-  using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast
-  hence "|A| \<le>o |Field(cardSuc |A| )|"
-  using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric
-  ordLeq_ordIso_trans by blast
-  thus "finite A" using * card_of_ordLeq_finite by blast
-next
-  assume "finite A"
-  then obtain n where "|A| =o natLeq_on n" using finite_iff_card_of_natLeq_on by blast
-  hence "cardSuc |A| =o cardSuc(natLeq_on n)"
-  using card_of_Card_order cardSuc_invar_ordIso natLeq_on_Card_order by blast
-  hence "cardSuc |A| =o natLeq_on(Suc n)"
-  using cardSuc_natLeq_on_Suc ordIso_transitive by blast
-  hence "cardSuc |A| =o |{0..<(Suc n)}|" using card_of_less ordIso_equivalence by blast
-  moreover have "|Field (cardSuc |A| ) | =o cardSuc |A|"
-  using card_of_Field_ordIso cardSuc_Card_order card_of_Card_order by blast
-  ultimately have "|Field (cardSuc |A| ) | =o |{0..<(Suc n)}|"
-  using ordIso_equivalence by blast
-  thus "finite (Field (cardSuc |A| ))"
-  using card_of_ordIso_finite finite_atLeastLessThan by blast
-qed
-
-
-lemma cardSuc_finite:
-assumes "Card_order r"
-shows "finite (Field (cardSuc r)) = finite (Field r)"
-proof-
-  let ?A = "Field r"
-  have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso)
-  hence "cardSuc |?A| =o cardSuc r" using assms
-  by (simp add: card_of_Card_order cardSuc_invar_ordIso)
-  moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|"
-  by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order)
-  moreover
-  {have "|Field (cardSuc r) | =o cardSuc r"
-   using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order)
-   hence "cardSuc r =o |Field (cardSuc r) |"
-   using ordIso_symmetric by blast
-  }
-  ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |"
-  using ordIso_transitive by blast
-  hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))"
-  using card_of_ordIso_finite by blast
-  thus ?thesis by (simp only: card_of_cardSuc_finite)
-qed
-
-
-lemma card_of_Plus_ordLeq_infinite_Field:
-assumes r: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
-and c: "Card_order r"
-shows "|A <+> B| \<le>o r"
-proof-
-  let ?r' = "cardSuc r"
-  have "Card_order ?r' \<and> infinite (Field ?r')" using assms
-  by (simp add: cardSuc_Card_order cardSuc_finite)
-  moreover have "|A| <o ?r'" and "|B| <o ?r'" using A B c
-  by (auto simp: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
-  ultimately have "|A <+> B| <o ?r'"
-  using card_of_Plus_ordLess_infinite_Field by blast
-  thus ?thesis using c r
-  by (simp add: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
-qed
-
-
-lemma card_of_Un_ordLeq_infinite_Field:
-assumes C: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
-and "Card_order r"
-shows "|A Un B| \<le>o r"
-using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq
-ordLeq_transitive by blast
-
-
-
-subsection {* Regular cardinals *}
-
-
-definition cofinal where
-"cofinal A r \<equiv>
- ALL a : Field r. EX b : A. a \<noteq> b \<and> (a,b) : r"
-
-
-definition regular where
-"regular r \<equiv>
- ALL K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r"
-
-
-definition relChain where
-"relChain r As \<equiv>
- ALL i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j"
-
-lemma regular_UNION:
-assumes r: "Card_order r"   "regular r"
-and As: "relChain r As"
-and Bsub: "B \<le> (UN i : Field r. As i)"
-and cardB: "|B| <o r"
-shows "EX i : Field r. B \<le> As i"
-proof-
-  let ?phi = "%b j. j : Field r \<and> b : As j"
-  have "ALL b : B. EX j. ?phi b j" using Bsub by blast
-  then obtain f where f: "!! b. b : B \<Longrightarrow> ?phi b (f b)"
-  using bchoice[of B ?phi] by blast
-  let ?K = "f ` B"
-  {assume 1: "!! i. i : Field r \<Longrightarrow> ~ B \<le> As i"
-   have 2: "cofinal ?K r"
-   unfolding cofinal_def proof auto
-     fix i assume i: "i : Field r"
-     with 1 obtain b where b: "b : B \<and> b \<notin> As i" by blast
-     hence "i \<noteq> f b \<and> ~ (f b,i) : r"
-     using As f unfolding relChain_def by auto
-     hence "i \<noteq> f b \<and> (i, f b) : r" using r
-     unfolding card_order_on_def well_order_on_def linear_order_on_def
-     total_on_def using i f b by auto
-     with b show "\<exists>b\<in>B. i \<noteq> f b \<and> (i, f b) \<in> r" by blast
-   qed
-   moreover have "?K \<le> Field r" using f by blast
-   ultimately have "|?K| =o r" using 2 r unfolding regular_def by blast
-   moreover
-   {
-    have "|?K| <=o |B|" using card_of_image .
-    hence "|?K| <o r" using cardB ordLeq_ordLess_trans by blast
-   }
-   ultimately have False using not_ordLess_ordIso by blast
-  }
-  thus ?thesis by blast
-qed
-
-
-lemma infinite_cardSuc_regular:
-assumes r_inf: "infinite (Field r)" and r_card: "Card_order r"
-shows "regular (cardSuc r)"
-proof-
-  let ?r' = "cardSuc r"
-  have r': "Card_order ?r'"
-  "!! p. Card_order p \<longrightarrow> (p \<le>o r) = (p <o ?r')"
-  using r_card by (auto simp: cardSuc_Card_order cardSuc_ordLeq_ordLess)
-  show ?thesis
-  unfolding regular_def proof auto
-    fix K assume 1: "K \<le> Field ?r'" and 2: "cofinal K ?r'"
-    hence "|K| \<le>o |Field ?r'|" by (simp only: card_of_mono1)
-    also have 22: "|Field ?r'| =o ?r'"
-    using r' by (simp add: card_of_Field_ordIso[of ?r'])
-    finally have "|K| \<le>o ?r'" .
-    moreover
-    {let ?L = "UN j : K. rel.underS ?r' j"
-     let ?J = "Field r"
-     have rJ: "r =o |?J|"
-     using r_card card_of_Field_ordIso ordIso_symmetric by blast
-     assume "|K| <o ?r'"
-     hence "|K| <=o r" using r' card_of_Card_order[of K] by blast
-     hence "|K| \<le>o |?J|" using rJ ordLeq_ordIso_trans by blast
-     moreover
-     {have "ALL j : K. |rel.underS ?r' j| <o ?r'"
-      using r' 1 by (auto simp: card_of_underS)
-      hence "ALL j : K. |rel.underS ?r' j| \<le>o r"
-      using r' card_of_Card_order by blast
-      hence "ALL j : K. |rel.underS ?r' j| \<le>o |?J|"
-      using rJ ordLeq_ordIso_trans by blast
-     }
-     ultimately have "|?L| \<le>o |?J|"
-     using r_inf card_of_UNION_ordLeq_infinite by blast
-     hence "|?L| \<le>o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast
-     hence "|?L| <o ?r'" using r' card_of_Card_order by blast
-     moreover
-     {
-      have "Field ?r' \<le> ?L"
-      using 2 unfolding rel.underS_def cofinal_def by auto
-      hence "|Field ?r'| \<le>o |?L|" by (simp add: card_of_mono1)
-      hence "?r' \<le>o |?L|"
-      using 22 ordIso_ordLeq_trans ordIso_symmetric by blast
-     }
-     ultimately have "|?L| <o |?L|" using ordLess_ordLeq_trans by blast
-     hence False using ordLess_irreflexive by blast
-    }
-    ultimately show "|K| =o ?r'"
-    unfolding ordLeq_iff_ordLess_or_ordIso by blast
-  qed
-qed
-
-lemma cardSuc_UNION:
-assumes r: "Card_order r" and "infinite (Field r)"
-and As: "relChain (cardSuc r) As"
-and Bsub: "B \<le> (UN i : Field (cardSuc r). As i)"
-and cardB: "|B| <=o r"
-shows "EX i : Field (cardSuc r). B \<le> As i"
-proof-
-  let ?r' = "cardSuc r"
-  have "Card_order ?r' \<and> |B| <o ?r'"
-  using r cardB cardSuc_ordLeq_ordLess cardSuc_Card_order
-  card_of_Card_order by blast
-  moreover have "regular ?r'"
-  using assms by(simp add: infinite_cardSuc_regular)
-  ultimately show ?thesis
-  using As Bsub cardB regular_UNION by blast
-qed
-
-
-subsection {* Others *}
-
-lemma card_of_infinite_diff_finite:
-assumes "infinite A" and "finite B"
-shows "|A - B| =o |A|"
-by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2)
-
-(* function space *)
-definition Func where
-"Func A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
-
-lemma Func_empty:
-"Func {} B = {\<lambda>x. undefined}"
-unfolding Func_def by auto
-
-lemma Func_elim:
-assumes "g \<in> Func A B" and "a \<in> A"
-shows "\<exists> b. b \<in> B \<and> g a = b"
-using assms unfolding Func_def by (cases "g a = undefined") auto
-
-definition curr where
-"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
-
-lemma curr_in:
-assumes f: "f \<in> Func (A <*> B) C"
-shows "curr A f \<in> Func A (Func B C)"
-using assms unfolding curr_def Func_def by auto
-
-lemma curr_inj:
-assumes "f1 \<in> Func (A <*> B) C" and "f2 \<in> Func (A <*> B) C"
-shows "curr A f1 = curr A f2 \<longleftrightarrow> f1 = f2"
-proof safe
-  assume c: "curr A f1 = curr A f2"
-  show "f1 = f2"
-  proof (rule ext, clarify)
-    fix a b show "f1 (a, b) = f2 (a, b)"
-    proof (cases "(a,b) \<in> A <*> B")
-      case False
-      thus ?thesis using assms unfolding Func_def by auto
-    next
-      case True hence a: "a \<in> A" and b: "b \<in> B" by auto
-      thus ?thesis
-      using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp
-    qed
-  qed
-qed
-
-lemma curr_surj:
-assumes "g \<in> Func A (Func B C)"
-shows "\<exists> f \<in> Func (A <*> B) C. curr A f = g"
-proof
-  let ?f = "\<lambda> ab. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
-  show "curr A ?f = g"
-  proof (rule ext)
-    fix a show "curr A ?f a = g a"
-    proof (cases "a \<in> A")
-      case False
-      hence "g a = undefined" using assms unfolding Func_def by auto
-      thus ?thesis unfolding curr_def using False by simp
-    next
-      case True
-      obtain g1 where "g1 \<in> Func B C" and "g a = g1"
-      using assms using Func_elim[OF assms True] by blast
-      thus ?thesis using True unfolding Func_def curr_def by auto
-    qed
-  qed
-  show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
-qed
-
-lemma bij_betw_curr:
-"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
-unfolding bij_betw_def inj_on_def image_def
-using curr_in curr_inj curr_surj by blast
-
-lemma card_of_Func_Times:
-"|Func (A <*> B) C| =o |Func A (Func B C)|"
-unfolding card_of_ordIso[symmetric]
-using bij_betw_curr by blast
-
-definition Func_map where
-"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
-
-lemma Func_map:
-assumes g: "g \<in> Func A2 A1" and f1: "f1 ` A1 \<subseteq> B1" and f2: "f2 ` B2 \<subseteq> A2"
-shows "Func_map B2 f1 f2 g \<in> Func B2 B1"
-using assms unfolding Func_def Func_map_def mem_Collect_eq by auto
-
-lemma Func_non_emp:
-assumes "B \<noteq> {}"
-shows "Func A B \<noteq> {}"
-proof-
-  obtain b where b: "b \<in> B" using assms by auto
-  hence "(\<lambda> a. if a \<in> A then b else undefined) \<in> Func A B" unfolding Func_def by auto
-  thus ?thesis by blast
-qed
-
-lemma Func_is_emp:
-"Func A B = {} \<longleftrightarrow> A \<noteq> {} \<and> B = {}" (is "?L \<longleftrightarrow> ?R")
-proof
-  assume L: ?L
-  moreover {assume "A = {}" hence False using L Func_empty by auto}
-  moreover {assume "B \<noteq> {}" hence False using L Func_non_emp by metis}
-  ultimately show ?R by blast
-next
-  assume R: ?R
-  moreover
-  {fix f assume "f \<in> Func A B"
-   moreover obtain a where "a \<in> A" using R by blast
-   ultimately obtain b where "b \<in> B" unfolding Func_def by(cases "f a = undefined", force+)
-   with R have False by auto
-  }
-  thus ?L by blast
-qed
-
-lemma Func_map_surj:
-assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \<subseteq> A2"
-and B2A2: "B2 = {} \<Longrightarrow> A2 = {}"
-shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1"
-proof(cases "B2 = {}")
-  case True
-  thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def)
-next
-  case False note B2 = False
-  show ?thesis
-  proof safe
-    fix h assume h: "h \<in> Func B2 B1"
-    def j1 \<equiv> "inv_into A1 f1"
-    have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
-    then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> f2 (k a2) = a2" by metis
-    {fix b2 assume b2: "b2 \<in> B2"
-     hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto
-     moreover have "k (f2 b2) \<in> B2" using b2 A2(2) k by auto
-     ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast
-    } note kk = this
-    obtain b22 where b22: "b22 \<in> B2" using B2 by auto
-    def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
-    have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
-    have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
-    using kk unfolding j2_def by auto
-    def g \<equiv> "Func_map A2 j1 j2 h"
-    have "Func_map B2 f1 f2 g = h"
-    proof (rule ext)
-      fix b2 show "Func_map B2 f1 f2 g b2 = h b2"
-      proof(cases "b2 \<in> B2")
-        case True
-        show ?thesis
-        proof (cases "h b2 = undefined")
-          case True
-          hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
-          show ?thesis using A2 f_inv_into_f[OF b1]
-            unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
-        qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
-          auto intro: f_inv_into_f)
-      qed(insert h, unfold Func_def Func_map_def, auto)
-    qed
-    moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
-    using inv_into_into j2A2 B1 A2 inv_into_into
-    unfolding j1_def image_def by fast+
-    ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
-    unfolding Func_map_def[abs_def] unfolding image_def by auto
-  qed(insert B1 Func_map[OF _ _ A2(2)], auto)
-qed
-
-lemma card_of_Pow_Func:
-"|Pow A| =o |Func A (UNIV::bool set)|"
-proof-
-  def F \<equiv> "\<lambda> A' a. if a \<in> A then (if a \<in> A' then True else False)
-                            else undefined"
-  have "bij_betw F (Pow A) (Func A (UNIV::bool set))"
-  unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI)
-    fix A1 A2 assume "A1 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2"
-    thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm)
-  next
-    show "F ` Pow A = Func A UNIV"
-    proof safe
-      fix f assume f: "f \<in> Func A (UNIV::bool set)"
-      show "f \<in> F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI)
-        let ?A1 = "{a \<in> A. f a = True}"
-        show "f = F ?A1" unfolding F_def apply(rule ext)
-        using f unfolding Func_def mem_Collect_eq by auto
-      qed auto
-    qed(unfold Func_def mem_Collect_eq F_def, auto)
-  qed
-  thus ?thesis unfolding card_of_ordIso[symmetric] by blast
-qed
-
-lemma card_of_Func_mono:
-fixes A1 A2 :: "'a set" and B :: "'b set"
-assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
-shows "|Func A1 B| \<le>o |Func A2 B|"
-proof-
-  obtain bb where bb: "bb \<in> B" using B by auto
-  def F \<equiv> "\<lambda> (f1::'a \<Rightarrow> 'b) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
-                                                else undefined"
-  show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI)
-    show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe
-      fix f g assume f: "f \<in> Func A1 B" and g: "g \<in> Func A1 B" and eq: "F f = F g"
-      show "f = g"
-      proof(rule ext)
-        fix a show "f a = g a"
-        proof(cases "a \<in> A1")
-          case True
-          thus ?thesis using eq A12 unfolding F_def fun_eq_iff
-          by (elim allE[of _ a]) auto
-        qed(insert f g, unfold Func_def, fastforce)
-      qed
-    qed
-  qed(insert bb, unfold Func_def F_def, force)
-qed
-
-lemma ordLeq_Func:
-assumes "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
-shows "|A| \<le>o |Func A B|"
-unfolding card_of_ordLeq[symmetric] proof(intro exI conjI)
-  let ?F = "\<lambda> aa a. if a \<in> A then (if a = aa then b1 else b2) else undefined"
-  show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto
-  show "?F ` A \<subseteq> Func A B" using assms unfolding Func_def by auto
-qed
-
-lemma infinite_Func:
-assumes A: "infinite A" and B: "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
-shows "infinite (Func A B)"
-using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite)
-
-lemma card_of_Func_UNIV:
-"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \<Rightarrow> 'b. range f \<subseteq> B}|"
-apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI)
-  let ?F = "\<lambda> f (a::'a). ((f a)::'b)"
-  show "bij_betw ?F {f. range f \<subseteq> B} (Func UNIV B)"
-  unfolding bij_betw_def inj_on_def proof safe
-    fix h :: "'a \<Rightarrow> 'b" assume h: "h \<in> Func UNIV B"
-    hence "\<forall> a. \<exists> b. h a = b" unfolding Func_def by auto
-    then obtain f where f: "\<forall> a. h a = f a" by metis
-    hence "range f \<subseteq> B" using h unfolding Func_def by auto
-    thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f unfolding image_def by auto
-  qed(unfold Func_def fun_eq_iff, auto)
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation_LFP.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -0,0 +1,2438 @@
+(*  Title:      HOL/Cardinals/Cardinal_Order_Relation_LFP.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Cardinal-order relations (LFP).
+*)
+
+header {* Cardinal-Order Relations (LFP)  *}
+
+theory Cardinal_Order_Relation_LFP
+imports Constructions_on_Wellorders_LFP
+begin
+
+
+text{* In this section, we define cardinal-order relations to be minim well-orders
+on their field.  Then we define the cardinal of a set to be {\em some} cardinal-order
+relation on that set, which will be unique up to order isomorphism.  Then we study
+the connection between cardinals and:
+\begin{itemize}
+\item standard set-theoretic constructions: products,
+sums, unions, lists, powersets, set-of finite sets operator;
+\item finiteness and infiniteness (in particular, with the numeric cardinal operator
+for finite sets, @{text "card"}, from the theory @{text "Finite_Sets.thy"}).
+\end{itemize}
+%
+On the way, we define the canonical $\omega$ cardinal and finite cardinals.  We also
+define (again, up to order isomorphism) the successor of a cardinal, and show that
+any cardinal admits a successor.
+
+Main results of this section are the existence of cardinal relations and the
+facts that, in the presence of infiniteness,
+most of the standard set-theoretic constructions (except for the powerset)
+{\em do not increase cardinality}.  In particular, e.g., the set of words/lists over
+any infinite set has the same cardinality (hence, is in bijection) with that set.
+*}
+
+
+subsection {* Cardinal orders *}
+
+
+text{* A cardinal order in our setting shall be a well-order {\em minim} w.r.t. the
+order-embedding relation, @{text "\<le>o"} (which is the same as being {\em minimal} w.r.t. the
+strict order-embedding relation, @{text "<o"}), among all the well-orders on its field.  *}
+
+definition card_order_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
+where
+"card_order_on A r \<equiv> well_order_on A r \<and> (\<forall>r'. well_order_on A r' \<longrightarrow> r \<le>o r')"
+
+
+abbreviation "Card_order r \<equiv> card_order_on (Field r) r"
+abbreviation "card_order r \<equiv> card_order_on UNIV r"
+
+
+lemma card_order_on_well_order_on:
+assumes "card_order_on A r"
+shows "well_order_on A r"
+using assms unfolding card_order_on_def by simp
+
+
+lemma card_order_on_Card_order:
+"card_order_on A r \<Longrightarrow> A = Field r \<and> Card_order r"
+unfolding card_order_on_def using rel.well_order_on_Field by blast
+
+
+text{* The existence of a cardinal relation on any given set (which will mean
+that any set has a cardinal) follows from two facts:
+\begin{itemize}
+\item Zermelo's theorem (proved in @{text "Zorn.thy"} as theorem @{text "well_order_on"}),
+which states that on any given set there exists a well-order;
+\item The well-founded-ness of @{text "<o"}, ensuring that then there exists a minimal
+such well-order, i.e., a cardinal order.
+\end{itemize}
+*}
+
+
+theorem card_order_on: "\<exists>r. card_order_on A r"
+proof-
+  obtain R where R_def: "R = {r. well_order_on A r}" by blast
+  have 1: "R \<noteq> {} \<and> (\<forall>r \<in> R. Well_order r)"
+  using well_order_on[of A] R_def rel.well_order_on_Well_order by blast
+  hence "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+  using  exists_minim_Well_order[of R] by auto
+  thus ?thesis using R_def unfolding card_order_on_def by auto
+qed
+
+
+lemma card_order_on_ordIso:
+assumes CO: "card_order_on A r" and CO': "card_order_on A r'"
+shows "r =o r'"
+using assms unfolding card_order_on_def
+using ordIso_iff_ordLeq by blast
+
+
+lemma Card_order_ordIso:
+assumes CO: "Card_order r" and ISO: "r' =o r"
+shows "Card_order r'"
+using ISO unfolding ordIso_def
+proof(unfold card_order_on_def, auto)
+  fix p' assume "well_order_on (Field r') p'"
+  hence 0: "Well_order p' \<and> Field p' = Field r'"
+  using rel.well_order_on_Well_order by blast
+  obtain f where 1: "iso r' r f" and 2: "Well_order r \<and> Well_order r'"
+  using ISO unfolding ordIso_def by auto
+  hence 3: "inj_on f (Field r') \<and> f ` (Field r') = Field r"
+  by (auto simp add: iso_iff embed_inj_on)
+  let ?p = "dir_image p' f"
+  have 4: "p' =o ?p \<and> Well_order ?p"
+  using 0 2 3 by (auto simp add: dir_image_ordIso Well_order_dir_image)
+  moreover have "Field ?p =  Field r"
+  using 0 3 by (auto simp add: dir_image_Field2 order_on_defs)
+  ultimately have "well_order_on (Field r) ?p" by auto
+  hence "r \<le>o ?p" using CO unfolding card_order_on_def by auto
+  thus "r' \<le>o p'"
+  using ISO 4 ordLeq_ordIso_trans ordIso_ordLeq_trans ordIso_symmetric by blast
+qed
+
+
+lemma Card_order_ordIso2:
+assumes CO: "Card_order r" and ISO: "r =o r'"
+shows "Card_order r'"
+using assms Card_order_ordIso ordIso_symmetric by blast
+
+
+subsection {* Cardinal of a set *}
+
+
+text{* We define the cardinal of set to be {\em some} cardinal order on that set.
+We shall prove that this notion is unique up to order isomorphism, meaning
+that order isomorphism shall be the true identity of cardinals.  *}
+
+
+definition card_of :: "'a set \<Rightarrow> 'a rel" ("|_|" )
+where "card_of A = (SOME r. card_order_on A r)"
+
+
+lemma card_of_card_order_on: "card_order_on A |A|"
+unfolding card_of_def by (auto simp add: card_order_on someI_ex)
+
+
+lemma card_of_well_order_on: "well_order_on A |A|"
+using card_of_card_order_on card_order_on_def by blast
+
+
+lemma Field_card_of: "Field |A| = A"
+using card_of_card_order_on[of A] unfolding card_order_on_def
+using rel.well_order_on_Field by blast
+
+
+lemma card_of_Card_order: "Card_order |A|"
+by (simp only: card_of_card_order_on Field_card_of)
+
+
+corollary ordIso_card_of_imp_Card_order:
+"r =o |A| \<Longrightarrow> Card_order r"
+using card_of_Card_order Card_order_ordIso by blast
+
+
+lemma card_of_Well_order: "Well_order |A|"
+using card_of_Card_order unfolding  card_order_on_def by auto
+
+
+lemma card_of_refl: "|A| =o |A|"
+using card_of_Well_order ordIso_reflexive by blast
+
+
+lemma card_of_least: "well_order_on A r \<Longrightarrow> |A| \<le>o r"
+using card_of_card_order_on unfolding card_order_on_def by blast
+
+
+lemma card_of_ordIso:
+"(\<exists>f. bij_betw f A B) = ( |A| =o |B| )"
+proof(auto)
+  fix f assume *: "bij_betw f A B"
+  then obtain r where "well_order_on B r \<and> |A| =o r"
+  using Well_order_iso_copy card_of_well_order_on by blast
+  hence "|B| \<le>o |A|" using card_of_least
+  ordLeq_ordIso_trans ordIso_symmetric by blast
+  moreover
+  {let ?g = "inv_into A f"
+   have "bij_betw ?g B A" using * bij_betw_inv_into by blast
+   then obtain r where "well_order_on A r \<and> |B| =o r"
+   using Well_order_iso_copy card_of_well_order_on by blast
+   hence "|A| \<le>o |B|" using card_of_least
+   ordLeq_ordIso_trans ordIso_symmetric by blast
+  }
+  ultimately show "|A| =o |B|" using ordIso_iff_ordLeq by blast
+next
+  assume "|A| =o |B|"
+  then obtain f where "iso ( |A| ) ( |B| ) f"
+  unfolding ordIso_def by auto
+  hence "bij_betw f A B" unfolding iso_def Field_card_of by simp
+  thus "\<exists>f. bij_betw f A B" by auto
+qed
+
+
+lemma card_of_ordLeq:
+"(\<exists>f. inj_on f A \<and> f ` A \<le> B) = ( |A| \<le>o |B| )"
+proof(auto)
+  fix f assume *: "inj_on f A" and **: "f ` A \<le> B"
+  {assume "|B| <o |A|"
+   hence "|B| \<le>o |A|" using ordLeq_iff_ordLess_or_ordIso by blast
+   then obtain g where "embed ( |B| ) ( |A| ) g"
+   unfolding ordLeq_def by auto
+   hence 1: "inj_on g B \<and> g ` B \<le> A" using embed_inj_on[of "|B|" "|A|" "g"]
+   card_of_Well_order[of "B"] Field_card_of[of "B"] Field_card_of[of "A"]
+   embed_Field[of "|B|" "|A|" g] by auto
+   obtain h where "bij_betw h A B"
+   using * ** 1 Cantor_Bernstein[of f] by fastforce
+   hence "|A| =o |B|" using card_of_ordIso by blast
+   hence "|A| \<le>o |B|" using ordIso_iff_ordLeq by auto
+  }
+  thus "|A| \<le>o |B|" using ordLess_or_ordLeq[of "|B|" "|A|"]
+  by (auto simp: card_of_Well_order)
+next
+  assume *: "|A| \<le>o |B|"
+  obtain f where "embed ( |A| ) ( |B| ) f"
+  using * unfolding ordLeq_def by auto
+  hence "inj_on f A \<and> f ` A \<le> B" using embed_inj_on[of "|A|" "|B|" f]
+  card_of_Well_order[of "A"] Field_card_of[of "A"] Field_card_of[of "B"]
+  embed_Field[of "|A|" "|B|" f] by auto
+  thus "\<exists>f. inj_on f A \<and> f ` A \<le> B" by auto
+qed
+
+
+lemma card_of_ordLeq2:
+"A \<noteq> {} \<Longrightarrow> (\<exists>g. g ` B = A) = ( |A| \<le>o |B| )"
+using card_of_ordLeq[of A B] inj_on_iff_surj[of A B] by auto
+
+
+lemma card_of_ordLess:
+"(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = ( |B| <o |A| )"
+proof-
+  have "(\<not>(\<exists>f. inj_on f A \<and> f ` A \<le> B)) = (\<not> |A| \<le>o |B| )"
+  using card_of_ordLeq by blast
+  also have "\<dots> = ( |B| <o |A| )"
+  using card_of_Well_order[of A] card_of_Well_order[of B]
+        not_ordLeq_iff_ordLess by blast
+  finally show ?thesis .
+qed
+
+
+lemma card_of_ordLess2:
+"B \<noteq> {} \<Longrightarrow> (\<not>(\<exists>f. f ` A = B)) = ( |A| <o |B| )"
+using card_of_ordLess[of B A] inj_on_iff_surj[of B A] by auto
+
+
+lemma card_of_ordIsoI:
+assumes "bij_betw f A B"
+shows "|A| =o |B|"
+using assms unfolding card_of_ordIso[symmetric] by auto
+
+
+lemma card_of_ordLeqI:
+assumes "inj_on f A" and "\<And> a. a \<in> A \<Longrightarrow> f a \<in> B"
+shows "|A| \<le>o |B|"
+using assms unfolding card_of_ordLeq[symmetric] by auto
+
+
+lemma card_of_unique:
+"card_order_on A r \<Longrightarrow> r =o |A|"
+by (simp only: card_order_on_ordIso card_of_card_order_on)
+
+
+lemma card_of_mono1:
+"A \<le> B \<Longrightarrow> |A| \<le>o |B|"
+using inj_on_id[of A] card_of_ordLeq[of A B] by fastforce
+
+
+lemma card_of_mono2:
+assumes "r \<le>o r'"
+shows "|Field r| \<le>o |Field r'|"
+proof-
+  obtain f where
+  1: "well_order_on (Field r) r \<and> well_order_on (Field r) r \<and> embed r r' f"
+  using assms unfolding ordLeq_def
+  by (auto simp add: rel.well_order_on_Well_order)
+  hence "inj_on f (Field r) \<and> f ` (Field r) \<le> Field r'"
+  by (auto simp add: embed_inj_on embed_Field)
+  thus "|Field r| \<le>o |Field r'|" using card_of_ordLeq by blast
+qed
+
+
+lemma card_of_cong: "r =o r' \<Longrightarrow> |Field r| =o |Field r'|"
+by (simp add: ordIso_iff_ordLeq card_of_mono2)
+
+
+lemma card_of_Field_ordLess: "Well_order r \<Longrightarrow> |Field r| \<le>o r"
+using card_of_least card_of_well_order_on rel.well_order_on_Well_order by blast
+
+
+lemma card_of_Field_ordIso:
+assumes "Card_order r"
+shows "|Field r| =o r"
+proof-
+  have "card_order_on (Field r) r"
+  using assms card_order_on_Card_order by blast
+  moreover have "card_order_on (Field r) |Field r|"
+  using card_of_card_order_on by blast
+  ultimately show ?thesis using card_order_on_ordIso by blast
+qed
+
+
+lemma Card_order_iff_ordIso_card_of:
+"Card_order r = (r =o |Field r| )"
+using ordIso_card_of_imp_Card_order card_of_Field_ordIso ordIso_symmetric by blast
+
+
+lemma Card_order_iff_ordLeq_card_of:
+"Card_order r = (r \<le>o |Field r| )"
+proof-
+  have "Card_order r = (r =o |Field r| )"
+  unfolding Card_order_iff_ordIso_card_of by simp
+  also have "... = (r \<le>o |Field r| \<and> |Field r| \<le>o r)"
+  unfolding ordIso_iff_ordLeq by simp
+  also have "... = (r \<le>o |Field r| )"
+  using card_of_Field_ordLess
+  by (auto simp: card_of_Field_ordLess ordLeq_Well_order_simp)
+  finally show ?thesis .
+qed
+
+
+lemma Card_order_iff_Restr_underS:
+assumes "Well_order r"
+shows "Card_order r = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o |Field r| )"
+using assms unfolding Card_order_iff_ordLeq_card_of
+using ordLeq_iff_ordLess_Restr card_of_Well_order by blast
+
+
+lemma card_of_underS:
+assumes r: "Card_order r" and a: "a : Field r"
+shows "|rel.underS r a| <o r"
+proof-
+  let ?A = "rel.underS r a"  let ?r' = "Restr r ?A"
+  have 1: "Well_order r"
+  using r unfolding card_order_on_def by simp
+  have "Well_order ?r'" using 1 Well_order_Restr by auto
+  moreover have "card_order_on (Field ?r') |Field ?r'|"
+  using card_of_card_order_on .
+  ultimately have "|Field ?r'| \<le>o ?r'"
+  unfolding card_order_on_def by simp
+  moreover have "Field ?r' = ?A"
+  using 1 wo_rel.underS_ofilter Field_Restr_ofilter
+  unfolding wo_rel_def by fastforce
+  ultimately have "|?A| \<le>o ?r'" by simp
+  also have "?r' <o |Field r|"
+  using 1 a r Card_order_iff_Restr_underS by blast
+  also have "|Field r| =o r"
+  using r ordIso_symmetric unfolding Card_order_iff_ordIso_card_of by auto
+  finally show ?thesis .
+qed
+
+
+lemma ordLess_Field:
+assumes "r <o r'"
+shows "|Field r| <o r'"
+proof-
+  have "well_order_on (Field r) r" using assms unfolding ordLess_def
+  by (auto simp add: rel.well_order_on_Well_order)
+  hence "|Field r| \<le>o r" using card_of_least by blast
+  thus ?thesis using assms ordLeq_ordLess_trans by blast
+qed
+
+
+lemma internalize_card_of_ordLeq:
+"( |A| \<le>o r) = (\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r)"
+proof
+  assume "|A| \<le>o r"
+  then obtain p where 1: "Field p \<le> Field r \<and> |A| =o p \<and> p \<le>o r"
+  using internalize_ordLeq[of "|A|" r] by blast
+  hence "Card_order p" using card_of_Card_order Card_order_ordIso2 by blast
+  hence "|Field p| =o p" using card_of_Field_ordIso by blast
+  hence "|A| =o |Field p| \<and> |Field p| \<le>o r"
+  using 1 ordIso_equivalence ordIso_ordLeq_trans by blast
+  thus "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r" using 1 by blast
+next
+  assume "\<exists>B \<le> Field r. |A| =o |B| \<and> |B| \<le>o r"
+  thus "|A| \<le>o r" using ordIso_ordLeq_trans by blast
+qed
+
+
+lemma internalize_card_of_ordLeq2:
+"( |A| \<le>o |C| ) = (\<exists>B \<le> C. |A| =o |B| \<and> |B| \<le>o |C| )"
+using internalize_card_of_ordLeq[of "A" "|C|"] Field_card_of[of C] by auto
+
+
+
+subsection {* Cardinals versus set operations on arbitrary sets *}
+
+
+text{* Here we embark in a long journey of simple results showing
+that the standard set-theoretic operations are well-behaved w.r.t. the notion of
+cardinal -- essentially, this means that they preserve the ``cardinal identity"
+@{text "=o"} and are monotonic w.r.t. @{text "\<le>o"}.
+*}
+
+
+lemma card_of_empty: "|{}| \<le>o |A|"
+using card_of_ordLeq inj_on_id by blast
+
+
+lemma card_of_empty1:
+assumes "Well_order r \<or> Card_order r"
+shows "|{}| \<le>o r"
+proof-
+  have "Well_order r" using assms unfolding card_order_on_def by auto
+  hence "|Field r| <=o r"
+  using assms card_of_Field_ordLess by blast
+  moreover have "|{}| \<le>o |Field r|" by (simp add: card_of_empty)
+  ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+corollary Card_order_empty:
+"Card_order r \<Longrightarrow> |{}| \<le>o r" by (simp add: card_of_empty1)
+
+
+lemma card_of_empty2:
+assumes LEQ: "|A| =o |{}|"
+shows "A = {}"
+using assms card_of_ordIso[of A] bij_betw_empty2 by blast
+
+
+lemma card_of_empty3:
+assumes LEQ: "|A| \<le>o |{}|"
+shows "A = {}"
+using assms
+by (simp add: ordIso_iff_ordLeq card_of_empty1 card_of_empty2
+              ordLeq_Well_order_simp)
+
+
+lemma card_of_empty_ordIso:
+"|{}::'a set| =o |{}::'b set|"
+using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
+
+
+lemma card_of_image:
+"|f ` A| <=o |A|"
+proof(cases "A = {}", simp add: card_of_empty)
+  assume "A ~= {}"
+  hence "f ` A ~= {}" by auto
+  thus "|f ` A| \<le>o |A|"
+  using card_of_ordLeq2[of "f ` A" A] by auto
+qed
+
+
+lemma surj_imp_ordLeq:
+assumes "B <= f ` A"
+shows "|B| <=o |A|"
+proof-
+  have "|B| <=o |f ` A|" using assms card_of_mono1 by auto
+  thus ?thesis using card_of_image ordLeq_transitive by blast
+qed
+
+
+lemma card_of_ordLeqI2:
+assumes "B \<subseteq> f ` A"
+shows "|B| \<le>o |A|"
+using assms by (metis surj_imp_ordLeq)
+
+
+lemma card_of_singl_ordLeq:
+assumes "A \<noteq> {}"
+shows "|{b}| \<le>o |A|"
+proof-
+  obtain a where *: "a \<in> A" using assms by auto
+  let ?h = "\<lambda> b'::'b. if b' = b then a else undefined"
+  have "inj_on ?h {b} \<and> ?h ` {b} \<le> A"
+  using * unfolding inj_on_def by auto
+  thus ?thesis using card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_singl_ordLeq:
+"\<lbrakk>Card_order r; Field r \<noteq> {}\<rbrakk> \<Longrightarrow> |{b}| \<le>o r"
+using card_of_singl_ordLeq[of "Field r" b]
+      card_of_Field_ordIso[of r] ordLeq_ordIso_trans by blast
+
+
+lemma card_of_Pow: "|A| <o |Pow A|"
+using card_of_ordLess2[of "Pow A" A]  Cantors_paradox[of A]
+      Pow_not_empty[of A] by auto
+
+
+lemma infinite_Pow:
+assumes "infinite A"
+shows "infinite (Pow A)"
+proof-
+  have "|A| \<le>o |Pow A|" by (metis card_of_Pow ordLess_imp_ordLeq)
+  thus ?thesis by (metis assms finite_Pow_iff)
+qed
+
+
+corollary Card_order_Pow:
+"Card_order r \<Longrightarrow> r <o |Pow(Field r)|"
+using card_of_Pow card_of_Field_ordIso ordIso_ordLess_trans ordIso_symmetric by blast
+
+
+corollary card_of_set_type: "|UNIV::'a set| <o |UNIV::'a set set|"
+using card_of_Pow[of "UNIV::'a set"] by simp
+
+
+lemma card_of_Plus1: "|A| \<le>o |A <+> B|"
+proof-
+  have "Inl ` A \<le> A <+> B" by auto
+  thus ?thesis using inj_Inl[of A] card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_Plus1:
+"Card_order r \<Longrightarrow> r \<le>o |(Field r) <+> B|"
+using card_of_Plus1 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Plus2: "|B| \<le>o |A <+> B|"
+proof-
+  have "Inr ` B \<le> A <+> B" by auto
+  thus ?thesis using inj_Inr[of B] card_of_ordLeq by blast
+qed
+
+
+corollary Card_order_Plus2:
+"Card_order r \<Longrightarrow> r \<le>o |A <+> (Field r)|"
+using card_of_Plus2 card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Plus_empty1: "|A| =o |A <+> {}|"
+proof-
+  have "bij_betw Inl A (A <+> {})" unfolding bij_betw_def inj_on_def by auto
+  thus ?thesis using card_of_ordIso by auto
+qed
+
+
+lemma card_of_Plus_empty2: "|A| =o |{} <+> A|"
+proof-
+  have "bij_betw Inr A ({} <+> A)" unfolding bij_betw_def inj_on_def by auto
+  thus ?thesis using card_of_ordIso by auto
+qed
+
+
+lemma card_of_Plus_commute: "|A <+> B| =o |B <+> A|"
+proof-
+  let ?f = "\<lambda>(c::'a + 'b). case c of Inl a \<Rightarrow> Inr a
+                                   | Inr b \<Rightarrow> Inl b"
+  have "bij_betw ?f (A <+> B) (B <+> A)"
+  unfolding bij_betw_def inj_on_def by force
+  thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_assoc:
+fixes A :: "'a set" and B :: "'b set" and C :: "'c set"
+shows "|(A <+> B) <+> C| =o |A <+> B <+> C|"
+proof -
+  def f \<equiv> "\<lambda>(k::('a + 'b) + 'c).
+  case k of Inl ab \<Rightarrow> (case ab of Inl a \<Rightarrow> Inl a
+                                 |Inr b \<Rightarrow> Inr (Inl b))
+           |Inr c \<Rightarrow> Inr (Inr c)"
+  have "A <+> B <+> C \<subseteq> f ` ((A <+> B) <+> C)"
+  proof
+    fix x assume x: "x \<in> A <+> B <+> C"
+    show "x \<in> f ` ((A <+> B) <+> C)"
+    proof(cases x)
+      case (Inl a)
+      hence "a \<in> A" "x = f (Inl (Inl a))"
+      using x unfolding f_def by auto
+      thus ?thesis by auto
+    next
+      case (Inr bc) note 1 = Inr show ?thesis
+      proof(cases bc)
+        case (Inl b)
+        hence "b \<in> B" "x = f (Inl (Inr b))"
+        using x 1 unfolding f_def by auto
+        thus ?thesis by auto
+      next
+        case (Inr c)
+        hence "c \<in> C" "x = f (Inr c)"
+        using x 1 unfolding f_def by auto
+        thus ?thesis by auto
+      qed
+    qed
+  qed
+  hence "bij_betw f ((A <+> B) <+> C) (A <+> B <+> C)"
+  unfolding bij_betw_def inj_on_def f_def by force
+  thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_mono1:
+assumes "|A| \<le>o |B|"
+shows "|A <+> C| \<le>o |B <+> C|"
+proof-
+  obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
+  using assms card_of_ordLeq[of A] by fastforce
+  obtain g where g_def:
+  "g = (\<lambda>d. case d of Inl a \<Rightarrow> Inl(f a) | Inr (c::'c) \<Rightarrow> Inr c)" by blast
+  have "inj_on g (A <+> C) \<and> g ` (A <+> C) \<le> (B <+> C)"
+  proof-
+    {fix d1 and d2 assume "d1 \<in> A <+> C \<and> d2 \<in> A <+> C" and
+                          "g d1 = g d2"
+     hence "d1 = d2" using 1 unfolding inj_on_def
+     by(case_tac d1, case_tac d2, auto simp add: g_def)
+    }
+    moreover
+    {fix d assume "d \<in> A <+> C"
+     hence "g d \<in> B <+> C"  using 1
+     by(case_tac d, auto simp add: g_def)
+    }
+    ultimately show ?thesis unfolding inj_on_def by auto
+  qed
+  thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary ordLeq_Plus_mono1:
+assumes "r \<le>o r'"
+shows "|(Field r) <+> C| \<le>o |(Field r') <+> C|"
+using assms card_of_mono2 card_of_Plus_mono1 by blast
+
+
+lemma card_of_Plus_mono2:
+assumes "|A| \<le>o |B|"
+shows "|C <+> A| \<le>o |C <+> B|"
+using assms card_of_Plus_mono1[of A B C]
+      card_of_Plus_commute[of C A]  card_of_Plus_commute[of B C]
+      ordIso_ordLeq_trans[of "|C <+> A|"] ordLeq_ordIso_trans[of "|C <+> A|"]
+by blast
+
+
+corollary ordLeq_Plus_mono2:
+assumes "r \<le>o r'"
+shows "|A <+> (Field r)| \<le>o |A <+> (Field r')|"
+using assms card_of_mono2 card_of_Plus_mono2 by blast
+
+
+lemma card_of_Plus_mono:
+assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
+shows "|A <+> C| \<le>o |B <+> D|"
+using assms card_of_Plus_mono1[of A B C] card_of_Plus_mono2[of C D B]
+      ordLeq_transitive[of "|A <+> C|"] by blast
+
+
+corollary ordLeq_Plus_mono:
+assumes "r \<le>o r'" and "p \<le>o p'"
+shows "|(Field r) <+> (Field p)| \<le>o |(Field r') <+> (Field p')|"
+using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Plus_mono by blast
+
+
+lemma card_of_Plus_cong1:
+assumes "|A| =o |B|"
+shows "|A <+> C| =o |B <+> C|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono1)
+
+
+corollary ordIso_Plus_cong1:
+assumes "r =o r'"
+shows "|(Field r) <+> C| =o |(Field r') <+> C|"
+using assms card_of_cong card_of_Plus_cong1 by blast
+
+
+lemma card_of_Plus_cong2:
+assumes "|A| =o |B|"
+shows "|C <+> A| =o |C <+> B|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono2)
+
+
+corollary ordIso_Plus_cong2:
+assumes "r =o r'"
+shows "|A <+> (Field r)| =o |A <+> (Field r')|"
+using assms card_of_cong card_of_Plus_cong2 by blast
+
+
+lemma card_of_Plus_cong:
+assumes "|A| =o |B|" and "|C| =o |D|"
+shows "|A <+> C| =o |B <+> D|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Plus_mono)
+
+
+corollary ordIso_Plus_cong:
+assumes "r =o r'" and "p =o p'"
+shows "|(Field r) <+> (Field p)| =o |(Field r') <+> (Field p')|"
+using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Plus_cong by blast
+
+
+lemma card_of_Un1:
+shows "|A| \<le>o |A \<union> B| "
+using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
+
+
+lemma card_of_diff:
+shows "|A - B| \<le>o |A|"
+using inj_on_id[of "A - B"] card_of_ordLeq[of "A - B" _] by fastforce
+
+
+lemma card_of_Un_Plus_ordLeq:
+"|A \<union> B| \<le>o |A <+> B|"
+proof-
+   let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
+   have "inj_on ?f (A \<union> B) \<and> ?f ` (A \<union> B) \<le> A <+> B"
+   unfolding inj_on_def by auto
+   thus ?thesis using card_of_ordLeq by blast
+qed
+
+
+lemma card_of_Times1:
+assumes "A \<noteq> {}"
+shows "|B| \<le>o |B \<times> A|"
+proof(cases "B = {}", simp add: card_of_empty)
+  assume *: "B \<noteq> {}"
+  have "fst `(B \<times> A) = B" unfolding image_def using assms by auto
+  thus ?thesis using inj_on_iff_surj[of B "B \<times> A"]
+                     card_of_ordLeq[of B "B \<times> A"] * by blast
+qed
+
+
+corollary Card_order_Times1:
+"\<lbrakk>Card_order r; B \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |(Field r) \<times> B|"
+using card_of_Times1[of B] card_of_Field_ordIso
+      ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Times_commute: "|A \<times> B| =o |B \<times> A|"
+proof-
+  let ?f = "\<lambda>(a::'a,b::'b). (b,a)"
+  have "bij_betw ?f (A \<times> B) (B \<times> A)"
+  unfolding bij_betw_def inj_on_def by auto
+  thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Times2:
+assumes "A \<noteq> {}"   shows "|B| \<le>o |A \<times> B|"
+using assms card_of_Times1[of A B] card_of_Times_commute[of B A]
+      ordLeq_ordIso_trans by blast
+
+
+corollary Card_order_Times2:
+"\<lbrakk>Card_order r; A \<noteq> {}\<rbrakk> \<Longrightarrow> r \<le>o |A \<times> (Field r)|"
+using card_of_Times2[of A] card_of_Field_ordIso
+      ordIso_ordLeq_trans ordIso_symmetric by blast
+
+
+lemma card_of_Times3: "|A| \<le>o |A \<times> A|"
+using card_of_Times1[of A]
+by(cases "A = {}", simp add: card_of_empty, blast)
+
+
+lemma card_of_Plus_Times_bool: "|A <+> A| =o |A \<times> (UNIV::bool set)|"
+proof-
+  let ?f = "\<lambda>c::'a + 'a. case c of Inl a \<Rightarrow> (a,True)
+                                  |Inr a \<Rightarrow> (a,False)"
+  have "bij_betw ?f (A <+> A) (A \<times> (UNIV::bool set))"
+  proof-
+    {fix  c1 and c2 assume "?f c1 = ?f c2"
+     hence "c1 = c2"
+     by(case_tac "c1", case_tac "c2", auto, case_tac "c2", auto)
+    }
+    moreover
+    {fix c assume "c \<in> A <+> A"
+     hence "?f c \<in> A \<times> (UNIV::bool set)"
+     by(case_tac c, auto)
+    }
+    moreover
+    {fix a bl assume *: "(a,bl) \<in> A \<times> (UNIV::bool set)"
+     have "(a,bl) \<in> ?f ` ( A <+> A)"
+     proof(cases bl)
+       assume bl hence "?f(Inl a) = (a,bl)" by auto
+       thus ?thesis using * by force
+     next
+       assume "\<not> bl" hence "?f(Inr a) = (a,bl)" by auto
+       thus ?thesis using * by force
+     qed
+    }
+    ultimately show ?thesis unfolding bij_betw_def inj_on_def by auto
+  qed
+  thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Times_mono1:
+assumes "|A| \<le>o |B|"
+shows "|A \<times> C| \<le>o |B \<times> C|"
+proof-
+  obtain f where 1: "inj_on f A \<and> f ` A \<le> B"
+  using assms card_of_ordLeq[of A] by fastforce
+  obtain g where g_def:
+  "g = (\<lambda>(a,c::'c). (f a,c))" by blast
+  have "inj_on g (A \<times> C) \<and> g ` (A \<times> C) \<le> (B \<times> C)"
+  using 1 unfolding inj_on_def using g_def by auto
+  thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary ordLeq_Times_mono1:
+assumes "r \<le>o r'"
+shows "|(Field r) \<times> C| \<le>o |(Field r') \<times> C|"
+using assms card_of_mono2 card_of_Times_mono1 by blast
+
+
+lemma card_of_Times_mono2:
+assumes "|A| \<le>o |B|"
+shows "|C \<times> A| \<le>o |C \<times> B|"
+using assms card_of_Times_mono1[of A B C]
+      card_of_Times_commute[of C A]  card_of_Times_commute[of B C]
+      ordIso_ordLeq_trans[of "|C \<times> A|"] ordLeq_ordIso_trans[of "|C \<times> A|"]
+by blast
+
+
+corollary ordLeq_Times_mono2:
+assumes "r \<le>o r'"
+shows "|A \<times> (Field r)| \<le>o |A \<times> (Field r')|"
+using assms card_of_mono2 card_of_Times_mono2 by blast
+
+
+lemma card_of_Times_cong1:
+assumes "|A| =o |B|"
+shows "|A \<times> C| =o |B \<times> C|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono1)
+
+
+lemma card_of_Times_cong2:
+assumes "|A| =o |B|"
+shows "|C \<times> A| =o |C \<times> B|"
+using assms by (simp add: ordIso_iff_ordLeq card_of_Times_mono2)
+
+
+corollary ordIso_Times_cong2:
+assumes "r =o r'"
+shows "|A \<times> (Field r)| =o |A \<times> (Field r')|"
+using assms card_of_cong card_of_Times_cong2 by blast
+
+
+lemma card_of_Sigma_mono1:
+assumes "\<forall>i \<in> I. |A i| \<le>o |B i|"
+shows "|SIGMA i : I. A i| \<le>o |SIGMA i : I. B i|"
+proof-
+  have "\<forall>i. i \<in> I \<longrightarrow> (\<exists>f. inj_on f (A i) \<and> f ` (A i) \<le> B i)"
+  using assms by (auto simp add: card_of_ordLeq)
+  with choice[of "\<lambda> i f. i \<in> I \<longrightarrow> inj_on f (A i) \<and> f ` (A i) \<le> B i"]
+  obtain F where 1: "\<forall>i \<in> I. inj_on (F i) (A i) \<and> (F i) ` (A i) \<le> B i" by metis
+  obtain g where g_def: "g = (\<lambda>(i,a::'b). (i,F i a))" by blast
+  have "inj_on g (Sigma I A) \<and> g ` (Sigma I A) \<le> (Sigma I B)"
+  using 1 unfolding inj_on_def using g_def by force
+  thus ?thesis using card_of_ordLeq by metis
+qed
+
+
+corollary card_of_Sigma_Times:
+"\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> |SIGMA i : I. A i| \<le>o |I \<times> B|"
+using card_of_Sigma_mono1[of I A "\<lambda>i. B"] .
+
+
+lemma card_of_UNION_Sigma:
+"|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
+using Ex_inj_on_UNION_Sigma[of I A] card_of_ordLeq by metis
+
+
+lemma card_of_bool:
+assumes "a1 \<noteq> a2"
+shows "|UNIV::bool set| =o |{a1,a2}|"
+proof-
+  let ?f = "\<lambda> bl. case bl of True \<Rightarrow> a1 | False \<Rightarrow> a2"
+  have "bij_betw ?f UNIV {a1,a2}"
+  proof-
+    {fix bl1 and bl2 assume "?f  bl1 = ?f bl2"
+     hence "bl1 = bl2" using assms by (case_tac bl1, case_tac bl2, auto)
+    }
+    moreover
+    {fix bl have "?f bl \<in> {a1,a2}" by (case_tac bl, auto)
+    }
+    moreover
+    {fix a assume *: "a \<in> {a1,a2}"
+     have "a \<in> ?f ` UNIV"
+     proof(cases "a = a1")
+       assume "a = a1"
+       hence "?f True = a" by auto  thus ?thesis by blast
+     next
+       assume "a \<noteq> a1" hence "a = a2" using * by auto
+       hence "?f False = a" by auto  thus ?thesis by blast
+     qed
+    }
+    ultimately show ?thesis unfolding bij_betw_def inj_on_def
+    by (metis image_subsetI order_eq_iff subsetI)
+  qed
+  thus ?thesis using card_of_ordIso by blast
+qed
+
+
+lemma card_of_Plus_Times_aux:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+        LEQ: "|A| \<le>o |B|"
+shows "|A <+> B| \<le>o |A \<times> B|"
+proof-
+  have 1: "|UNIV::bool set| \<le>o |A|"
+  using A2 card_of_mono1[of "{a1,a2}"] card_of_bool[of a1 a2]
+        ordIso_ordLeq_trans[of "|UNIV::bool set|"] by metis
+  (*  *)
+  have "|A <+> B| \<le>o |B <+> B|"
+  using LEQ card_of_Plus_mono1 by blast
+  moreover have "|B <+> B| =o |B \<times> (UNIV::bool set)|"
+  using card_of_Plus_Times_bool by blast
+  moreover have "|B \<times> (UNIV::bool set)| \<le>o |B \<times> A|"
+  using 1 by (simp add: card_of_Times_mono2)
+  moreover have " |B \<times> A| =o |A \<times> B|"
+  using card_of_Times_commute by blast
+  ultimately show "|A <+> B| \<le>o |A \<times> B|"
+  using ordLeq_ordIso_trans[of "|A <+> B|" "|B <+> B|" "|B \<times> (UNIV::bool set)|"]
+        ordLeq_transitive[of "|A <+> B|" "|B \<times> (UNIV::bool set)|" "|B \<times> A|"]
+        ordLeq_ordIso_trans[of "|A <+> B|" "|B \<times> A|" "|A \<times> B|"]
+  by blast
+qed
+
+
+lemma card_of_Plus_Times:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+        B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
+shows "|A <+> B| \<le>o |A \<times> B|"
+proof-
+  {assume "|A| \<le>o |B|"
+   hence ?thesis using assms by (auto simp add: card_of_Plus_Times_aux)
+  }
+  moreover
+  {assume "|B| \<le>o |A|"
+   hence "|B <+> A| \<le>o |B \<times> A|"
+   using assms by (auto simp add: card_of_Plus_Times_aux)
+   hence ?thesis
+   using card_of_Plus_commute card_of_Times_commute
+         ordIso_ordLeq_trans ordLeq_ordIso_trans by metis
+  }
+  ultimately show ?thesis
+  using card_of_Well_order[of A] card_of_Well_order[of B]
+        ordLeq_total[of "|A|"] by metis
+qed
+
+
+lemma card_of_ordLeq_finite:
+assumes "|A| \<le>o |B|" and "finite B"
+shows "finite A"
+using assms unfolding ordLeq_def
+using embed_inj_on[of "|A|" "|B|"]  embed_Field[of "|A|" "|B|"]
+      Field_card_of[of "A"] Field_card_of[of "B"] inj_on_finite[of _ "A" "B"] by fastforce
+
+
+lemma card_of_ordLeq_infinite:
+assumes "|A| \<le>o |B|" and "infinite A"
+shows "infinite B"
+using assms card_of_ordLeq_finite by auto
+
+
+lemma card_of_ordIso_finite:
+assumes "|A| =o |B|"
+shows "finite A = finite B"
+using assms unfolding ordIso_def iso_def[abs_def]
+by (auto simp: bij_betw_finite Field_card_of)
+
+
+lemma card_of_ordIso_finite_Field:
+assumes "Card_order r" and "r =o |A|"
+shows "finite(Field r) = finite A"
+using assms card_of_Field_ordIso card_of_ordIso_finite ordIso_equivalence by blast
+
+
+subsection {* Cardinals versus set operations involving infinite sets *}
+
+
+text{* Here we show that, for infinite sets, most set-theoretic constructions
+do not increase the cardinality.  The cornerstone for this is
+theorem @{text "Card_order_Times_same_infinite"}, which states that self-product
+does not increase cardinality -- the proof of this fact adapts a standard
+set-theoretic argument, as presented, e.g., in the proof of theorem 1.5.11
+at page 47 in \cite{card-book}. Then everything else follows fairly easily.  *}
+
+
+lemma infinite_iff_card_of_nat:
+"infinite A = ( |UNIV::nat set| \<le>o |A| )"
+by (auto simp add: infinite_iff_countable_subset card_of_ordLeq)
+
+
+lemma finite_iff_cardOf_nat:
+"finite A = ( |A| <o |UNIV :: nat set| )"
+using infinite_iff_card_of_nat[of A]
+not_ordLeq_iff_ordLess[of "|A|" "|UNIV :: nat set|"]
+by (fastforce simp: card_of_Well_order)
+
+lemma finite_ordLess_infinite2:
+assumes "finite A" and "infinite B"
+shows "|A| <o |B|"
+using assms
+finite_ordLess_infinite[of "|A|" "|B|"]
+card_of_Well_order[of A] card_of_Well_order[of B]
+Field_card_of[of A] Field_card_of[of B] by auto
+
+
+text{* The next two results correspond to the ZF fact that all infinite cardinals are
+limit ordinals: *}
+
+lemma Card_order_infinite_not_under:
+assumes CARD: "Card_order r" and INF: "infinite (Field r)"
+shows "\<not> (\<exists>a. Field r = rel.under r a)"
+proof(auto)
+  have 0: "Well_order r \<and> wo_rel r \<and> Refl r"
+  using CARD unfolding wo_rel_def card_order_on_def order_on_defs by auto
+  fix a assume *: "Field r = rel.under r a"
+  show False
+  proof(cases "a \<in> Field r")
+    assume Case1: "a \<notin> Field r"
+    hence "rel.under r a = {}" unfolding Field_def rel.under_def by auto
+    thus False using INF *  by auto
+  next
+    let ?r' = "Restr r (rel.underS r a)"
+    assume Case2: "a \<in> Field r"
+    hence 1: "rel.under r a = rel.underS r a \<union> {a} \<and> a \<notin> rel.underS r a"
+    using 0 rel.Refl_under_underS rel.underS_notIn by fastforce
+    have 2: "wo_rel.ofilter r (rel.underS r a) \<and> rel.underS r a < Field r"
+    using 0 wo_rel.underS_ofilter * 1 Case2 by auto
+    hence "?r' <o r" using 0 using ofilter_ordLess by blast
+    moreover
+    have "Field ?r' = rel.underS r a \<and> Well_order ?r'"
+    using  2 0 Field_Restr_ofilter[of r] Well_order_Restr[of r] by blast
+    ultimately have "|rel.underS r a| <o r" using ordLess_Field[of ?r'] by auto
+    moreover have "|rel.under r a| =o r" using * CARD card_of_Field_ordIso[of r] by auto
+    ultimately have "|rel.underS r a| <o |rel.under r a|"
+    using ordIso_symmetric ordLess_ordIso_trans by blast
+    moreover
+    {have "\<exists>f. bij_betw f (rel.under r a) (rel.underS r a)"
+     using infinite_imp_bij_betw[of "Field r" a] INF * 1 by auto
+     hence "|rel.under r a| =o |rel.underS r a|" using card_of_ordIso by blast
+    }
+    ultimately show False using not_ordLess_ordIso ordIso_symmetric by blast
+  qed
+qed
+
+
+lemma infinite_Card_order_limit:
+assumes r: "Card_order r" and "infinite (Field r)"
+and a: "a : Field r"
+shows "EX b : Field r. a \<noteq> b \<and> (a,b) : r"
+proof-
+  have "Field r \<noteq> rel.under r a"
+  using assms Card_order_infinite_not_under by blast
+  moreover have "rel.under r a \<le> Field r"
+  using rel.under_Field .
+  ultimately have "rel.under r a < Field r" by blast
+  then obtain b where 1: "b : Field r \<and> ~ (b,a) : r"
+  unfolding rel.under_def by blast
+  moreover have ba: "b \<noteq> a"
+  using 1 r unfolding card_order_on_def well_order_on_def
+  linear_order_on_def partial_order_on_def preorder_on_def refl_on_def by auto
+  ultimately have "(a,b) : r"
+  using a r unfolding card_order_on_def well_order_on_def linear_order_on_def
+  total_on_def by blast
+  thus ?thesis using 1 ba by auto
+qed
+
+
+theorem Card_order_Times_same_infinite:
+assumes CO: "Card_order r" and INF: "infinite(Field r)"
+shows "|Field r \<times> Field r| \<le>o r"
+proof-
+  obtain phi where phi_def:
+  "phi = (\<lambda>r::'a rel. Card_order r \<and> infinite(Field r) \<and>
+                      \<not> |Field r \<times> Field r| \<le>o r )" by blast
+  have temp1: "\<forall>r. phi r \<longrightarrow> Well_order r"
+  unfolding phi_def card_order_on_def by auto
+  have Ft: "\<not>(\<exists>r. phi r)"
+  proof
+    assume "\<exists>r. phi r"
+    hence "{r. phi r} \<noteq> {} \<and> {r. phi r} \<le> {r. Well_order r}"
+    using temp1 by auto
+    then obtain r where 1: "phi r" and 2: "\<forall>r'. phi r' \<longrightarrow> r \<le>o r'" and
+                   3: "Card_order r \<and> Well_order r"
+    using exists_minim_Well_order[of "{r. phi r}"] temp1 phi_def by blast
+    let ?A = "Field r"  let ?r' = "bsqr r"
+    have 4: "Well_order ?r' \<and> Field ?r' = ?A \<times> ?A \<and> |?A| =o r"
+    using 3 bsqr_Well_order Field_bsqr card_of_Field_ordIso by blast
+    have 5: "Card_order |?A \<times> ?A| \<and> Well_order |?A \<times> ?A|"
+    using card_of_Card_order card_of_Well_order by blast
+    (*  *)
+    have "r <o |?A \<times> ?A|"
+    using 1 3 5 ordLess_or_ordLeq unfolding phi_def by blast
+    moreover have "|?A \<times> ?A| \<le>o ?r'"
+    using card_of_least[of "?A \<times> ?A"] 4 by auto
+    ultimately have "r <o ?r'" using ordLess_ordLeq_trans by auto
+    then obtain f where 6: "embed r ?r' f" and 7: "\<not> bij_betw f ?A (?A \<times> ?A)"
+    unfolding ordLess_def embedS_def[abs_def]
+    by (auto simp add: Field_bsqr)
+    let ?B = "f ` ?A"
+    have "|?A| =o |?B|"
+    using 3 6 embed_inj_on inj_on_imp_bij_betw card_of_ordIso by blast
+    hence 8: "r =o |?B|" using 4 ordIso_transitive ordIso_symmetric by blast
+    (*  *)
+    have "wo_rel.ofilter ?r' ?B"
+    using 6 embed_Field_ofilter 3 4 by blast
+    hence "wo_rel.ofilter ?r' ?B \<and> ?B \<noteq> ?A \<times> ?A \<and> ?B \<noteq> Field ?r'"
+    using 7 unfolding bij_betw_def using 6 3 embed_inj_on 4 by auto
+    hence temp2: "wo_rel.ofilter ?r' ?B \<and> ?B < ?A \<times> ?A"
+    using 4 wo_rel_def[of ?r'] wo_rel.ofilter_def[of ?r' ?B] by blast
+    have "\<not> (\<exists>a. Field r = rel.under r a)"
+    using 1 unfolding phi_def using Card_order_infinite_not_under[of r] by auto
+    then obtain A1 where temp3: "wo_rel.ofilter r A1 \<and> A1 < ?A" and 9: "?B \<le> A1 \<times> A1"
+    using temp2 3 bsqr_ofilter[of r ?B] by blast
+    hence "|?B| \<le>o |A1 \<times> A1|" using card_of_mono1 by blast
+    hence 10: "r \<le>o |A1 \<times> A1|" using 8 ordIso_ordLeq_trans by blast
+    let ?r1 = "Restr r A1"
+    have "?r1 <o r" using temp3 ofilter_ordLess 3 by blast
+    moreover
+    {have "well_order_on A1 ?r1" using 3 temp3 well_order_on_Restr by blast
+     hence "|A1| \<le>o ?r1" using 3 Well_order_Restr card_of_least by blast
+    }
+    ultimately have 11: "|A1| <o r" using ordLeq_ordLess_trans by blast
+    (*  *)
+    have "infinite (Field r)" using 1 unfolding phi_def by simp
+    hence "infinite ?B" using 8 3 card_of_ordIso_finite_Field[of r ?B] by blast
+    hence "infinite A1" using 9 infinite_super finite_cartesian_product by blast
+    moreover have temp4: "Field |A1| = A1 \<and> Well_order |A1| \<and> Card_order |A1|"
+    using card_of_Card_order[of A1] card_of_Well_order[of A1]
+    by (simp add: Field_card_of)
+    moreover have "\<not> r \<le>o | A1 |"
+    using temp4 11 3 using not_ordLeq_iff_ordLess by blast
+    ultimately have "infinite(Field |A1| ) \<and> Card_order |A1| \<and> \<not> r \<le>o | A1 |"
+    by (simp add: card_of_card_order_on)
+    hence "|Field |A1| \<times> Field |A1| | \<le>o |A1|"
+    using 2 unfolding phi_def by blast
+    hence "|A1 \<times> A1 | \<le>o |A1|" using temp4 by auto
+    hence "r \<le>o |A1|" using 10 ordLeq_transitive by blast
+    thus False using 11 not_ordLess_ordLeq by auto
+  qed
+  thus ?thesis using assms unfolding phi_def by blast
+qed
+
+
+corollary card_of_Times_same_infinite:
+assumes "infinite A"
+shows "|A \<times> A| =o |A|"
+proof-
+  let ?r = "|A|"
+  have "Field ?r = A \<and> Card_order ?r"
+  using Field_card_of card_of_Card_order[of A] by fastforce
+  hence "|A \<times> A| \<le>o |A|"
+  using Card_order_Times_same_infinite[of ?r] assms by auto
+  thus ?thesis using card_of_Times3 ordIso_iff_ordLeq by blast
+qed
+
+
+lemma card_of_Times_infinite:
+assumes INF: "infinite A" and NE: "B \<noteq> {}" and LEQ: "|B| \<le>o |A|"
+shows "|A \<times> B| =o |A| \<and> |B \<times> A| =o |A|"
+proof-
+  have "|A| \<le>o |A \<times> B| \<and> |A| \<le>o |B \<times> A|"
+  using assms by (simp add: card_of_Times1 card_of_Times2)
+  moreover
+  {have "|A \<times> B| \<le>o |A \<times> A| \<and> |B \<times> A| \<le>o |A \<times> A|"
+   using LEQ card_of_Times_mono1 card_of_Times_mono2 by blast
+   moreover have "|A \<times> A| =o |A|" using INF card_of_Times_same_infinite by blast
+   ultimately have "|A \<times> B| \<le>o |A| \<and> |B \<times> A| \<le>o |A|"
+   using ordLeq_ordIso_trans[of "|A \<times> B|"] ordLeq_ordIso_trans[of "|B \<times> A|"] by auto
+  }
+  ultimately show ?thesis by (simp add: ordIso_iff_ordLeq)
+qed
+
+
+corollary card_of_Times_infinite_simps:
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A \<times> B| =o |A|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |A \<times> B|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |B \<times> A| =o |A|"
+"\<lbrakk>infinite A; B \<noteq> {}; |B| \<le>o |A|\<rbrakk> \<Longrightarrow> |A| =o |B \<times> A|"
+by (auto simp add: card_of_Times_infinite ordIso_symmetric)
+
+
+corollary Card_order_Times_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+        NE: "Field p \<noteq> {}" and LEQ: "p \<le>o r"
+shows "| (Field r) \<times> (Field p) | =o r \<and> | (Field p) \<times> (Field r) | =o r"
+proof-
+  have "|Field r \<times> Field p| =o |Field r| \<and> |Field p \<times> Field r| =o |Field r|"
+  using assms by (simp add: card_of_Times_infinite card_of_mono2)
+  thus ?thesis
+  using assms card_of_Field_ordIso[of r]
+        ordIso_transitive[of "|Field r \<times> Field p|"]
+        ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+
+lemma card_of_Sigma_ordLeq_infinite:
+assumes INF: "infinite B" and
+        LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
+shows "|SIGMA i : I. A i| \<le>o |B|"
+proof(cases "I = {}", simp add: card_of_empty)
+  assume *: "I \<noteq> {}"
+  have "|SIGMA i : I. A i| \<le>o |I \<times> B|"
+  using LEQ card_of_Sigma_Times by blast
+  moreover have "|I \<times> B| =o |B|"
+  using INF * LEQ_I by (auto simp add: card_of_Times_infinite)
+  ultimately show ?thesis using ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Sigma_ordLeq_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+        LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
+shows "|SIGMA i : I. A i| \<le>o r"
+proof-
+  let ?B  = "Field r"
+  have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
+  ordIso_symmetric by blast
+  hence "|I| \<le>o |?B|"  "\<forall>i \<in> I. |A i| \<le>o |?B|"
+  using LEQ_I LEQ ordLeq_ordIso_trans by blast+
+  hence  "|SIGMA i : I. A i| \<le>o |?B|" using INF LEQ
+  card_of_Sigma_ordLeq_infinite by blast
+  thus ?thesis using 1 ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Times_ordLeq_infinite_Field:
+"\<lbrakk>infinite (Field r); |A| \<le>o r; |B| \<le>o r; Card_order r\<rbrakk>
+ \<Longrightarrow> |A <*> B| \<le>o r"
+by(simp add: card_of_Sigma_ordLeq_infinite_Field)
+
+
+lemma card_of_UNION_ordLeq_infinite:
+assumes INF: "infinite B" and
+        LEQ_I: "|I| \<le>o |B|" and LEQ: "\<forall>i \<in> I. |A i| \<le>o |B|"
+shows "|\<Union> i \<in> I. A i| \<le>o |B|"
+proof(cases "I = {}", simp add: card_of_empty)
+  assume *: "I \<noteq> {}"
+  have "|\<Union> i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
+  using card_of_UNION_Sigma by blast
+  moreover have "|SIGMA i : I. A i| \<le>o |B|"
+  using assms card_of_Sigma_ordLeq_infinite by blast
+  ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+corollary card_of_UNION_ordLeq_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+        LEQ_I: "|I| \<le>o r" and LEQ: "\<forall>i \<in> I. |A i| \<le>o r"
+shows "|\<Union> i \<in> I. A i| \<le>o r"
+proof-
+  let ?B  = "Field r"
+  have 1: "r =o |?B| \<and> |?B| =o r" using r card_of_Field_ordIso
+  ordIso_symmetric by blast
+  hence "|I| \<le>o |?B|"  "\<forall>i \<in> I. |A i| \<le>o |?B|"
+  using LEQ_I LEQ ordLeq_ordIso_trans by blast+
+  hence  "|\<Union> i \<in> I. A i| \<le>o |?B|" using INF LEQ
+  card_of_UNION_ordLeq_infinite by blast
+  thus ?thesis using 1 ordLeq_ordIso_trans by blast
+qed
+
+
+lemma card_of_Plus_infinite1:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A <+> B| =o |A|"
+proof(cases "B = {}", simp add: card_of_Plus_empty1 card_of_Plus_empty2 ordIso_symmetric)
+  let ?Inl = "Inl::'a \<Rightarrow> 'a + 'b"  let ?Inr = "Inr::'b \<Rightarrow> 'a + 'b"
+  assume *: "B \<noteq> {}"
+  then obtain b1 where 1: "b1 \<in> B" by blast
+  show ?thesis
+  proof(cases "B = {b1}")
+    assume Case1: "B = {b1}"
+    have 2: "bij_betw ?Inl A ((?Inl ` A))"
+    unfolding bij_betw_def inj_on_def by auto
+    hence 3: "infinite (?Inl ` A)"
+    using INF bij_betw_finite[of ?Inl A] by blast
+    let ?A' = "?Inl ` A \<union> {?Inr b1}"
+    obtain g where "bij_betw g (?Inl ` A) ?A'"
+    using 3 infinite_imp_bij_betw2[of "?Inl ` A"] by auto
+    moreover have "?A' = A <+> B" using Case1 by blast
+    ultimately have "bij_betw g (?Inl ` A) (A <+> B)" by simp
+    hence "bij_betw (g o ?Inl) A (A <+> B)"
+    using 2 by (auto simp add: bij_betw_trans)
+    thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+  next
+    assume Case2: "B \<noteq> {b1}"
+    with * 1 obtain b2 where 3: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B" by fastforce
+    obtain f where "inj_on f B \<and> f ` B \<le> A"
+    using LEQ card_of_ordLeq[of B] by fastforce
+    with 3 have "f b1 \<noteq> f b2 \<and> {f b1, f b2} \<le> A"
+    unfolding inj_on_def by auto
+    with 3 have "|A <+> B| \<le>o |A \<times> B|"
+    by (auto simp add: card_of_Plus_Times)
+    moreover have "|A \<times> B| =o |A|"
+    using assms * by (simp add: card_of_Times_infinite_simps)
+    ultimately have "|A <+> B| \<le>o |A|" using ordLeq_ordIso_trans by metis
+    thus ?thesis using card_of_Plus1 ordIso_iff_ordLeq by blast
+  qed
+qed
+
+
+lemma card_of_Plus_infinite2:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|B <+> A| =o |A|"
+using assms card_of_Plus_commute card_of_Plus_infinite1
+ordIso_equivalence by blast
+
+
+lemma card_of_Plus_infinite:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A <+> B| =o |A| \<and> |B <+> A| =o |A|"
+using assms by (auto simp: card_of_Plus_infinite1 card_of_Plus_infinite2)
+
+
+corollary Card_order_Plus_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+        LEQ: "p \<le>o r"
+shows "| (Field r) <+> (Field p) | =o r \<and> | (Field p) <+> (Field r) | =o r"
+proof-
+  have "| Field r <+> Field p | =o | Field r | \<and>
+        | Field p <+> Field r | =o | Field r |"
+  using assms by (simp add: card_of_Plus_infinite card_of_mono2)
+  thus ?thesis
+  using assms card_of_Field_ordIso[of r]
+        ordIso_transitive[of "|Field r <+> Field p|"]
+        ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+
+lemma card_of_Un_infinite:
+assumes INF: "infinite A" and LEQ: "|B| \<le>o |A|"
+shows "|A \<union> B| =o |A| \<and> |B \<union> A| =o |A|"
+proof-
+  have "|A \<union> B| \<le>o |A <+> B|" by (rule card_of_Un_Plus_ordLeq)
+  moreover have "|A <+> B| =o |A|"
+  using assms by (metis card_of_Plus_infinite)
+  ultimately have "|A \<union> B| \<le>o |A|" using ordLeq_ordIso_trans by blast
+  hence "|A \<union> B| =o |A|" using card_of_Un1 ordIso_iff_ordLeq by blast
+  thus ?thesis using Un_commute[of B A] by auto
+qed
+
+
+lemma card_of_Un_diff_infinite:
+assumes INF: "infinite A" and LESS: "|B| <o |A|"
+shows "|A - B| =o |A|"
+proof-
+  obtain C where C_def: "C = A - B" by blast
+  have "|A \<union> B| =o |A|"
+  using assms ordLeq_iff_ordLess_or_ordIso card_of_Un_infinite by blast
+  moreover have "C \<union> B = A \<union> B" unfolding C_def by auto
+  ultimately have 1: "|C \<union> B| =o |A|" by auto
+  (*  *)
+  {assume *: "|C| \<le>o |B|"
+   moreover
+   {assume **: "finite B"
+    hence "finite C"
+    using card_of_ordLeq_finite * by blast
+    hence False using ** INF card_of_ordIso_finite 1 by blast
+   }
+   hence "infinite B" by auto
+   ultimately have False
+   using card_of_Un_infinite 1 ordIso_equivalence(1,3) LESS not_ordLess_ordIso by metis
+  }
+  hence 2: "|B| \<le>o |C|" using card_of_Well_order ordLeq_total by blast
+  {assume *: "finite C"
+    hence "finite B" using card_of_ordLeq_finite 2 by blast
+    hence False using * INF card_of_ordIso_finite 1 by blast
+  }
+  hence "infinite C" by auto
+  hence "|C| =o |A|"
+  using  card_of_Un_infinite 1 2 ordIso_equivalence(1,3) by metis
+  thus ?thesis unfolding C_def .
+qed
+
+
+lemma card_of_Plus_ordLess_infinite:
+assumes INF: "infinite C" and
+        LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A <+> B| <o |C|"
+proof(cases "A = {} \<or> B = {}")
+  assume Case1: "A = {} \<or> B = {}"
+  hence "|A| =o |A <+> B| \<or> |B| =o |A <+> B|"
+  using card_of_Plus_empty1 card_of_Plus_empty2 by blast
+  hence "|A <+> B| =o |A| \<or> |A <+> B| =o |B|"
+  using ordIso_symmetric[of "|A|"] ordIso_symmetric[of "|B|"] by blast
+  thus ?thesis using LESS1 LESS2
+       ordIso_ordLess_trans[of "|A <+> B|" "|A|"]
+       ordIso_ordLess_trans[of "|A <+> B|" "|B|"] by blast
+next
+  assume Case2: "\<not>(A = {} \<or> B = {})"
+  {assume *: "|C| \<le>o |A <+> B|"
+   hence "infinite (A <+> B)" using INF card_of_ordLeq_finite by blast
+   hence 1: "infinite A \<or> infinite B" using finite_Plus by blast
+   {assume Case21: "|A| \<le>o |B|"
+    hence "infinite B" using 1 card_of_ordLeq_finite by blast
+    hence "|A <+> B| =o |B|" using Case2 Case21
+    by (auto simp add: card_of_Plus_infinite)
+    hence False using LESS2 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+   }
+   moreover
+   {assume Case22: "|B| \<le>o |A|"
+    hence "infinite A" using 1 card_of_ordLeq_finite by blast
+    hence "|A <+> B| =o |A|" using Case2 Case22
+    by (auto simp add: card_of_Plus_infinite)
+    hence False using LESS1 not_ordLess_ordLeq * ordLeq_ordIso_trans by blast
+   }
+   ultimately have False using ordLeq_total card_of_Well_order[of A]
+   card_of_Well_order[of B] by blast
+  }
+  thus ?thesis using ordLess_or_ordLeq[of "|A <+> B|" "|C|"]
+  card_of_Well_order[of "A <+> B"] card_of_Well_order[of "C"] by auto
+qed
+
+
+lemma card_of_Plus_ordLess_infinite_Field:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+        LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A <+> B| <o r"
+proof-
+  let ?C  = "Field r"
+  have 1: "r =o |?C| \<and> |?C| =o r" using r card_of_Field_ordIso
+  ordIso_symmetric by blast
+  hence "|A| <o |?C|"  "|B| <o |?C|"
+  using LESS1 LESS2 ordLess_ordIso_trans by blast+
+  hence  "|A <+> B| <o |?C|" using INF
+  card_of_Plus_ordLess_infinite by blast
+  thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+
+lemma infinite_card_of_insert:
+assumes "infinite A"
+shows "|insert a A| =o |A|"
+proof-
+  have iA: "insert a A = A \<union> {a}" by simp
+  show ?thesis
+  using infinite_imp_bij_betw2[OF assms] unfolding iA
+  by (metis bij_betw_inv card_of_ordIso)
+qed
+
+
+subsection {* Cardinals versus lists  *}
+
+
+text{* The next is an auxiliary operator, which shall be used for inductive
+proofs of facts concerning the cardinality of @{text "List"} : *}
+
+definition nlists :: "'a set \<Rightarrow> nat \<Rightarrow> 'a list set"
+where "nlists A n \<equiv> {l. set l \<le> A \<and> length l = n}"
+
+
+lemma lists_def2: "lists A = {l. set l \<le> A}"
+using in_listsI by blast
+
+
+lemma lists_UNION_nlists: "lists A = (\<Union> n. nlists A n)"
+unfolding lists_def2 nlists_def by blast
+
+
+lemma card_of_lists: "|A| \<le>o |lists A|"
+proof-
+  let ?h = "\<lambda> a. [a]"
+  have "inj_on ?h A \<and> ?h ` A \<le> lists A"
+  unfolding inj_on_def lists_def2 by auto
+  thus ?thesis by (metis card_of_ordLeq)
+qed
+
+
+lemma nlists_0: "nlists A 0 = {[]}"
+unfolding nlists_def by auto
+
+
+lemma nlists_not_empty:
+assumes "A \<noteq> {}"
+shows "nlists A n \<noteq> {}"
+proof(induct n, simp add: nlists_0)
+  fix n assume "nlists A n \<noteq> {}"
+  then obtain a and l where "a \<in> A \<and> l \<in> nlists A n" using assms by auto
+  hence "a # l \<in> nlists A (Suc n)" unfolding nlists_def by auto
+  thus "nlists A (Suc n) \<noteq> {}" by auto
+qed
+
+
+lemma Nil_in_lists: "[] \<in> lists A"
+unfolding lists_def2 by auto
+
+
+lemma lists_not_empty: "lists A \<noteq> {}"
+using Nil_in_lists by blast
+
+
+lemma card_of_nlists_Succ: "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
+proof-
+  let ?B = "A \<times> (nlists A n)"   let ?h = "\<lambda>(a,l). a # l"
+  have "inj_on ?h ?B \<and> ?h ` ?B \<le> nlists A (Suc n)"
+  unfolding inj_on_def nlists_def by auto
+  moreover have "nlists A (Suc n) \<le> ?h ` ?B"
+  proof(auto)
+    fix l assume "l \<in> nlists A (Suc n)"
+    hence 1: "length l = Suc n \<and> set l \<le> A" unfolding nlists_def by auto
+    then obtain a and l' where 2: "l = a # l'" by (auto simp: length_Suc_conv)
+    hence "a \<in> A \<and> set l' \<le> A \<and> length l' = n" using 1 by auto
+    thus "l \<in> ?h ` ?B"  using 2 unfolding nlists_def by auto
+  qed
+  ultimately have "bij_betw ?h ?B (nlists A (Suc n))"
+  unfolding bij_betw_def by auto
+  thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+
+lemma card_of_nlists_infinite:
+assumes "infinite A"
+shows "|nlists A n| \<le>o |A|"
+proof(induct n)
+  have "A \<noteq> {}" using assms by auto
+  thus "|nlists A 0| \<le>o |A|" by (simp add: nlists_0 card_of_singl_ordLeq)
+next
+  fix n assume IH: "|nlists A n| \<le>o |A|"
+  have "|nlists A (Suc n)| =o |A \<times> (nlists A n)|"
+  using card_of_nlists_Succ by blast
+  moreover
+  {have "nlists A n \<noteq> {}" using assms nlists_not_empty[of A] by blast
+   hence "|A \<times> (nlists A n)| =o |A|"
+   using assms IH by (auto simp add: card_of_Times_infinite)
+  }
+  ultimately show "|nlists A (Suc n)| \<le>o |A|"
+  using ordIso_transitive ordIso_iff_ordLeq by blast
+qed
+
+
+lemma card_of_lists_infinite:
+assumes "infinite A"
+shows "|lists A| =o |A|"
+proof-
+  have "|lists A| \<le>o |A|"
+  using assms
+  by (auto simp add: lists_UNION_nlists card_of_UNION_ordLeq_infinite
+                     infinite_iff_card_of_nat card_of_nlists_infinite)
+  thus ?thesis using card_of_lists ordIso_iff_ordLeq by blast
+qed
+
+
+lemma Card_order_lists_infinite:
+assumes "Card_order r" and "infinite(Field r)"
+shows "|lists(Field r)| =o r"
+using assms card_of_lists_infinite card_of_Field_ordIso ordIso_transitive by blast
+
+
+
+subsection {* The cardinal $\omega$ and the finite cardinals  *}
+
+
+text{* The cardinal $\omega$, of natural numbers, shall be the standard non-strict
+order relation on
+@{text "nat"}, that we abbreviate by @{text "natLeq"}.  The finite cardinals
+shall be the restrictions of these relations to the numbers smaller than
+fixed numbers @{text "n"}, that we abbreviate by @{text "natLeq_on n"}.  *}
+
+abbreviation "(natLeq::(nat * nat) set) \<equiv> {(x,y). x \<le> y}"
+abbreviation "(natLess::(nat * nat) set) \<equiv> {(x,y). x < y}"
+
+abbreviation natLeq_on :: "nat \<Rightarrow> (nat * nat) set"
+where "natLeq_on n \<equiv> {(x,y). x < n \<and> y < n \<and> x \<le> y}"
+
+lemma infinite_cartesian_product:
+assumes "infinite A" "infinite B"
+shows "infinite (A \<times> B)"
+proof
+  assume "finite (A \<times> B)"
+  from assms(1) have "A \<noteq> {}" by auto
+  with `finite (A \<times> B)` have "finite B" using finite_cartesian_productD2 by auto
+  with assms(2) show False by simp
+qed
+
+
+
+subsubsection {* First as well-orders *}
+
+
+lemma Field_natLeq: "Field natLeq = (UNIV::nat set)"
+by(unfold Field_def, auto)
+
+
+lemma natLeq_Refl: "Refl natLeq"
+unfolding refl_on_def Field_def by auto
+
+
+lemma natLeq_trans: "trans natLeq"
+unfolding trans_def by auto
+
+
+lemma natLeq_Preorder: "Preorder natLeq"
+unfolding preorder_on_def
+by (auto simp add: natLeq_Refl natLeq_trans)
+
+
+lemma natLeq_antisym: "antisym natLeq"
+unfolding antisym_def by auto
+
+
+lemma natLeq_Partial_order: "Partial_order natLeq"
+unfolding partial_order_on_def
+by (auto simp add: natLeq_Preorder natLeq_antisym)
+
+
+lemma natLeq_Total: "Total natLeq"
+unfolding total_on_def by auto
+
+
+lemma natLeq_Linear_order: "Linear_order natLeq"
+unfolding linear_order_on_def
+by (auto simp add: natLeq_Partial_order natLeq_Total)
+
+
+lemma natLeq_natLess_Id: "natLess = natLeq - Id"
+by auto
+
+
+lemma natLeq_Well_order: "Well_order natLeq"
+unfolding well_order_on_def
+using natLeq_Linear_order wf_less natLeq_natLess_Id by auto
+
+
+corollary natLeq_well_order_on: "well_order_on UNIV natLeq"
+using natLeq_Well_order Field_natLeq by auto
+
+
+lemma natLeq_wo_rel: "wo_rel natLeq"
+unfolding wo_rel_def using natLeq_Well_order .
+
+
+lemma natLeq_UNIV_ofilter: "wo_rel.ofilter natLeq UNIV"
+using natLeq_wo_rel Field_natLeq wo_rel.Field_ofilter[of natLeq] by auto
+
+
+lemma closed_nat_set_iff:
+assumes "\<forall>(m::nat) n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
+shows "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
+proof-
+  {assume "A \<noteq> UNIV" hence "\<exists>n. n \<notin> A" by blast
+   moreover obtain n where n_def: "n = (LEAST n. n \<notin> A)" by blast
+   ultimately have 1: "n \<notin> A \<and> (\<forall>m. m < n \<longrightarrow> m \<in> A)"
+   using LeastI_ex[of "\<lambda> n. n \<notin> A"] n_def Least_le[of "\<lambda> n. n \<notin> A"] by fastforce
+   have "A = {0 ..< n}"
+   proof(auto simp add: 1)
+     fix m assume *: "m \<in> A"
+     {assume "n \<le> m" with assms * have "n \<in> A" by blast
+      hence False using 1 by auto
+     }
+     thus "m < n" by fastforce
+   qed
+   hence "\<exists>n. A = {0 ..< n}" by blast
+  }
+  thus ?thesis by blast
+qed
+
+
+lemma Field_natLeq_on: "Field (natLeq_on n) = {0 ..< n}"
+unfolding Field_def by auto
+
+
+lemma natLeq_underS_less: "rel.underS natLeq n = {0 ..< n}"
+unfolding rel.underS_def by auto
+
+
+lemma Restr_natLeq: "Restr natLeq {0 ..< n} = natLeq_on n"
+by auto
+
+
+lemma Restr_natLeq2:
+"Restr natLeq (rel.underS natLeq n) = natLeq_on n"
+by (auto simp add: Restr_natLeq natLeq_underS_less)
+
+
+lemma natLeq_on_Well_order: "Well_order(natLeq_on n)"
+using Restr_natLeq[of n] natLeq_Well_order
+      Well_order_Restr[of natLeq "{0..<n}"] by auto
+
+
+corollary natLeq_on_well_order_on: "well_order_on {0 ..< n} (natLeq_on n)"
+using natLeq_on_Well_order Field_natLeq_on by auto
+
+
+lemma natLeq_on_wo_rel: "wo_rel(natLeq_on n)"
+unfolding wo_rel_def using natLeq_on_Well_order .
+
+
+lemma natLeq_on_ofilter_less_eq:
+"n \<le> m \<Longrightarrow> wo_rel.ofilter (natLeq_on m) {0 ..< n}"
+by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def,
+    simp add: Field_natLeq_on, unfold rel.under_def, auto)
+
+
+lemma natLeq_on_ofilter_iff:
+"wo_rel.ofilter (natLeq_on m) A = (\<exists>n \<le> m. A = {0 ..< n})"
+proof(rule iffI)
+  assume *: "wo_rel.ofilter (natLeq_on m) A"
+  hence 1: "A \<le> {0..<m}"
+  by (auto simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def Field_natLeq_on)
+  hence "\<forall>n1 n2. n2 \<in> A \<and> n1 \<le> n2 \<longrightarrow> n1 \<in> A"
+  using * by(fastforce simp add: natLeq_on_wo_rel wo_rel.ofilter_def rel.under_def)
+  hence "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
+  thus "\<exists>n \<le> m. A = {0 ..< n}" using 1 atLeastLessThan_less_eq by blast
+next
+  assume "(\<exists>n\<le>m. A = {0 ..< n})"
+  thus "wo_rel.ofilter (natLeq_on m) A" by (auto simp add: natLeq_on_ofilter_less_eq)
+qed
+
+
+
+subsubsection {* Then as cardinals *}
+
+
+lemma natLeq_Card_order: "Card_order natLeq"
+proof(auto simp add: natLeq_Well_order
+      Card_order_iff_Restr_underS Restr_natLeq2, simp add:  Field_natLeq)
+  fix n have "finite(Field (natLeq_on n))"
+  unfolding Field_natLeq_on by auto
+  moreover have "infinite(UNIV::nat set)" by auto
+  ultimately show "natLeq_on n <o |UNIV::nat set|"
+  using finite_ordLess_infinite[of "natLeq_on n" "|UNIV::nat set|"]
+        Field_card_of[of "UNIV::nat set"]
+        card_of_Well_order[of "UNIV::nat set"] natLeq_on_Well_order[of n] by auto
+qed
+
+
+corollary card_of_Field_natLeq:
+"|Field natLeq| =o natLeq"
+using Field_natLeq natLeq_Card_order Card_order_iff_ordIso_card_of[of natLeq]
+      ordIso_symmetric[of natLeq] by blast
+
+
+corollary card_of_nat:
+"|UNIV::nat set| =o natLeq"
+using Field_natLeq card_of_Field_natLeq by auto
+
+
+corollary infinite_iff_natLeq_ordLeq:
+"infinite A = ( natLeq \<le>o |A| )"
+using infinite_iff_card_of_nat[of A] card_of_nat
+      ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
+
+
+corollary finite_iff_ordLess_natLeq:
+"finite A = ( |A| <o natLeq)"
+using infinite_iff_natLeq_ordLeq not_ordLeq_iff_ordLess
+      card_of_Well_order natLeq_Well_order by blast
+
+
+lemma ordIso_natLeq_on_imp_finite:
+"|A| =o natLeq_on n \<Longrightarrow> finite A"
+unfolding ordIso_def iso_def[abs_def]
+by (auto simp: Field_natLeq_on bij_betw_finite Field_card_of)
+
+
+lemma natLeq_on_Card_order: "Card_order (natLeq_on n)"
+proof(unfold card_order_on_def,
+      auto simp add: natLeq_on_Well_order, simp add: Field_natLeq_on)
+  fix r assume "well_order_on {0..<n} r"
+  thus "natLeq_on n \<le>o r"
+  using finite_atLeastLessThan natLeq_on_well_order_on
+        finite_well_order_on_ordIso ordIso_iff_ordLeq by blast
+qed
+
+
+corollary card_of_Field_natLeq_on:
+"|Field (natLeq_on n)| =o natLeq_on n"
+using Field_natLeq_on natLeq_on_Card_order
+      Card_order_iff_ordIso_card_of[of "natLeq_on n"]
+      ordIso_symmetric[of "natLeq_on n"] by blast
+
+
+corollary card_of_less:
+"|{0 ..< n}| =o natLeq_on n"
+using Field_natLeq_on card_of_Field_natLeq_on by auto
+
+
+lemma natLeq_on_ordLeq_less_eq:
+"((natLeq_on m) \<le>o (natLeq_on n)) = (m \<le> n)"
+proof
+  assume "natLeq_on m \<le>o natLeq_on n"
+  then obtain f where "inj_on f {0..<m} \<and> f ` {0..<m} \<le> {0..<n}"
+  unfolding ordLeq_def using
+    embed_inj_on[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on]
+     embed_Field[OF natLeq_on_Well_order[of m], of "natLeq_on n", unfolded Field_natLeq_on] by blast
+  thus "m \<le> n" using atLeastLessThan_less_eq2 by blast
+next
+  assume "m \<le> n"
+  hence "inj_on id {0..<m} \<and> id ` {0..<m} \<le> {0..<n}" unfolding inj_on_def by auto
+  hence "|{0..<m}| \<le>o |{0..<n}|" using card_of_ordLeq by blast
+  thus "natLeq_on m \<le>o natLeq_on n"
+  using card_of_less ordIso_ordLeq_trans ordLeq_ordIso_trans ordIso_symmetric by blast
+qed
+
+
+lemma natLeq_on_ordLeq_less:
+"((natLeq_on m) <o (natLeq_on n)) = (m < n)"
+using not_ordLeq_iff_ordLess[of "natLeq_on m" "natLeq_on n"]
+natLeq_on_Well_order natLeq_on_ordLeq_less_eq by auto
+
+
+
+subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
+
+
+lemma finite_card_of_iff_card2:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| \<le>o |B| ) = (card A \<le> card B)"
+using assms card_of_ordLeq[of A B] inj_on_iff_card_le[of A B] by blast
+
+
+lemma finite_imp_card_of_natLeq_on:
+assumes "finite A"
+shows "|A| =o natLeq_on (card A)"
+proof-
+  obtain h where "bij_betw h A {0 ..< card A}"
+  using assms ex_bij_betw_finite_nat by blast
+  thus ?thesis using card_of_ordIso card_of_less ordIso_equivalence by blast
+qed
+
+
+lemma finite_iff_card_of_natLeq_on:
+"finite A = (\<exists>n. |A| =o natLeq_on n)"
+using finite_imp_card_of_natLeq_on[of A]
+by(auto simp add: ordIso_natLeq_on_imp_finite)
+
+
+
+subsection {* The successor of a cardinal *}
+
+
+text{* First we define @{text "isCardSuc r r'"}, the notion of @{text "r'"}
+being a successor cardinal of @{text "r"}. Although the definition does
+not require @{text "r"} to be a cardinal, only this case will be meaningful.  *}
+
+
+definition isCardSuc :: "'a rel \<Rightarrow> 'a set rel \<Rightarrow> bool"
+where
+"isCardSuc r r' \<equiv>
+ Card_order r' \<and> r <o r' \<and>
+ (\<forall>(r''::'a set rel). Card_order r'' \<and> r <o r'' \<longrightarrow> r' \<le>o r'')"
+
+
+text{* Now we introduce the cardinal-successor operator @{text "cardSuc"},
+by picking {\em some} cardinal-order relation fulfilling @{text "isCardSuc"}.
+Again, the picked item shall be proved unique up to order-isomorphism. *}
+
+
+definition cardSuc :: "'a rel \<Rightarrow> 'a set rel"
+where
+"cardSuc r \<equiv> SOME r'. isCardSuc r r'"
+
+
+lemma exists_minim_Card_order:
+"\<lbrakk>R \<noteq> {}; \<forall>r \<in> R. Card_order r\<rbrakk> \<Longrightarrow> \<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+unfolding card_order_on_def using exists_minim_Well_order by blast
+
+
+lemma exists_isCardSuc:
+assumes "Card_order r"
+shows "\<exists>r'. isCardSuc r r'"
+proof-
+  let ?R = "{(r'::'a set rel). Card_order r' \<and> r <o r'}"
+  have "|Pow(Field r)| \<in> ?R \<and> (\<forall>r \<in> ?R. Card_order r)" using assms
+  by (simp add: card_of_Card_order Card_order_Pow)
+  then obtain r where "r \<in> ?R \<and> (\<forall>r' \<in> ?R. r \<le>o r')"
+  using exists_minim_Card_order[of ?R] by blast
+  thus ?thesis unfolding isCardSuc_def by auto
+qed
+
+
+lemma cardSuc_isCardSuc:
+assumes "Card_order r"
+shows "isCardSuc r (cardSuc r)"
+unfolding cardSuc_def using assms
+by (simp add: exists_isCardSuc someI_ex)
+
+
+lemma cardSuc_Card_order:
+"Card_order r \<Longrightarrow> Card_order(cardSuc r)"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+lemma cardSuc_greater:
+"Card_order r \<Longrightarrow> r <o cardSuc r"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+lemma cardSuc_ordLeq:
+"Card_order r \<Longrightarrow> r \<le>o cardSuc r"
+using cardSuc_greater ordLeq_iff_ordLess_or_ordIso by blast
+
+
+text{* The minimality property of @{text "cardSuc"} originally present in its definition
+is local to the type @{text "'a set rel"}, i.e., that of @{text "cardSuc r"}:  *}
+
+lemma cardSuc_least_aux:
+"\<lbrakk>Card_order (r::'a rel); Card_order (r'::'a set rel); r <o r'\<rbrakk> \<Longrightarrow> cardSuc r \<le>o r'"
+using cardSuc_isCardSuc unfolding isCardSuc_def by blast
+
+
+text{* But from this we can infer general minimality: *}
+
+lemma cardSuc_least:
+assumes CARD: "Card_order r" and CARD': "Card_order r'" and LESS: "r <o r'"
+shows "cardSuc r \<le>o r'"
+proof-
+  let ?p = "cardSuc r"
+  have 0: "Well_order ?p \<and> Well_order r'"
+  using assms cardSuc_Card_order unfolding card_order_on_def by blast
+  {assume "r' <o ?p"
+   then obtain r'' where 1: "Field r'' < Field ?p" and 2: "r' =o r'' \<and> r'' <o ?p"
+   using internalize_ordLess[of r' ?p] by blast
+   (*  *)
+   have "Card_order r''" using CARD' Card_order_ordIso2 2 by blast
+   moreover have "r <o r''" using LESS 2 ordLess_ordIso_trans by blast
+   ultimately have "?p \<le>o r''" using cardSuc_least_aux CARD by blast
+   hence False using 2 not_ordLess_ordLeq by blast
+  }
+  thus ?thesis using 0 ordLess_or_ordLeq by blast
+qed
+
+
+lemma cardSuc_ordLess_ordLeq:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(r <o r') = (cardSuc r \<le>o r')"
+proof(auto simp add: assms cardSuc_least)
+  assume "cardSuc r \<le>o r'"
+  thus "r <o r'" using assms cardSuc_greater ordLess_ordLeq_trans by blast
+qed
+
+
+lemma cardSuc_ordLeq_ordLess:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(r' <o cardSuc r) = (r' \<le>o r)"
+proof-
+  have "Well_order r \<and> Well_order r'"
+  using assms unfolding card_order_on_def by auto
+  moreover have "Well_order(cardSuc r)"
+  using assms cardSuc_Card_order card_order_on_def by blast
+  ultimately show ?thesis
+  using assms cardSuc_ordLess_ordLeq[of r r']
+  not_ordLeq_iff_ordLess[of r r'] not_ordLeq_iff_ordLess[of r' "cardSuc r"] by blast
+qed
+
+
+lemma cardSuc_mono_ordLeq:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(cardSuc r \<le>o cardSuc r') = (r \<le>o r')"
+using assms cardSuc_ordLeq_ordLess cardSuc_ordLess_ordLeq cardSuc_Card_order by blast
+
+
+lemma cardSuc_invar_ordIso:
+assumes CARD: "Card_order r" and CARD': "Card_order r'"
+shows "(cardSuc r =o cardSuc r') = (r =o r')"
+proof-
+  have 0: "Well_order r \<and> Well_order r' \<and> Well_order(cardSuc r) \<and> Well_order(cardSuc r')"
+  using assms by (simp add: card_order_on_well_order_on cardSuc_Card_order)
+  thus ?thesis
+  using ordIso_iff_ordLeq[of r r'] ordIso_iff_ordLeq
+  using cardSuc_mono_ordLeq[of r r'] cardSuc_mono_ordLeq[of r' r] assms by blast
+qed
+
+
+lemma cardSuc_natLeq_on_Suc:
+"cardSuc(natLeq_on n) =o natLeq_on(Suc n)"
+proof-
+  obtain r r' p where r_def: "r = natLeq_on n" and
+                      r'_def: "r' = cardSuc(natLeq_on n)"  and
+                      p_def: "p = natLeq_on(Suc n)" by blast
+  (* Preliminary facts:  *)
+  have CARD: "Card_order r \<and> Card_order r' \<and> Card_order p" unfolding r_def r'_def p_def
+  using cardSuc_ordLess_ordLeq natLeq_on_Card_order cardSuc_Card_order by blast
+  hence WELL: "Well_order r \<and> Well_order r' \<and>  Well_order p"
+  unfolding card_order_on_def by force
+  have FIELD: "Field r = {0..<n} \<and> Field p = {0..<(Suc n)}"
+  unfolding r_def p_def Field_natLeq_on by simp
+  hence FIN: "finite (Field r)" by force
+  have "r <o r'" using CARD unfolding r_def r'_def using cardSuc_greater by blast
+  hence "|Field r| <o r'" using CARD card_of_Field_ordIso ordIso_ordLess_trans by blast
+  hence LESS: "|Field r| <o |Field r'|"
+  using CARD card_of_Field_ordIso ordLess_ordIso_trans ordIso_symmetric by blast
+  (* Main proof: *)
+  have "r' \<le>o p" using CARD unfolding r_def r'_def p_def
+  using natLeq_on_ordLeq_less cardSuc_ordLess_ordLeq by blast
+  moreover have "p \<le>o r'"
+  proof-
+    {assume "r' <o p"
+     then obtain f where 0: "embedS r' p f" unfolding ordLess_def by force
+     let ?q = "Restr p (f ` Field r')"
+     have 1: "embed r' p f" using 0 unfolding embedS_def by force
+     hence 2: "f ` Field r' < {0..<(Suc n)}"
+     using WELL FIELD 0 by (auto simp add: embedS_iff)
+     have "wo_rel.ofilter p (f ` Field r')" using embed_Field_ofilter 1 WELL by blast
+     then obtain m where "m \<le> Suc n" and 3: "f ` (Field r') = {0..<m}"
+     unfolding p_def by (auto simp add: natLeq_on_ofilter_iff)
+     hence 4: "m \<le> n" using 2 by force
+     (*  *)
+     have "bij_betw f (Field r') (f ` (Field r'))"
+     using 1 WELL embed_inj_on unfolding bij_betw_def by force
+     moreover have "finite(f ` (Field r'))" using 3 finite_atLeastLessThan[of 0 m] by force
+     ultimately have 5: "finite (Field r') \<and> card(Field r') = card (f ` (Field r'))"
+     using bij_betw_same_card bij_betw_finite by metis
+     hence "card(Field r') \<le> card(Field r)" using 3 4 FIELD by force
+     hence "|Field r'| \<le>o |Field r|" using FIN 5 finite_card_of_iff_card2 by blast
+     hence False using LESS not_ordLess_ordLeq by auto
+    }
+    thus ?thesis using WELL CARD by (fastforce simp: not_ordLess_iff_ordLeq)
+  qed
+  ultimately show ?thesis using ordIso_iff_ordLeq unfolding r'_def p_def by blast
+qed
+
+
+lemma card_of_cardSuc_finite:
+"finite(Field(cardSuc |A| )) = finite A"
+proof
+  assume *: "finite (Field (cardSuc |A| ))"
+  have 0: "|Field(cardSuc |A| )| =o cardSuc |A|"
+  using card_of_Card_order cardSuc_Card_order card_of_Field_ordIso by blast
+  hence "|A| \<le>o |Field(cardSuc |A| )|"
+  using card_of_Card_order[of A] cardSuc_ordLeq[of "|A|"] ordIso_symmetric
+  ordLeq_ordIso_trans by blast
+  thus "finite A" using * card_of_ordLeq_finite by blast
+next
+  assume "finite A"
+  then obtain n where "|A| =o natLeq_on n" using finite_iff_card_of_natLeq_on by blast
+  hence "cardSuc |A| =o cardSuc(natLeq_on n)"
+  using card_of_Card_order cardSuc_invar_ordIso natLeq_on_Card_order by blast
+  hence "cardSuc |A| =o natLeq_on(Suc n)"
+  using cardSuc_natLeq_on_Suc ordIso_transitive by blast
+  hence "cardSuc |A| =o |{0..<(Suc n)}|" using card_of_less ordIso_equivalence by blast
+  moreover have "|Field (cardSuc |A| ) | =o cardSuc |A|"
+  using card_of_Field_ordIso cardSuc_Card_order card_of_Card_order by blast
+  ultimately have "|Field (cardSuc |A| ) | =o |{0..<(Suc n)}|"
+  using ordIso_equivalence by blast
+  thus "finite (Field (cardSuc |A| ))"
+  using card_of_ordIso_finite finite_atLeastLessThan by blast
+qed
+
+
+lemma cardSuc_finite:
+assumes "Card_order r"
+shows "finite (Field (cardSuc r)) = finite (Field r)"
+proof-
+  let ?A = "Field r"
+  have "|?A| =o r" using assms by (simp add: card_of_Field_ordIso)
+  hence "cardSuc |?A| =o cardSuc r" using assms
+  by (simp add: card_of_Card_order cardSuc_invar_ordIso)
+  moreover have "|Field (cardSuc |?A| ) | =o cardSuc |?A|"
+  by (simp add: card_of_card_order_on Field_card_of card_of_Field_ordIso cardSuc_Card_order)
+  moreover
+  {have "|Field (cardSuc r) | =o cardSuc r"
+   using assms by (simp add: card_of_Field_ordIso cardSuc_Card_order)
+   hence "cardSuc r =o |Field (cardSuc r) |"
+   using ordIso_symmetric by blast
+  }
+  ultimately have "|Field (cardSuc |?A| ) | =o |Field (cardSuc r) |"
+  using ordIso_transitive by blast
+  hence "finite (Field (cardSuc |?A| )) = finite (Field (cardSuc r))"
+  using card_of_ordIso_finite by blast
+  thus ?thesis by (simp only: card_of_cardSuc_finite)
+qed
+
+
+lemma card_of_Plus_ordLeq_infinite_Field:
+assumes r: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
+and c: "Card_order r"
+shows "|A <+> B| \<le>o r"
+proof-
+  let ?r' = "cardSuc r"
+  have "Card_order ?r' \<and> infinite (Field ?r')" using assms
+  by (simp add: cardSuc_Card_order cardSuc_finite)
+  moreover have "|A| <o ?r'" and "|B| <o ?r'" using A B c
+  by (auto simp: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
+  ultimately have "|A <+> B| <o ?r'"
+  using card_of_Plus_ordLess_infinite_Field by blast
+  thus ?thesis using c r
+  by (simp add: card_of_card_order_on Field_card_of cardSuc_ordLeq_ordLess)
+qed
+
+
+lemma card_of_Un_ordLeq_infinite_Field:
+assumes C: "infinite (Field r)" and A: "|A| \<le>o r" and B: "|B| \<le>o r"
+and "Card_order r"
+shows "|A Un B| \<le>o r"
+using assms card_of_Plus_ordLeq_infinite_Field card_of_Un_Plus_ordLeq
+ordLeq_transitive by blast
+
+
+
+subsection {* Regular cardinals *}
+
+
+definition cofinal where
+"cofinal A r \<equiv>
+ ALL a : Field r. EX b : A. a \<noteq> b \<and> (a,b) : r"
+
+
+definition regular where
+"regular r \<equiv>
+ ALL K. K \<le> Field r \<and> cofinal K r \<longrightarrow> |K| =o r"
+
+
+definition relChain where
+"relChain r As \<equiv>
+ ALL i j. (i,j) \<in> r \<longrightarrow> As i \<le> As j"
+
+lemma regular_UNION:
+assumes r: "Card_order r"   "regular r"
+and As: "relChain r As"
+and Bsub: "B \<le> (UN i : Field r. As i)"
+and cardB: "|B| <o r"
+shows "EX i : Field r. B \<le> As i"
+proof-
+  let ?phi = "%b j. j : Field r \<and> b : As j"
+  have "ALL b : B. EX j. ?phi b j" using Bsub by blast
+  then obtain f where f: "!! b. b : B \<Longrightarrow> ?phi b (f b)"
+  using bchoice[of B ?phi] by blast
+  let ?K = "f ` B"
+  {assume 1: "!! i. i : Field r \<Longrightarrow> ~ B \<le> As i"
+   have 2: "cofinal ?K r"
+   unfolding cofinal_def proof auto
+     fix i assume i: "i : Field r"
+     with 1 obtain b where b: "b : B \<and> b \<notin> As i" by blast
+     hence "i \<noteq> f b \<and> ~ (f b,i) : r"
+     using As f unfolding relChain_def by auto
+     hence "i \<noteq> f b \<and> (i, f b) : r" using r
+     unfolding card_order_on_def well_order_on_def linear_order_on_def
+     total_on_def using i f b by auto
+     with b show "\<exists>b\<in>B. i \<noteq> f b \<and> (i, f b) \<in> r" by blast
+   qed
+   moreover have "?K \<le> Field r" using f by blast
+   ultimately have "|?K| =o r" using 2 r unfolding regular_def by blast
+   moreover
+   {
+    have "|?K| <=o |B|" using card_of_image .
+    hence "|?K| <o r" using cardB ordLeq_ordLess_trans by blast
+   }
+   ultimately have False using not_ordLess_ordIso by blast
+  }
+  thus ?thesis by blast
+qed
+
+
+lemma infinite_cardSuc_regular:
+assumes r_inf: "infinite (Field r)" and r_card: "Card_order r"
+shows "regular (cardSuc r)"
+proof-
+  let ?r' = "cardSuc r"
+  have r': "Card_order ?r'"
+  "!! p. Card_order p \<longrightarrow> (p \<le>o r) = (p <o ?r')"
+  using r_card by (auto simp: cardSuc_Card_order cardSuc_ordLeq_ordLess)
+  show ?thesis
+  unfolding regular_def proof auto
+    fix K assume 1: "K \<le> Field ?r'" and 2: "cofinal K ?r'"
+    hence "|K| \<le>o |Field ?r'|" by (simp only: card_of_mono1)
+    also have 22: "|Field ?r'| =o ?r'"
+    using r' by (simp add: card_of_Field_ordIso[of ?r'])
+    finally have "|K| \<le>o ?r'" .
+    moreover
+    {let ?L = "UN j : K. rel.underS ?r' j"
+     let ?J = "Field r"
+     have rJ: "r =o |?J|"
+     using r_card card_of_Field_ordIso ordIso_symmetric by blast
+     assume "|K| <o ?r'"
+     hence "|K| <=o r" using r' card_of_Card_order[of K] by blast
+     hence "|K| \<le>o |?J|" using rJ ordLeq_ordIso_trans by blast
+     moreover
+     {have "ALL j : K. |rel.underS ?r' j| <o ?r'"
+      using r' 1 by (auto simp: card_of_underS)
+      hence "ALL j : K. |rel.underS ?r' j| \<le>o r"
+      using r' card_of_Card_order by blast
+      hence "ALL j : K. |rel.underS ?r' j| \<le>o |?J|"
+      using rJ ordLeq_ordIso_trans by blast
+     }
+     ultimately have "|?L| \<le>o |?J|"
+     using r_inf card_of_UNION_ordLeq_infinite by blast
+     hence "|?L| \<le>o r" using rJ ordIso_symmetric ordLeq_ordIso_trans by blast
+     hence "|?L| <o ?r'" using r' card_of_Card_order by blast
+     moreover
+     {
+      have "Field ?r' \<le> ?L"
+      using 2 unfolding rel.underS_def cofinal_def by auto
+      hence "|Field ?r'| \<le>o |?L|" by (simp add: card_of_mono1)
+      hence "?r' \<le>o |?L|"
+      using 22 ordIso_ordLeq_trans ordIso_symmetric by blast
+     }
+     ultimately have "|?L| <o |?L|" using ordLess_ordLeq_trans by blast
+     hence False using ordLess_irreflexive by blast
+    }
+    ultimately show "|K| =o ?r'"
+    unfolding ordLeq_iff_ordLess_or_ordIso by blast
+  qed
+qed
+
+lemma cardSuc_UNION:
+assumes r: "Card_order r" and "infinite (Field r)"
+and As: "relChain (cardSuc r) As"
+and Bsub: "B \<le> (UN i : Field (cardSuc r). As i)"
+and cardB: "|B| <=o r"
+shows "EX i : Field (cardSuc r). B \<le> As i"
+proof-
+  let ?r' = "cardSuc r"
+  have "Card_order ?r' \<and> |B| <o ?r'"
+  using r cardB cardSuc_ordLeq_ordLess cardSuc_Card_order
+  card_of_Card_order by blast
+  moreover have "regular ?r'"
+  using assms by(simp add: infinite_cardSuc_regular)
+  ultimately show ?thesis
+  using As Bsub cardB regular_UNION by blast
+qed
+
+
+subsection {* Others *}
+
+lemma card_of_infinite_diff_finite:
+assumes "infinite A" and "finite B"
+shows "|A - B| =o |A|"
+by (metis assms card_of_Un_diff_infinite finite_ordLess_infinite2)
+
+(* function space *)
+definition Func where
+"Func A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
+
+lemma Func_empty:
+"Func {} B = {\<lambda>x. undefined}"
+unfolding Func_def by auto
+
+lemma Func_elim:
+assumes "g \<in> Func A B" and "a \<in> A"
+shows "\<exists> b. b \<in> B \<and> g a = b"
+using assms unfolding Func_def by (cases "g a = undefined") auto
+
+definition curr where
+"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
+
+lemma curr_in:
+assumes f: "f \<in> Func (A <*> B) C"
+shows "curr A f \<in> Func A (Func B C)"
+using assms unfolding curr_def Func_def by auto
+
+lemma curr_inj:
+assumes "f1 \<in> Func (A <*> B) C" and "f2 \<in> Func (A <*> B) C"
+shows "curr A f1 = curr A f2 \<longleftrightarrow> f1 = f2"
+proof safe
+  assume c: "curr A f1 = curr A f2"
+  show "f1 = f2"
+  proof (rule ext, clarify)
+    fix a b show "f1 (a, b) = f2 (a, b)"
+    proof (cases "(a,b) \<in> A <*> B")
+      case False
+      thus ?thesis using assms unfolding Func_def by auto
+    next
+      case True hence a: "a \<in> A" and b: "b \<in> B" by auto
+      thus ?thesis
+      using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp
+    qed
+  qed
+qed
+
+lemma curr_surj:
+assumes "g \<in> Func A (Func B C)"
+shows "\<exists> f \<in> Func (A <*> B) C. curr A f = g"
+proof
+  let ?f = "\<lambda> ab. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
+  show "curr A ?f = g"
+  proof (rule ext)
+    fix a show "curr A ?f a = g a"
+    proof (cases "a \<in> A")
+      case False
+      hence "g a = undefined" using assms unfolding Func_def by auto
+      thus ?thesis unfolding curr_def using False by simp
+    next
+      case True
+      obtain g1 where "g1 \<in> Func B C" and "g a = g1"
+      using assms using Func_elim[OF assms True] by blast
+      thus ?thesis using True unfolding Func_def curr_def by auto
+    qed
+  qed
+  show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
+qed
+
+lemma bij_betw_curr:
+"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
+unfolding bij_betw_def inj_on_def image_def
+using curr_in curr_inj curr_surj by blast
+
+lemma card_of_Func_Times:
+"|Func (A <*> B) C| =o |Func A (Func B C)|"
+unfolding card_of_ordIso[symmetric]
+using bij_betw_curr by blast
+
+definition Func_map where
+"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
+
+lemma Func_map:
+assumes g: "g \<in> Func A2 A1" and f1: "f1 ` A1 \<subseteq> B1" and f2: "f2 ` B2 \<subseteq> A2"
+shows "Func_map B2 f1 f2 g \<in> Func B2 B1"
+using assms unfolding Func_def Func_map_def mem_Collect_eq by auto
+
+lemma Func_non_emp:
+assumes "B \<noteq> {}"
+shows "Func A B \<noteq> {}"
+proof-
+  obtain b where b: "b \<in> B" using assms by auto
+  hence "(\<lambda> a. if a \<in> A then b else undefined) \<in> Func A B" unfolding Func_def by auto
+  thus ?thesis by blast
+qed
+
+lemma Func_is_emp:
+"Func A B = {} \<longleftrightarrow> A \<noteq> {} \<and> B = {}" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume L: ?L
+  moreover {assume "A = {}" hence False using L Func_empty by auto}
+  moreover {assume "B \<noteq> {}" hence False using L Func_non_emp by metis}
+  ultimately show ?R by blast
+next
+  assume R: ?R
+  moreover
+  {fix f assume "f \<in> Func A B"
+   moreover obtain a where "a \<in> A" using R by blast
+   ultimately obtain b where "b \<in> B" unfolding Func_def by(cases "f a = undefined", force+)
+   with R have False by auto
+  }
+  thus ?L by blast
+qed
+
+lemma Func_map_surj:
+assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \<subseteq> A2"
+and B2A2: "B2 = {} \<Longrightarrow> A2 = {}"
+shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1"
+proof(cases "B2 = {}")
+  case True
+  thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def)
+next
+  case False note B2 = False
+  show ?thesis
+  proof safe
+    fix h assume h: "h \<in> Func B2 B1"
+    def j1 \<equiv> "inv_into A1 f1"
+    have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
+    then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> f2 (k a2) = a2" by metis
+    {fix b2 assume b2: "b2 \<in> B2"
+     hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto
+     moreover have "k (f2 b2) \<in> B2" using b2 A2(2) k by auto
+     ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast
+    } note kk = this
+    obtain b22 where b22: "b22 \<in> B2" using B2 by auto
+    def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
+    have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
+    have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
+    using kk unfolding j2_def by auto
+    def g \<equiv> "Func_map A2 j1 j2 h"
+    have "Func_map B2 f1 f2 g = h"
+    proof (rule ext)
+      fix b2 show "Func_map B2 f1 f2 g b2 = h b2"
+      proof(cases "b2 \<in> B2")
+        case True
+        show ?thesis
+        proof (cases "h b2 = undefined")
+          case True
+          hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
+          show ?thesis using A2 f_inv_into_f[OF b1]
+            unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
+        qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
+          auto intro: f_inv_into_f)
+      qed(insert h, unfold Func_def Func_map_def, auto)
+    qed
+    moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
+    using inv_into_into j2A2 B1 A2 inv_into_into
+    unfolding j1_def image_def by fast+
+    ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
+    unfolding Func_map_def[abs_def] unfolding image_def by auto
+  qed(insert B1 Func_map[OF _ _ A2(2)], auto)
+qed
+
+lemma card_of_Pow_Func:
+"|Pow A| =o |Func A (UNIV::bool set)|"
+proof-
+  def F \<equiv> "\<lambda> A' a. if a \<in> A then (if a \<in> A' then True else False)
+                            else undefined"
+  have "bij_betw F (Pow A) (Func A (UNIV::bool set))"
+  unfolding bij_betw_def inj_on_def proof (intro ballI impI conjI)
+    fix A1 A2 assume "A1 \<in> Pow A" "A2 \<in> Pow A" "F A1 = F A2"
+    thus "A1 = A2" unfolding F_def Pow_def fun_eq_iff by (auto split: split_if_asm)
+  next
+    show "F ` Pow A = Func A UNIV"
+    proof safe
+      fix f assume f: "f \<in> Func A (UNIV::bool set)"
+      show "f \<in> F ` Pow A" unfolding image_def mem_Collect_eq proof(intro bexI)
+        let ?A1 = "{a \<in> A. f a = True}"
+        show "f = F ?A1" unfolding F_def apply(rule ext)
+        using f unfolding Func_def mem_Collect_eq by auto
+      qed auto
+    qed(unfold Func_def mem_Collect_eq F_def, auto)
+  qed
+  thus ?thesis unfolding card_of_ordIso[symmetric] by blast
+qed
+
+lemma card_of_Func_mono:
+fixes A1 A2 :: "'a set" and B :: "'b set"
+assumes A12: "A1 \<subseteq> A2" and B: "B \<noteq> {}"
+shows "|Func A1 B| \<le>o |Func A2 B|"
+proof-
+  obtain bb where bb: "bb \<in> B" using B by auto
+  def F \<equiv> "\<lambda> (f1::'a \<Rightarrow> 'b) a. if a \<in> A2 then (if a \<in> A1 then f1 a else bb)
+                                                else undefined"
+  show ?thesis unfolding card_of_ordLeq[symmetric] proof(intro exI[of _ F] conjI)
+    show "inj_on F (Func A1 B)" unfolding inj_on_def proof safe
+      fix f g assume f: "f \<in> Func A1 B" and g: "g \<in> Func A1 B" and eq: "F f = F g"
+      show "f = g"
+      proof(rule ext)
+        fix a show "f a = g a"
+        proof(cases "a \<in> A1")
+          case True
+          thus ?thesis using eq A12 unfolding F_def fun_eq_iff
+          by (elim allE[of _ a]) auto
+        qed(insert f g, unfold Func_def, fastforce)
+      qed
+    qed
+  qed(insert bb, unfold Func_def F_def, force)
+qed
+
+lemma ordLeq_Func:
+assumes "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
+shows "|A| \<le>o |Func A B|"
+unfolding card_of_ordLeq[symmetric] proof(intro exI conjI)
+  let ?F = "\<lambda> aa a. if a \<in> A then (if a = aa then b1 else b2) else undefined"
+  show "inj_on ?F A" using assms unfolding inj_on_def fun_eq_iff by auto
+  show "?F ` A \<subseteq> Func A B" using assms unfolding Func_def by auto
+qed
+
+lemma infinite_Func:
+assumes A: "infinite A" and B: "{b1,b2} \<subseteq> B" "b1 \<noteq> b2"
+shows "infinite (Func A B)"
+using ordLeq_Func[OF B] by (metis A card_of_ordLeq_finite)
+
+lemma card_of_Func_UNIV:
+"|Func (UNIV::'a set) (B::'b set)| =o |{f::'a \<Rightarrow> 'b. range f \<subseteq> B}|"
+apply(rule ordIso_symmetric) proof(intro card_of_ordIsoI)
+  let ?F = "\<lambda> f (a::'a). ((f a)::'b)"
+  show "bij_betw ?F {f. range f \<subseteq> B} (Func UNIV B)"
+  unfolding bij_betw_def inj_on_def proof safe
+    fix h :: "'a \<Rightarrow> 'b" assume h: "h \<in> Func UNIV B"
+    hence "\<forall> a. \<exists> b. h a = b" unfolding Func_def by auto
+    then obtain f where f: "\<forall> a. h a = f a" by metis
+    hence "range f \<subseteq> B" using h unfolding Func_def by auto
+    thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f unfolding image_def by auto
+  qed(unfold Func_def fun_eq_iff, auto)
+qed
+
+end
--- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy	Mon Nov 18 18:04:44 2013 +0100
+++ b/src/HOL/Cardinals/Constructions_on_Wellorders.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -8,7 +8,7 @@
 header {* Constructions on Wellorders *}
 
 theory Constructions_on_Wellorders
-imports Constructions_on_Wellorders_Base Wellorder_Embedding
+imports Constructions_on_Wellorders_LFP Wellorder_Embedding
 begin
 
 declare
--- a/src/HOL/Cardinals/Constructions_on_Wellorders_Base.thy	Mon Nov 18 18:04:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1633 +0,0 @@
-(*  Title:      HOL/Cardinals/Constructions_on_Wellorders_Base.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Constructions on wellorders (base).
-*)
-
-header {* Constructions on Wellorders (Base) *}
-
-theory Constructions_on_Wellorders_Base
-imports Wellorder_Embedding_Base
-begin
-
-
-text {* In this section, we study basic constructions on well-orders, such as restriction to
-a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders,
-and bounded square.  We also define between well-orders
-the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"}),
-@{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"}), and
-@{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}).  We study the
-connections between these relations, order filters, and the aforementioned constructions.
-A main result of this section is that @{text "<o"} is well-founded.  *}
-
-
-subsection {* Restriction to a set  *}
-
-
-abbreviation Restr :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a rel"
-where "Restr r A \<equiv> r Int (A \<times> A)"
-
-
-lemma Restr_subset:
-"A \<le> B \<Longrightarrow> Restr (Restr r B) A = Restr r A"
-by blast
-
-
-lemma Restr_Field: "Restr r (Field r) = r"
-unfolding Field_def by auto
-
-
-lemma Refl_Restr: "Refl r \<Longrightarrow> Refl(Restr r A)"
-unfolding refl_on_def Field_def by auto
-
-
-lemma antisym_Restr:
-"antisym r \<Longrightarrow> antisym(Restr r A)"
-unfolding antisym_def Field_def by auto
-
-
-lemma Total_Restr:
-"Total r \<Longrightarrow> Total(Restr r A)"
-unfolding total_on_def Field_def by auto
-
-
-lemma trans_Restr:
-"trans r \<Longrightarrow> trans(Restr r A)"
-unfolding trans_def Field_def by blast
-
-
-lemma Preorder_Restr:
-"Preorder r \<Longrightarrow> Preorder(Restr r A)"
-unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr)
-
-
-lemma Partial_order_Restr:
-"Partial_order r \<Longrightarrow> Partial_order(Restr r A)"
-unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr)
-
-
-lemma Linear_order_Restr:
-"Linear_order r \<Longrightarrow> Linear_order(Restr r A)"
-unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr)
-
-
-lemma Well_order_Restr:
-assumes "Well_order r"
-shows "Well_order(Restr r A)"
-proof-
-  have "Restr r A - Id \<le> r - Id" using Restr_subset by blast
-  hence "wf(Restr r A - Id)" using assms
-  using well_order_on_def wf_subset by blast
-  thus ?thesis using assms unfolding well_order_on_def
-  by (simp add: Linear_order_Restr)
-qed
-
-
-lemma Field_Restr_subset: "Field(Restr r A) \<le> A"
-by (auto simp add: Field_def)
-
-
-lemma Refl_Field_Restr:
-"Refl r \<Longrightarrow> Field(Restr r A) = (Field r) Int A"
-by (auto simp add: refl_on_def Field_def)
-
-
-lemma Refl_Field_Restr2:
-"\<lbrakk>Refl r; A \<le> Field r\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
-by (auto simp add: Refl_Field_Restr)
-
-
-lemma well_order_on_Restr:
-assumes WELL: "Well_order r" and SUB: "A \<le> Field r"
-shows "well_order_on A (Restr r A)"
-using assms
-using Well_order_Restr[of r A] Refl_Field_Restr2[of r A]
-     order_on_defs[of "Field r" r] by auto
-
-
-subsection {* Order filters versus restrictions and embeddings  *}
-
-
-lemma Field_Restr_ofilter:
-"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
-by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2)
-
-
-lemma ofilter_Restr_under:
-assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \<in> A"
-shows "rel.under (Restr r A) a = rel.under r a"
-using assms wo_rel_def
-proof(auto simp add: wo_rel.ofilter_def rel.under_def)
-  fix b assume *: "a \<in> A" and "(b,a) \<in> r"
-  hence "b \<in> rel.under r a \<and> a \<in> Field r"
-  unfolding rel.under_def using Field_def by fastforce
-  thus "b \<in> A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
-qed
-
-
-lemma ofilter_embed:
-assumes "Well_order r"
-shows "wo_rel.ofilter r A = (A \<le> Field r \<and> embed (Restr r A) r id)"
-proof
-  assume *: "wo_rel.ofilter r A"
-  show "A \<le> Field r \<and> embed (Restr r A) r id"
-  proof(unfold embed_def, auto)
-    fix a assume "a \<in> A" thus "a \<in> Field r" using assms *
-    by (auto simp add: wo_rel_def wo_rel.ofilter_def)
-  next
-    fix a assume "a \<in> Field (Restr r A)"
-    thus "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" using assms *
-    by (simp add: ofilter_Restr_under Field_Restr_ofilter)
-  qed
-next
-  assume *: "A \<le> Field r \<and> embed (Restr r A) r id"
-  hence "Field(Restr r A) \<le> Field r"
-  using assms  embed_Field[of "Restr r A" r id] id_def
-        Well_order_Restr[of r] by auto
-  {fix a assume "a \<in> A"
-   hence "a \<in> Field(Restr r A)" using * assms
-   by (simp add: order_on_defs Refl_Field_Restr2)
-   hence "bij_betw id (rel.under (Restr r A) a) (rel.under r a)"
-   using * unfolding embed_def by auto
-   hence "rel.under r a \<le> rel.under (Restr r A) a"
-   unfolding bij_betw_def by auto
-   also have "\<dots> \<le> Field(Restr r A)" by (simp add: rel.under_Field)
-   also have "\<dots> \<le> A" by (simp add: Field_Restr_subset)
-   finally have "rel.under r a \<le> A" .
-  }
-  thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def)
-qed
-
-
-lemma ofilter_Restr_Int:
-assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A"
-shows "wo_rel.ofilter (Restr r B) (A Int B)"
-proof-
-  let ?rB = "Restr r B"
-  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
-  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
-  hence Field: "Field ?rB = Field r Int B"
-  using Refl_Field_Restr by blast
-  have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
-  by (simp add: Well_order_Restr wo_rel_def)
-  (* Main proof *)
-  show ?thesis using WellB assms
-  proof(auto simp add: wo_rel.ofilter_def rel.under_def)
-    fix a assume "a \<in> A" and *: "a \<in> B"
-    hence "a \<in> Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def)
-    with * show "a \<in> Field ?rB" using Field by auto
-  next
-    fix a b assume "a \<in> A" and "(b,a) \<in> r"
-    thus "b \<in> A" using Well OFA by (auto simp add: wo_rel.ofilter_def rel.under_def)
-  qed
-qed
-
-
-lemma ofilter_Restr_subset:
-assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \<le> B"
-shows "wo_rel.ofilter (Restr r B) A"
-proof-
-  have "A Int B = A" using SUB by blast
-  thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto
-qed
-
-
-lemma ofilter_subset_embed:
-assumes WELL: "Well_order r" and
-        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
-shows "(A \<le> B) = (embed (Restr r A) (Restr r B) id)"
-proof-
-  let ?rA = "Restr r A"  let ?rB = "Restr r B"
-  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
-  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
-  hence FieldA: "Field ?rA = Field r Int A"
-  using Refl_Field_Restr by blast
-  have FieldB: "Field ?rB = Field r Int B"
-  using Refl Refl_Field_Restr by blast
-  have WellA: "wo_rel ?rA \<and> Well_order ?rA" using WELL
-  by (simp add: Well_order_Restr wo_rel_def)
-  have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
-  by (simp add: Well_order_Restr wo_rel_def)
-  (* Main proof *)
-  show ?thesis
-  proof
-    assume *: "A \<le> B"
-    hence "wo_rel.ofilter (Restr r B) A" using assms
-    by (simp add: ofilter_Restr_subset)
-    hence "embed (Restr ?rB A) (Restr r B) id"
-    using WellB ofilter_embed[of "?rB" A] by auto
-    thus "embed (Restr r A) (Restr r B) id"
-    using * by (simp add: Restr_subset)
-  next
-    assume *: "embed (Restr r A) (Restr r B) id"
-    {fix a assume **: "a \<in> A"
-     hence "a \<in> Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def)
-     with ** FieldA have "a \<in> Field ?rA" by auto
-     hence "a \<in> Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto
-     hence "a \<in> B" using FieldB by auto
-    }
-    thus "A \<le> B" by blast
-  qed
-qed
-
-
-lemma ofilter_subset_embedS_iso:
-assumes WELL: "Well_order r" and
-        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
-shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \<and>
-       ((A = B) = (iso (Restr r A) (Restr r B) id))"
-proof-
-  let ?rA = "Restr r A"  let ?rB = "Restr r B"
-  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
-  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
-  hence "Field ?rA = Field r Int A"
-  using Refl_Field_Restr by blast
-  hence FieldA: "Field ?rA = A" using OFA Well
-  by (auto simp add: wo_rel.ofilter_def)
-  have "Field ?rB = Field r Int B"
-  using Refl Refl_Field_Restr by blast
-  hence FieldB: "Field ?rB = B" using OFB Well
-  by (auto simp add: wo_rel.ofilter_def)
-  (* Main proof *)
-  show ?thesis unfolding embedS_def iso_def
-  using assms ofilter_subset_embed[of r A B]
-        FieldA FieldB bij_betw_id_iff[of A B] by auto
-qed
-
-
-lemma ofilter_subset_embedS:
-assumes WELL: "Well_order r" and
-        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
-shows "(A < B) = embedS (Restr r A) (Restr r B) id"
-using assms by (simp add: ofilter_subset_embedS_iso)
-
-
-lemma embed_implies_iso_Restr:
-assumes WELL: "Well_order r" and WELL': "Well_order r'" and
-        EMB: "embed r' r f"
-shows "iso r' (Restr r (f ` (Field r'))) f"
-proof-
-  let ?A' = "Field r'"
-  let ?r'' = "Restr r (f ` ?A')"
-  have 0: "Well_order ?r''" using WELL Well_order_Restr by blast
-  have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter  by blast
-  hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast
-  hence "bij_betw f ?A' (Field ?r'')"
-  using EMB embed_inj_on WELL' unfolding bij_betw_def by blast
-  moreover
-  {have "\<forall>a b. (a,b) \<in> r' \<longrightarrow> a \<in> Field r' \<and> b \<in> Field r'"
-   unfolding Field_def by auto
-   hence "compat r' ?r'' f"
-   using assms embed_iff_compat_inj_on_ofilter
-   unfolding compat_def by blast
-  }
-  ultimately show ?thesis using WELL' 0 iso_iff3 by blast
-qed
-
-
-subsection {* The strict inclusion on proper ofilters is well-founded *}
-
-
-definition ofilterIncl :: "'a rel \<Rightarrow> 'a set rel"
-where
-"ofilterIncl r \<equiv> {(A,B). wo_rel.ofilter r A \<and> A \<noteq> Field r \<and>
-                         wo_rel.ofilter r B \<and> B \<noteq> Field r \<and> A < B}"
-
-
-lemma wf_ofilterIncl:
-assumes WELL: "Well_order r"
-shows "wf(ofilterIncl r)"
-proof-
-  have Well: "wo_rel r" using WELL by (simp add: wo_rel_def)
-  hence Lo: "Linear_order r" by (simp add: wo_rel.LIN)
-  let ?h = "(\<lambda> A. wo_rel.suc r A)"
-  let ?rS = "r - Id"
-  have "wf ?rS" using WELL by (simp add: order_on_defs)
-  moreover
-  have "compat (ofilterIncl r) ?rS ?h"
-  proof(unfold compat_def ofilterIncl_def,
-        intro allI impI, simp, elim conjE)
-    fix A B
-    assume *: "wo_rel.ofilter r A" "A \<noteq> Field r" and
-           **: "wo_rel.ofilter r B" "B \<noteq> Field r" and ***: "A < B"
-    then obtain a and b where 0: "a \<in> Field r \<and> b \<in> Field r" and
-                         1: "A = rel.underS r a \<and> B = rel.underS r b"
-    using Well by (auto simp add: wo_rel.ofilter_underS_Field)
-    hence "a \<noteq> b" using *** by auto
-    moreover
-    have "(a,b) \<in> r" using 0 1 Lo ***
-    by (auto simp add: rel.underS_incl_iff)
-    moreover
-    have "a = wo_rel.suc r A \<and> b = wo_rel.suc r B"
-    using Well 0 1 by (simp add: wo_rel.suc_underS)
-    ultimately
-    show "(wo_rel.suc r A, wo_rel.suc r B) \<in> r \<and> wo_rel.suc r A \<noteq> wo_rel.suc r B"
-    by simp
-  qed
-  ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf)
-qed
-
-
-
-subsection {* Ordering the  well-orders by existence of embeddings *}
-
-
-text {* We define three relations between well-orders:
-\begin{itemize}
-\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"});
-\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"});
-\item @{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}).
-\end{itemize}
-%
-The prefix "ord" and the index "o" in these names stand for "ordinal-like".
-These relations shall be proved to be inter-connected in a similar fashion as the trio
-@{text "\<le>"}, @{text "<"}, @{text "="} associated to a total order on a set.
-*}
-
-
-definition ordLeq :: "('a rel * 'a' rel) set"
-where
-"ordLeq = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embed r r' f)}"
-
-
-abbreviation ordLeq2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<=o" 50)
-where "r <=o r' \<equiv> (r,r') \<in> ordLeq"
-
-
-abbreviation ordLeq3 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "\<le>o" 50)
-where "r \<le>o r' \<equiv> r <=o r'"
-
-
-definition ordLess :: "('a rel * 'a' rel) set"
-where
-"ordLess = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embedS r r' f)}"
-
-abbreviation ordLess2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<o" 50)
-where "r <o r' \<equiv> (r,r') \<in> ordLess"
-
-
-definition ordIso :: "('a rel * 'a' rel) set"
-where
-"ordIso = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. iso r r' f)}"
-
-abbreviation ordIso2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "=o" 50)
-where "r =o r' \<equiv> (r,r') \<in> ordIso"
-
-
-lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def
-
-lemma ordLeq_Well_order_simp:
-assumes "r \<le>o r'"
-shows "Well_order r \<and> Well_order r'"
-using assms unfolding ordLeq_def by simp
-
-
-lemma ordLess_Well_order_simp:
-assumes "r <o r'"
-shows "Well_order r \<and> Well_order r'"
-using assms unfolding ordLess_def by simp
-
-
-lemma ordIso_Well_order_simp:
-assumes "r =o r'"
-shows "Well_order r \<and> Well_order r'"
-using assms unfolding ordIso_def by simp
-
-
-text{* Notice that the relations @{text "\<le>o"}, @{text "<o"}, @{text "=o"} connect well-orders
-on potentially {\em distinct} types. However, some of the lemmas below, including the next one,
-restrict implicitly the type of these relations to @{text "(('a rel) * ('a rel)) set"} , i.e.,
-to @{text "'a rel rel"}.  *}
-
-
-lemma ordLeq_reflexive:
-"Well_order r \<Longrightarrow> r \<le>o r"
-unfolding ordLeq_def using id_embed[of r] by blast
-
-
-lemma ordLeq_transitive[trans]:
-assumes *: "r \<le>o r'" and **: "r' \<le>o r''"
-shows "r \<le>o r''"
-proof-
-  obtain f and f'
-  where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
-        "embed r r' f" and "embed r' r'' f'"
-  using * ** unfolding ordLeq_def by blast
-  hence "embed r r'' (f' o f)"
-  using comp_embed[of r r' f r'' f'] by auto
-  thus "r \<le>o r''" unfolding ordLeq_def using 1 by auto
-qed
-
-
-lemma ordLeq_total:
-"\<lbrakk>Well_order r; Well_order r'\<rbrakk> \<Longrightarrow> r \<le>o r' \<or> r' \<le>o r"
-unfolding ordLeq_def using wellorders_totally_ordered by blast
-
-
-lemma ordIso_reflexive:
-"Well_order r \<Longrightarrow> r =o r"
-unfolding ordIso_def using id_iso[of r] by blast
-
-
-lemma ordIso_transitive[trans]:
-assumes *: "r =o r'" and **: "r' =o r''"
-shows "r =o r''"
-proof-
-  obtain f and f'
-  where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
-        "iso r r' f" and 3: "iso r' r'' f'"
-  using * ** unfolding ordIso_def by auto
-  hence "iso r r'' (f' o f)"
-  using comp_iso[of r r' f r'' f'] by auto
-  thus "r =o r''" unfolding ordIso_def using 1 by auto
-qed
-
-
-lemma ordIso_symmetric:
-assumes *: "r =o r'"
-shows "r' =o r"
-proof-
-  obtain f where 1: "Well_order r \<and> Well_order r'" and
-                 2: "embed r r' f \<and> bij_betw f (Field r) (Field r')"
-  using * by (auto simp add: ordIso_def iso_def)
-  let ?f' = "inv_into (Field r) f"
-  have "embed r' r ?f' \<and> bij_betw ?f' (Field r') (Field r)"
-  using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw)
-  thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def)
-qed
-
-
-lemma ordLeq_ordLess_trans[trans]:
-assumes "r \<le>o r'" and " r' <o r''"
-shows "r <o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLeq_def ordLess_def by auto
-  thus ?thesis using assms unfolding ordLeq_def ordLess_def
-  using embed_comp_embedS by blast
-qed
-
-
-lemma ordLess_ordLeq_trans[trans]:
-assumes "r <o r'" and " r' \<le>o r''"
-shows "r <o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLeq_def ordLess_def by auto
-  thus ?thesis using assms unfolding ordLeq_def ordLess_def
-  using embedS_comp_embed by blast
-qed
-
-
-lemma ordLeq_ordIso_trans[trans]:
-assumes "r \<le>o r'" and " r' =o r''"
-shows "r \<le>o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLeq_def ordIso_def by auto
-  thus ?thesis using assms unfolding ordLeq_def ordIso_def
-  using embed_comp_iso by blast
-qed
-
-
-lemma ordIso_ordLeq_trans[trans]:
-assumes "r =o r'" and " r' \<le>o r''"
-shows "r \<le>o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLeq_def ordIso_def by auto
-  thus ?thesis using assms unfolding ordLeq_def ordIso_def
-  using iso_comp_embed by blast
-qed
-
-
-lemma ordLess_ordIso_trans[trans]:
-assumes "r <o r'" and " r' =o r''"
-shows "r <o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLess_def ordIso_def by auto
-  thus ?thesis using assms unfolding ordLess_def ordIso_def
-  using embedS_comp_iso by blast
-qed
-
-
-lemma ordIso_ordLess_trans[trans]:
-assumes "r =o r'" and " r' <o r''"
-shows "r <o r''"
-proof-
-  have "Well_order r \<and> Well_order r''"
-  using assms unfolding ordLess_def ordIso_def by auto
-  thus ?thesis using assms unfolding ordLess_def ordIso_def
-  using iso_comp_embedS by blast
-qed
-
-
-lemma ordLess_not_embed:
-assumes "r <o r'"
-shows "\<not>(\<exists>f'. embed r' r f')"
-proof-
-  obtain f where 1: "Well_order r \<and> Well_order r'" and 2: "embed r r' f" and
-                 3: " \<not> bij_betw f (Field r) (Field r')"
-  using assms unfolding ordLess_def by (auto simp add: embedS_def)
-  {fix f' assume *: "embed r' r f'"
-   hence "bij_betw f (Field r) (Field r')" using 1 2
-   by (simp add: embed_bothWays_Field_bij_betw)
-   with 3 have False by contradiction
-  }
-  thus ?thesis by blast
-qed
-
-
-lemma ordLess_Field:
-assumes OL: "r1 <o r2" and EMB: "embed r1 r2 f"
-shows "\<not> (f`(Field r1) = Field r2)"
-proof-
-  let ?A1 = "Field r1"  let ?A2 = "Field r2"
-  obtain g where
-  0: "Well_order r1 \<and> Well_order r2" and
-  1: "embed r1 r2 g \<and> \<not>(bij_betw g ?A1 ?A2)"
-  using OL unfolding ordLess_def by (auto simp add: embedS_def)
-  hence "\<forall>a \<in> ?A1. f a = g a"
-  using 0 EMB embed_unique[of r1] by auto
-  hence "\<not>(bij_betw f ?A1 ?A2)"
-  using 1 bij_betw_cong[of ?A1] by blast
-  moreover
-  have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on)
-  ultimately show ?thesis by (simp add: bij_betw_def)
-qed
-
-
-lemma ordLess_iff:
-"r <o r' = (Well_order r \<and> Well_order r' \<and> \<not>(\<exists>f'. embed r' r f'))"
-proof
-  assume *: "r <o r'"
-  hence "\<not>(\<exists>f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp
-  with * show "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
-  unfolding ordLess_def by auto
-next
-  assume *: "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
-  then obtain f where 1: "embed r r' f"
-  using wellorders_totally_ordered[of r r'] by blast
-  moreover
-  {assume "bij_betw f (Field r) (Field r')"
-   with * 1 have "embed r' r (inv_into (Field r) f) "
-   using inv_into_Field_embed_bij_betw[of r r' f] by auto
-   with * have False by blast
-  }
-  ultimately show "(r,r') \<in> ordLess"
-  unfolding ordLess_def using * by (fastforce simp add: embedS_def)
-qed
-
-
-lemma ordLess_irreflexive: "\<not> r <o r"
-proof
-  assume "r <o r"
-  hence "Well_order r \<and>  \<not>(\<exists>f. embed r r f)"
-  unfolding ordLess_iff ..
-  moreover have "embed r r id" using id_embed[of r] .
-  ultimately show False by blast
-qed
-
-
-lemma ordLeq_iff_ordLess_or_ordIso:
-"r \<le>o r' = (r <o r' \<or> r =o r')"
-unfolding ordRels_def embedS_defs iso_defs by blast
-
-
-lemma ordIso_iff_ordLeq:
-"(r =o r') = (r \<le>o r' \<and> r' \<le>o r)"
-proof
-  assume "r =o r'"
-  then obtain f where 1: "Well_order r \<and> Well_order r' \<and>
-                     embed r r' f \<and> bij_betw f (Field r) (Field r')"
-  unfolding ordIso_def iso_defs by auto
-  hence "embed r r' f \<and> embed r' r (inv_into (Field r) f)"
-  by (simp add: inv_into_Field_embed_bij_betw)
-  thus  "r \<le>o r' \<and> r' \<le>o r"
-  unfolding ordLeq_def using 1 by auto
-next
-  assume "r \<le>o r' \<and> r' \<le>o r"
-  then obtain f and g where 1: "Well_order r \<and> Well_order r' \<and>
-                           embed r r' f \<and> embed r' r g"
-  unfolding ordLeq_def by auto
-  hence "iso r r' f" by (auto simp add: embed_bothWays_iso)
-  thus "r =o r'" unfolding ordIso_def using 1 by auto
-qed
-
-
-lemma not_ordLess_ordLeq:
-"r <o r' \<Longrightarrow> \<not> r' \<le>o r"
-using ordLess_ordLeq_trans ordLess_irreflexive by blast
-
-
-lemma ordLess_or_ordLeq:
-assumes WELL: "Well_order r" and WELL': "Well_order r'"
-shows "r <o r' \<or> r' \<le>o r"
-proof-
-  have "r \<le>o r' \<or> r' \<le>o r"
-  using assms by (simp add: ordLeq_total)
-  moreover
-  {assume "\<not> r <o r' \<and> r \<le>o r'"
-   hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast
-   hence "r' \<le>o r" using ordIso_symmetric ordIso_iff_ordLeq by blast
-  }
-  ultimately show ?thesis by blast
-qed
-
-
-lemma not_ordLess_ordIso:
-"r <o r' \<Longrightarrow> \<not> r =o r'"
-using assms ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast
-
-
-lemma not_ordLeq_iff_ordLess:
-assumes WELL: "Well_order r" and WELL': "Well_order r'"
-shows "(\<not> r' \<le>o r) = (r <o r')"
-using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
-
-
-lemma not_ordLess_iff_ordLeq:
-assumes WELL: "Well_order r" and WELL': "Well_order r'"
-shows "(\<not> r' <o r) = (r \<le>o r')"
-using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
-
-
-lemma ordLess_transitive[trans]:
-"\<lbrakk>r <o r'; r' <o r''\<rbrakk> \<Longrightarrow> r <o r''"
-using assms ordLess_ordLeq_trans ordLeq_iff_ordLess_or_ordIso by blast
-
-
-corollary ordLess_trans: "trans ordLess"
-unfolding trans_def using ordLess_transitive by blast
-
-
-lemmas ordIso_equivalence = ordIso_transitive ordIso_reflexive ordIso_symmetric
-
-
-lemma ordIso_imp_ordLeq:
-"r =o r' \<Longrightarrow> r \<le>o r'"
-using ordIso_iff_ordLeq by blast
-
-
-lemma ordLess_imp_ordLeq:
-"r <o r' \<Longrightarrow> r \<le>o r'"
-using ordLeq_iff_ordLess_or_ordIso by blast
-
-
-lemma ofilter_subset_ordLeq:
-assumes WELL: "Well_order r" and
-        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
-shows "(A \<le> B) = (Restr r A \<le>o Restr r B)"
-proof
-  assume "A \<le> B"
-  thus "Restr r A \<le>o Restr r B"
-  unfolding ordLeq_def using assms
-  Well_order_Restr Well_order_Restr ofilter_subset_embed by blast
-next
-  assume *: "Restr r A \<le>o Restr r B"
-  then obtain f where "embed (Restr r A) (Restr r B) f"
-  unfolding ordLeq_def by blast
-  {assume "B < A"
-   hence "Restr r B <o Restr r A"
-   unfolding ordLess_def using assms
-   Well_order_Restr Well_order_Restr ofilter_subset_embedS by blast
-   hence False using * not_ordLess_ordLeq by blast
-  }
-  thus "A \<le> B" using OFA OFB WELL
-  wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast
-qed
-
-
-lemma ofilter_subset_ordLess:
-assumes WELL: "Well_order r" and
-        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
-shows "(A < B) = (Restr r A <o Restr r B)"
-proof-
-  let ?rA = "Restr r A" let ?rB = "Restr r B"
-  have 1: "Well_order ?rA \<and> Well_order ?rB"
-  using WELL Well_order_Restr by blast
-  have "(A < B) = (\<not> B \<le> A)" using assms
-  wo_rel_def wo_rel.ofilter_linord[of r A B] by blast
-  also have "\<dots> = (\<not> Restr r B \<le>o Restr r A)"
-  using assms ofilter_subset_ordLeq by blast
-  also have "\<dots> = (Restr r A <o Restr r B)"
-  using 1 not_ordLeq_iff_ordLess by blast
-  finally show ?thesis .
-qed
-
-
-lemma ofilter_ordLess:
-"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> (A < Field r) = (Restr r A <o r)"
-by (simp add: ofilter_subset_ordLess wo_rel.Field_ofilter
-    wo_rel_def Restr_Field)
-
-
-corollary underS_Restr_ordLess:
-assumes "Well_order r" and "Field r \<noteq> {}"
-shows "Restr r (rel.underS r a) <o r"
-proof-
-  have "rel.underS r a < Field r" using assms
-  by (simp add: rel.underS_Field3)
-  thus ?thesis using assms
-  by (simp add: ofilter_ordLess wo_rel.underS_ofilter wo_rel_def)
-qed
-
-
-lemma embed_ordLess_ofilterIncl:
-assumes
-  OL12: "r1 <o r2" and OL23: "r2 <o r3" and
-  EMB13: "embed r1 r3 f13" and EMB23: "embed r2 r3 f23"
-shows "(f13`(Field r1), f23`(Field r2)) \<in> (ofilterIncl r3)"
-proof-
-  have OL13: "r1 <o r3"
-  using OL12 OL23 using ordLess_transitive by auto
-  let ?A1 = "Field r1"  let ?A2 ="Field r2" let ?A3 ="Field r3"
-  obtain f12 g23 where
-  0: "Well_order r1 \<and> Well_order r2 \<and> Well_order r3" and
-  1: "embed r1 r2 f12 \<and> \<not>(bij_betw f12 ?A1 ?A2)" and
-  2: "embed r2 r3 g23 \<and> \<not>(bij_betw g23 ?A2 ?A3)"
-  using OL12 OL23 by (auto simp add: ordLess_def embedS_def)
-  hence "\<forall>a \<in> ?A2. f23 a = g23 a"
-  using EMB23 embed_unique[of r2 r3] by blast
-  hence 3: "\<not>(bij_betw f23 ?A2 ?A3)"
-  using 2 bij_betw_cong[of ?A2 f23 g23] by blast
-  (*  *)
-  have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \<and> f12 ` ?A1 \<noteq> ?A2"
-  using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field)
-  have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \<and> f23 ` ?A2 \<noteq> ?A3"
-  using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field)
-  have 6: "wo_rel.ofilter r3 (f13 ` ?A1)  \<and> f13 ` ?A1 \<noteq> ?A3"
-  using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field)
-  (*  *)
-  have "f12 ` ?A1 < ?A2"
-  using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def)
-  moreover have "inj_on f23 ?A2"
-  using EMB23 0 by (simp add: wo_rel_def embed_inj_on)
-  ultimately
-  have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset)
-  moreover
-  {have "embed r1 r3 (f23 o f12)"
-   using 1 EMB23 0 by (auto simp add: comp_embed)
-   hence "\<forall>a \<in> ?A1. f23(f12 a) = f13 a"
-   using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto
-   hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force
-  }
-  ultimately
-  have "f13 ` ?A1 < f23 ` ?A2" by simp
-  (*  *)
-  with 5 6 show ?thesis
-  unfolding ofilterIncl_def by auto
-qed
-
-
-lemma ordLess_iff_ordIso_Restr:
-assumes WELL: "Well_order r" and WELL': "Well_order r'"
-shows "(r' <o r) = (\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a))"
-proof(auto)
-  fix a assume *: "a \<in> Field r" and **: "r' =o Restr r (rel.underS r a)"
-  hence "Restr r (rel.underS r a) <o r" using WELL underS_Restr_ordLess[of r] by blast
-  thus "r' <o r" using ** ordIso_ordLess_trans by blast
-next
-  assume "r' <o r"
-  then obtain f where 1: "Well_order r \<and> Well_order r'" and
-                      2: "embed r' r f \<and> f ` (Field r') \<noteq> Field r"
-  unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast
-  hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast
-  then obtain a where 3: "a \<in> Field r" and 4: "rel.underS r a = f ` (Field r')"
-  using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def)
-  have "iso r' (Restr r (f ` (Field r'))) f"
-  using embed_implies_iso_Restr 2 assms by blast
-  moreover have "Well_order (Restr r (f ` (Field r')))"
-  using WELL Well_order_Restr by blast
-  ultimately have "r' =o Restr r (f ` (Field r'))"
-  using WELL' unfolding ordIso_def by auto
-  hence "r' =o Restr r (rel.underS r a)" using 4 by auto
-  thus "\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a)" using 3 by auto
-qed
-
-
-lemma internalize_ordLess:
-"(r' <o r) = (\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r)"
-proof
-  assume *: "r' <o r"
-  hence 0: "Well_order r \<and> Well_order r'" unfolding ordLess_def by auto
-  with * obtain a where 1: "a \<in> Field r" and 2: "r' =o Restr r (rel.underS r a)"
-  using ordLess_iff_ordIso_Restr by blast
-  let ?p = "Restr r (rel.underS r a)"
-  have "wo_rel.ofilter r (rel.underS r a)" using 0
-  by (simp add: wo_rel_def wo_rel.underS_ofilter)
-  hence "Field ?p = rel.underS r a" using 0 Field_Restr_ofilter by blast
-  hence "Field ?p < Field r" using rel.underS_Field2 1 by fastforce
-  moreover have "?p <o r" using underS_Restr_ordLess[of r a] 0 1 by blast
-  ultimately
-  show "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r" using 2 by blast
-next
-  assume "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r"
-  thus "r' <o r" using ordIso_ordLess_trans by blast
-qed
-
-
-lemma internalize_ordLeq:
-"(r' \<le>o r) = (\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r)"
-proof
-  assume *: "r' \<le>o r"
-  moreover
-  {assume "r' <o r"
-   then obtain p where "Field p < Field r \<and> r' =o p \<and> p <o r"
-   using internalize_ordLess[of r' r] by blast
-   hence "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
-   using ordLeq_iff_ordLess_or_ordIso by blast
-  }
-  moreover
-  have "r \<le>o r" using * ordLeq_def ordLeq_reflexive by blast
-  ultimately show "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
-  using ordLeq_iff_ordLess_or_ordIso by blast
-next
-  assume "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
-  thus "r' \<le>o r" using ordIso_ordLeq_trans by blast
-qed
-
-
-lemma ordLeq_iff_ordLess_Restr:
-assumes WELL: "Well_order r" and WELL': "Well_order r'"
-shows "(r \<le>o r') = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r')"
-proof(auto)
-  assume *: "r \<le>o r'"
-  fix a assume "a \<in> Field r"
-  hence "Restr r (rel.underS r a) <o r"
-  using WELL underS_Restr_ordLess[of r] by blast
-  thus "Restr r (rel.underS r a) <o r'"
-  using * ordLess_ordLeq_trans by blast
-next
-  assume *: "\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r'"
-  {assume "r' <o r"
-   then obtain a where "a \<in> Field r \<and> r' =o Restr r (rel.underS r a)"
-   using assms ordLess_iff_ordIso_Restr by blast
-   hence False using * not_ordLess_ordIso ordIso_symmetric by blast
-  }
-  thus "r \<le>o r'" using ordLess_or_ordLeq assms by blast
-qed
-
-
-lemma finite_ordLess_infinite:
-assumes WELL: "Well_order r" and WELL': "Well_order r'" and
-        FIN: "finite(Field r)" and INF: "infinite(Field r')"
-shows "r <o r'"
-proof-
-  {assume "r' \<le>o r"
-   then obtain h where "inj_on h (Field r') \<and> h ` (Field r') \<le> Field r"
-   unfolding ordLeq_def using assms embed_inj_on embed_Field by blast
-   hence False using finite_imageD finite_subset FIN INF by metis
-  }
-  thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast
-qed
-
-
-lemma finite_well_order_on_ordIso:
-assumes FIN: "finite A" and
-        WELL: "well_order_on A r" and WELL': "well_order_on A r'"
-shows "r =o r'"
-proof-
-  have 0: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
-  using assms rel.well_order_on_Well_order by blast
-  moreover
-  have "\<forall>r r'. well_order_on A r \<and> well_order_on A r' \<and> r \<le>o r'
-                  \<longrightarrow> r =o r'"
-  proof(clarify)
-    fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'"
-    have 2: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
-    using * ** rel.well_order_on_Well_order by blast
-    assume "r \<le>o r'"
-    then obtain f where 1: "embed r r' f" and
-                        "inj_on f A \<and> f ` A \<le> A"
-    unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast
-    hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast
-    thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto
-  qed
-  ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast
-qed
-
-
-subsection{* @{text "<o"} is well-founded *}
-
-
-text {* Of course, it only makes sense to state that the @{text "<o"} is well-founded
-on the restricted type @{text "'a rel rel"}.  We prove this by first showing that, for any set
-of well-orders all embedded in a fixed well-order, the function mapping each well-order
-in the set to an order filter of the fixed well-order is compatible w.r.t. to @{text "<o"} versus
-{\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded. *}
-
-
-definition ord_to_filter :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a set"
-where "ord_to_filter r0 r \<equiv> (SOME f. embed r r0 f) ` (Field r)"
-
-
-lemma ord_to_filter_compat:
-"compat (ordLess Int (ordLess^-1``{r0} \<times> ordLess^-1``{r0}))
-        (ofilterIncl r0)
-        (ord_to_filter r0)"
-proof(unfold compat_def ord_to_filter_def, clarify)
-  fix r1::"'a rel" and r2::"'a rel"
-  let ?A1 = "Field r1"  let ?A2 ="Field r2" let ?A0 ="Field r0"
-  let ?phi10 = "\<lambda> f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f"
-  let ?phi20 = "\<lambda> f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f"
-  assume *: "r1 <o r0" "r2 <o r0" and **: "r1 <o r2"
-  hence "(\<exists>f. ?phi10 f) \<and> (\<exists>f. ?phi20 f)"
-  by (auto simp add: ordLess_def embedS_def)
-  hence "?phi10 ?f10 \<and> ?phi20 ?f20" by (auto simp add: someI_ex)
-  thus "(?f10 ` ?A1, ?f20 ` ?A2) \<in> ofilterIncl r0"
-  using * ** by (simp add: embed_ordLess_ofilterIncl)
-qed
-
-
-theorem wf_ordLess: "wf ordLess"
-proof-
-  {fix r0 :: "('a \<times> 'a) set"
-   (* need to annotate here!*)
-   let ?ordLess = "ordLess::('d rel * 'd rel) set"
-   let ?R = "?ordLess Int (?ordLess^-1``{r0} \<times> ?ordLess^-1``{r0})"
-   {assume Case1: "Well_order r0"
-    hence "wf ?R"
-    using wf_ofilterIncl[of r0]
-          compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"]
-          ord_to_filter_compat[of r0] by auto
-   }
-   moreover
-   {assume Case2: "\<not> Well_order r0"
-    hence "?R = {}" unfolding ordLess_def by auto
-    hence "wf ?R" using wf_empty by simp
-   }
-   ultimately have "wf ?R" by blast
-  }
-  thus ?thesis by (simp add: trans_wf_iff ordLess_trans)
-qed
-
-corollary exists_minim_Well_order:
-assumes NE: "R \<noteq> {}" and WELL: "\<forall>r \<in> R. Well_order r"
-shows "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
-proof-
-  obtain r where "r \<in> R \<and> (\<forall>r' \<in> R. \<not> r' <o r)"
-  using NE spec[OF spec[OF subst[OF wf_eq_minimal, of "%x. x", OF wf_ordLess]], of _ R]
-    equals0I[of R] by blast
-  with not_ordLeq_iff_ordLess WELL show ?thesis by blast
-qed
-
-
-
-subsection {* Copy via direct images  *}
-
-
-text{* The direct image operator is the dual of the inverse image operator @{text "inv_image"}
-from @{text "Relation.thy"}.  It is useful for transporting a well-order between
-different types. *}
-
-
-definition dir_image :: "'a rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> 'a' rel"
-where
-"dir_image r f = {(f a, f b)| a b. (a,b) \<in> r}"
-
-
-lemma dir_image_Field:
-"Field(dir_image r f) \<le> f ` (Field r)"
-unfolding dir_image_def Field_def by auto
-
-
-lemma dir_image_minus_Id:
-"inj_on f (Field r) \<Longrightarrow> (dir_image r f) - Id = dir_image (r - Id) f"
-unfolding inj_on_def Field_def dir_image_def by auto
-
-
-lemma Refl_dir_image:
-assumes "Refl r"
-shows "Refl(dir_image r f)"
-proof-
-  {fix a' b'
-   assume "(a',b') \<in> dir_image r f"
-   then obtain a b where 1: "a' = f a \<and> b' = f b \<and> (a,b) \<in> r"
-   unfolding dir_image_def by blast
-   hence "a \<in> Field r \<and> b \<in> Field r" using Field_def by fastforce
-   hence "(a,a) \<in> r \<and> (b,b) \<in> r" using assms by (simp add: refl_on_def)
-   with 1 have "(a',a') \<in> dir_image r f \<and> (b',b') \<in> dir_image r f"
-   unfolding dir_image_def by auto
-  }
-  thus ?thesis
-  by(unfold refl_on_def Field_def Domain_def Range_def, auto)
-qed
-
-
-lemma trans_dir_image:
-assumes TRANS: "trans r" and INJ: "inj_on f (Field r)"
-shows "trans(dir_image r f)"
-proof(unfold trans_def, auto)
-  fix a' b' c'
-  assume "(a',b') \<in> dir_image r f" "(b',c') \<in> dir_image r f"
-  then obtain a b1 b2 c where 1: "a' = f a \<and> b' = f b1 \<and> b' = f b2 \<and> c' = f c" and
-                         2: "(a,b1) \<in> r \<and> (b2,c) \<in> r"
-  unfolding dir_image_def by blast
-  hence "b1 \<in> Field r \<and> b2 \<in> Field r"
-  unfolding Field_def by auto
-  hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto
-  hence "(a,c): r" using 2 TRANS unfolding trans_def by blast
-  thus "(a',c') \<in> dir_image r f"
-  unfolding dir_image_def using 1 by auto
-qed
-
-
-lemma Preorder_dir_image:
-"\<lbrakk>Preorder r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Preorder (dir_image r f)"
-by (simp add: preorder_on_def Refl_dir_image trans_dir_image)
-
-
-lemma antisym_dir_image:
-assumes AN: "antisym r" and INJ: "inj_on f (Field r)"
-shows "antisym(dir_image r f)"
-proof(unfold antisym_def, auto)
-  fix a' b'
-  assume "(a',b') \<in> dir_image r f" "(b',a') \<in> dir_image r f"
-  then obtain a1 b1 a2 b2 where 1: "a' = f a1 \<and> a' = f a2 \<and> b' = f b1 \<and> b' = f b2" and
-                           2: "(a1,b1) \<in> r \<and> (b2,a2) \<in> r " and
-                           3: "{a1,a2,b1,b2} \<le> Field r"
-  unfolding dir_image_def Field_def by blast
-  hence "a1 = a2 \<and> b1 = b2" using INJ unfolding inj_on_def by auto
-  hence "a1 = b2" using 2 AN unfolding antisym_def by auto
-  thus "a' = b'" using 1 by auto
-qed
-
-
-lemma Partial_order_dir_image:
-"\<lbrakk>Partial_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Partial_order (dir_image r f)"
-by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image)
-
-
-lemma Total_dir_image:
-assumes TOT: "Total r" and INJ: "inj_on f (Field r)"
-shows "Total(dir_image r f)"
-proof(unfold total_on_def, intro ballI impI)
-  fix a' b'
-  assume "a' \<in> Field (dir_image r f)" "b' \<in> Field (dir_image r f)"
-  then obtain a and b where 1: "a \<in> Field r \<and> b \<in> Field r \<and> f a = a' \<and> f b = b'"
-  using dir_image_Field[of r f] by blast
-  moreover assume "a' \<noteq> b'"
-  ultimately have "a \<noteq> b" using INJ unfolding inj_on_def by auto
-  hence "(a,b) \<in> r \<or> (b,a) \<in> r" using 1 TOT unfolding total_on_def by auto
-  thus "(a',b') \<in> dir_image r f \<or> (b',a') \<in> dir_image r f"
-  using 1 unfolding dir_image_def by auto
-qed
-
-
-lemma Linear_order_dir_image:
-"\<lbrakk>Linear_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Linear_order (dir_image r f)"
-by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image)
-
-
-lemma wf_dir_image:
-assumes WF: "wf r" and INJ: "inj_on f (Field r)"
-shows "wf(dir_image r f)"
-proof(unfold wf_eq_minimal2, intro allI impI, elim conjE)
-  fix A'::"'b set"
-  assume SUB: "A' \<le> Field(dir_image r f)" and NE: "A' \<noteq> {}"
-  obtain A where A_def: "A = {a \<in> Field r. f a \<in> A'}" by blast
-  have "A \<noteq> {} \<and> A \<le> Field r"
-  using A_def dir_image_Field[of r f] SUB NE by blast
-  then obtain a where 1: "a \<in> A \<and> (\<forall>b \<in> A. (b,a) \<notin> r)"
-  using WF unfolding wf_eq_minimal2 by metis
-  have "\<forall>b' \<in> A'. (b',f a) \<notin> dir_image r f"
-  proof(clarify)
-    fix b' assume *: "b' \<in> A'" and **: "(b',f a) \<in> dir_image r f"
-    obtain b1 a1 where 2: "b' = f b1 \<and> f a = f a1" and
-                       3: "(b1,a1) \<in> r \<and> {a1,b1} \<le> Field r"
-    using ** unfolding dir_image_def Field_def by blast
-    hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto
-    hence "b1 \<in> A \<and> (b1,a) \<in> r" using 2 3 A_def * by auto
-    with 1 show False by auto
-  qed
-  thus "\<exists>a'\<in>A'. \<forall>b'\<in>A'. (b', a') \<notin> dir_image r f"
-  using A_def 1 by blast
-qed
-
-
-lemma Well_order_dir_image:
-"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Well_order (dir_image r f)"
-using assms unfolding well_order_on_def
-using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f]
-  dir_image_minus_Id[of f r]
-  subset_inj_on[of f "Field r" "Field(r - Id)"]
-  mono_Field[of "r - Id" r] by auto
-
-
-lemma dir_image_Field2:
-"Refl r \<Longrightarrow> Field(dir_image r f) = f ` (Field r)"
-unfolding Field_def dir_image_def refl_on_def Domain_def Range_def by blast
-
-
-lemma dir_image_bij_betw:
-"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> bij_betw f (Field r) (Field (dir_image r f))"
-unfolding bij_betw_def
-by (simp add: dir_image_Field2 order_on_defs)
-
-
-lemma dir_image_compat:
-"compat r (dir_image r f) f"
-unfolding compat_def dir_image_def by auto
-
-
-lemma dir_image_iso:
-"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk>  \<Longrightarrow> iso r (dir_image r f) f"
-using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast
-
-
-lemma dir_image_ordIso:
-"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk>  \<Longrightarrow> r =o dir_image r f"
-unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast
-
-
-lemma Well_order_iso_copy:
-assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'"
-shows "\<exists>r'. well_order_on A' r' \<and> r =o r'"
-proof-
-   let ?r' = "dir_image r f"
-   have 1: "A = Field r \<and> Well_order r"
-   using WELL rel.well_order_on_Well_order by blast
-   hence 2: "iso r ?r' f"
-   using dir_image_iso using BIJ unfolding bij_betw_def by auto
-   hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast
-   hence "Field ?r' = A'"
-   using 1 BIJ unfolding bij_betw_def by auto
-   moreover have "Well_order ?r'"
-   using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast
-   ultimately show ?thesis unfolding ordIso_def using 1 2 by blast
-qed
-
-
-
-subsection {* Bounded square  *}
-
-
-text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
-order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
-following criteria (in this order):
-\begin{itemize}
-\item compare the maximums;
-\item compare the first components;
-\item compare the second components.
-\end{itemize}
-%
-The only application of this construction that we are aware of is
-at proving that the square of an infinite set has the same cardinal
-as that set. The essential property required there (and which is ensured by this
-construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
-in a product of proper filters on the original relation (assumed to be a well-order). *}
-
-
-definition bsqr :: "'a rel => ('a * 'a)rel"
-where
-"bsqr r = {((a1,a2),(b1,b2)).
-           {a1,a2,b1,b2} \<le> Field r \<and>
-           (a1 = b1 \<and> a2 = b2 \<or>
-            (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1  \<and> (a2,b2) \<in> r - Id
-           )}"
-
-
-lemma Field_bsqr:
-"Field (bsqr r) = Field r \<times> Field r"
-proof
-  show "Field (bsqr r) \<le> Field r \<times> Field r"
-  proof-
-    {fix a1 a2 assume "(a1,a2) \<in> Field (bsqr r)"
-     moreover
-     have "\<And> b1 b2. ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r \<Longrightarrow>
-                      a1 \<in> Field r \<and> a2 \<in> Field r" unfolding bsqr_def by auto
-     ultimately have "a1 \<in> Field r \<and> a2 \<in> Field r" unfolding Field_def by auto
-    }
-    thus ?thesis unfolding Field_def by force
-  qed
-next
-  show "Field r \<times> Field r \<le> Field (bsqr r)"
-  proof(auto)
-    fix a1 a2 assume "a1 \<in> Field r" and "a2 \<in> Field r"
-    hence "((a1,a2),(a1,a2)) \<in> bsqr r" unfolding bsqr_def by blast
-    thus "(a1,a2) \<in> Field (bsqr r)" unfolding Field_def by auto
-  qed
-qed
-
-
-lemma bsqr_Refl: "Refl(bsqr r)"
-by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def)
-
-
-lemma bsqr_Trans:
-assumes "Well_order r"
-shows "trans (bsqr r)"
-proof(unfold trans_def, auto)
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Trans: "trans r" using wo_rel.TRANS by auto
-  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
-  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
-  (* Main proof *)
-  fix a1 a2 b1 b2 c1 c2
-  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(c1,c2)) \<in> bsqr r"
-  hence 0: "{a1,a2,b1,b2,c1,c2} \<le> Field r" unfolding bsqr_def by auto
-  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-  using * unfolding bsqr_def by auto
-  have 2: "b1 = c1 \<and> b2 = c2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
-  using ** unfolding bsqr_def by auto
-  show "((a1,a2),(c1,c2)) \<in> bsqr r"
-  proof-
-    {assume Case1: "a1 = b1 \<and> a2 = b2"
-     hence ?thesis using ** by simp
-    }
-    moreover
-    {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
-     {assume Case21: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      using Case2 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2"
-      hence ?thesis using Case2 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
-     {assume Case31: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence ?thesis using Case3 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
-      hence "(a1,c1) \<in> r - Id"
-      using Case3 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1"
-      hence ?thesis using Case3 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-     {assume Case41: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence ?thesis using Case4 0 unfolding bsqr_def by force
-     }
-     moreover
-     {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
-      hence ?thesis using Case4 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
-      hence "(a2,c2) \<in> r - Id"
-      using Case4 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    ultimately show ?thesis using 0 1 by auto
-  qed
-qed
-
-
-lemma bsqr_antisym:
-assumes "Well_order r"
-shows "antisym (bsqr r)"
-proof(unfold antisym_def, clarify)
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Trans: "trans r" using wo_rel.TRANS by auto
-  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
-  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
-  hence IrrS: "\<forall>a b. \<not>((a,b) \<in> r - Id \<and> (b,a) \<in> r - Id)"
-  using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast
-  (* Main proof *)
-  fix a1 a2 b1 b2
-  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(a1,a2)) \<in> bsqr r"
-  hence 0: "{a1,a2,b1,b2} \<le> Field r" unfolding bsqr_def by auto
-  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-  using * unfolding bsqr_def by auto
-  have 2: "b1 = a1 \<and> b2 = a2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> (b1,a1) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> b1 = a1 \<and> (b2,a2) \<in> r - Id"
-  using ** unfolding bsqr_def by auto
-  show "a1 = b1 \<and> a2 = b2"
-  proof-
-    {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
-     {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-      hence False using Case1 IrrS by blast
-     }
-     moreover
-     {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2"
-      hence False using Case1 by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
-     {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-       hence False using Case2 by auto
-     }
-     moreover
-     {assume Case22: "(b1,a1) \<in> r - Id"
-      hence False using Case2 IrrS by blast
-     }
-     moreover
-     {assume Case23: "b1 = a1"
-      hence False using Case2 by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-     moreover
-     {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-      hence False using Case3 by auto
-     }
-     moreover
-     {assume Case32: "(b1,a1) \<in> r - Id"
-      hence False using Case3 by auto
-     }
-     moreover
-     {assume Case33: "(b2,a2) \<in> r - Id"
-      hence False using Case3 IrrS by blast
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    ultimately show ?thesis using 0 1 by blast
-  qed
-qed
-
-
-lemma bsqr_Total:
-assumes "Well_order r"
-shows "Total(bsqr r)"
-proof-
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Total: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
-  using wo_rel.TOTALS by auto
-  (* Main proof *)
-  {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \<le> Field(bsqr r)"
-   hence 0: "a1 \<in> Field r \<and> a2 \<in> Field r \<and> b1 \<in> Field r \<and> b2 \<in> Field r"
-   using Field_bsqr by blast
-   have "((a1,a2) = (b1,b2) \<or> ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r)"
-   proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0)
-       (* Why didn't clarsimp simp add: Well 0 do the same job? *)
-     assume Case1: "(a1,a2) \<in> r"
-     hence 1: "wo_rel.max2 r a1 a2 = a2"
-     using Well 0 by (simp add: wo_rel.max2_equals2)
-     show ?thesis
-     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
-       assume Case11: "(b1,b2) \<in> r"
-       hence 2: "wo_rel.max2 r b1 b2 = b2"
-       using Well 0 by (simp add: wo_rel.max2_equals2)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-         assume Case111: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
-       next
-         assume Case112: "a2 = b2"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case1121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto
-         next
-           assume Case1122: "a1 = b1"
-           thus ?thesis using Case112 by auto
-         qed
-       qed
-     next
-       assume Case12: "(b2,b1) \<in> r"
-       hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0)
-         assume Case121: "(a2,b1) \<in> r - Id \<or> (b1,a2) \<in> r - Id"
-         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
-       next
-         assume Case122: "a2 = b1"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case1221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto
-         next
-           assume Case1222: "a1 = b1"
-           show ?thesis
-           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-             assume Case12221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-             thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto
-           next
-             assume Case12222: "a2 = b2"
-             thus ?thesis using Case122 Case1222 by auto
-           qed
-         qed
-       qed
-     qed
-   next
-     assume Case2: "(a2,a1) \<in> r"
-     hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1)
-     show ?thesis
-     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
-       assume Case21: "(b1,b2) \<in> r"
-       hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0)
-         assume Case211: "(a1,b2) \<in> r - Id \<or> (b2,a1) \<in> r - Id"
-         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
-       next
-         assume Case212: "a1 = b2"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case2121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto
-         next
-           assume Case2122: "a1 = b1"
-           show ?thesis
-           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-             assume Case21221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-             thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto
-           next
-             assume Case21222: "a2 = b2"
-             thus ?thesis using Case2122 Case212 by auto
-           qed
-         qed
-       qed
-     next
-       assume Case22: "(b2,b1) \<in> r"
-       hence 3: "wo_rel.max2 r b1 b2 = b1"  using Well 0 by (simp add: wo_rel.max2_equals1)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-         assume Case221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
-       next
-         assume Case222: "a1 = b1"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-           assume Case2221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-           thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto
-         next
-           assume Case2222: "a2 = b2"
-           thus ?thesis using Case222 by auto
-         qed
-       qed
-     qed
-   qed
-  }
-  thus ?thesis unfolding total_on_def by fast
-qed
-
-
-lemma bsqr_Linear_order:
-assumes "Well_order r"
-shows "Linear_order(bsqr r)"
-unfolding order_on_defs
-using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast
-
-
-lemma bsqr_Well_order:
-assumes "Well_order r"
-shows "Well_order(bsqr r)"
-using assms
-proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI)
-  have 0: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
-  using assms well_order_on_def Linear_order_Well_order_iff by blast
-  fix D assume *: "D \<le> Field (bsqr r)" and **: "D \<noteq> {}"
-  hence 1: "D \<le> Field r \<times> Field r" unfolding Field_bsqr by simp
-  (*  *)
-  obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \<in> D}" by blast
-  have "M \<noteq> {}" using 1 M_def ** by auto
-  moreover
-  have "M \<le> Field r" unfolding M_def
-  using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
-  ultimately obtain m where m_min: "m \<in> M \<and> (\<forall>a \<in> M. (m,a) \<in> r)"
-  using 0 by blast
-  (*  *)
-  obtain A1 where A1_def: "A1 = {a1. \<exists>a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
-  have "A1 \<le> Field r" unfolding A1_def using 1 by auto
-  moreover have "A1 \<noteq> {}" unfolding A1_def using m_min unfolding M_def by blast
-  ultimately obtain a1 where a1_min: "a1 \<in> A1 \<and> (\<forall>a \<in> A1. (a1,a) \<in> r)"
-  using 0 by blast
-  (*  *)
-  obtain A2 where A2_def: "A2 = {a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
-  have "A2 \<le> Field r" unfolding A2_def using 1 by auto
-  moreover have "A2 \<noteq> {}" unfolding A2_def
-  using m_min a1_min unfolding A1_def M_def by blast
-  ultimately obtain a2 where a2_min: "a2 \<in> A2 \<and> (\<forall>a \<in> A2. (a2,a) \<in> r)"
-  using 0 by blast
-  (*   *)
-  have 2: "wo_rel.max2 r a1 a2 = m"
-  using a1_min a2_min unfolding A1_def A2_def by auto
-  have 3: "(a1,a2) \<in> D" using a2_min unfolding A2_def by auto
-  (*  *)
-  moreover
-  {fix b1 b2 assume ***: "(b1,b2) \<in> D"
-   hence 4: "{a1,a2,b1,b2} \<le> Field r" using 1 3 by blast
-   have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
-   using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto
-   have "((a1,a2),(b1,b2)) \<in> bsqr r"
-   proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2")
-     assume Case1: "wo_rel.max2 r a1 a2 \<noteq> wo_rel.max2 r b1 b2"
-     thus ?thesis unfolding bsqr_def using 4 5 by auto
-   next
-     assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
-     hence "b1 \<in> A1" unfolding A1_def using 2 *** by auto
-     hence 6: "(a1,b1) \<in> r" using a1_min by auto
-     show ?thesis
-     proof(cases "a1 = b1")
-       assume Case21: "a1 \<noteq> b1"
-       thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto
-     next
-       assume Case22: "a1 = b1"
-       hence "b2 \<in> A2" unfolding A2_def using 2 *** Case2 by auto
-       hence 7: "(a2,b2) \<in> r" using a2_min by auto
-       thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto
-     qed
-   qed
-  }
-  (*  *)
-  ultimately show "\<exists>d \<in> D. \<forall>d' \<in> D. (d,d') \<in> bsqr r" by fastforce
-qed
-
-
-lemma bsqr_max2:
-assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \<in> bsqr r"
-shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
-proof-
-  have "{(a1,a2),(b1,b2)} \<le> Field(bsqr r)"
-  using LEQ unfolding Field_def by auto
-  hence "{a1,a2,b1,b2} \<le> Field r" unfolding Field_bsqr by auto
-  hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \<le> Field r"
-  using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
-  moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r \<or> wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
-  using LEQ unfolding bsqr_def by auto
-  ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto
-qed
-
-
-lemma bsqr_ofilter:
-assumes WELL: "Well_order r" and
-        OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \<times> Field r" and
-        NE: "\<not> (\<exists>a. Field r = rel.under r a)"
-shows "\<exists>A. wo_rel.ofilter r A \<and> A < Field r \<and> D \<le> A \<times> A"
-proof-
-  let ?r' = "bsqr r"
-  have Well: "wo_rel r" using WELL wo_rel_def by blast
-  hence Trans: "trans r" using wo_rel.TRANS by blast
-  have Well': "Well_order ?r' \<and> wo_rel ?r'"
-  using WELL bsqr_Well_order wo_rel_def by blast
-  (*  *)
-  have "D < Field ?r'" unfolding Field_bsqr using SUB .
-  with OF obtain a1 and a2 where
-  "(a1,a2) \<in> Field ?r'" and 1: "D = rel.underS ?r' (a1,a2)"
-  using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto
-  hence 2: "{a1,a2} \<le> Field r" unfolding Field_bsqr by auto
-  let ?m = "wo_rel.max2 r a1 a2"
-  have "D \<le> (rel.under r ?m) \<times> (rel.under r ?m)"
-  proof(unfold 1)
-    {fix b1 b2
-     let ?n = "wo_rel.max2 r b1 b2"
-     assume "(b1,b2) \<in> rel.underS ?r' (a1,a2)"
-     hence 3: "((b1,b2),(a1,a2)) \<in> ?r'"
-     unfolding rel.underS_def by blast
-     hence "(?n,?m) \<in> r" using WELL by (simp add: bsqr_max2)
-     moreover
-     {have "(b1,b2) \<in> Field ?r'" using 3 unfolding Field_def by auto
-      hence "{b1,b2} \<le> Field r" unfolding Field_bsqr by auto
-      hence "(b1,?n) \<in> r \<and> (b2,?n) \<in> r"
-      using Well by (simp add: wo_rel.max2_greater)
-     }
-     ultimately have "(b1,?m) \<in> r \<and> (b2,?m) \<in> r"
-     using Trans trans_def[of r] by blast
-     hence "(b1,b2) \<in> (rel.under r ?m) \<times> (rel.under r ?m)" unfolding rel.under_def by simp}
-     thus "rel.underS ?r' (a1,a2) \<le> (rel.under r ?m) \<times> (rel.under r ?m)" by auto
-  qed
-  moreover have "wo_rel.ofilter r (rel.under r ?m)"
-  using Well by (simp add: wo_rel.under_ofilter)
-  moreover have "rel.under r ?m < Field r"
-  using NE rel.under_Field[of r ?m] by blast
-  ultimately show ?thesis by blast
-qed
-
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Constructions_on_Wellorders_LFP.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -0,0 +1,1633 @@
+(*  Title:      HOL/Cardinals/Constructions_on_Wellorders_LFP.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Constructions on wellorders (LFP).
+*)
+
+header {* Constructions on Wellorders (LFP) *}
+
+theory Constructions_on_Wellorders_LFP
+imports Wellorder_Embedding_LFP
+begin
+
+
+text {* In this section, we study basic constructions on well-orders, such as restriction to
+a set/order filter, copy via direct images, ordinal-like sum of disjoint well-orders,
+and bounded square.  We also define between well-orders
+the relations @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"}),
+@{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"}), and
+@{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}).  We study the
+connections between these relations, order filters, and the aforementioned constructions.
+A main result of this section is that @{text "<o"} is well-founded.  *}
+
+
+subsection {* Restriction to a set  *}
+
+
+abbreviation Restr :: "'a rel \<Rightarrow> 'a set \<Rightarrow> 'a rel"
+where "Restr r A \<equiv> r Int (A \<times> A)"
+
+
+lemma Restr_subset:
+"A \<le> B \<Longrightarrow> Restr (Restr r B) A = Restr r A"
+by blast
+
+
+lemma Restr_Field: "Restr r (Field r) = r"
+unfolding Field_def by auto
+
+
+lemma Refl_Restr: "Refl r \<Longrightarrow> Refl(Restr r A)"
+unfolding refl_on_def Field_def by auto
+
+
+lemma antisym_Restr:
+"antisym r \<Longrightarrow> antisym(Restr r A)"
+unfolding antisym_def Field_def by auto
+
+
+lemma Total_Restr:
+"Total r \<Longrightarrow> Total(Restr r A)"
+unfolding total_on_def Field_def by auto
+
+
+lemma trans_Restr:
+"trans r \<Longrightarrow> trans(Restr r A)"
+unfolding trans_def Field_def by blast
+
+
+lemma Preorder_Restr:
+"Preorder r \<Longrightarrow> Preorder(Restr r A)"
+unfolding preorder_on_def by (simp add: Refl_Restr trans_Restr)
+
+
+lemma Partial_order_Restr:
+"Partial_order r \<Longrightarrow> Partial_order(Restr r A)"
+unfolding partial_order_on_def by (simp add: Preorder_Restr antisym_Restr)
+
+
+lemma Linear_order_Restr:
+"Linear_order r \<Longrightarrow> Linear_order(Restr r A)"
+unfolding linear_order_on_def by (simp add: Partial_order_Restr Total_Restr)
+
+
+lemma Well_order_Restr:
+assumes "Well_order r"
+shows "Well_order(Restr r A)"
+proof-
+  have "Restr r A - Id \<le> r - Id" using Restr_subset by blast
+  hence "wf(Restr r A - Id)" using assms
+  using well_order_on_def wf_subset by blast
+  thus ?thesis using assms unfolding well_order_on_def
+  by (simp add: Linear_order_Restr)
+qed
+
+
+lemma Field_Restr_subset: "Field(Restr r A) \<le> A"
+by (auto simp add: Field_def)
+
+
+lemma Refl_Field_Restr:
+"Refl r \<Longrightarrow> Field(Restr r A) = (Field r) Int A"
+by (auto simp add: refl_on_def Field_def)
+
+
+lemma Refl_Field_Restr2:
+"\<lbrakk>Refl r; A \<le> Field r\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
+by (auto simp add: Refl_Field_Restr)
+
+
+lemma well_order_on_Restr:
+assumes WELL: "Well_order r" and SUB: "A \<le> Field r"
+shows "well_order_on A (Restr r A)"
+using assms
+using Well_order_Restr[of r A] Refl_Field_Restr2[of r A]
+     order_on_defs[of "Field r" r] by auto
+
+
+subsection {* Order filters versus restrictions and embeddings  *}
+
+
+lemma Field_Restr_ofilter:
+"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> Field(Restr r A) = A"
+by (auto simp add: wo_rel_def wo_rel.ofilter_def wo_rel.REFL Refl_Field_Restr2)
+
+
+lemma ofilter_Restr_under:
+assumes WELL: "Well_order r" and OF: "wo_rel.ofilter r A" and IN: "a \<in> A"
+shows "rel.under (Restr r A) a = rel.under r a"
+using assms wo_rel_def
+proof(auto simp add: wo_rel.ofilter_def rel.under_def)
+  fix b assume *: "a \<in> A" and "(b,a) \<in> r"
+  hence "b \<in> rel.under r a \<and> a \<in> Field r"
+  unfolding rel.under_def using Field_def by fastforce
+  thus "b \<in> A" using * assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+qed
+
+
+lemma ofilter_embed:
+assumes "Well_order r"
+shows "wo_rel.ofilter r A = (A \<le> Field r \<and> embed (Restr r A) r id)"
+proof
+  assume *: "wo_rel.ofilter r A"
+  show "A \<le> Field r \<and> embed (Restr r A) r id"
+  proof(unfold embed_def, auto)
+    fix a assume "a \<in> A" thus "a \<in> Field r" using assms *
+    by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+  next
+    fix a assume "a \<in> Field (Restr r A)"
+    thus "bij_betw id (rel.under (Restr r A) a) (rel.under r a)" using assms *
+    by (simp add: ofilter_Restr_under Field_Restr_ofilter)
+  qed
+next
+  assume *: "A \<le> Field r \<and> embed (Restr r A) r id"
+  hence "Field(Restr r A) \<le> Field r"
+  using assms  embed_Field[of "Restr r A" r id] id_def
+        Well_order_Restr[of r] by auto
+  {fix a assume "a \<in> A"
+   hence "a \<in> Field(Restr r A)" using * assms
+   by (simp add: order_on_defs Refl_Field_Restr2)
+   hence "bij_betw id (rel.under (Restr r A) a) (rel.under r a)"
+   using * unfolding embed_def by auto
+   hence "rel.under r a \<le> rel.under (Restr r A) a"
+   unfolding bij_betw_def by auto
+   also have "\<dots> \<le> Field(Restr r A)" by (simp add: rel.under_Field)
+   also have "\<dots> \<le> A" by (simp add: Field_Restr_subset)
+   finally have "rel.under r a \<le> A" .
+  }
+  thus "wo_rel.ofilter r A" using assms * by (simp add: wo_rel_def wo_rel.ofilter_def)
+qed
+
+
+lemma ofilter_Restr_Int:
+assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A"
+shows "wo_rel.ofilter (Restr r B) (A Int B)"
+proof-
+  let ?rB = "Restr r B"
+  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+  hence Field: "Field ?rB = Field r Int B"
+  using Refl_Field_Restr by blast
+  have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
+  by (simp add: Well_order_Restr wo_rel_def)
+  (* Main proof *)
+  show ?thesis using WellB assms
+  proof(auto simp add: wo_rel.ofilter_def rel.under_def)
+    fix a assume "a \<in> A" and *: "a \<in> B"
+    hence "a \<in> Field r" using OFA Well by (auto simp add: wo_rel.ofilter_def)
+    with * show "a \<in> Field ?rB" using Field by auto
+  next
+    fix a b assume "a \<in> A" and "(b,a) \<in> r"
+    thus "b \<in> A" using Well OFA by (auto simp add: wo_rel.ofilter_def rel.under_def)
+  qed
+qed
+
+
+lemma ofilter_Restr_subset:
+assumes WELL: "Well_order r" and OFA: "wo_rel.ofilter r A" and SUB: "A \<le> B"
+shows "wo_rel.ofilter (Restr r B) A"
+proof-
+  have "A Int B = A" using SUB by blast
+  thus ?thesis using assms ofilter_Restr_Int[of r A B] by auto
+qed
+
+
+lemma ofilter_subset_embed:
+assumes WELL: "Well_order r" and
+        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A \<le> B) = (embed (Restr r A) (Restr r B) id)"
+proof-
+  let ?rA = "Restr r A"  let ?rB = "Restr r B"
+  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+  hence FieldA: "Field ?rA = Field r Int A"
+  using Refl_Field_Restr by blast
+  have FieldB: "Field ?rB = Field r Int B"
+  using Refl Refl_Field_Restr by blast
+  have WellA: "wo_rel ?rA \<and> Well_order ?rA" using WELL
+  by (simp add: Well_order_Restr wo_rel_def)
+  have WellB: "wo_rel ?rB \<and> Well_order ?rB" using WELL
+  by (simp add: Well_order_Restr wo_rel_def)
+  (* Main proof *)
+  show ?thesis
+  proof
+    assume *: "A \<le> B"
+    hence "wo_rel.ofilter (Restr r B) A" using assms
+    by (simp add: ofilter_Restr_subset)
+    hence "embed (Restr ?rB A) (Restr r B) id"
+    using WellB ofilter_embed[of "?rB" A] by auto
+    thus "embed (Restr r A) (Restr r B) id"
+    using * by (simp add: Restr_subset)
+  next
+    assume *: "embed (Restr r A) (Restr r B) id"
+    {fix a assume **: "a \<in> A"
+     hence "a \<in> Field r" using Well OFA by (auto simp add: wo_rel.ofilter_def)
+     with ** FieldA have "a \<in> Field ?rA" by auto
+     hence "a \<in> Field ?rB" using * WellA embed_Field[of ?rA ?rB id] by auto
+     hence "a \<in> B" using FieldB by auto
+    }
+    thus "A \<le> B" by blast
+  qed
+qed
+
+
+lemma ofilter_subset_embedS_iso:
+assumes WELL: "Well_order r" and
+        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "((A < B) = (embedS (Restr r A) (Restr r B) id)) \<and>
+       ((A = B) = (iso (Restr r A) (Restr r B) id))"
+proof-
+  let ?rA = "Restr r A"  let ?rB = "Restr r B"
+  have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+  hence Refl: "Refl r" by (simp add: wo_rel.REFL)
+  hence "Field ?rA = Field r Int A"
+  using Refl_Field_Restr by blast
+  hence FieldA: "Field ?rA = A" using OFA Well
+  by (auto simp add: wo_rel.ofilter_def)
+  have "Field ?rB = Field r Int B"
+  using Refl Refl_Field_Restr by blast
+  hence FieldB: "Field ?rB = B" using OFB Well
+  by (auto simp add: wo_rel.ofilter_def)
+  (* Main proof *)
+  show ?thesis unfolding embedS_def iso_def
+  using assms ofilter_subset_embed[of r A B]
+        FieldA FieldB bij_betw_id_iff[of A B] by auto
+qed
+
+
+lemma ofilter_subset_embedS:
+assumes WELL: "Well_order r" and
+        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A < B) = embedS (Restr r A) (Restr r B) id"
+using assms by (simp add: ofilter_subset_embedS_iso)
+
+
+lemma embed_implies_iso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+        EMB: "embed r' r f"
+shows "iso r' (Restr r (f ` (Field r'))) f"
+proof-
+  let ?A' = "Field r'"
+  let ?r'' = "Restr r (f ` ?A')"
+  have 0: "Well_order ?r''" using WELL Well_order_Restr by blast
+  have 1: "wo_rel.ofilter r (f ` ?A')" using assms embed_Field_ofilter  by blast
+  hence "Field ?r'' = f ` (Field r')" using WELL Field_Restr_ofilter by blast
+  hence "bij_betw f ?A' (Field ?r'')"
+  using EMB embed_inj_on WELL' unfolding bij_betw_def by blast
+  moreover
+  {have "\<forall>a b. (a,b) \<in> r' \<longrightarrow> a \<in> Field r' \<and> b \<in> Field r'"
+   unfolding Field_def by auto
+   hence "compat r' ?r'' f"
+   using assms embed_iff_compat_inj_on_ofilter
+   unfolding compat_def by blast
+  }
+  ultimately show ?thesis using WELL' 0 iso_iff3 by blast
+qed
+
+
+subsection {* The strict inclusion on proper ofilters is well-founded *}
+
+
+definition ofilterIncl :: "'a rel \<Rightarrow> 'a set rel"
+where
+"ofilterIncl r \<equiv> {(A,B). wo_rel.ofilter r A \<and> A \<noteq> Field r \<and>
+                         wo_rel.ofilter r B \<and> B \<noteq> Field r \<and> A < B}"
+
+
+lemma wf_ofilterIncl:
+assumes WELL: "Well_order r"
+shows "wf(ofilterIncl r)"
+proof-
+  have Well: "wo_rel r" using WELL by (simp add: wo_rel_def)
+  hence Lo: "Linear_order r" by (simp add: wo_rel.LIN)
+  let ?h = "(\<lambda> A. wo_rel.suc r A)"
+  let ?rS = "r - Id"
+  have "wf ?rS" using WELL by (simp add: order_on_defs)
+  moreover
+  have "compat (ofilterIncl r) ?rS ?h"
+  proof(unfold compat_def ofilterIncl_def,
+        intro allI impI, simp, elim conjE)
+    fix A B
+    assume *: "wo_rel.ofilter r A" "A \<noteq> Field r" and
+           **: "wo_rel.ofilter r B" "B \<noteq> Field r" and ***: "A < B"
+    then obtain a and b where 0: "a \<in> Field r \<and> b \<in> Field r" and
+                         1: "A = rel.underS r a \<and> B = rel.underS r b"
+    using Well by (auto simp add: wo_rel.ofilter_underS_Field)
+    hence "a \<noteq> b" using *** by auto
+    moreover
+    have "(a,b) \<in> r" using 0 1 Lo ***
+    by (auto simp add: rel.underS_incl_iff)
+    moreover
+    have "a = wo_rel.suc r A \<and> b = wo_rel.suc r B"
+    using Well 0 1 by (simp add: wo_rel.suc_underS)
+    ultimately
+    show "(wo_rel.suc r A, wo_rel.suc r B) \<in> r \<and> wo_rel.suc r A \<noteq> wo_rel.suc r B"
+    by simp
+  qed
+  ultimately show "wf (ofilterIncl r)" by (simp add: compat_wf)
+qed
+
+
+
+subsection {* Ordering the  well-orders by existence of embeddings *}
+
+
+text {* We define three relations between well-orders:
+\begin{itemize}
+\item @{text "ordLeq"}, of being embedded (abbreviated @{text "\<le>o"});
+\item @{text "ordLess"}, of being strictly embedded (abbreviated @{text "<o"});
+\item @{text "ordIso"}, of being isomorphic (abbreviated @{text "=o"}).
+\end{itemize}
+%
+The prefix "ord" and the index "o" in these names stand for "ordinal-like".
+These relations shall be proved to be inter-connected in a similar fashion as the trio
+@{text "\<le>"}, @{text "<"}, @{text "="} associated to a total order on a set.
+*}
+
+
+definition ordLeq :: "('a rel * 'a' rel) set"
+where
+"ordLeq = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embed r r' f)}"
+
+
+abbreviation ordLeq2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<=o" 50)
+where "r <=o r' \<equiv> (r,r') \<in> ordLeq"
+
+
+abbreviation ordLeq3 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "\<le>o" 50)
+where "r \<le>o r' \<equiv> r <=o r'"
+
+
+definition ordLess :: "('a rel * 'a' rel) set"
+where
+"ordLess = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. embedS r r' f)}"
+
+abbreviation ordLess2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "<o" 50)
+where "r <o r' \<equiv> (r,r') \<in> ordLess"
+
+
+definition ordIso :: "('a rel * 'a' rel) set"
+where
+"ordIso = {(r,r'). Well_order r \<and> Well_order r' \<and> (\<exists>f. iso r r' f)}"
+
+abbreviation ordIso2 :: "'a rel \<Rightarrow> 'a' rel \<Rightarrow> bool" (infix "=o" 50)
+where "r =o r' \<equiv> (r,r') \<in> ordIso"
+
+
+lemmas ordRels_def = ordLeq_def ordLess_def ordIso_def
+
+lemma ordLeq_Well_order_simp:
+assumes "r \<le>o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordLeq_def by simp
+
+
+lemma ordLess_Well_order_simp:
+assumes "r <o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordLess_def by simp
+
+
+lemma ordIso_Well_order_simp:
+assumes "r =o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordIso_def by simp
+
+
+text{* Notice that the relations @{text "\<le>o"}, @{text "<o"}, @{text "=o"} connect well-orders
+on potentially {\em distinct} types. However, some of the lemmas below, including the next one,
+restrict implicitly the type of these relations to @{text "(('a rel) * ('a rel)) set"} , i.e.,
+to @{text "'a rel rel"}.  *}
+
+
+lemma ordLeq_reflexive:
+"Well_order r \<Longrightarrow> r \<le>o r"
+unfolding ordLeq_def using id_embed[of r] by blast
+
+
+lemma ordLeq_transitive[trans]:
+assumes *: "r \<le>o r'" and **: "r' \<le>o r''"
+shows "r \<le>o r''"
+proof-
+  obtain f and f'
+  where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
+        "embed r r' f" and "embed r' r'' f'"
+  using * ** unfolding ordLeq_def by blast
+  hence "embed r r'' (f' o f)"
+  using comp_embed[of r r' f r'' f'] by auto
+  thus "r \<le>o r''" unfolding ordLeq_def using 1 by auto
+qed
+
+
+lemma ordLeq_total:
+"\<lbrakk>Well_order r; Well_order r'\<rbrakk> \<Longrightarrow> r \<le>o r' \<or> r' \<le>o r"
+unfolding ordLeq_def using wellorders_totally_ordered by blast
+
+
+lemma ordIso_reflexive:
+"Well_order r \<Longrightarrow> r =o r"
+unfolding ordIso_def using id_iso[of r] by blast
+
+
+lemma ordIso_transitive[trans]:
+assumes *: "r =o r'" and **: "r' =o r''"
+shows "r =o r''"
+proof-
+  obtain f and f'
+  where 1: "Well_order r \<and> Well_order r' \<and> Well_order r''" and
+        "iso r r' f" and 3: "iso r' r'' f'"
+  using * ** unfolding ordIso_def by auto
+  hence "iso r r'' (f' o f)"
+  using comp_iso[of r r' f r'' f'] by auto
+  thus "r =o r''" unfolding ordIso_def using 1 by auto
+qed
+
+
+lemma ordIso_symmetric:
+assumes *: "r =o r'"
+shows "r' =o r"
+proof-
+  obtain f where 1: "Well_order r \<and> Well_order r'" and
+                 2: "embed r r' f \<and> bij_betw f (Field r) (Field r')"
+  using * by (auto simp add: ordIso_def iso_def)
+  let ?f' = "inv_into (Field r) f"
+  have "embed r' r ?f' \<and> bij_betw ?f' (Field r') (Field r)"
+  using 1 2 by (simp add: bij_betw_inv_into inv_into_Field_embed_bij_betw)
+  thus "r' =o r" unfolding ordIso_def using 1 by (auto simp add: iso_def)
+qed
+
+
+lemma ordLeq_ordLess_trans[trans]:
+assumes "r \<le>o r'" and " r' <o r''"
+shows "r <o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLeq_def ordLess_def by auto
+  thus ?thesis using assms unfolding ordLeq_def ordLess_def
+  using embed_comp_embedS by blast
+qed
+
+
+lemma ordLess_ordLeq_trans[trans]:
+assumes "r <o r'" and " r' \<le>o r''"
+shows "r <o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLeq_def ordLess_def by auto
+  thus ?thesis using assms unfolding ordLeq_def ordLess_def
+  using embedS_comp_embed by blast
+qed
+
+
+lemma ordLeq_ordIso_trans[trans]:
+assumes "r \<le>o r'" and " r' =o r''"
+shows "r \<le>o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLeq_def ordIso_def by auto
+  thus ?thesis using assms unfolding ordLeq_def ordIso_def
+  using embed_comp_iso by blast
+qed
+
+
+lemma ordIso_ordLeq_trans[trans]:
+assumes "r =o r'" and " r' \<le>o r''"
+shows "r \<le>o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLeq_def ordIso_def by auto
+  thus ?thesis using assms unfolding ordLeq_def ordIso_def
+  using iso_comp_embed by blast
+qed
+
+
+lemma ordLess_ordIso_trans[trans]:
+assumes "r <o r'" and " r' =o r''"
+shows "r <o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLess_def ordIso_def by auto
+  thus ?thesis using assms unfolding ordLess_def ordIso_def
+  using embedS_comp_iso by blast
+qed
+
+
+lemma ordIso_ordLess_trans[trans]:
+assumes "r =o r'" and " r' <o r''"
+shows "r <o r''"
+proof-
+  have "Well_order r \<and> Well_order r''"
+  using assms unfolding ordLess_def ordIso_def by auto
+  thus ?thesis using assms unfolding ordLess_def ordIso_def
+  using iso_comp_embedS by blast
+qed
+
+
+lemma ordLess_not_embed:
+assumes "r <o r'"
+shows "\<not>(\<exists>f'. embed r' r f')"
+proof-
+  obtain f where 1: "Well_order r \<and> Well_order r'" and 2: "embed r r' f" and
+                 3: " \<not> bij_betw f (Field r) (Field r')"
+  using assms unfolding ordLess_def by (auto simp add: embedS_def)
+  {fix f' assume *: "embed r' r f'"
+   hence "bij_betw f (Field r) (Field r')" using 1 2
+   by (simp add: embed_bothWays_Field_bij_betw)
+   with 3 have False by contradiction
+  }
+  thus ?thesis by blast
+qed
+
+
+lemma ordLess_Field:
+assumes OL: "r1 <o r2" and EMB: "embed r1 r2 f"
+shows "\<not> (f`(Field r1) = Field r2)"
+proof-
+  let ?A1 = "Field r1"  let ?A2 = "Field r2"
+  obtain g where
+  0: "Well_order r1 \<and> Well_order r2" and
+  1: "embed r1 r2 g \<and> \<not>(bij_betw g ?A1 ?A2)"
+  using OL unfolding ordLess_def by (auto simp add: embedS_def)
+  hence "\<forall>a \<in> ?A1. f a = g a"
+  using 0 EMB embed_unique[of r1] by auto
+  hence "\<not>(bij_betw f ?A1 ?A2)"
+  using 1 bij_betw_cong[of ?A1] by blast
+  moreover
+  have "inj_on f ?A1" using EMB 0 by (simp add: embed_inj_on)
+  ultimately show ?thesis by (simp add: bij_betw_def)
+qed
+
+
+lemma ordLess_iff:
+"r <o r' = (Well_order r \<and> Well_order r' \<and> \<not>(\<exists>f'. embed r' r f'))"
+proof
+  assume *: "r <o r'"
+  hence "\<not>(\<exists>f'. embed r' r f')" using ordLess_not_embed[of r r'] by simp
+  with * show "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
+  unfolding ordLess_def by auto
+next
+  assume *: "Well_order r \<and> Well_order r' \<and> \<not> (\<exists>f'. embed r' r f')"
+  then obtain f where 1: "embed r r' f"
+  using wellorders_totally_ordered[of r r'] by blast
+  moreover
+  {assume "bij_betw f (Field r) (Field r')"
+   with * 1 have "embed r' r (inv_into (Field r) f) "
+   using inv_into_Field_embed_bij_betw[of r r' f] by auto
+   with * have False by blast
+  }
+  ultimately show "(r,r') \<in> ordLess"
+  unfolding ordLess_def using * by (fastforce simp add: embedS_def)
+qed
+
+
+lemma ordLess_irreflexive: "\<not> r <o r"
+proof
+  assume "r <o r"
+  hence "Well_order r \<and>  \<not>(\<exists>f. embed r r f)"
+  unfolding ordLess_iff ..
+  moreover have "embed r r id" using id_embed[of r] .
+  ultimately show False by blast
+qed
+
+
+lemma ordLeq_iff_ordLess_or_ordIso:
+"r \<le>o r' = (r <o r' \<or> r =o r')"
+unfolding ordRels_def embedS_defs iso_defs by blast
+
+
+lemma ordIso_iff_ordLeq:
+"(r =o r') = (r \<le>o r' \<and> r' \<le>o r)"
+proof
+  assume "r =o r'"
+  then obtain f where 1: "Well_order r \<and> Well_order r' \<and>
+                     embed r r' f \<and> bij_betw f (Field r) (Field r')"
+  unfolding ordIso_def iso_defs by auto
+  hence "embed r r' f \<and> embed r' r (inv_into (Field r) f)"
+  by (simp add: inv_into_Field_embed_bij_betw)
+  thus  "r \<le>o r' \<and> r' \<le>o r"
+  unfolding ordLeq_def using 1 by auto
+next
+  assume "r \<le>o r' \<and> r' \<le>o r"
+  then obtain f and g where 1: "Well_order r \<and> Well_order r' \<and>
+                           embed r r' f \<and> embed r' r g"
+  unfolding ordLeq_def by auto
+  hence "iso r r' f" by (auto simp add: embed_bothWays_iso)
+  thus "r =o r'" unfolding ordIso_def using 1 by auto
+qed
+
+
+lemma not_ordLess_ordLeq:
+"r <o r' \<Longrightarrow> \<not> r' \<le>o r"
+using ordLess_ordLeq_trans ordLess_irreflexive by blast
+
+
+lemma ordLess_or_ordLeq:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r <o r' \<or> r' \<le>o r"
+proof-
+  have "r \<le>o r' \<or> r' \<le>o r"
+  using assms by (simp add: ordLeq_total)
+  moreover
+  {assume "\<not> r <o r' \<and> r \<le>o r'"
+   hence "r =o r'" using ordLeq_iff_ordLess_or_ordIso by blast
+   hence "r' \<le>o r" using ordIso_symmetric ordIso_iff_ordLeq by blast
+  }
+  ultimately show ?thesis by blast
+qed
+
+
+lemma not_ordLess_ordIso:
+"r <o r' \<Longrightarrow> \<not> r =o r'"
+using assms ordLess_ordIso_trans ordIso_symmetric ordLess_irreflexive by blast
+
+
+lemma not_ordLeq_iff_ordLess:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<not> r' \<le>o r) = (r <o r')"
+using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
+
+
+lemma not_ordLess_iff_ordLeq:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<not> r' <o r) = (r \<le>o r')"
+using assms not_ordLess_ordLeq ordLess_or_ordLeq by blast
+
+
+lemma ordLess_transitive[trans]:
+"\<lbrakk>r <o r'; r' <o r''\<rbrakk> \<Longrightarrow> r <o r''"
+using assms ordLess_ordLeq_trans ordLeq_iff_ordLess_or_ordIso by blast
+
+
+corollary ordLess_trans: "trans ordLess"
+unfolding trans_def using ordLess_transitive by blast
+
+
+lemmas ordIso_equivalence = ordIso_transitive ordIso_reflexive ordIso_symmetric
+
+
+lemma ordIso_imp_ordLeq:
+"r =o r' \<Longrightarrow> r \<le>o r'"
+using ordIso_iff_ordLeq by blast
+
+
+lemma ordLess_imp_ordLeq:
+"r <o r' \<Longrightarrow> r \<le>o r'"
+using ordLeq_iff_ordLess_or_ordIso by blast
+
+
+lemma ofilter_subset_ordLeq:
+assumes WELL: "Well_order r" and
+        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A \<le> B) = (Restr r A \<le>o Restr r B)"
+proof
+  assume "A \<le> B"
+  thus "Restr r A \<le>o Restr r B"
+  unfolding ordLeq_def using assms
+  Well_order_Restr Well_order_Restr ofilter_subset_embed by blast
+next
+  assume *: "Restr r A \<le>o Restr r B"
+  then obtain f where "embed (Restr r A) (Restr r B) f"
+  unfolding ordLeq_def by blast
+  {assume "B < A"
+   hence "Restr r B <o Restr r A"
+   unfolding ordLess_def using assms
+   Well_order_Restr Well_order_Restr ofilter_subset_embedS by blast
+   hence False using * not_ordLess_ordLeq by blast
+  }
+  thus "A \<le> B" using OFA OFB WELL
+  wo_rel_def[of r] wo_rel.ofilter_linord[of r A B] by blast
+qed
+
+
+lemma ofilter_subset_ordLess:
+assumes WELL: "Well_order r" and
+        OFA: "wo_rel.ofilter r A" and OFB: "wo_rel.ofilter r B"
+shows "(A < B) = (Restr r A <o Restr r B)"
+proof-
+  let ?rA = "Restr r A" let ?rB = "Restr r B"
+  have 1: "Well_order ?rA \<and> Well_order ?rB"
+  using WELL Well_order_Restr by blast
+  have "(A < B) = (\<not> B \<le> A)" using assms
+  wo_rel_def wo_rel.ofilter_linord[of r A B] by blast
+  also have "\<dots> = (\<not> Restr r B \<le>o Restr r A)"
+  using assms ofilter_subset_ordLeq by blast
+  also have "\<dots> = (Restr r A <o Restr r B)"
+  using 1 not_ordLeq_iff_ordLess by blast
+  finally show ?thesis .
+qed
+
+
+lemma ofilter_ordLess:
+"\<lbrakk>Well_order r; wo_rel.ofilter r A\<rbrakk> \<Longrightarrow> (A < Field r) = (Restr r A <o r)"
+by (simp add: ofilter_subset_ordLess wo_rel.Field_ofilter
+    wo_rel_def Restr_Field)
+
+
+corollary underS_Restr_ordLess:
+assumes "Well_order r" and "Field r \<noteq> {}"
+shows "Restr r (rel.underS r a) <o r"
+proof-
+  have "rel.underS r a < Field r" using assms
+  by (simp add: rel.underS_Field3)
+  thus ?thesis using assms
+  by (simp add: ofilter_ordLess wo_rel.underS_ofilter wo_rel_def)
+qed
+
+
+lemma embed_ordLess_ofilterIncl:
+assumes
+  OL12: "r1 <o r2" and OL23: "r2 <o r3" and
+  EMB13: "embed r1 r3 f13" and EMB23: "embed r2 r3 f23"
+shows "(f13`(Field r1), f23`(Field r2)) \<in> (ofilterIncl r3)"
+proof-
+  have OL13: "r1 <o r3"
+  using OL12 OL23 using ordLess_transitive by auto
+  let ?A1 = "Field r1"  let ?A2 ="Field r2" let ?A3 ="Field r3"
+  obtain f12 g23 where
+  0: "Well_order r1 \<and> Well_order r2 \<and> Well_order r3" and
+  1: "embed r1 r2 f12 \<and> \<not>(bij_betw f12 ?A1 ?A2)" and
+  2: "embed r2 r3 g23 \<and> \<not>(bij_betw g23 ?A2 ?A3)"
+  using OL12 OL23 by (auto simp add: ordLess_def embedS_def)
+  hence "\<forall>a \<in> ?A2. f23 a = g23 a"
+  using EMB23 embed_unique[of r2 r3] by blast
+  hence 3: "\<not>(bij_betw f23 ?A2 ?A3)"
+  using 2 bij_betw_cong[of ?A2 f23 g23] by blast
+  (*  *)
+  have 4: "wo_rel.ofilter r2 (f12 ` ?A1) \<and> f12 ` ?A1 \<noteq> ?A2"
+  using 0 1 OL12 by (simp add: embed_Field_ofilter ordLess_Field)
+  have 5: "wo_rel.ofilter r3 (f23 ` ?A2) \<and> f23 ` ?A2 \<noteq> ?A3"
+  using 0 EMB23 OL23 by (simp add: embed_Field_ofilter ordLess_Field)
+  have 6: "wo_rel.ofilter r3 (f13 ` ?A1)  \<and> f13 ` ?A1 \<noteq> ?A3"
+  using 0 EMB13 OL13 by (simp add: embed_Field_ofilter ordLess_Field)
+  (*  *)
+  have "f12 ` ?A1 < ?A2"
+  using 0 4 by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+  moreover have "inj_on f23 ?A2"
+  using EMB23 0 by (simp add: wo_rel_def embed_inj_on)
+  ultimately
+  have "f23 ` (f12 ` ?A1) < f23 ` ?A2" by (simp add: inj_on_strict_subset)
+  moreover
+  {have "embed r1 r3 (f23 o f12)"
+   using 1 EMB23 0 by (auto simp add: comp_embed)
+   hence "\<forall>a \<in> ?A1. f23(f12 a) = f13 a"
+   using EMB13 0 embed_unique[of r1 r3 "f23 o f12" f13] by auto
+   hence "f23 ` (f12 ` ?A1) = f13 ` ?A1" by force
+  }
+  ultimately
+  have "f13 ` ?A1 < f23 ` ?A2" by simp
+  (*  *)
+  with 5 6 show ?thesis
+  unfolding ofilterIncl_def by auto
+qed
+
+
+lemma ordLess_iff_ordIso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(r' <o r) = (\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a))"
+proof(auto)
+  fix a assume *: "a \<in> Field r" and **: "r' =o Restr r (rel.underS r a)"
+  hence "Restr r (rel.underS r a) <o r" using WELL underS_Restr_ordLess[of r] by blast
+  thus "r' <o r" using ** ordIso_ordLess_trans by blast
+next
+  assume "r' <o r"
+  then obtain f where 1: "Well_order r \<and> Well_order r'" and
+                      2: "embed r' r f \<and> f ` (Field r') \<noteq> Field r"
+  unfolding ordLess_def embedS_def[abs_def] bij_betw_def using embed_inj_on by blast
+  hence "wo_rel.ofilter r (f ` (Field r'))" using embed_Field_ofilter by blast
+  then obtain a where 3: "a \<in> Field r" and 4: "rel.underS r a = f ` (Field r')"
+  using 1 2 by (auto simp add: wo_rel.ofilter_underS_Field wo_rel_def)
+  have "iso r' (Restr r (f ` (Field r'))) f"
+  using embed_implies_iso_Restr 2 assms by blast
+  moreover have "Well_order (Restr r (f ` (Field r')))"
+  using WELL Well_order_Restr by blast
+  ultimately have "r' =o Restr r (f ` (Field r'))"
+  using WELL' unfolding ordIso_def by auto
+  hence "r' =o Restr r (rel.underS r a)" using 4 by auto
+  thus "\<exists>a \<in> Field r. r' =o Restr r (rel.underS r a)" using 3 by auto
+qed
+
+
+lemma internalize_ordLess:
+"(r' <o r) = (\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r)"
+proof
+  assume *: "r' <o r"
+  hence 0: "Well_order r \<and> Well_order r'" unfolding ordLess_def by auto
+  with * obtain a where 1: "a \<in> Field r" and 2: "r' =o Restr r (rel.underS r a)"
+  using ordLess_iff_ordIso_Restr by blast
+  let ?p = "Restr r (rel.underS r a)"
+  have "wo_rel.ofilter r (rel.underS r a)" using 0
+  by (simp add: wo_rel_def wo_rel.underS_ofilter)
+  hence "Field ?p = rel.underS r a" using 0 Field_Restr_ofilter by blast
+  hence "Field ?p < Field r" using rel.underS_Field2 1 by fastforce
+  moreover have "?p <o r" using underS_Restr_ordLess[of r a] 0 1 by blast
+  ultimately
+  show "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r" using 2 by blast
+next
+  assume "\<exists>p. Field p < Field r \<and> r' =o p \<and> p <o r"
+  thus "r' <o r" using ordIso_ordLess_trans by blast
+qed
+
+
+lemma internalize_ordLeq:
+"(r' \<le>o r) = (\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r)"
+proof
+  assume *: "r' \<le>o r"
+  moreover
+  {assume "r' <o r"
+   then obtain p where "Field p < Field r \<and> r' =o p \<and> p <o r"
+   using internalize_ordLess[of r' r] by blast
+   hence "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+   using ordLeq_iff_ordLess_or_ordIso by blast
+  }
+  moreover
+  have "r \<le>o r" using * ordLeq_def ordLeq_reflexive by blast
+  ultimately show "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+  using ordLeq_iff_ordLess_or_ordIso by blast
+next
+  assume "\<exists>p. Field p \<le> Field r \<and> r' =o p \<and> p \<le>o r"
+  thus "r' \<le>o r" using ordIso_ordLeq_trans by blast
+qed
+
+
+lemma ordLeq_iff_ordLess_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(r \<le>o r') = (\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r')"
+proof(auto)
+  assume *: "r \<le>o r'"
+  fix a assume "a \<in> Field r"
+  hence "Restr r (rel.underS r a) <o r"
+  using WELL underS_Restr_ordLess[of r] by blast
+  thus "Restr r (rel.underS r a) <o r'"
+  using * ordLess_ordLeq_trans by blast
+next
+  assume *: "\<forall>a \<in> Field r. Restr r (rel.underS r a) <o r'"
+  {assume "r' <o r"
+   then obtain a where "a \<in> Field r \<and> r' =o Restr r (rel.underS r a)"
+   using assms ordLess_iff_ordIso_Restr by blast
+   hence False using * not_ordLess_ordIso ordIso_symmetric by blast
+  }
+  thus "r \<le>o r'" using ordLess_or_ordLeq assms by blast
+qed
+
+
+lemma finite_ordLess_infinite:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and
+        FIN: "finite(Field r)" and INF: "infinite(Field r')"
+shows "r <o r'"
+proof-
+  {assume "r' \<le>o r"
+   then obtain h where "inj_on h (Field r') \<and> h ` (Field r') \<le> Field r"
+   unfolding ordLeq_def using assms embed_inj_on embed_Field by blast
+   hence False using finite_imageD finite_subset FIN INF by metis
+  }
+  thus ?thesis using WELL WELL' ordLess_or_ordLeq by blast
+qed
+
+
+lemma finite_well_order_on_ordIso:
+assumes FIN: "finite A" and
+        WELL: "well_order_on A r" and WELL': "well_order_on A r'"
+shows "r =o r'"
+proof-
+  have 0: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
+  using assms rel.well_order_on_Well_order by blast
+  moreover
+  have "\<forall>r r'. well_order_on A r \<and> well_order_on A r' \<and> r \<le>o r'
+                  \<longrightarrow> r =o r'"
+  proof(clarify)
+    fix r r' assume *: "well_order_on A r" and **: "well_order_on A r'"
+    have 2: "Well_order r \<and> Well_order r' \<and> Field r = A \<and> Field r' = A"
+    using * ** rel.well_order_on_Well_order by blast
+    assume "r \<le>o r'"
+    then obtain f where 1: "embed r r' f" and
+                        "inj_on f A \<and> f ` A \<le> A"
+    unfolding ordLeq_def using 2 embed_inj_on embed_Field by blast
+    hence "bij_betw f A A" unfolding bij_betw_def using FIN endo_inj_surj by blast
+    thus "r =o r'" unfolding ordIso_def iso_def[abs_def] using 1 2 by auto
+  qed
+  ultimately show ?thesis using assms ordLeq_total ordIso_symmetric by blast
+qed
+
+
+subsection{* @{text "<o"} is well-founded *}
+
+
+text {* Of course, it only makes sense to state that the @{text "<o"} is well-founded
+on the restricted type @{text "'a rel rel"}.  We prove this by first showing that, for any set
+of well-orders all embedded in a fixed well-order, the function mapping each well-order
+in the set to an order filter of the fixed well-order is compatible w.r.t. to @{text "<o"} versus
+{\em strict inclusion}; and we already know that strict inclusion of order filters is well-founded. *}
+
+
+definition ord_to_filter :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a set"
+where "ord_to_filter r0 r \<equiv> (SOME f. embed r r0 f) ` (Field r)"
+
+
+lemma ord_to_filter_compat:
+"compat (ordLess Int (ordLess^-1``{r0} \<times> ordLess^-1``{r0}))
+        (ofilterIncl r0)
+        (ord_to_filter r0)"
+proof(unfold compat_def ord_to_filter_def, clarify)
+  fix r1::"'a rel" and r2::"'a rel"
+  let ?A1 = "Field r1"  let ?A2 ="Field r2" let ?A0 ="Field r0"
+  let ?phi10 = "\<lambda> f10. embed r1 r0 f10" let ?f10 = "SOME f. ?phi10 f"
+  let ?phi20 = "\<lambda> f20. embed r2 r0 f20" let ?f20 = "SOME f. ?phi20 f"
+  assume *: "r1 <o r0" "r2 <o r0" and **: "r1 <o r2"
+  hence "(\<exists>f. ?phi10 f) \<and> (\<exists>f. ?phi20 f)"
+  by (auto simp add: ordLess_def embedS_def)
+  hence "?phi10 ?f10 \<and> ?phi20 ?f20" by (auto simp add: someI_ex)
+  thus "(?f10 ` ?A1, ?f20 ` ?A2) \<in> ofilterIncl r0"
+  using * ** by (simp add: embed_ordLess_ofilterIncl)
+qed
+
+
+theorem wf_ordLess: "wf ordLess"
+proof-
+  {fix r0 :: "('a \<times> 'a) set"
+   (* need to annotate here!*)
+   let ?ordLess = "ordLess::('d rel * 'd rel) set"
+   let ?R = "?ordLess Int (?ordLess^-1``{r0} \<times> ?ordLess^-1``{r0})"
+   {assume Case1: "Well_order r0"
+    hence "wf ?R"
+    using wf_ofilterIncl[of r0]
+          compat_wf[of ?R "ofilterIncl r0" "ord_to_filter r0"]
+          ord_to_filter_compat[of r0] by auto
+   }
+   moreover
+   {assume Case2: "\<not> Well_order r0"
+    hence "?R = {}" unfolding ordLess_def by auto
+    hence "wf ?R" using wf_empty by simp
+   }
+   ultimately have "wf ?R" by blast
+  }
+  thus ?thesis by (simp add: trans_wf_iff ordLess_trans)
+qed
+
+corollary exists_minim_Well_order:
+assumes NE: "R \<noteq> {}" and WELL: "\<forall>r \<in> R. Well_order r"
+shows "\<exists>r \<in> R. \<forall>r' \<in> R. r \<le>o r'"
+proof-
+  obtain r where "r \<in> R \<and> (\<forall>r' \<in> R. \<not> r' <o r)"
+  using NE spec[OF spec[OF subst[OF wf_eq_minimal, of "%x. x", OF wf_ordLess]], of _ R]
+    equals0I[of R] by blast
+  with not_ordLeq_iff_ordLess WELL show ?thesis by blast
+qed
+
+
+
+subsection {* Copy via direct images  *}
+
+
+text{* The direct image operator is the dual of the inverse image operator @{text "inv_image"}
+from @{text "Relation.thy"}.  It is useful for transporting a well-order between
+different types. *}
+
+
+definition dir_image :: "'a rel \<Rightarrow> ('a \<Rightarrow> 'a') \<Rightarrow> 'a' rel"
+where
+"dir_image r f = {(f a, f b)| a b. (a,b) \<in> r}"
+
+
+lemma dir_image_Field:
+"Field(dir_image r f) \<le> f ` (Field r)"
+unfolding dir_image_def Field_def by auto
+
+
+lemma dir_image_minus_Id:
+"inj_on f (Field r) \<Longrightarrow> (dir_image r f) - Id = dir_image (r - Id) f"
+unfolding inj_on_def Field_def dir_image_def by auto
+
+
+lemma Refl_dir_image:
+assumes "Refl r"
+shows "Refl(dir_image r f)"
+proof-
+  {fix a' b'
+   assume "(a',b') \<in> dir_image r f"
+   then obtain a b where 1: "a' = f a \<and> b' = f b \<and> (a,b) \<in> r"
+   unfolding dir_image_def by blast
+   hence "a \<in> Field r \<and> b \<in> Field r" using Field_def by fastforce
+   hence "(a,a) \<in> r \<and> (b,b) \<in> r" using assms by (simp add: refl_on_def)
+   with 1 have "(a',a') \<in> dir_image r f \<and> (b',b') \<in> dir_image r f"
+   unfolding dir_image_def by auto
+  }
+  thus ?thesis
+  by(unfold refl_on_def Field_def Domain_def Range_def, auto)
+qed
+
+
+lemma trans_dir_image:
+assumes TRANS: "trans r" and INJ: "inj_on f (Field r)"
+shows "trans(dir_image r f)"
+proof(unfold trans_def, auto)
+  fix a' b' c'
+  assume "(a',b') \<in> dir_image r f" "(b',c') \<in> dir_image r f"
+  then obtain a b1 b2 c where 1: "a' = f a \<and> b' = f b1 \<and> b' = f b2 \<and> c' = f c" and
+                         2: "(a,b1) \<in> r \<and> (b2,c) \<in> r"
+  unfolding dir_image_def by blast
+  hence "b1 \<in> Field r \<and> b2 \<in> Field r"
+  unfolding Field_def by auto
+  hence "b1 = b2" using 1 INJ unfolding inj_on_def by auto
+  hence "(a,c): r" using 2 TRANS unfolding trans_def by blast
+  thus "(a',c') \<in> dir_image r f"
+  unfolding dir_image_def using 1 by auto
+qed
+
+
+lemma Preorder_dir_image:
+"\<lbrakk>Preorder r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Preorder (dir_image r f)"
+by (simp add: preorder_on_def Refl_dir_image trans_dir_image)
+
+
+lemma antisym_dir_image:
+assumes AN: "antisym r" and INJ: "inj_on f (Field r)"
+shows "antisym(dir_image r f)"
+proof(unfold antisym_def, auto)
+  fix a' b'
+  assume "(a',b') \<in> dir_image r f" "(b',a') \<in> dir_image r f"
+  then obtain a1 b1 a2 b2 where 1: "a' = f a1 \<and> a' = f a2 \<and> b' = f b1 \<and> b' = f b2" and
+                           2: "(a1,b1) \<in> r \<and> (b2,a2) \<in> r " and
+                           3: "{a1,a2,b1,b2} \<le> Field r"
+  unfolding dir_image_def Field_def by blast
+  hence "a1 = a2 \<and> b1 = b2" using INJ unfolding inj_on_def by auto
+  hence "a1 = b2" using 2 AN unfolding antisym_def by auto
+  thus "a' = b'" using 1 by auto
+qed
+
+
+lemma Partial_order_dir_image:
+"\<lbrakk>Partial_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Partial_order (dir_image r f)"
+by (simp add: partial_order_on_def Preorder_dir_image antisym_dir_image)
+
+
+lemma Total_dir_image:
+assumes TOT: "Total r" and INJ: "inj_on f (Field r)"
+shows "Total(dir_image r f)"
+proof(unfold total_on_def, intro ballI impI)
+  fix a' b'
+  assume "a' \<in> Field (dir_image r f)" "b' \<in> Field (dir_image r f)"
+  then obtain a and b where 1: "a \<in> Field r \<and> b \<in> Field r \<and> f a = a' \<and> f b = b'"
+  using dir_image_Field[of r f] by blast
+  moreover assume "a' \<noteq> b'"
+  ultimately have "a \<noteq> b" using INJ unfolding inj_on_def by auto
+  hence "(a,b) \<in> r \<or> (b,a) \<in> r" using 1 TOT unfolding total_on_def by auto
+  thus "(a',b') \<in> dir_image r f \<or> (b',a') \<in> dir_image r f"
+  using 1 unfolding dir_image_def by auto
+qed
+
+
+lemma Linear_order_dir_image:
+"\<lbrakk>Linear_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Linear_order (dir_image r f)"
+by (simp add: linear_order_on_def Partial_order_dir_image Total_dir_image)
+
+
+lemma wf_dir_image:
+assumes WF: "wf r" and INJ: "inj_on f (Field r)"
+shows "wf(dir_image r f)"
+proof(unfold wf_eq_minimal2, intro allI impI, elim conjE)
+  fix A'::"'b set"
+  assume SUB: "A' \<le> Field(dir_image r f)" and NE: "A' \<noteq> {}"
+  obtain A where A_def: "A = {a \<in> Field r. f a \<in> A'}" by blast
+  have "A \<noteq> {} \<and> A \<le> Field r"
+  using A_def dir_image_Field[of r f] SUB NE by blast
+  then obtain a where 1: "a \<in> A \<and> (\<forall>b \<in> A. (b,a) \<notin> r)"
+  using WF unfolding wf_eq_minimal2 by metis
+  have "\<forall>b' \<in> A'. (b',f a) \<notin> dir_image r f"
+  proof(clarify)
+    fix b' assume *: "b' \<in> A'" and **: "(b',f a) \<in> dir_image r f"
+    obtain b1 a1 where 2: "b' = f b1 \<and> f a = f a1" and
+                       3: "(b1,a1) \<in> r \<and> {a1,b1} \<le> Field r"
+    using ** unfolding dir_image_def Field_def by blast
+    hence "a = a1" using 1 A_def INJ unfolding inj_on_def by auto
+    hence "b1 \<in> A \<and> (b1,a) \<in> r" using 2 3 A_def * by auto
+    with 1 show False by auto
+  qed
+  thus "\<exists>a'\<in>A'. \<forall>b'\<in>A'. (b', a') \<notin> dir_image r f"
+  using A_def 1 by blast
+qed
+
+
+lemma Well_order_dir_image:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> Well_order (dir_image r f)"
+using assms unfolding well_order_on_def
+using Linear_order_dir_image[of r f] wf_dir_image[of "r - Id" f]
+  dir_image_minus_Id[of f r]
+  subset_inj_on[of f "Field r" "Field(r - Id)"]
+  mono_Field[of "r - Id" r] by auto
+
+
+lemma dir_image_Field2:
+"Refl r \<Longrightarrow> Field(dir_image r f) = f ` (Field r)"
+unfolding Field_def dir_image_def refl_on_def Domain_def Range_def by blast
+
+
+lemma dir_image_bij_betw:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk> \<Longrightarrow> bij_betw f (Field r) (Field (dir_image r f))"
+unfolding bij_betw_def
+by (simp add: dir_image_Field2 order_on_defs)
+
+
+lemma dir_image_compat:
+"compat r (dir_image r f) f"
+unfolding compat_def dir_image_def by auto
+
+
+lemma dir_image_iso:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk>  \<Longrightarrow> iso r (dir_image r f) f"
+using iso_iff3 dir_image_compat dir_image_bij_betw Well_order_dir_image by blast
+
+
+lemma dir_image_ordIso:
+"\<lbrakk>Well_order r; inj_on f (Field r)\<rbrakk>  \<Longrightarrow> r =o dir_image r f"
+unfolding ordIso_def using dir_image_iso Well_order_dir_image by blast
+
+
+lemma Well_order_iso_copy:
+assumes WELL: "well_order_on A r" and BIJ: "bij_betw f A A'"
+shows "\<exists>r'. well_order_on A' r' \<and> r =o r'"
+proof-
+   let ?r' = "dir_image r f"
+   have 1: "A = Field r \<and> Well_order r"
+   using WELL rel.well_order_on_Well_order by blast
+   hence 2: "iso r ?r' f"
+   using dir_image_iso using BIJ unfolding bij_betw_def by auto
+   hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast
+   hence "Field ?r' = A'"
+   using 1 BIJ unfolding bij_betw_def by auto
+   moreover have "Well_order ?r'"
+   using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast
+   ultimately show ?thesis unfolding ordIso_def using 1 2 by blast
+qed
+
+
+
+subsection {* Bounded square  *}
+
+
+text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
+order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
+following criteria (in this order):
+\begin{itemize}
+\item compare the maximums;
+\item compare the first components;
+\item compare the second components.
+\end{itemize}
+%
+The only application of this construction that we are aware of is
+at proving that the square of an infinite set has the same cardinal
+as that set. The essential property required there (and which is ensured by this
+construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
+in a product of proper filters on the original relation (assumed to be a well-order). *}
+
+
+definition bsqr :: "'a rel => ('a * 'a)rel"
+where
+"bsqr r = {((a1,a2),(b1,b2)).
+           {a1,a2,b1,b2} \<le> Field r \<and>
+           (a1 = b1 \<and> a2 = b2 \<or>
+            (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1  \<and> (a2,b2) \<in> r - Id
+           )}"
+
+
+lemma Field_bsqr:
+"Field (bsqr r) = Field r \<times> Field r"
+proof
+  show "Field (bsqr r) \<le> Field r \<times> Field r"
+  proof-
+    {fix a1 a2 assume "(a1,a2) \<in> Field (bsqr r)"
+     moreover
+     have "\<And> b1 b2. ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r \<Longrightarrow>
+                      a1 \<in> Field r \<and> a2 \<in> Field r" unfolding bsqr_def by auto
+     ultimately have "a1 \<in> Field r \<and> a2 \<in> Field r" unfolding Field_def by auto
+    }
+    thus ?thesis unfolding Field_def by force
+  qed
+next
+  show "Field r \<times> Field r \<le> Field (bsqr r)"
+  proof(auto)
+    fix a1 a2 assume "a1 \<in> Field r" and "a2 \<in> Field r"
+    hence "((a1,a2),(a1,a2)) \<in> bsqr r" unfolding bsqr_def by blast
+    thus "(a1,a2) \<in> Field (bsqr r)" unfolding Field_def by auto
+  qed
+qed
+
+
+lemma bsqr_Refl: "Refl(bsqr r)"
+by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def)
+
+
+lemma bsqr_Trans:
+assumes "Well_order r"
+shows "trans (bsqr r)"
+proof(unfold trans_def, auto)
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Trans: "trans r" using wo_rel.TRANS by auto
+  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+  (* Main proof *)
+  fix a1 a2 b1 b2 c1 c2
+  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(c1,c2)) \<in> bsqr r"
+  hence 0: "{a1,a2,b1,b2,c1,c2} \<le> Field r" unfolding bsqr_def by auto
+  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+  using * unfolding bsqr_def by auto
+  have 2: "b1 = c1 \<and> b2 = c2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+  using ** unfolding bsqr_def by auto
+  show "((a1,a2),(c1,c2)) \<in> bsqr r"
+  proof-
+    {assume Case1: "a1 = b1 \<and> a2 = b2"
+     hence ?thesis using ** by simp
+    }
+    moreover
+    {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+     {assume Case21: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      using Case2 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2"
+      hence ?thesis using Case2 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+     {assume Case31: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence ?thesis using Case3 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+      hence "(a1,c1) \<in> r - Id"
+      using Case3 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1"
+      hence ?thesis using Case3 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+     {assume Case41: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence ?thesis using Case4 0 unfolding bsqr_def by force
+     }
+     moreover
+     {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+      hence ?thesis using Case4 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+      hence "(a2,c2) \<in> r - Id"
+      using Case4 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    ultimately show ?thesis using 0 1 by auto
+  qed
+qed
+
+
+lemma bsqr_antisym:
+assumes "Well_order r"
+shows "antisym (bsqr r)"
+proof(unfold antisym_def, clarify)
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Trans: "trans r" using wo_rel.TRANS by auto
+  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+  hence IrrS: "\<forall>a b. \<not>((a,b) \<in> r - Id \<and> (b,a) \<in> r - Id)"
+  using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast
+  (* Main proof *)
+  fix a1 a2 b1 b2
+  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(a1,a2)) \<in> bsqr r"
+  hence 0: "{a1,a2,b1,b2} \<le> Field r" unfolding bsqr_def by auto
+  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+  using * unfolding bsqr_def by auto
+  have 2: "b1 = a1 \<and> b2 = a2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> (b1,a1) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> b1 = a1 \<and> (b2,a2) \<in> r - Id"
+  using ** unfolding bsqr_def by auto
+  show "a1 = b1 \<and> a2 = b2"
+  proof-
+    {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+     {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+      hence False using Case1 IrrS by blast
+     }
+     moreover
+     {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2"
+      hence False using Case1 by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+     {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+       hence False using Case2 by auto
+     }
+     moreover
+     {assume Case22: "(b1,a1) \<in> r - Id"
+      hence False using Case2 IrrS by blast
+     }
+     moreover
+     {assume Case23: "b1 = a1"
+      hence False using Case2 by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+     moreover
+     {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+      hence False using Case3 by auto
+     }
+     moreover
+     {assume Case32: "(b1,a1) \<in> r - Id"
+      hence False using Case3 by auto
+     }
+     moreover
+     {assume Case33: "(b2,a2) \<in> r - Id"
+      hence False using Case3 IrrS by blast
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    ultimately show ?thesis using 0 1 by blast
+  qed
+qed
+
+
+lemma bsqr_Total:
+assumes "Well_order r"
+shows "Total(bsqr r)"
+proof-
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Total: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
+  using wo_rel.TOTALS by auto
+  (* Main proof *)
+  {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \<le> Field(bsqr r)"
+   hence 0: "a1 \<in> Field r \<and> a2 \<in> Field r \<and> b1 \<in> Field r \<and> b2 \<in> Field r"
+   using Field_bsqr by blast
+   have "((a1,a2) = (b1,b2) \<or> ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r)"
+   proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0)
+       (* Why didn't clarsimp simp add: Well 0 do the same job? *)
+     assume Case1: "(a1,a2) \<in> r"
+     hence 1: "wo_rel.max2 r a1 a2 = a2"
+     using Well 0 by (simp add: wo_rel.max2_equals2)
+     show ?thesis
+     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+       assume Case11: "(b1,b2) \<in> r"
+       hence 2: "wo_rel.max2 r b1 b2 = b2"
+       using Well 0 by (simp add: wo_rel.max2_equals2)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+         assume Case111: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+       next
+         assume Case112: "a2 = b2"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case1121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto
+         next
+           assume Case1122: "a1 = b1"
+           thus ?thesis using Case112 by auto
+         qed
+       qed
+     next
+       assume Case12: "(b2,b1) \<in> r"
+       hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0)
+         assume Case121: "(a2,b1) \<in> r - Id \<or> (b1,a2) \<in> r - Id"
+         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+       next
+         assume Case122: "a2 = b1"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case1221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto
+         next
+           assume Case1222: "a1 = b1"
+           show ?thesis
+           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+             assume Case12221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+             thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto
+           next
+             assume Case12222: "a2 = b2"
+             thus ?thesis using Case122 Case1222 by auto
+           qed
+         qed
+       qed
+     qed
+   next
+     assume Case2: "(a2,a1) \<in> r"
+     hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1)
+     show ?thesis
+     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+       assume Case21: "(b1,b2) \<in> r"
+       hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0)
+         assume Case211: "(a1,b2) \<in> r - Id \<or> (b2,a1) \<in> r - Id"
+         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+       next
+         assume Case212: "a1 = b2"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case2121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto
+         next
+           assume Case2122: "a1 = b1"
+           show ?thesis
+           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+             assume Case21221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+             thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto
+           next
+             assume Case21222: "a2 = b2"
+             thus ?thesis using Case2122 Case212 by auto
+           qed
+         qed
+       qed
+     next
+       assume Case22: "(b2,b1) \<in> r"
+       hence 3: "wo_rel.max2 r b1 b2 = b1"  using Well 0 by (simp add: wo_rel.max2_equals1)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+         assume Case221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+       next
+         assume Case222: "a1 = b1"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+           assume Case2221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+           thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto
+         next
+           assume Case2222: "a2 = b2"
+           thus ?thesis using Case222 by auto
+         qed
+       qed
+     qed
+   qed
+  }
+  thus ?thesis unfolding total_on_def by fast
+qed
+
+
+lemma bsqr_Linear_order:
+assumes "Well_order r"
+shows "Linear_order(bsqr r)"
+unfolding order_on_defs
+using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast
+
+
+lemma bsqr_Well_order:
+assumes "Well_order r"
+shows "Well_order(bsqr r)"
+using assms
+proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI)
+  have 0: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+  using assms well_order_on_def Linear_order_Well_order_iff by blast
+  fix D assume *: "D \<le> Field (bsqr r)" and **: "D \<noteq> {}"
+  hence 1: "D \<le> Field r \<times> Field r" unfolding Field_bsqr by simp
+  (*  *)
+  obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \<in> D}" by blast
+  have "M \<noteq> {}" using 1 M_def ** by auto
+  moreover
+  have "M \<le> Field r" unfolding M_def
+  using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+  ultimately obtain m where m_min: "m \<in> M \<and> (\<forall>a \<in> M. (m,a) \<in> r)"
+  using 0 by blast
+  (*  *)
+  obtain A1 where A1_def: "A1 = {a1. \<exists>a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+  have "A1 \<le> Field r" unfolding A1_def using 1 by auto
+  moreover have "A1 \<noteq> {}" unfolding A1_def using m_min unfolding M_def by blast
+  ultimately obtain a1 where a1_min: "a1 \<in> A1 \<and> (\<forall>a \<in> A1. (a1,a) \<in> r)"
+  using 0 by blast
+  (*  *)
+  obtain A2 where A2_def: "A2 = {a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+  have "A2 \<le> Field r" unfolding A2_def using 1 by auto
+  moreover have "A2 \<noteq> {}" unfolding A2_def
+  using m_min a1_min unfolding A1_def M_def by blast
+  ultimately obtain a2 where a2_min: "a2 \<in> A2 \<and> (\<forall>a \<in> A2. (a2,a) \<in> r)"
+  using 0 by blast
+  (*   *)
+  have 2: "wo_rel.max2 r a1 a2 = m"
+  using a1_min a2_min unfolding A1_def A2_def by auto
+  have 3: "(a1,a2) \<in> D" using a2_min unfolding A2_def by auto
+  (*  *)
+  moreover
+  {fix b1 b2 assume ***: "(b1,b2) \<in> D"
+   hence 4: "{a1,a2,b1,b2} \<le> Field r" using 1 3 by blast
+   have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+   using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto
+   have "((a1,a2),(b1,b2)) \<in> bsqr r"
+   proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2")
+     assume Case1: "wo_rel.max2 r a1 a2 \<noteq> wo_rel.max2 r b1 b2"
+     thus ?thesis unfolding bsqr_def using 4 5 by auto
+   next
+     assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+     hence "b1 \<in> A1" unfolding A1_def using 2 *** by auto
+     hence 6: "(a1,b1) \<in> r" using a1_min by auto
+     show ?thesis
+     proof(cases "a1 = b1")
+       assume Case21: "a1 \<noteq> b1"
+       thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto
+     next
+       assume Case22: "a1 = b1"
+       hence "b2 \<in> A2" unfolding A2_def using 2 *** Case2 by auto
+       hence 7: "(a2,b2) \<in> r" using a2_min by auto
+       thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto
+     qed
+   qed
+  }
+  (*  *)
+  ultimately show "\<exists>d \<in> D. \<forall>d' \<in> D. (d,d') \<in> bsqr r" by fastforce
+qed
+
+
+lemma bsqr_max2:
+assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \<in> bsqr r"
+shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+proof-
+  have "{(a1,a2),(b1,b2)} \<le> Field(bsqr r)"
+  using LEQ unfolding Field_def by auto
+  hence "{a1,a2,b1,b2} \<le> Field r" unfolding Field_bsqr by auto
+  hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \<le> Field r"
+  using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+  moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r \<or> wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+  using LEQ unfolding bsqr_def by auto
+  ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto
+qed
+
+
+lemma bsqr_ofilter:
+assumes WELL: "Well_order r" and
+        OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \<times> Field r" and
+        NE: "\<not> (\<exists>a. Field r = rel.under r a)"
+shows "\<exists>A. wo_rel.ofilter r A \<and> A < Field r \<and> D \<le> A \<times> A"
+proof-
+  let ?r' = "bsqr r"
+  have Well: "wo_rel r" using WELL wo_rel_def by blast
+  hence Trans: "trans r" using wo_rel.TRANS by blast
+  have Well': "Well_order ?r' \<and> wo_rel ?r'"
+  using WELL bsqr_Well_order wo_rel_def by blast
+  (*  *)
+  have "D < Field ?r'" unfolding Field_bsqr using SUB .
+  with OF obtain a1 and a2 where
+  "(a1,a2) \<in> Field ?r'" and 1: "D = rel.underS ?r' (a1,a2)"
+  using Well' wo_rel.ofilter_underS_Field[of ?r' D] by auto
+  hence 2: "{a1,a2} \<le> Field r" unfolding Field_bsqr by auto
+  let ?m = "wo_rel.max2 r a1 a2"
+  have "D \<le> (rel.under r ?m) \<times> (rel.under r ?m)"
+  proof(unfold 1)
+    {fix b1 b2
+     let ?n = "wo_rel.max2 r b1 b2"
+     assume "(b1,b2) \<in> rel.underS ?r' (a1,a2)"
+     hence 3: "((b1,b2),(a1,a2)) \<in> ?r'"
+     unfolding rel.underS_def by blast
+     hence "(?n,?m) \<in> r" using WELL by (simp add: bsqr_max2)
+     moreover
+     {have "(b1,b2) \<in> Field ?r'" using 3 unfolding Field_def by auto
+      hence "{b1,b2} \<le> Field r" unfolding Field_bsqr by auto
+      hence "(b1,?n) \<in> r \<and> (b2,?n) \<in> r"
+      using Well by (simp add: wo_rel.max2_greater)
+     }
+     ultimately have "(b1,?m) \<in> r \<and> (b2,?m) \<in> r"
+     using Trans trans_def[of r] by blast
+     hence "(b1,b2) \<in> (rel.under r ?m) \<times> (rel.under r ?m)" unfolding rel.under_def by simp}
+     thus "rel.underS ?r' (a1,a2) \<le> (rel.under r ?m) \<times> (rel.under r ?m)" by auto
+  qed
+  moreover have "wo_rel.ofilter r (rel.under r ?m)"
+  using Well by (simp add: wo_rel.under_ofilter)
+  moreover have "rel.under r ?m < Field r"
+  using NE rel.under_Field[of r ?m] by blast
+  ultimately show ?thesis by blast
+qed
+
+
+end
--- a/src/HOL/Cardinals/Fun_More.thy	Mon Nov 18 18:04:44 2013 +0100
+++ b/src/HOL/Cardinals/Fun_More.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -8,7 +8,7 @@
 header {* More on Injections, Bijections and Inverses *}
 
 theory Fun_More
-imports Fun_More_Base
+imports Fun_More_LFP
 begin
 
 
--- a/src/HOL/Cardinals/Fun_More_Base.thy	Mon Nov 18 18:04:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,252 +0,0 @@
-(*  Title:      HOL/Cardinals/Fun_More_Base.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-More on injections, bijections and inverses (base).
-*)
-
-header {* More on Injections, Bijections and Inverses (Base) *}
-
-theory Fun_More_Base
-imports "~~/src/HOL/Library/Infinite_Set"
-begin
-
-
-text {* This section proves more facts (additional to those in @{text "Fun.thy"},
-@{text "Hilbert_Choice.thy"}, @{text "Finite_Set.thy"} and @{text "Infinite_Set.thy"}),
-mainly concerning injections, bijections, inverses and (numeric) cardinals of
-finite sets. *}
-
-
-subsection {* Purely functional properties  *}
-
-
-(*2*)lemma bij_betw_id_iff:
-"(A = B) = (bij_betw id A B)"
-by(simp add: bij_betw_def)
-
-
-(*2*)lemma bij_betw_byWitness:
-assumes LEFT: "\<forall>a \<in> A. f'(f a) = a" and
-        RIGHT: "\<forall>a' \<in> A'. f(f' a') = a'" and
-        IM1: "f ` A \<le> A'" and IM2: "f' ` A' \<le> A"
-shows "bij_betw f A A'"
-using assms
-proof(unfold bij_betw_def inj_on_def, safe)
-  fix a b assume *: "a \<in> A" "b \<in> A" and **: "f a = f b"
-  have "a = f'(f a) \<and> b = f'(f b)" using * LEFT by simp
-  with ** show "a = b" by simp
-next
-  fix a' assume *: "a' \<in> A'"
-  hence "f' a' \<in> A" using IM2 by blast
-  moreover
-  have "a' = f(f' a')" using * RIGHT by simp
-  ultimately show "a' \<in> f ` A" by blast
-qed
-
-
-(*3*)corollary notIn_Un_bij_betw:
-assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'" and
-       BIJ: "bij_betw f A A'"
-shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
-proof-
-  have "bij_betw f {b} {f b}"
-  unfolding bij_betw_def inj_on_def by simp
-  with assms show ?thesis
-  using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
-qed
-
-
-(*1*)lemma notIn_Un_bij_betw3:
-assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'"
-shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
-proof
-  assume "bij_betw f A A'"
-  thus "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
-  using assms notIn_Un_bij_betw[of b A f A'] by blast
-next
-  assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
-  have "f ` A = A'"
-  proof(auto)
-    fix a assume **: "a \<in> A"
-    hence "f a \<in> A' \<union> {f b}" using * unfolding bij_betw_def by blast
-    moreover
-    {assume "f a = f b"
-     hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast
-     with NIN ** have False by blast
-    }
-    ultimately show "f a \<in> A'" by blast
-  next
-    fix a' assume **: "a' \<in> A'"
-    hence "a' \<in> f`(A \<union> {b})"
-    using * by (auto simp add: bij_betw_def)
-    then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
-    moreover
-    {assume "a = b" with 1 ** NIN' have False by blast
-    }
-    ultimately have "a \<in> A" by blast
-    with 1 show "a' \<in> f ` A" by blast
-  qed
-  thus "bij_betw f A A'" using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
-qed
-
-
-subsection {* Properties involving finite and infinite sets *}
-
-
-(*3*)lemma inj_on_finite:
-assumes "inj_on f A" "f ` A \<le> B" "finite B"
-shows "finite A"
-using assms infinite_super by (fast dest: finite_imageD)
-
-
-(*3*)lemma infinite_imp_bij_betw:
-assumes INF: "infinite A"
-shows "\<exists>h. bij_betw h A (A - {a})"
-proof(cases "a \<in> A")
-  assume Case1: "a \<notin> A"  hence "A - {a} = A" by blast
-  thus ?thesis using bij_betw_id[of A] by auto
-next
-  assume Case2: "a \<in> A"
-  have "infinite (A - {a})" using INF infinite_remove by auto
-  with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \<Rightarrow> 'a"
-  where 1: "inj f" and 2: "f ` UNIV \<le> A - {a}" by blast
-  obtain g where g_def: "g = (\<lambda> n. if n = 0 then a else f (Suc n))" by blast
-  obtain A' where A'_def: "A' = g ` UNIV" by blast
-  have temp: "\<forall>y. f y \<noteq> a" using 2 by blast
-  have 3: "inj_on g UNIV \<and> g ` UNIV \<le> A \<and> a \<in> g ` UNIV"
-  proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI,
-        case_tac "x = 0", auto simp add: 2)
-    fix y  assume "a = (if y = 0 then a else f (Suc y))"
-    thus "y = 0" using temp by (case_tac "y = 0", auto)
-  next
-    fix x y
-    assume "f (Suc x) = (if y = 0 then a else f (Suc y))"
-    thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto)
-  next
-    fix n show "f (Suc n) \<in> A" using 2 by blast
-  qed
-  hence 4: "bij_betw g UNIV A' \<and> a \<in> A' \<and> A' \<le> A"
-  using inj_on_imp_bij_betw[of g] unfolding A'_def by auto
-  hence 5: "bij_betw (inv g) A' UNIV"
-  by (auto simp add: bij_betw_inv_into)
-  (*  *)
-  obtain n where "g n = a" using 3 by auto
-  hence 6: "bij_betw g (UNIV - {n}) (A' - {a})"
-  using 3 4 unfolding A'_def
-  by clarify (rule bij_betw_subset, auto simp: image_set_diff)
-  (*  *)
-  obtain v where v_def: "v = (\<lambda> m. if m < n then m else Suc m)" by blast
-  have 7: "bij_betw v UNIV (UNIV - {n})"
-  proof(unfold bij_betw_def inj_on_def, intro conjI, clarify)
-    fix m1 m2 assume "v m1 = v m2"
-    thus "m1 = m2"
-    by(case_tac "m1 < n", case_tac "m2 < n",
-       auto simp add: inj_on_def v_def, case_tac "m2 < n", auto)
-  next
-    show "v ` UNIV = UNIV - {n}"
-    proof(auto simp add: v_def)
-      fix m assume *: "m \<noteq> n" and **: "m \<notin> Suc ` {m'. \<not> m' < n}"
-      {assume "n \<le> m" with * have 71: "Suc n \<le> m" by auto
-       then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto
-       with 71 have "n \<le> m'" by auto
-       with 72 ** have False by auto
-      }
-      thus "m < n" by force
-    qed
-  qed
-  (*  *)
-  obtain h' where h'_def: "h' = g o v o (inv g)" by blast
-  hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6
-  by (auto simp add: bij_betw_trans)
-  (*  *)
-  obtain h where h_def: "h = (\<lambda> b. if b \<in> A' then h' b else b)" by blast
-  have "\<forall>b \<in> A'. h b = h' b" unfolding h_def by auto
-  hence "bij_betw h  A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto
-  moreover
-  {have "\<forall>b \<in> A - A'. h b = b" unfolding h_def by auto
-   hence "bij_betw h  (A - A') (A - A')"
-   using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto
-  }
-  moreover
-  have "(A' Int (A - A') = {} \<and> A' \<union> (A - A') = A) \<and>
-        ((A' - {a}) Int (A - A') = {} \<and> (A' - {a}) \<union> (A - A') = A - {a})"
-  using 4 by blast
-  ultimately have "bij_betw h A (A - {a})"
-  using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp
-  thus ?thesis by blast
-qed
-
-
-(*3*)lemma infinite_imp_bij_betw2:
-assumes INF: "infinite A"
-shows "\<exists>h. bij_betw h A (A \<union> {a})"
-proof(cases "a \<in> A")
-  assume Case1: "a \<in> A"  hence "A \<union> {a} = A" by blast
-  thus ?thesis using bij_betw_id[of A] by auto
-next
-  let ?A' = "A \<union> {a}"
-  assume Case2: "a \<notin> A" hence "A = ?A' - {a}" by blast
-  moreover have "infinite ?A'" using INF by auto
-  ultimately obtain f where "bij_betw f ?A' A"
-  using infinite_imp_bij_betw[of ?A' a] by auto
-  hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast
-  thus ?thesis by auto
-qed
-
-
-subsection {* Properties involving Hilbert choice *}
-
-
-(*2*)lemma bij_betw_inv_into_left:
-assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
-shows "(inv_into A f) (f a) = a"
-using assms unfolding bij_betw_def
-by clarify (rule inv_into_f_f)
-
-(*2*)lemma bij_betw_inv_into_right:
-assumes "bij_betw f A A'" "a' \<in> A'"
-shows "f(inv_into A f a') = a'"
-using assms unfolding bij_betw_def using f_inv_into_f by force
-
-
-(*1*)lemma bij_betw_inv_into_LEFT:
-assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A"
-shows "(inv_into A f)`(f ` B) = B"
-using assms unfolding bij_betw_def using inv_into_image_cancel by force
-
-
-(*1*)lemma bij_betw_inv_into_LEFT_RIGHT:
-assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A" and
-        IM: "f ` B = B'"
-shows "(inv_into A f) ` B' = B"
-using assms bij_betw_inv_into_LEFT[of f A A' B] by fast
-
-
-(*1*)lemma bij_betw_inv_into_subset:
-assumes BIJ: "bij_betw f A A'" and
-        SUB: "B \<le> A" and IM: "f ` B = B'"
-shows "bij_betw (inv_into A f) B' B"
-using assms unfolding bij_betw_def
-by (auto intro: inj_on_inv_into)
-
-
-subsection {* Other facts  *}
-
-
-(*2*)lemma atLeastLessThan_less_eq:
-"({0..<m} \<le> {0..<n}) = ((m::nat) \<le> n)"
-unfolding ivl_subset by arith
-
-
-(*2*)lemma atLeastLessThan_less_eq2:
-assumes "inj_on f {0..<(m::nat)} \<and> f ` {0..<m} \<le> {0..<n}"
-shows "m \<le> n"
-using assms
-using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n]
-      card_atLeastLessThan[of m] card_atLeastLessThan[of n]
-      card_inj_on_le[of f "{0 ..< m}" "{0 ..< n}"] by auto
-
-
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Fun_More_LFP.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -0,0 +1,252 @@
+(*  Title:      HOL/Cardinals/Fun_More_LFP.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+More on injections, bijections and inverses (LFP).
+*)
+
+header {* More on Injections, Bijections and Inverses (LFP) *}
+
+theory Fun_More_LFP
+imports "~~/src/HOL/Library/Infinite_Set"
+begin
+
+
+text {* This section proves more facts (additional to those in @{text "Fun.thy"},
+@{text "Hilbert_Choice.thy"}, @{text "Finite_Set.thy"} and @{text "Infinite_Set.thy"}),
+mainly concerning injections, bijections, inverses and (numeric) cardinals of
+finite sets. *}
+
+
+subsection {* Purely functional properties  *}
+
+
+(*2*)lemma bij_betw_id_iff:
+"(A = B) = (bij_betw id A B)"
+by(simp add: bij_betw_def)
+
+
+(*2*)lemma bij_betw_byWitness:
+assumes LEFT: "\<forall>a \<in> A. f'(f a) = a" and
+        RIGHT: "\<forall>a' \<in> A'. f(f' a') = a'" and
+        IM1: "f ` A \<le> A'" and IM2: "f' ` A' \<le> A"
+shows "bij_betw f A A'"
+using assms
+proof(unfold bij_betw_def inj_on_def, safe)
+  fix a b assume *: "a \<in> A" "b \<in> A" and **: "f a = f b"
+  have "a = f'(f a) \<and> b = f'(f b)" using * LEFT by simp
+  with ** show "a = b" by simp
+next
+  fix a' assume *: "a' \<in> A'"
+  hence "f' a' \<in> A" using IM2 by blast
+  moreover
+  have "a' = f(f' a')" using * RIGHT by simp
+  ultimately show "a' \<in> f ` A" by blast
+qed
+
+
+(*3*)corollary notIn_Un_bij_betw:
+assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'" and
+       BIJ: "bij_betw f A A'"
+shows "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+proof-
+  have "bij_betw f {b} {f b}"
+  unfolding bij_betw_def inj_on_def by simp
+  with assms show ?thesis
+  using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast
+qed
+
+
+(*1*)lemma notIn_Un_bij_betw3:
+assumes NIN: "b \<notin> A" and NIN': "f b \<notin> A'"
+shows "bij_betw f A A' = bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+proof
+  assume "bij_betw f A A'"
+  thus "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+  using assms notIn_Un_bij_betw[of b A f A'] by blast
+next
+  assume *: "bij_betw f (A \<union> {b}) (A' \<union> {f b})"
+  have "f ` A = A'"
+  proof(auto)
+    fix a assume **: "a \<in> A"
+    hence "f a \<in> A' \<union> {f b}" using * unfolding bij_betw_def by blast
+    moreover
+    {assume "f a = f b"
+     hence "a = b" using * ** unfolding bij_betw_def inj_on_def by blast
+     with NIN ** have False by blast
+    }
+    ultimately show "f a \<in> A'" by blast
+  next
+    fix a' assume **: "a' \<in> A'"
+    hence "a' \<in> f`(A \<union> {b})"
+    using * by (auto simp add: bij_betw_def)
+    then obtain a where 1: "a \<in> A \<union> {b} \<and> f a = a'" by blast
+    moreover
+    {assume "a = b" with 1 ** NIN' have False by blast
+    }
+    ultimately have "a \<in> A" by blast
+    with 1 show "a' \<in> f ` A" by blast
+  qed
+  thus "bij_betw f A A'" using * bij_betw_subset[of f "A \<union> {b}" _ A] by blast
+qed
+
+
+subsection {* Properties involving finite and infinite sets *}
+
+
+(*3*)lemma inj_on_finite:
+assumes "inj_on f A" "f ` A \<le> B" "finite B"
+shows "finite A"
+using assms infinite_super by (fast dest: finite_imageD)
+
+
+(*3*)lemma infinite_imp_bij_betw:
+assumes INF: "infinite A"
+shows "\<exists>h. bij_betw h A (A - {a})"
+proof(cases "a \<in> A")
+  assume Case1: "a \<notin> A"  hence "A - {a} = A" by blast
+  thus ?thesis using bij_betw_id[of A] by auto
+next
+  assume Case2: "a \<in> A"
+  have "infinite (A - {a})" using INF infinite_remove by auto
+  with infinite_iff_countable_subset[of "A - {a}"] obtain f::"nat \<Rightarrow> 'a"
+  where 1: "inj f" and 2: "f ` UNIV \<le> A - {a}" by blast
+  obtain g where g_def: "g = (\<lambda> n. if n = 0 then a else f (Suc n))" by blast
+  obtain A' where A'_def: "A' = g ` UNIV" by blast
+  have temp: "\<forall>y. f y \<noteq> a" using 2 by blast
+  have 3: "inj_on g UNIV \<and> g ` UNIV \<le> A \<and> a \<in> g ` UNIV"
+  proof(auto simp add: Case2 g_def, unfold inj_on_def, intro ballI impI,
+        case_tac "x = 0", auto simp add: 2)
+    fix y  assume "a = (if y = 0 then a else f (Suc y))"
+    thus "y = 0" using temp by (case_tac "y = 0", auto)
+  next
+    fix x y
+    assume "f (Suc x) = (if y = 0 then a else f (Suc y))"
+    thus "x = y" using 1 temp unfolding inj_on_def by (case_tac "y = 0", auto)
+  next
+    fix n show "f (Suc n) \<in> A" using 2 by blast
+  qed
+  hence 4: "bij_betw g UNIV A' \<and> a \<in> A' \<and> A' \<le> A"
+  using inj_on_imp_bij_betw[of g] unfolding A'_def by auto
+  hence 5: "bij_betw (inv g) A' UNIV"
+  by (auto simp add: bij_betw_inv_into)
+  (*  *)
+  obtain n where "g n = a" using 3 by auto
+  hence 6: "bij_betw g (UNIV - {n}) (A' - {a})"
+  using 3 4 unfolding A'_def
+  by clarify (rule bij_betw_subset, auto simp: image_set_diff)
+  (*  *)
+  obtain v where v_def: "v = (\<lambda> m. if m < n then m else Suc m)" by blast
+  have 7: "bij_betw v UNIV (UNIV - {n})"
+  proof(unfold bij_betw_def inj_on_def, intro conjI, clarify)
+    fix m1 m2 assume "v m1 = v m2"
+    thus "m1 = m2"
+    by(case_tac "m1 < n", case_tac "m2 < n",
+       auto simp add: inj_on_def v_def, case_tac "m2 < n", auto)
+  next
+    show "v ` UNIV = UNIV - {n}"
+    proof(auto simp add: v_def)
+      fix m assume *: "m \<noteq> n" and **: "m \<notin> Suc ` {m'. \<not> m' < n}"
+      {assume "n \<le> m" with * have 71: "Suc n \<le> m" by auto
+       then obtain m' where 72: "m = Suc m'" using Suc_le_D by auto
+       with 71 have "n \<le> m'" by auto
+       with 72 ** have False by auto
+      }
+      thus "m < n" by force
+    qed
+  qed
+  (*  *)
+  obtain h' where h'_def: "h' = g o v o (inv g)" by blast
+  hence 8: "bij_betw h' A' (A' - {a})" using 5 7 6
+  by (auto simp add: bij_betw_trans)
+  (*  *)
+  obtain h where h_def: "h = (\<lambda> b. if b \<in> A' then h' b else b)" by blast
+  have "\<forall>b \<in> A'. h b = h' b" unfolding h_def by auto
+  hence "bij_betw h  A' (A' - {a})" using 8 bij_betw_cong[of A' h] by auto
+  moreover
+  {have "\<forall>b \<in> A - A'. h b = b" unfolding h_def by auto
+   hence "bij_betw h  (A - A') (A - A')"
+   using bij_betw_cong[of "A - A'" h id] bij_betw_id[of "A - A'"] by auto
+  }
+  moreover
+  have "(A' Int (A - A') = {} \<and> A' \<union> (A - A') = A) \<and>
+        ((A' - {a}) Int (A - A') = {} \<and> (A' - {a}) \<union> (A - A') = A - {a})"
+  using 4 by blast
+  ultimately have "bij_betw h A (A - {a})"
+  using bij_betw_combine[of h A' "A' - {a}" "A - A'" "A - A'"] by simp
+  thus ?thesis by blast
+qed
+
+
+(*3*)lemma infinite_imp_bij_betw2:
+assumes INF: "infinite A"
+shows "\<exists>h. bij_betw h A (A \<union> {a})"
+proof(cases "a \<in> A")
+  assume Case1: "a \<in> A"  hence "A \<union> {a} = A" by blast
+  thus ?thesis using bij_betw_id[of A] by auto
+next
+  let ?A' = "A \<union> {a}"
+  assume Case2: "a \<notin> A" hence "A = ?A' - {a}" by blast
+  moreover have "infinite ?A'" using INF by auto
+  ultimately obtain f where "bij_betw f ?A' A"
+  using infinite_imp_bij_betw[of ?A' a] by auto
+  hence "bij_betw(inv_into ?A' f) A ?A'" using bij_betw_inv_into by blast
+  thus ?thesis by auto
+qed
+
+
+subsection {* Properties involving Hilbert choice *}
+
+
+(*2*)lemma bij_betw_inv_into_left:
+assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
+shows "(inv_into A f) (f a) = a"
+using assms unfolding bij_betw_def
+by clarify (rule inv_into_f_f)
+
+(*2*)lemma bij_betw_inv_into_right:
+assumes "bij_betw f A A'" "a' \<in> A'"
+shows "f(inv_into A f a') = a'"
+using assms unfolding bij_betw_def using f_inv_into_f by force
+
+
+(*1*)lemma bij_betw_inv_into_LEFT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A"
+shows "(inv_into A f)`(f ` B) = B"
+using assms unfolding bij_betw_def using inv_into_image_cancel by force
+
+
+(*1*)lemma bij_betw_inv_into_LEFT_RIGHT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B \<le> A" and
+        IM: "f ` B = B'"
+shows "(inv_into A f) ` B' = B"
+using assms bij_betw_inv_into_LEFT[of f A A' B] by fast
+
+
+(*1*)lemma bij_betw_inv_into_subset:
+assumes BIJ: "bij_betw f A A'" and
+        SUB: "B \<le> A" and IM: "f ` B = B'"
+shows "bij_betw (inv_into A f) B' B"
+using assms unfolding bij_betw_def
+by (auto intro: inj_on_inv_into)
+
+
+subsection {* Other facts  *}
+
+
+(*2*)lemma atLeastLessThan_less_eq:
+"({0..<m} \<le> {0..<n}) = ((m::nat) \<le> n)"
+unfolding ivl_subset by arith
+
+
+(*2*)lemma atLeastLessThan_less_eq2:
+assumes "inj_on f {0..<(m::nat)} \<and> f ` {0..<m} \<le> {0..<n}"
+shows "m \<le> n"
+using assms
+using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n]
+      card_atLeastLessThan[of m] card_atLeastLessThan[of n]
+      card_inj_on_le[of f "{0 ..< m}" "{0 ..< n}"] by auto
+
+
+
+end
--- a/src/HOL/Cardinals/Order_Relation_More.thy	Mon Nov 18 18:04:44 2013 +0100
+++ b/src/HOL/Cardinals/Order_Relation_More.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -8,7 +8,7 @@
 header {* Basics on Order-Like Relations *}
 
 theory Order_Relation_More
-imports Order_Relation_More_Base
+imports Order_Relation_More_LFP
 begin
 
 
--- a/src/HOL/Cardinals/Order_Relation_More_Base.thy	Mon Nov 18 18:04:44 2013 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,286 +0,0 @@
-(*  Title:      HOL/Cardinals/Order_Relation_More_Base.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Basics on order-like relations (base).
-*)
-
-header {* Basics on Order-Like Relations (Base) *}
-
-theory Order_Relation_More_Base
-imports "~~/src/HOL/Library/Order_Relation"
-begin
-
-
-text{* In this section, we develop basic concepts and results pertaining
-to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or
-total relations.  The development is placed on top of the definitions
-from the theory @{text "Order_Relation"}.  We also
-further define upper and lower bounds operators. *}
-
-
-locale rel = fixes r :: "'a rel"
-
-text{* The following context encompasses all this section, except
-for its last subsection. In other words, for the rest of this section except its last
-subsection, we consider a fixed relation @{text "r"}. *}
-
-context rel
-begin
-
-
-subsection {* Auxiliaries *}
-
-
-lemma refl_on_domain:
-"\<lbrakk>refl_on A r; (a,b) : r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
-by(auto simp add: refl_on_def)
-
-
-corollary well_order_on_domain:
-"\<lbrakk>well_order_on A r; (a,b) \<in> r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
-by(simp add: refl_on_domain order_on_defs)
-
-
-lemma well_order_on_Field:
-"well_order_on A r \<Longrightarrow> A = Field r"
-by(auto simp add: refl_on_def Field_def order_on_defs)
-
-
-lemma well_order_on_Well_order:
-"well_order_on A r \<Longrightarrow> A = Field r \<and> Well_order r"
-using well_order_on_Field by simp
-
-
-lemma Total_subset_Id:
-assumes TOT: "Total r" and SUB: "r \<le> Id"
-shows "r = {} \<or> (\<exists>a. r = {(a,a)})"
-proof-
-  {assume "r \<noteq> {}"
-   then obtain a b where 1: "(a,b) \<in> r" by fast
-   hence "a = b" using SUB by blast
-   hence 2: "(a,a) \<in> r" using 1 by simp
-   {fix c d assume "(c,d) \<in> r"
-    hence "{a,c,d} \<le> Field r" using 1 unfolding Field_def by blast
-    hence "((a,c) \<in> r \<or> (c,a) \<in> r \<or> a = c) \<and>
-           ((a,d) \<in> r \<or> (d,a) \<in> r \<or> a = d)"
-    using TOT unfolding total_on_def by blast
-    hence "a = c \<and> a = d" using SUB by blast
-   }
-   hence "r \<le> {(a,a)}" by auto
-   with 2 have "\<exists>a. r = {(a,a)}" by blast
-  }
-  thus ?thesis by blast
-qed
-
-
-lemma Linear_order_in_diff_Id:
-assumes LI: "Linear_order r" and
-        IN1: "a \<in> Field r" and IN2: "b \<in> Field r"
-shows "((a,b) \<in> r) = ((b,a) \<notin> r - Id)"
-using assms unfolding order_on_defs total_on_def antisym_def Id_def refl_on_def by force
-
-
-subsection {* The upper and lower bounds operators  *}
-
-
-text{* Here we define upper (``above") and lower (``below") bounds operators.
-We think of @{text "r"} as a {\em non-strict} relation.  The suffix ``S"
-at the names of some operators indicates that the bounds are strict -- e.g.,
-@{text "underS a"} is the set of all strict lower bounds of @{text "a"} (w.r.t. @{text "r"}).
-Capitalization of the first letter in the name reminds that the operator acts on sets, rather
-than on individual elements. *}
-
-definition under::"'a \<Rightarrow> 'a set"
-where "under a \<equiv> {b. (b,a) \<in> r}"
-
-definition underS::"'a \<Rightarrow> 'a set"
-where "underS a \<equiv> {b. b \<noteq> a \<and> (b,a) \<in> r}"
-
-definition Under::"'a set \<Rightarrow> 'a set"
-where "Under A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (b,a) \<in> r}"
-
-definition UnderS::"'a set \<Rightarrow> 'a set"
-where "UnderS A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (b,a) \<in> r}"
-
-definition above::"'a \<Rightarrow> 'a set"
-where "above a \<equiv> {b. (a,b) \<in> r}"
-
-definition aboveS::"'a \<Rightarrow> 'a set"
-where "aboveS a \<equiv> {b. b \<noteq> a \<and> (a,b) \<in> r}"
-
-definition Above::"'a set \<Rightarrow> 'a set"
-where "Above A \<equiv> {b \<in> Field r. \<forall>a \<in> A. (a,b) \<in> r}"
-
-definition AboveS::"'a set \<Rightarrow> 'a set"
-where "AboveS A \<equiv> {b \<in> Field r. \<forall>a \<in> A. b \<noteq> a \<and> (a,b) \<in> r}"
-(*  *)
-
-text{* Note:  In the definitions of @{text "Above[S]"} and @{text "Under[S]"},
-  we bounded comprehension by @{text "Field r"} in order to properly cover
-  the case of @{text "A"} being empty. *}
-
-
-lemma UnderS_subset_Under: "UnderS A \<le> Under A"
-by(auto simp add: UnderS_def Under_def)
-
-
-lemma underS_subset_under: "underS a \<le> under a"
-by(auto simp add: underS_def under_def)
-
-
-lemma underS_notIn: "a \<notin> underS a"
-by(simp add: underS_def)
-
-
-lemma Refl_under_in: "\<lbrakk>Refl r; a \<in> Field r\<rbrakk> \<Longrightarrow> a \<in> under a"
-by(simp add: refl_on_def under_def)
-
-
-lemma AboveS_disjoint: "A Int (AboveS A) = {}"
-by(auto simp add: AboveS_def)
-
-
-lemma in_AboveS_underS: "a \<in> Field r \<Longrightarrow> a \<in> AboveS (underS a)"
-by(auto simp add: AboveS_def underS_def)
-
-
-lemma Refl_under_underS:
-assumes "Refl r" "a \<in> Field r"
-shows "under a = underS a \<union> {a}"
-unfolding under_def underS_def
-using assms refl_on_def[of _ r] by fastforce
-
-
-lemma underS_empty: "a \<notin> Field r \<Longrightarrow> underS a = {}"
-by (auto simp: Field_def underS_def)
-
-
-lemma under_Field: "under a \<le> Field r"
-by(unfold under_def Field_def, auto)
-
-
-lemma underS_Field: "underS a \<le> Field r"
-by(unfold underS_def Field_def, auto)
-
-
-lemma underS_Field2:
-"a \<in> Field r \<Longrightarrow> underS a < Field r"
-using assms underS_notIn underS_Field by blast
-
-
-lemma underS_Field3:
-"Field r \<noteq> {} \<Longrightarrow> underS a < Field r"
-by(cases "a \<in> Field r", simp add: underS_Field2, auto simp add: underS_empty)
-
-
-lemma Under_Field: "Under A \<le> Field r"
-by(unfold Under_def Field_def, auto)
-
-
-lemma UnderS_Field: "UnderS A \<le> Field r"
-by(unfold UnderS_def Field_def, auto)
-
-
-lemma AboveS_Field: "AboveS A \<le> Field r"
-by(unfold AboveS_def Field_def, auto)
-
-
-lemma under_incr:
-assumes TRANS: "trans r" and REL: "(a,b) \<in> r"
-shows "under a \<le> under b"
-proof(unfold under_def, auto)
-  fix x assume "(x,a) \<in> r"
-  with REL TRANS trans_def[of r]
-  show "(x,b) \<in> r" by blast
-qed
-
-
-lemma underS_incr:
-assumes TRANS: "trans r" and ANTISYM: "antisym r" and
-        REL: "(a,b) \<in> r"
-shows "underS a \<le> underS b"
-proof(unfold underS_def, auto)
-  assume *: "b \<noteq> a" and **: "(b,a) \<in> r"
-  with ANTISYM antisym_def[of r] REL
-  show False by blast
-next
-  fix x assume "x \<noteq> a" "(x,a) \<in> r"
-  with REL TRANS trans_def[of r]
-  show "(x,b) \<in> r" by blast
-qed
-
-
-lemma underS_incl_iff:
-assumes LO: "Linear_order r" and
-        INa: "a \<in> Field r" and INb: "b \<in> Field r"
-shows "(underS a \<le> underS b) = ((a,b) \<in> r)"
-proof
-  assume "(a,b) \<in> r"
-  thus "underS a \<le> underS b" using LO
-  by (simp add: order_on_defs underS_incr)
-next
-  assume *: "underS a \<le> underS b"
-  {assume "a = b"
-   hence "(a,b) \<in> r" using assms
-   by (simp add: order_on_defs refl_on_def)
-  }
-  moreover
-  {assume "a \<noteq> b \<and> (b,a) \<in> r"
-   hence "b \<in> underS a" unfolding underS_def by blast
-   hence "b \<in> underS b" using * by blast
-   hence False by (simp add: underS_notIn)
-  }
-  ultimately
-  show "(a,b) \<in> r" using assms
-  order_on_defs[of "Field r" r] total_on_def[of "Field r" r] by blast
-qed
-
-
-lemma under_Under_trans:
-assumes TRANS: "trans r" and
-        IN1: "a \<in> under b" and IN2: "b \<in> Under C"
-shows "a \<in> Under C"
-proof-
-  have "(a,b) \<in> r \<and> (\<forall>c \<in> C. (b,c) \<in> r)"
-  using IN1 IN2 under_def Under_def by blast
-  hence "\<forall>c \<in> C. (a,c) \<in> r"
-  using TRANS trans_def[of r] by blast
-  moreover
-  have "a \<in> Field r" using IN1 unfolding Field_def under_def by blast
-  ultimately
-  show ?thesis unfolding Under_def by blast
-qed
-
-
-lemma under_UnderS_trans:
-assumes TRANS: "trans r" and ANTISYM: "antisym r" and
-        IN1: "a \<in> under b" and IN2: "b \<in> UnderS C"
-shows "a \<in> UnderS C"
-proof-
-  from IN2 have "b \<in> Under C"
-  using UnderS_subset_Under[of C] by blast
-  with assms under_Under_trans
-  have "a \<in> Under C" by blast
-  (*  *)
-  moreover
-  have "a \<notin> C"
-  proof
-    assume *: "a \<in> C"
-    have 1: "(a,b) \<in> r"
-    using IN1 under_def[of b] by auto
-    have "\<forall>c \<in> C. b \<noteq> c \<and> (b,c) \<in> r"
-    using IN2 UnderS_def[of C] by blast
-    with * have "b \<noteq> a \<and> (b,a) \<in> r" by blast
-    with 1 ANTISYM antisym_def[of r]
-    show False by blast
-  qed
-  (*  *)
-  ultimately
-  show ?thesis unfolding UnderS_def Under_def by fast
-qed
-
-
-end  (* context rel *)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Order_Relation_More_LFP.thy	Mon Nov 18 18:04:44 2013 +0100
@@ -0,0 +1,286 @@
+(*  Title:      HOL/Cardinals/Order_Relation_More_LFP.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Basics on order-like relations (LFP).
+*)
+
+header {* Basics on Order-Like Relations (LFP) *}
+
+theory Order_Relation_More_LFP
+imports "~~/src/HOL/Library/Order_Relation"
+begin
+
+
+text{* In this section, we develop basic concepts and results pertaining
+to order-like relations, i.e., to reflexive and/or transitive and/or symmetric and/or
+total relations.  The development is placed on top of the definitions
+from the theory @{text "Order_Relation"}.  We also
+further define upper and lower bounds operators. *}
+
+
+locale rel = fixes r :: "'a rel"
+
+text{* The following context encompasses all this section, except
+for its last subsection. In other words, for the rest of this section except its last
+subsection, we consider a fixed relation @{text "r"}. *}
+
+context rel
+begin
+
+
+subsection {* Auxiliaries *}
+
+
+lemma refl_on_domain:
+"\<lbrakk>refl_on A r; (a,b) : r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
+by(auto simp add: refl_on_def)
+
+
+corollary well_order_on_domain:
+"\<lbrakk>well_order_on A r; (a,b) \<in> r\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> A"
+by(simp add: refl_on_domain order_on_defs)
+
+
+lemma well_order_on_Field:
+"well_order_on A r \<Longrightarrow> A = Field r"
+by(auto simp add: refl_on_def Field_def order_on_defs)
+
+
+lemma well_order_on_Well_order:
+"well_order_on A r \<Longrightarrow> A = Field r \<and> Well_order r"
+using well_order_on_Field by simp
+
+
+lemma Total_subset_Id:
+assumes TOT: "Total r" and SUB: "r \<le> Id"
+shows "r = {} \<or> (\<exists>a. r = {(a,a)})"
+proof-
+  {assume "r \<noteq> {}"
+   then obtain a b where 1: "(a,b) \<in> r" by fast
+   hence "a = b" using SUB by blast
+   hence 2: "(a,a) \<in> r" using 1 by simp
+   {fix c d assume "(c,d) \<in> r"
+    hence