--- a/Admin/isatest/isatest-stats Wed Sep 12 05:21:47 2012 +0200
+++ b/Admin/isatest/isatest-stats Wed Sep 12 05:29:21 2012 +0200
@@ -97,6 +97,7 @@
HOL-Binomial-Heaps
HOL-Binomial-Queues
HOL-BytecodeLogicJmlTypes
+ HOL-Cardinals
HOL-Category
HOL-Category2
HOL-Cauchy
@@ -146,7 +147,6 @@
HOL-Nominal-Lam-ml-Normalization
HOL-Nominal-SequentInvertibility
HOL-Ordinal
- HOL-Ordinals_and_Cardinals
HOL-POPLmark-deBruijn
HOL-Perfect-Number-Thm
HOL-Polynomials
--- a/NEWS Wed Sep 12 05:21:47 2012 +0200
+++ b/NEWS Wed Sep 12 05:29:21 2012 +0200
@@ -100,8 +100,8 @@
* HOL/Codatatype: New (co)datatype package with support for mixed,
nested recursion and interesting non-free datatypes.
-* HOL/Ordinals_and_Cardinals: Theories of ordinals and cardinals
-(supersedes the AFP entry of the same name).
+* HOL/Cardinals: Theories of ordinals and cardinals
+(supersedes the AFP entry "Ordinals_and_Cardinals").
* Library/Debug.thy and Library/Parallel.thy: debugging and parallel
execution for code generated towards Isabelle/ML.
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,878 @@
+(* Title: HOL/Cardinals/Cardinal_Arithmetic.thy
+ Author: Dmitriy Traytel, TU Muenchen
+ Copyright 2012
+
+Cardinal arithmetic.
+*)
+
+header {* Cardinal Arithmetic *}
+
+theory Cardinal_Arithmetic
+imports Cardinal_Order_Relation_Base
+begin
+
+text {*
+ The following collection of lemmas should be seen as an user interface to the HOL Theory
+ of cardinals. It is not expected to be complete in any sense, since its
+ development was driven by demand arising from the development of the (co)datatype package.
+*}
+
+(*library candidate*)
+lemma dir_image: "\<lbrakk>\<And>x y. (f x = f y) = (x = y); Card_order r\<rbrakk> \<Longrightarrow> r =o dir_image r f"
+by (rule dir_image_ordIso) (auto simp add: inj_on_def card_order_on_def)
+
+(*should supersede a weaker lemma from the library*)
+lemma dir_image_Field: "Field (dir_image r f) = f ` Field r"
+unfolding dir_image_def Field_def Range_def Domain_def by fastforce
+
+lemma card_order_dir_image:
+ assumes bij: "bij f" and co: "card_order r"
+ shows "card_order (dir_image r f)"
+proof -
+ from assms have "Field (dir_image r f) = UNIV"
+ using card_order_on_Card_order[of UNIV r] unfolding bij_def dir_image_Field by auto
+ moreover from bij have "\<And>x y. (f x = f y) = (x = y)" unfolding bij_def inj_on_def by auto
+ with co have "Card_order (dir_image r f)"
+ using card_order_on_Card_order[of UNIV r] Card_order_ordIso2[OF _ dir_image] by blast
+ ultimately show ?thesis by auto
+qed
+
+(*library candidate*)
+lemma ordIso_refl: "Card_order r \<Longrightarrow> r =o r"
+by (rule card_order_on_ordIso)
+
+(*library candidate*)
+lemma ordLeq_refl: "Card_order r \<Longrightarrow> r \<le>o r"
+by (rule ordIso_imp_ordLeq, rule card_order_on_ordIso)
+
+(*library candidate*)
+lemma card_of_ordIso_subst: "A = B \<Longrightarrow> |A| =o |B|"
+by (simp only: ordIso_refl card_of_Card_order)
+
+(*library candidate*)
+lemma card_of_Times_Plus_distrib:
+ "|A <*> (B <+> C)| =o |A <*> B <+> A <*> C|" (is "|?RHS| =o |?LHS|")
+proof -
+ let ?f = "\<lambda>(a, bc). case bc of Inl b \<Rightarrow> Inl (a, b) | Inr c \<Rightarrow> Inr (a, c)"
+ have "bij_betw ?f ?RHS ?LHS" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+(*library candidate*)
+lemma Field_card_order: "card_order r \<Longrightarrow> Field r = UNIV"
+using card_order_on_Card_order[of UNIV r] by simp
+
+subsection {* Zero *}
+
+definition czero where
+ "czero = card_of {}"
+
+lemma czero_ordIso:
+ "czero =o czero"
+using card_of_empty_ordIso by (simp add: czero_def)
+
+lemma card_of_ordIso_czero_iff_empty:
+ "|A| =o (czero :: 'a rel) \<longleftrightarrow> A = ({} :: 'a set)"
+unfolding czero_def by (rule iffI[OF card_of_empty2]) (auto simp: card_of_refl)
+
+(* A "not czero" Cardinal predicate *)
+abbreviation Cnotzero where
+ "Cnotzero (r :: 'a rel) \<equiv> \<not>(r =o (czero :: 'a rel)) \<and> Card_order r"
+
+(*helper*)
+lemma Cnotzero_imp_not_empty: "Cnotzero r \<Longrightarrow> Field r \<noteq> {}"
+by (metis Card_order_iff_ordIso_card_of czero_def)
+
+lemma czeroI:
+ "\<lbrakk>Card_order r; Field r = {}\<rbrakk> \<Longrightarrow> r =o czero"
+using Cnotzero_imp_not_empty ordIso_transitive[OF _ czero_ordIso] by blast
+
+lemma czeroE:
+ "r =o czero \<Longrightarrow> Field r = {}"
+unfolding czero_def
+by (drule card_of_cong) (simp only: Field_card_of card_of_empty2)
+
+lemma Cnotzero_mono:
+ "\<lbrakk>Cnotzero r; Card_order q; r \<le>o q\<rbrakk> \<Longrightarrow> Cnotzero q"
+apply (rule ccontr)
+apply auto
+apply (drule czeroE)
+apply (erule notE)
+apply (erule czeroI)
+apply (drule card_of_mono2)
+apply (simp only: card_of_empty3)
+done
+
+subsection {* Infinite cardinals *}
+
+definition cinfinite where
+ "cinfinite r = infinite (Field r)"
+
+abbreviation Cinfinite where
+ "Cinfinite r \<equiv> cinfinite r \<and> Card_order r"
+
+lemma natLeq_ordLeq_cinfinite:
+ assumes inf: "Cinfinite r"
+ shows "natLeq \<le>o r"
+proof -
+ from inf have "natLeq \<le>o |Field r|" by (simp add: cinfinite_def infinite_iff_natLeq_ordLeq)
+ also from inf have "|Field r| =o r" by (simp add: card_of_unique ordIso_symmetric)
+ finally show ?thesis .
+qed
+
+lemma cinfinite_not_czero: "cinfinite r \<Longrightarrow> \<not> (r =o (czero :: 'a rel))"
+unfolding cinfinite_def by (metis czeroE finite.emptyI)
+
+lemma Cinfinite_Cnotzero: "Cinfinite r \<Longrightarrow> Cnotzero r"
+by (metis cinfinite_not_czero)
+
+lemma Cinfinite_cong: "\<lbrakk>r1 =o r2; Cinfinite r1\<rbrakk> \<Longrightarrow> Cinfinite r2"
+by (metis Card_order_ordIso2 card_of_mono2 card_of_ordLeq_infinite cinfinite_def ordIso_iff_ordLeq)
+
+lemma cinfinite_mono: "\<lbrakk>r1 \<le>o r2; cinfinite r1\<rbrakk> \<Longrightarrow> cinfinite r2"
+by (metis card_of_mono2 card_of_ordLeq_infinite cinfinite_def)
+
+
+subsection {* Binary sum *}
+
+definition csum (infixr "+c" 65) where
+ "r1 +c r2 \<equiv> |Field r1 <+> Field r2|"
+
+lemma Card_order_csum:
+ "Card_order (r1 +c r2)"
+unfolding csum_def by (simp add: card_of_Card_order)
+
+lemma csum_Cnotzero1:
+ "Cnotzero r1 \<Longrightarrow> Cnotzero (r1 +c r2)"
+unfolding csum_def
+by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE)
+
+lemma csum_Cnotzero2:
+ "Cnotzero r2 \<Longrightarrow> Cnotzero (r1 +c r2)"
+unfolding csum_def
+by (metis Cnotzero_imp_not_empty Field_card_of Plus_eq_empty_conv card_of_card_order_on czeroE)
+
+lemma card_order_csum:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 +c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis unfolding csum_def by (auto simp: card_of_card_order_on)
+qed
+
+lemma cinfinite_csum:
+ "cinfinite r1 \<or> cinfinite r2 \<Longrightarrow> cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (auto simp: Field_card_of)
+
+lemma Cinfinite_csum:
+ "Cinfinite r1 \<or> Cinfinite r2 \<Longrightarrow> Cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff)
+
+lemma Cinfinite_csum_strong:
+ "\<lbrakk>Cinfinite r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 +c r2)"
+by (metis Cinfinite_csum)
+
+lemma Cinfinite_csum1:
+ "Cinfinite r1 \<Longrightarrow> Cinfinite (r1 +c r2)"
+unfolding cinfinite_def csum_def by (metis Field_card_of card_of_Card_order finite_Plus_iff)
+
+lemma csum_cong: "\<lbrakk>p1 =o r1; p2 =o r2\<rbrakk> \<Longrightarrow> p1 +c p2 =o r1 +c r2"
+by (simp only: csum_def ordIso_Plus_cong)
+
+lemma csum_cong1: "p1 =o r1 \<Longrightarrow> p1 +c q =o r1 +c q"
+by (simp only: csum_def ordIso_Plus_cong1)
+
+lemma csum_cong2: "p2 =o r2 \<Longrightarrow> q +c p2 =o q +c r2"
+by (simp only: csum_def ordIso_Plus_cong2)
+
+lemma csum_mono: "\<lbrakk>p1 \<le>o r1; p2 \<le>o r2\<rbrakk> \<Longrightarrow> p1 +c p2 \<le>o r1 +c r2"
+by (simp only: csum_def ordLeq_Plus_mono)
+
+lemma csum_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 +c q \<le>o r1 +c q"
+by (simp only: csum_def ordLeq_Plus_mono1)
+
+lemma csum_mono2: "p2 \<le>o r2 \<Longrightarrow> q +c p2 \<le>o q +c r2"
+by (simp only: csum_def ordLeq_Plus_mono2)
+
+lemma ordLeq_csum1: "Card_order p1 \<Longrightarrow> p1 \<le>o p1 +c p2"
+by (simp only: csum_def Card_order_Plus1)
+
+lemma ordLeq_csum2: "Card_order p2 \<Longrightarrow> p2 \<le>o p1 +c p2"
+by (simp only: csum_def Card_order_Plus2)
+
+lemma csum_com: "p1 +c p2 =o p2 +c p1"
+by (simp only: csum_def card_of_Plus_commute)
+
+lemma csum_assoc: "(p1 +c p2) +c p3 =o p1 +c p2 +c p3"
+by (simp only: csum_def Field_card_of card_of_Plus_assoc)
+
+lemma Plus_csum: "|A <+> B| =o |A| +c |B|"
+by (simp only: csum_def Field_card_of card_of_refl)
+
+lemma Un_csum: "|A \<union> B| \<le>o |A| +c |B|"
+using ordLeq_ordIso_trans[OF card_of_Un_Plus_ordLeq Plus_csum] by blast
+
+
+subsection {* One *}
+
+definition cone where
+ "cone = card_of {()}"
+
+lemma Card_order_cone: "Card_order cone"
+unfolding cone_def by (rule card_of_Card_order)
+
+lemma single_cone:
+ "|{x}| =o cone"
+proof -
+ let ?f = "\<lambda>x. ()"
+ have "bij_betw ?f {x} {()}" unfolding bij_betw_def by auto
+ thus ?thesis unfolding cone_def using card_of_ordIso by blast
+qed
+
+lemma cone_not_czero: "\<not> (cone =o czero)"
+unfolding czero_def cone_def by (metis empty_not_insert card_of_empty3[of "{()}"] ordIso_iff_ordLeq)
+
+lemma cone_Cnotzero: "Cnotzero cone"
+by (simp add: cone_not_czero Card_order_cone)
+
+lemma cone_ordLeq_Cnotzero: "Cnotzero r \<Longrightarrow> cone \<le>o r"
+unfolding cone_def by (metis Card_order_singl_ordLeq czeroI)
+
+
+subsection{* Two *}
+
+definition ctwo where
+ "ctwo = |UNIV :: bool set|"
+
+lemma Card_order_ctwo: "Card_order ctwo"
+unfolding ctwo_def by (rule card_of_Card_order)
+
+lemma cone_ordLeq_ctwo: "cone \<le>o ctwo"
+unfolding cone_def ctwo_def card_of_ordLeq[symmetric] by auto
+
+lemma ctwo_not_czero: "\<not> (ctwo =o czero)"
+using card_of_empty3[of "UNIV :: bool set"] ordIso_iff_ordLeq
+unfolding czero_def ctwo_def by (metis UNIV_not_empty)
+
+lemma ctwo_Cnotzero: "Cnotzero ctwo"
+by (simp add: ctwo_not_czero Card_order_ctwo)
+
+
+subsection {* Family sum *}
+
+definition Csum where
+ "Csum r rs \<equiv> |SIGMA i : Field r. Field (rs i)|"
+
+(* Similar setup to the one for SIGMA from theory Big_Operators: *)
+syntax "_Csum" ::
+ "pttrn => ('a * 'a) set => 'b * 'b set => (('a * 'b) * ('a * 'b)) set"
+ ("(3CSUM _:_. _)" [0, 51, 10] 10)
+
+translations
+ "CSUM i:r. rs" == "CONST Csum r (%i. rs)"
+
+lemma SIGMA_CSUM: "|SIGMA i : I. As i| = (CSUM i : |I|. |As i| )"
+by (auto simp: Csum_def Field_card_of)
+
+(* NB: Always, under the cardinal operator,
+operations on sets are reduced automatically to operations on cardinals.
+This should make cardinal reasoning more direct and natural. *)
+
+
+subsection {* Product *}
+
+definition cprod (infixr "*c" 80) where
+ "r1 *c r2 = |Field r1 <*> Field r2|"
+
+lemma Times_cprod: "|A \<times> B| =o |A| *c |B|"
+by (simp only: cprod_def Field_card_of card_of_refl)
+
+lemma card_order_cprod:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 *c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis by (auto simp: cprod_def card_of_card_order_on)
+qed
+
+lemma Card_order_cprod: "Card_order (r1 *c r2)"
+by (simp only: cprod_def Field_card_of card_of_card_order_on)
+
+lemma cprod_cong2: "p2 =o r2 \<Longrightarrow> q *c p2 =o q *c r2"
+by (simp only: cprod_def ordIso_Times_cong2)
+
+lemma cprod_mono1: "p1 \<le>o r1 \<Longrightarrow> p1 *c q \<le>o r1 *c q"
+by (simp only: cprod_def ordLeq_Times_mono1)
+
+lemma cprod_mono2: "p2 \<le>o r2 \<Longrightarrow> q *c p2 \<le>o q *c r2"
+by (simp only: cprod_def ordLeq_Times_mono2)
+
+lemma ordLeq_cprod1: "\<lbrakk>Card_order p1; Cnotzero p2\<rbrakk> \<Longrightarrow> p1 \<le>o p1 *c p2"
+unfolding cprod_def by (metis Card_order_Times1 czeroI)
+
+lemma ordLeq_cprod2: "\<lbrakk>Cnotzero p1; Card_order p2\<rbrakk> \<Longrightarrow> p2 \<le>o p1 *c p2"
+unfolding cprod_def by (metis Card_order_Times2 czeroI)
+
+lemma cinfinite_cprod: "\<lbrakk>cinfinite r1; cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"
+by (simp add: cinfinite_def cprod_def Field_card_of infinite_cartesian_product)
+
+lemma cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> cinfinite (r1 *c r2)"
+by (metis cinfinite_mono ordLeq_cprod2)
+
+lemma Cinfinite_cprod2: "\<lbrakk>Cnotzero r1; Cinfinite r2\<rbrakk> \<Longrightarrow> Cinfinite (r1 *c r2)"
+by (blast intro: cinfinite_cprod2 Card_order_cprod)
+
+lemma cprod_com: "p1 *c p2 =o p2 *c p1"
+by (simp only: cprod_def card_of_Times_commute)
+
+lemma card_of_Csum_Times:
+ "\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> (CSUM i : |I|. |A i| ) \<le>o |I| *c |B|"
+by (simp only: Csum_def cprod_def Field_card_of card_of_Sigma_Times)
+
+lemma card_of_Csum_Times':
+ assumes "Card_order r" "\<forall>i \<in> I. |A i| \<le>o r"
+ shows "(CSUM i : |I|. |A i| ) \<le>o |I| *c r"
+proof -
+ from assms(1) have *: "r =o |Field r|" by (simp add: card_of_unique)
+ with assms(2) have "\<forall>i \<in> I. |A i| \<le>o |Field r|" by (blast intro: ordLeq_ordIso_trans)
+ hence "(CSUM i : |I|. |A i| ) \<le>o |I| *c |Field r|" by (simp only: card_of_Csum_Times)
+ also from * have "|I| *c |Field r| \<le>o |I| *c r"
+ by (simp only: Field_card_of card_of_refl cprod_def ordIso_imp_ordLeq)
+ finally show ?thesis .
+qed
+
+lemma cprod_csum_distrib1: "r1 *c r2 +c r1 *c r3 =o r1 *c (r2 +c r3)"
+unfolding csum_def cprod_def by (simp add: Field_card_of card_of_Times_Plus_distrib ordIso_symmetric)
+
+lemma csum_absorb2': "\<lbrakk>Card_order r2; r1 \<le>o r2; cinfinite r1 \<or> cinfinite r2\<rbrakk> \<Longrightarrow> r1 +c r2 =o r2"
+unfolding csum_def by (metis Card_order_Plus_infinite cinfinite_def cinfinite_mono)
+
+lemma csum_absorb1':
+ assumes card: "Card_order r2"
+ and r12: "r1 \<le>o r2" and cr12: "cinfinite r1 \<or> cinfinite r2"
+ shows "r2 +c r1 =o r2"
+by (rule ordIso_transitive, rule csum_com, rule csum_absorb2', (simp only: assms)+)
+
+lemma csum_absorb1: "\<lbrakk>Cinfinite r2; r1 \<le>o r2\<rbrakk> \<Longrightarrow> r2 +c r1 =o r2"
+by (rule csum_absorb1') auto
+
+lemma cprod_infinite1': "\<lbrakk>Cinfinite r; Cnotzero p; p \<le>o r\<rbrakk> \<Longrightarrow> r *c p =o r"
+unfolding cinfinite_def cprod_def
+by (rule Card_order_Times_infinite[THEN conjunct1]) (blast intro: czeroI)+
+
+lemma cprod_infinite: "Cinfinite r \<Longrightarrow> r *c r =o r"
+using cprod_infinite1' Cinfinite_Cnotzero ordLeq_refl by blast
+
+
+subsection {* Exponentiation *}
+
+definition cexp (infixr "^c" 80) where
+ "r1 ^c r2 \<equiv> |Func (Field r2) (Field r1)|"
+
+definition ccexp (infixr "^^c" 80) where
+ "r1 ^^c r2 \<equiv> |Pfunc (Field r2) (Field r1)|"
+
+lemma cexp_ordLeq_ccexp: "r1 ^c r2 \<le>o r1 ^^c r2"
+unfolding cexp_def ccexp_def by (rule card_of_mono1) (rule Func_Pfunc)
+
+lemma card_order_ccexp:
+ assumes "card_order r1" "card_order r2"
+ shows "card_order (r1 ^^c r2)"
+proof -
+ have "Field r1 = UNIV" "Field r2 = UNIV" using assms card_order_on_Card_order by auto
+ thus ?thesis unfolding ccexp_def Pfunc_def
+ by (auto simp: card_of_card_order_on split: option.split)
+qed
+
+lemma Card_order_cexp: "Card_order (r1 ^c r2)"
+unfolding cexp_def by (rule card_of_Card_order)
+
+lemma cexp_mono':
+ assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"
+ and n1: "Field p1 \<noteq> {} \<or> cone \<le>o r1 ^c r2"
+ and n2: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+proof(cases "Field p1 = {}")
+ case True
+ hence "|Field (p1 ^c p2)| \<le>o cone"
+ unfolding czero_def cone_def cexp_def Field_card_of
+ by (cases "Field p2 = {}", auto intro: card_of_ordLeqI2 simp: Func_empty)
+ (metis Func_is_emp card_of_empty ex_in_conv)
+ hence "p1 ^c p2 \<le>o cone" by (simp add: Field_card_of cexp_def)
+ thus ?thesis using True n1 ordLeq_transitive by auto
+next
+ case False
+ have 1: "|Field p1| \<le>o |Field r1|" and 2: "|Field p2| \<le>o |Field r2|"
+ using 1 2 by (auto simp: card_of_mono2)
+ obtain f1 where f1: "f1 ` Field r1 = Field p1"
+ using 1 unfolding card_of_ordLeq2[OF False, symmetric] by auto
+ obtain f2 where f2: "inj_on f2 (Field p2)" "f2 ` Field p2 \<subseteq> Field r2"
+ using 2 unfolding card_of_ordLeq[symmetric] by blast
+ have 0: "Func_map (Field p2) f1 f2 ` (Field (r1 ^c r2)) = Field (p1 ^c p2)"
+ unfolding cexp_def Field_card_of using Func_map_surj[OF f1 f2 n2, symmetric] .
+ have 00: "Field (p1 ^c p2) \<noteq> {}" unfolding cexp_def Field_card_of Func_is_emp
+ using False by simp
+ show ?thesis
+ using 0 card_of_ordLeq2[OF 00] unfolding cexp_def Field_card_of by blast
+qed
+
+lemma cexp_mono:
+ assumes 1: "p1 \<le>o r1" and 2: "p2 \<le>o r2"
+ and n1: "Cnotzero p1 \<or> cone \<le>o r1 ^c r2"
+ and n2: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
+ shows "p1 ^c p2 \<le>o r1 ^c r2"
+proof (rule cexp_mono'[OF 1 2])
+ show "Field p1 \<noteq> {} \<or> cone \<le>o r1 ^c r2"
+ proof (cases "Cnotzero p1")
+ case True show ?thesis using Cnotzero_imp_not_empty[OF True] by (rule disjI1)
+ next
+ case False with n1 show ?thesis by blast
+ qed
+qed (rule czeroI[OF card, THEN n2, THEN czeroE])
+
+lemma cexp_mono1:
+ assumes 1: "p1 \<le>o r1"
+ and n1: "Cnotzero p1 \<or> cone \<le>o r1 ^c q" and q: "Card_order q"
+ shows "p1 ^c q \<le>o r1 ^c q"
+using ordLeq_refl[OF q] by (rule cexp_mono[OF 1 _ n1]) (auto simp: q)
+
+lemma cexp_mono1_Cnotzero: "\<lbrakk>p1 \<le>o r1; Cnotzero p1; Card_order q\<rbrakk> \<Longrightarrow> p1 ^c q \<le>o r1 ^c q"
+by (simp add: cexp_mono1)
+
+lemma cexp_mono1_cone_ordLeq: "\<lbrakk>p1 \<le>o r1; cone \<le>o r1 ^c q; Card_order q\<rbrakk> \<Longrightarrow> p1 ^c q \<le>o r1 ^c q"
+using assms by (simp add: cexp_mono1)
+
+lemma cexp_mono2':
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n1: "Field q \<noteq> {} \<or> cone \<le>o q ^c r2"
+ and n2: "Field p2 = {} \<Longrightarrow> Field r2 = {}"
+ shows "q ^c p2 \<le>o q ^c r2"
+using ordLeq_refl[OF q] by (rule cexp_mono'[OF _ 2 n1 n2]) auto
+
+lemma cexp_mono2:
+ assumes 2: "p2 \<le>o r2" and q: "Card_order q"
+ and n1: "Cnotzero q \<or> cone \<le>o q ^c r2"
+ and n2: "p2 =o czero \<Longrightarrow> r2 =o czero" and card: "Card_order p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+using ordLeq_refl[OF q] by (rule cexp_mono[OF _ 2 n1 n2 card]) auto
+
+lemma cexp_mono2_Cnotzero:
+ assumes "p2 \<le>o r2" "Cnotzero q" and n2: "Cnotzero p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+proof (rule cexp_mono)
+ assume *: "p2 =o czero"
+ have False using n2 czeroI czeroE[OF *] by blast
+ thus "r2 =o czero" by blast
+qed (auto simp add: assms ordLeq_refl)
+
+lemma cexp_cong:
+ assumes 1: "p1 =o r1" and 2: "p2 =o r2"
+ and p1: "Cnotzero p1 \<or> cone \<le>o r1 ^c r2" and Cr: "Card_order r2"
+ and r1: "Cnotzero r1 \<or> cone \<le>o p1 ^c p2" and Cp: "Card_order p2"
+ shows "p1 ^c p2 =o r1 ^c r2"
+proof -
+ obtain f where "bij_betw f (Field p2) (Field r2)"
+ using 2 card_of_ordIso[of "Field p2" "Field r2"] card_of_cong by auto
+ hence 0: "Field p2 = {} \<longleftrightarrow> Field r2 = {}" unfolding bij_betw_def by auto
+ have r: "p2 =o czero \<Longrightarrow> r2 =o czero"
+ and p: "r2 =o czero \<Longrightarrow> p2 =o czero"
+ using 0 Cr Cp czeroE czeroI by auto
+ show ?thesis using 0 1 2 unfolding ordIso_iff_ordLeq
+ using r p cexp_mono[OF _ _ p1 _ Cp] cexp_mono[OF _ _ r1 _ Cr]
+ by blast
+qed
+
+lemma cexp_cong1:
+ assumes 1: "p1 =o r1" and q: "Card_order q"
+ and p1: "Cnotzero p1 \<or> cone \<le>o r1 ^c q"
+ and r1: "Cnotzero r1 \<or> cone \<le>o p1 ^c q"
+ shows "p1 ^c q =o r1 ^c q"
+by (rule cexp_cong[OF 1 _ p1 q r1 q]) (rule ordIso_refl[OF q])
+
+lemma cexp_cong1_Cnotzero:
+ assumes "p1 =o r1" "Card_order q" "Cnotzero p1" "Cnotzero r1"
+ shows "p1 ^c q =o r1 ^c q"
+by (rule cexp_cong1, auto simp add: assms)
+
+lemma cexp_cong2:
+ assumes 2: "p2 =o r2" and q: "Card_order q"
+ and p: "Card_order p2" and r: "Card_order r2"
+ shows "Cnotzero q \<or> (cone \<le>o q ^c p2 \<and> cone \<le>o q ^c r2) \<Longrightarrow>
+ q ^c p2 =o q ^c r2"
+by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl q p r)
+
+lemma cexp_cong2_Cnotzero:
+ assumes 2: "p2 =o r2" and q: "Cnotzero q"
+ and p: "Card_order p2"
+ shows "q ^c p2 =o q ^c r2"
+by (rule cexp_cong[OF _ 2]) (auto simp only: ordIso_refl Card_order_ordIso2[OF p 2] q p)
+
+lemma cexp_czero: "r ^c czero =o cone"
+unfolding cexp_def czero_def Field_card_of Func_empty by (rule single_cone)
+
+lemma cexp_cone:
+ assumes "Card_order r"
+ shows "r ^c cone =o r"
+proof -
+ have "r ^c cone =o |Field r|"
+ unfolding cexp_def cone_def Field_card_of Func_empty
+ card_of_ordIso[symmetric] bij_betw_def Func_def inj_on_def image_def
+ by (rule exI[of _ "\<lambda>f. case f () of Some a \<Rightarrow> a"]) auto
+ also have "|Field r| =o r" by (rule card_of_Field_ordIso[OF assms])
+ finally show ?thesis .
+qed
+
+lemma cexp_cprod:
+ assumes r1: "Cnotzero r1"
+ shows "(r1 ^c r2) ^c r3 =o r1 ^c (r2 *c r3)" (is "?L =o ?R")
+proof -
+ have "?L =o r1 ^c (r3 *c r2)"
+ unfolding cprod_def cexp_def Field_card_of
+ using card_of_Func_Times by(rule ordIso_symmetric)
+ also have "r1 ^c (r3 *c r2) =o ?R"
+ apply(rule cexp_cong2) using cprod_com r1 by (auto simp: Card_order_cprod)
+ finally show ?thesis .
+qed
+
+lemma cexp_cprod_ordLeq:
+ assumes r1: "Cnotzero r1" and r2: "Cinfinite r2"
+ and r3: "Cnotzero r3" "r3 \<le>o r2"
+ shows "(r1 ^c r2) ^c r3 =o r1 ^c r2" (is "?L =o ?R")
+proof-
+ have "?L =o r1 ^c (r2 *c r3)" using cexp_cprod[OF r1] .
+ also have "r1 ^c (r2 *c r3) =o ?R"
+ apply(rule cexp_cong2)
+ apply(rule cprod_infinite1'[OF r2 r3]) using r1 r2 by (fastforce simp: Card_order_cprod)+
+ finally show ?thesis .
+qed
+
+lemma Cnotzero_UNIV: "Cnotzero |UNIV|"
+by (auto simp: card_of_Card_order card_of_ordIso_czero_iff_empty)
+
+lemma Pow_cexp_ctwo:
+ "|Pow A| =o ctwo ^c |A|"
+unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+
+lemma ordLess_ctwo_cexp:
+ assumes "Card_order r"
+ shows "r <o ctwo ^c r"
+proof -
+ have "r <o |Pow (Field r)|" using assms by (rule Card_order_Pow)
+ also have "|Pow (Field r)| =o ctwo ^c r"
+ unfolding ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+ finally show ?thesis .
+qed
+
+lemma ordLeq_cexp1:
+ assumes "Cnotzero r" "Card_order q"
+ shows "q \<le>o q ^c r"
+proof (cases "q =o (czero :: 'a rel)")
+ case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)
+next
+ case False
+ thus ?thesis
+ apply -
+ apply (rule ordIso_ordLeq_trans)
+ apply (rule ordIso_symmetric)
+ apply (rule cexp_cone)
+ apply (rule assms(2))
+ apply (rule cexp_mono2)
+ apply (rule cone_ordLeq_Cnotzero)
+ apply (rule assms(1))
+ apply (rule assms(2))
+ apply (rule disjI1)
+ apply (rule conjI)
+ apply (rule notI)
+ apply (erule notE)
+ apply (rule ordIso_transitive)
+ apply assumption
+ apply (rule czero_ordIso)
+ apply (rule assms(2))
+ apply (rule notE)
+ apply (rule cone_not_czero)
+ apply assumption
+ apply (rule Card_order_cone)
+ done
+qed
+
+lemma ordLeq_cexp2:
+ assumes "ctwo \<le>o q" "Card_order r"
+ shows "r \<le>o q ^c r"
+proof (cases "r =o (czero :: 'a rel)")
+ case True thus ?thesis by (simp only: card_of_empty cexp_def czero_def ordIso_ordLeq_trans)
+next
+ case False thus ?thesis
+ apply -
+ apply (rule ordLess_imp_ordLeq)
+ apply (rule ordLess_ordLeq_trans)
+ apply (rule ordLess_ctwo_cexp)
+ apply (rule assms(2))
+ apply (rule cexp_mono1)
+ apply (rule assms(1))
+ apply (rule disjI1)
+ apply (rule ctwo_Cnotzero)
+ apply (rule assms(2))
+ done
+qed
+
+lemma Cnotzero_cexp:
+ assumes "Cnotzero q" "Card_order r"
+ shows "Cnotzero (q ^c r)"
+proof (cases "r =o czero")
+ case False thus ?thesis
+ apply -
+ apply (rule Cnotzero_mono)
+ apply (rule assms(1))
+ apply (rule Card_order_cexp)
+ apply (rule ordLeq_cexp1)
+ apply (rule conjI)
+ apply (rule notI)
+ apply (erule notE)
+ apply (rule ordIso_transitive)
+ apply assumption
+ apply (rule czero_ordIso)
+ apply (rule assms(2))
+ apply (rule conjunct2)
+ apply (rule assms(1))
+ done
+next
+ case True thus ?thesis
+ apply -
+ apply (rule Cnotzero_mono)
+ apply (rule cone_Cnotzero)
+ apply (rule Card_order_cexp)
+ apply (rule ordIso_imp_ordLeq)
+ apply (rule ordIso_symmetric)
+ apply (rule ordIso_transitive)
+ apply (rule cexp_cong2)
+ apply assumption
+ apply (rule conjunct2)
+ apply (rule assms(1))
+ apply (rule assms(2))
+ apply (simp only: card_of_Card_order czero_def)
+ apply (rule disjI1)
+ apply (rule assms(1))
+ apply (rule cexp_czero)
+ done
+qed
+
+lemma Cinfinite_ctwo_cexp:
+ "Cinfinite r \<Longrightarrow> Cinfinite (ctwo ^c r)"
+unfolding ctwo_def cexp_def cinfinite_def Field_card_of
+by (rule conjI, rule infinite_Func, auto, rule card_of_card_order_on)
+
+lemma cinfinite_cexp: "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> cinfinite (q ^c r)"
+by (metis assms cinfinite_mono ordLeq_cexp2)
+
+lemma cinfinite_ccexp: "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> cinfinite (q ^^c r)"
+by (rule cinfinite_mono[OF cexp_ordLeq_ccexp cinfinite_cexp])
+
+lemma Cinfinite_cexp:
+ "\<lbrakk>ctwo \<le>o q; Cinfinite r\<rbrakk> \<Longrightarrow> Cinfinite (q ^c r)"
+by (simp add: cinfinite_cexp Card_order_cexp)
+
+lemma ctwo_ordLess_natLeq:
+ "ctwo <o natLeq"
+unfolding ctwo_def using finite_iff_ordLess_natLeq finite_UNIV by fast
+
+lemma ctwo_ordLess_Cinfinite: "Cinfinite r \<Longrightarrow> ctwo <o r"
+by (metis ctwo_ordLess_natLeq natLeq_ordLeq_cinfinite ordLess_ordLeq_trans)
+
+lemma ctwo_ordLeq_Cinfinite:
+ assumes "Cinfinite r"
+ shows "ctwo \<le>o r"
+by (rule ordLess_imp_ordLeq[OF ctwo_ordLess_Cinfinite[OF assms]])
+
+lemma Cinfinite_ordLess_cexp:
+ assumes r: "Cinfinite r"
+ shows "r <o r ^c r"
+proof -
+ have "r <o ctwo ^c r" using r by (simp only: ordLess_ctwo_cexp)
+ also have "ctwo ^c r \<le>o r ^c r"
+ by (rule cexp_mono1[OF ctwo_ordLeq_Cinfinite]) (auto simp: r ctwo_not_czero Card_order_ctwo)
+ finally show ?thesis .
+qed
+
+lemma infinite_ordLeq_cexp:
+ assumes "Cinfinite r"
+ shows "r \<le>o r ^c r"
+by (rule ordLess_imp_ordLeq[OF Cinfinite_ordLess_cexp[OF assms]])
+
+lemma cone_ordLeq_iff_Field:
+ assumes "cone \<le>o r"
+ shows "Field r \<noteq> {}"
+proof (rule ccontr)
+ assume "\<not> Field r \<noteq> {}"
+ hence "Field r = {}" by simp
+ thus False using card_of_empty3
+ card_of_mono2[OF assms] Cnotzero_imp_not_empty[OF cone_Cnotzero] by auto
+qed
+
+lemma cone_ordLeq_cexp: "cone \<le>o r1 \<Longrightarrow> cone \<le>o r1 ^c r2"
+by (simp add: cexp_def cone_def Func_non_emp card_of_singl_ordLeq cone_ordLeq_iff_Field)
+
+lemma Card_order_czero: "Card_order czero"
+by (simp only: card_of_Card_order czero_def)
+
+lemma cexp_mono2'':
+ assumes 2: "p2 \<le>o r2"
+ and n1: "Cnotzero q"
+ and n2: "Card_order p2"
+ shows "q ^c p2 \<le>o q ^c r2"
+proof (cases "p2 =o (czero :: 'a rel)")
+ case True
+ hence "q ^c p2 =o q ^c (czero :: 'a rel)" using n1 n2 cexp_cong2 Card_order_czero by blast
+ also have "q ^c (czero :: 'a rel) =o cone" using cexp_czero by blast
+ also have "cone \<le>o q ^c r2" using cone_ordLeq_cexp cone_ordLeq_Cnotzero n1 by blast
+ finally show ?thesis .
+next
+ case False thus ?thesis using assms cexp_mono2' czeroI by metis
+qed
+
+lemma Un_Cinfinite_bound: "\<lbrakk>|A| \<le>o r; |B| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |A \<union> B| \<le>o r"
+by (auto simp add: cinfinite_def card_of_Un_ordLeq_infinite_Field)
+
+lemma UNION_Cinfinite_bound: "\<lbrakk>|I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r; Cinfinite r\<rbrakk> \<Longrightarrow> |\<Union>i \<in> I. A i| \<le>o r"
+by (auto simp add: card_of_UNION_ordLeq_infinite_Field cinfinite_def)
+
+lemma csum_cinfinite_bound:
+ assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"
+ shows "p +c q \<le>o r"
+proof -
+ from assms(1-4) have "|Field p| \<le>o r" "|Field q| \<le>o r"
+ unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+
+ with assms show ?thesis unfolding cinfinite_def csum_def
+ by (blast intro: card_of_Plus_ordLeq_infinite_Field)
+qed
+
+lemma csum_cexp: "\<lbrakk>Cinfinite r1; Cinfinite r2; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow>
+ q ^c r1 +c q ^c r2 \<le>o q ^c (r1 +c r2)"
+apply (rule csum_cinfinite_bound)
+apply (rule cexp_mono2)
+apply (rule ordLeq_csum1)
+apply (erule conjunct2)
+apply assumption
+apply (rule disjI2)
+apply (rule ordLeq_transitive)
+apply (rule cone_ordLeq_ctwo)
+apply (rule ordLeq_transitive)
+apply assumption
+apply (rule ordLeq_cexp1)
+apply (rule Cinfinite_Cnotzero)
+apply (rule Cinfinite_csum)
+apply (rule disjI1)
+apply assumption
+apply assumption
+apply (rule notE)
+apply (rule cinfinite_not_czero[of r1])
+apply (erule conjunct1)
+apply assumption
+apply (erule conjunct2)
+apply (rule cexp_mono2)
+apply (rule ordLeq_csum2)
+apply (erule conjunct2)
+apply assumption
+apply (rule disjI2)
+apply (rule ordLeq_transitive)
+apply (rule cone_ordLeq_ctwo)
+apply (rule ordLeq_transitive)
+apply assumption
+apply (rule ordLeq_cexp1)
+apply (rule Cinfinite_Cnotzero)
+apply (rule Cinfinite_csum)
+apply (rule disjI1)
+apply assumption
+apply assumption
+apply (rule notE)
+apply (rule cinfinite_not_czero[of r2])
+apply (erule conjunct1)
+apply assumption
+apply (erule conjunct2)
+apply (rule Card_order_cexp)
+apply (rule Card_order_cexp)
+apply (rule Cinfinite_cexp)
+apply assumption
+apply (rule Cinfinite_csum)
+by (rule disjI1)
+
+lemma csum_cexp': "\<lbrakk>Cinfinite r; Card_order q; ctwo \<le>o q\<rbrakk> \<Longrightarrow> q +c r \<le>o q ^c r"
+apply (rule csum_cinfinite_bound)
+ apply (metis Cinfinite_Cnotzero ordLeq_cexp1)
+ apply (metis ordLeq_cexp2)
+ apply blast+
+by (metis Cinfinite_cexp)
+
+lemma cprod_cinfinite_bound:
+ assumes "p \<le>o r" "q \<le>o r" "Card_order p" "Card_order q" "Cinfinite r"
+ shows "p *c q \<le>o r"
+proof -
+ from assms(1-4) have "|Field p| \<le>o r" "|Field q| \<le>o r"
+ unfolding card_order_on_def using card_of_least ordLeq_transitive by blast+
+ with assms show ?thesis unfolding cinfinite_def cprod_def
+ by (blast intro: card_of_Times_ordLeq_infinite_Field)
+qed
+
+lemma cprod_csum_cexp:
+ "r1 *c r2 \<le>o (r1 +c r2) ^c ctwo"
+unfolding cprod_def csum_def cexp_def ctwo_def Field_card_of
+proof -
+ let ?f = "\<lambda>(a, b). %x. if x then Some (Inl a) else Some (Inr b)"
+ have "inj_on ?f (Field r1 \<times> Field r2)" (is "inj_on _ ?LHS")
+ by (auto simp: inj_on_def fun_eq_iff split: bool.split)
+ moreover
+ have "?f ` ?LHS \<subseteq> Func (UNIV :: bool set) (Field r1 <+> Field r2)" (is "_ \<subseteq> ?RHS")
+ by (auto simp: Func_def)
+ ultimately show "|?LHS| \<le>o |?RHS|" using card_of_ordLeq by blast
+qed
+
+lemma card_of_Sigma_ordLeq_Cinfinite:
+ "\<lbrakk>Cinfinite r; |I| \<le>o r; \<forall>i \<in> I. |A i| \<le>o r\<rbrakk> \<Longrightarrow> |SIGMA i : I. A i| \<le>o r"
+unfolding cinfinite_def by (blast intro: card_of_Sigma_ordLeq_infinite_Field)
+
+
+(* cardSuc *)
+
+lemma Cinfinite_cardSuc: "Cinfinite r \<Longrightarrow> Cinfinite (cardSuc r)"
+by (simp add: cinfinite_def cardSuc_Card_order cardSuc_finite)
+
+lemma cardSuc_UNION_Cinfinite:
+ assumes "Cinfinite r" "relChain (cardSuc r) As" "B \<le> (UN i : Field (cardSuc r). As i)" "|B| <=o r"
+ shows "EX i : Field (cardSuc r). B \<le> As i"
+using cardSuc_UNION assms unfolding cinfinite_def by blast
+
+subsection {* Powerset *}
+
+definition cpow where "cpow r = |Pow (Field r)|"
+
+lemma card_order_cpow: "card_order r \<Longrightarrow> card_order (cpow r)"
+by (simp only: cpow_def Field_card_order Pow_UNIV card_of_card_order_on)
+
+lemma cpow_greater_eq: "Card_order r \<Longrightarrow> r \<le>o cpow r"
+by (rule ordLess_imp_ordLeq) (simp only: cpow_def Card_order_Pow)
+
+lemma Card_order_cpow: "Card_order (cpow r)"
+unfolding cpow_def by (rule card_of_Card_order)
+
+lemma Cinfinite_cpow: "Cinfinite r \<Longrightarrow> Cinfinite (cpow r)"
+unfolding cpow_def cinfinite_def by (metis Field_card_of card_of_Card_order infinite_Pow)
+
+lemma cardSuc_ordLeq_cpow: "Card_order r \<Longrightarrow> cardSuc r \<le>o cpow r"
+unfolding cpow_def by (metis Card_order_Pow cardSuc_ordLess_ordLeq card_of_Card_order)
+
+lemma cpow_cexp_ctwo: "cpow r =o ctwo ^c r"
+unfolding cpow_def ctwo_def cexp_def Field_card_of by (rule card_of_Pow_Func)
+
+subsection {* Lists *}
+
+definition clists where "clists r = |lists (Field r)|"
+
+lemma clists_Cinfinite: "Cinfinite r \<Longrightarrow> clists r =o r"
+unfolding cinfinite_def clists_def by (blast intro: Card_order_lists_infinite)
+
+lemma Card_order_clists: "Card_order (clists r)"
+unfolding clists_def by (rule card_of_Card_order)
+
+lemma Cnotzero_clists: "Cnotzero (clists r)"
+by (simp add: clists_def card_of_ordIso_czero_iff_empty lists_not_empty) (rule card_of_Card_order)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,1097 @@
+(* Title: HOL/Cardinals/Cardinal_Order_Relation.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Cardinal-order relations.
+*)
+
+header {* Cardinal-Order Relations *}
+
+theory Cardinal_Order_Relation
+imports Cardinal_Order_Relation_Base Constructions_on_Wellorders
+begin
+
+declare
+ card_order_on_well_order_on[simp]
+ card_of_card_order_on[simp]
+ card_of_well_order_on[simp]
+ Field_card_of[simp]
+ card_of_Card_order[simp]
+ card_of_Well_order[simp]
+ card_of_least[simp]
+ card_of_unique[simp]
+ card_of_mono1[simp]
+ card_of_mono2[simp]
+ card_of_cong[simp]
+ card_of_Field_ordLess[simp]
+ card_of_Field_ordIso[simp]
+ card_of_underS[simp]
+ ordLess_Field[simp]
+ card_of_empty[simp]
+ card_of_empty1[simp]
+ card_of_image[simp]
+ card_of_singl_ordLeq[simp]
+ Card_order_singl_ordLeq[simp]
+ card_of_Pow[simp]
+ Card_order_Pow[simp]
+ card_of_set_type[simp]
+ card_of_Plus1[simp]
+ Card_order_Plus1[simp]
+ card_of_Plus2[simp]
+ Card_order_Plus2[simp]
+ card_of_Plus_mono1[simp]
+ card_of_Plus_mono2[simp]
+ card_of_Plus_mono[simp]
+ card_of_Plus_cong2[simp]
+ card_of_Plus_cong[simp]
+ card_of_Un1[simp]
+ card_of_diff[simp]
+ card_of_Un_Plus_ordLeq[simp]
+ card_of_Times1[simp]
+ card_of_Times2[simp]
+ card_of_Times3[simp]
+ card_of_Times_mono1[simp]
+ card_of_Times_mono2[simp]
+ card_of_Times_cong1[simp]
+ card_of_Times_cong2[simp]
+ card_of_ordIso_finite[simp]
+ finite_ordLess_infinite2[simp]
+ card_of_Times_same_infinite[simp]
+ card_of_Times_infinite_simps[simp]
+ card_of_Plus_infinite1[simp]
+ card_of_Plus_infinite2[simp]
+ card_of_Plus_ordLess_infinite[simp]
+ card_of_Plus_ordLess_infinite_Field[simp]
+ card_of_lists_infinite[simp]
+ infinite_cartesian_product[simp]
+ cardSuc_Card_order[simp]
+ cardSuc_greater[simp]
+ cardSuc_ordLeq[simp]
+ cardSuc_ordLeq_ordLess[simp]
+ cardSuc_mono_ordLeq[simp]
+ cardSuc_invar_ordIso[simp]
+ card_of_cardSuc_finite[simp]
+ cardSuc_finite[simp]
+ card_of_Plus_ordLeq_infinite_Field[simp]
+ curr_in[intro, simp]
+ Func_empty[simp]
+ Func_map_empty[simp]
+ Func_is_emp[simp]
+
+
+subsection {* Cardinal of a set *}
+
+lemma card_of_inj_rel: assumes INJ: "!! x y y'. \<lbrakk>(x,y) : R; (x,y') : R\<rbrakk> \<Longrightarrow> y = y'"
+shows "|{y. EX x. (x,y) : R}| <=o |{x. EX y. (x,y) : R}|"
+proof-
+ let ?Y = "{y. EX x. (x,y) : R}" let ?X = "{x. EX y. (x,y) : R}"
+ let ?f = "% y. SOME x. (x,y) : R"
+ have "?f ` ?Y <= ?X" using someI by force (* FIXME: takes a bit long *)
+ moreover have "inj_on ?f ?Y"
+ unfolding inj_on_def proof(auto)
+ fix y1 x1 y2 x2
+ assume *: "(x1, y1) \<in> R" "(x2, y2) \<in> R" and **: "?f y1 = ?f y2"
+ hence "(?f y1,y1) : R" using someI[of "% x. (x,y1) : R"] by auto
+ moreover have "(?f y2,y2) : R" using * someI[of "% x. (x,y2) : R"] by auto
+ ultimately show "y1 = y2" using ** INJ by auto
+ qed
+ ultimately show "|?Y| <=o |?X|" using card_of_ordLeq by blast
+qed
+
+lemma card_of_unique2: "\<lbrakk>card_order_on B r; bij_betw f A B\<rbrakk> \<Longrightarrow> r =o |A|"
+using card_of_ordIso card_of_unique ordIso_equivalence by blast
+
+lemma internalize_card_of_ordLess:
+"( |A| <o r) = (\<exists>B < Field r. |A| =o |B| \<and> |B| <o r)"
+proof
+ assume "|A| <o r"
+ then obtain p where 1: "Field p < Field r \<and> |A| =o p \<and> p <o r"
+ using internalize_ordLess[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| <o r"
+ using 1 ordIso_equivalence ordIso_ordLess_trans by blast
+ thus "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r" using 1 by blast
+next
+ assume "\<exists>B < Field r. |A| =o |B| \<and> |B| <o r"
+ thus "|A| <o r" using ordIso_ordLess_trans by blast
+qed
+
+lemma internalize_card_of_ordLess2:
+"( |A| <o |C| ) = (\<exists>B < C. |A| =o |B| \<and> |B| <o |C| )"
+using internalize_card_of_ordLess[of "A" "|C|"] Field_card_of[of C] by auto
+
+lemma Card_order_omax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Card_order r"
+shows "Card_order (omax R)"
+proof-
+ have "\<forall>r\<in>R. Well_order r"
+ using assms unfolding card_order_on_def by simp
+ thus ?thesis using assms apply - apply(drule omax_in) by auto
+qed
+
+lemma Card_order_omax2:
+assumes "finite I" and "I \<noteq> {}"
+shows "Card_order (omax {|A i| | i. i \<in> I})"
+proof-
+ let ?R = "{|A i| | i. i \<in> I}"
+ have "finite ?R" and "?R \<noteq> {}" using assms by auto
+ moreover have "\<forall>r\<in>?R. Card_order r"
+ using card_of_Card_order by auto
+ ultimately show ?thesis by(rule Card_order_omax)
+qed
+
+
+subsection {* Cardinals versus set operations on arbitrary sets *}
+
+lemma subset_ordLeq_strict:
+assumes "A \<le> B" and "|A| <o |B|"
+shows "A < B"
+proof-
+ {assume "\<not>(A < B)"
+ hence "A = B" using assms(1) by blast
+ hence False using assms(2) not_ordLess_ordIso card_of_refl by blast
+ }
+ thus ?thesis by blast
+qed
+
+corollary subset_ordLeq_diff:
+assumes "A \<le> B" and "|A| <o |B|"
+shows "B - A \<noteq> {}"
+using assms subset_ordLeq_strict by blast
+
+lemma card_of_empty4:
+"|{}::'b set| <o |A::'a set| = (A \<noteq> {})"
+proof(intro iffI notI)
+ assume *: "|{}::'b set| <o |A|" and "A = {}"
+ hence "|A| =o |{}::'b set|"
+ using card_of_ordIso unfolding bij_betw_def inj_on_def by blast
+ hence "|{}::'b set| =o |A|" using ordIso_symmetric by blast
+ with * show False using not_ordLess_ordIso[of "|{}::'b set|" "|A|"] by blast
+next
+ assume "A \<noteq> {}"
+ hence "(\<not> (\<exists>f. inj_on f A \<and> f ` A \<subseteq> {}))"
+ unfolding inj_on_def by blast
+ thus "| {} | <o | A |"
+ using card_of_ordLess by blast
+qed
+
+lemma card_of_empty5:
+"|A| <o |B| \<Longrightarrow> B \<noteq> {}"
+using card_of_empty not_ordLess_ordLeq by blast
+
+lemma Well_order_card_of_empty:
+"Well_order r \<Longrightarrow> |{}| \<le>o r" by simp
+
+lemma card_of_UNIV[simp]:
+"|A :: 'a set| \<le>o |UNIV :: 'a set|"
+using card_of_mono1[of A] by simp
+
+lemma card_of_UNIV2[simp]:
+"Card_order r \<Longrightarrow> (r :: 'a rel) \<le>o |UNIV :: 'a set|"
+using card_of_UNIV[of "Field r"] card_of_Field_ordIso
+ ordIso_symmetric ordIso_ordLeq_trans by blast
+
+lemma card_of_Pow_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|Pow A| \<le>o |Pow B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (image f) (Pow A) \<and> (image f) ` (Pow A) \<le> (Pow B)"
+ by (auto simp add: inj_on_image_Pow image_Pow_mono)
+ thus ?thesis using card_of_ordLeq[of "Pow A"] by metis
+qed
+
+lemma ordIso_Pow_mono[simp]:
+assumes "r \<le>o r'"
+shows "|Pow(Field r)| \<le>o |Pow(Field r')|"
+using assms card_of_mono2 card_of_Pow_mono by blast
+
+lemma card_of_Pow_cong[simp]:
+assumes "|A| =o |B|"
+shows "|Pow A| =o |Pow B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (image f) (Pow A) (Pow B)"
+ by (auto simp add: bij_betw_image_Pow)
+ thus ?thesis using card_of_ordIso[of "Pow A"] by auto
+qed
+
+lemma ordIso_Pow_cong[simp]:
+assumes "r =o r'"
+shows "|Pow(Field r)| =o |Pow(Field r')|"
+using assms card_of_cong card_of_Pow_cong by blast
+
+corollary Card_order_Plus_empty1:
+"Card_order r \<Longrightarrow> r =o |(Field r) <+> {}|"
+using card_of_Plus_empty1 card_of_Field_ordIso ordIso_equivalence by blast
+
+corollary Card_order_Plus_empty2:
+"Card_order r \<Longrightarrow> r =o |{} <+> (Field r)|"
+using card_of_Plus_empty2 card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma Card_order_Un1:
+shows "Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<union> B| "
+using card_of_Un1 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
+
+lemma card_of_Un2[simp]:
+shows "|A| \<le>o |B \<union> A|"
+using inj_on_id[of A] card_of_ordLeq[of A _] by fastforce
+
+lemma Card_order_Un2:
+shows "Card_order r \<Longrightarrow> |Field r| \<le>o |A \<union> (Field r)| "
+using card_of_Un2 card_of_Field_ordIso ordIso_symmetric ordIso_ordLeq_trans by auto
+
+lemma Un_Plus_bij_betw:
+assumes "A Int B = {}"
+shows "\<exists>f. bij_betw f (A \<union> B) (A <+> B)"
+proof-
+ let ?f = "\<lambda> c. if c \<in> A then Inl c else Inr c"
+ have "bij_betw ?f (A \<union> B) (A <+> B)"
+ using assms by(unfold bij_betw_def inj_on_def, auto)
+ thus ?thesis by blast
+qed
+
+lemma card_of_Un_Plus_ordIso:
+assumes "A Int B = {}"
+shows "|A \<union> B| =o |A <+> B|"
+using assms card_of_ordIso[of "A \<union> B"] Un_Plus_bij_betw[of A B] by auto
+
+lemma card_of_Un_Plus_ordIso1:
+"|A \<union> B| =o |A <+> (B - A)|"
+using card_of_Un_Plus_ordIso[of A "B - A"] by auto
+
+lemma card_of_Un_Plus_ordIso2:
+"|A \<union> B| =o |(A - B) <+> B|"
+using card_of_Un_Plus_ordIso[of "A - B" B] by auto
+
+lemma card_of_Times_singl1: "|A| =o |A \<times> {b}|"
+proof-
+ have "bij_betw fst (A \<times> {b}) A" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+corollary Card_order_Times_singl1:
+"Card_order r \<Longrightarrow> r =o |(Field r) \<times> {b}|"
+using card_of_Times_singl1[of _ b] card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma card_of_Times_singl2: "|A| =o |{b} \<times> A|"
+proof-
+ have "bij_betw snd ({b} \<times> A) A" unfolding bij_betw_def inj_on_def by force
+ thus ?thesis using card_of_ordIso ordIso_symmetric by blast
+qed
+
+corollary Card_order_Times_singl2:
+"Card_order r \<Longrightarrow> r =o |{a} \<times> (Field r)|"
+using card_of_Times_singl2[of _ a] card_of_Field_ordIso ordIso_equivalence by blast
+
+lemma card_of_Times_assoc: "|(A \<times> B) \<times> C| =o |A \<times> B \<times> C|"
+proof -
+ let ?f = "\<lambda>((a,b),c). (a,(b,c))"
+ have "A \<times> B \<times> C \<subseteq> ?f ` ((A \<times> B) \<times> C)"
+ proof
+ fix x assume "x \<in> A \<times> B \<times> C"
+ then obtain a b c where *: "a \<in> A" "b \<in> B" "c \<in> C" "x = (a, b, c)" by blast
+ let ?x = "((a, b), c)"
+ from * have "?x \<in> (A \<times> B) \<times> C" "x = ?f ?x" by auto
+ thus "x \<in> ?f ` ((A \<times> B) \<times> C)" by blast
+ qed
+ hence "bij_betw ?f ((A \<times> B) \<times> C) (A \<times> B \<times> C)"
+ unfolding bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+corollary Card_order_Times3:
+"Card_order r \<Longrightarrow> |Field r| \<le>o |(Field r) \<times> (Field r)|"
+using card_of_Times3 card_of_Field_ordIso
+ ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma card_of_Times_mono[simp]:
+assumes "|A| \<le>o |B|" and "|C| \<le>o |D|"
+shows "|A \<times> C| \<le>o |B \<times> D|"
+using assms card_of_Times_mono1[of A B C] card_of_Times_mono2[of C D B]
+ ordLeq_transitive[of "|A \<times> C|"] by blast
+
+corollary ordLeq_Times_mono:
+assumes "r \<le>o r'" and "p \<le>o p'"
+shows "|(Field r) \<times> (Field p)| \<le>o |(Field r') \<times> (Field p')|"
+using assms card_of_mono2[of r r'] card_of_mono2[of p p'] card_of_Times_mono by blast
+
+corollary ordIso_Times_cong1[simp]:
+assumes "r =o r'"
+shows "|(Field r) \<times> C| =o |(Field r') \<times> C|"
+using assms card_of_cong card_of_Times_cong1 by blast
+
+lemma card_of_Times_cong[simp]:
+assumes "|A| =o |B|" and "|C| =o |D|"
+shows "|A \<times> C| =o |B \<times> D|"
+using assms
+by (auto simp add: ordIso_iff_ordLeq)
+
+corollary ordIso_Times_cong:
+assumes "r =o r'" and "p =o p'"
+shows "|(Field r) \<times> (Field p)| =o |(Field r') \<times> (Field p')|"
+using assms card_of_cong[of r r'] card_of_cong[of p p'] card_of_Times_cong by blast
+
+lemma card_of_Sigma_mono2:
+assumes "inj_on f (I::'i set)" and "f ` I \<le> (J::'j set)"
+shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| \<le>o |SIGMA j : J. A j|"
+proof-
+ let ?LEFT = "SIGMA i : I. A (f i)"
+ let ?RIGHT = "SIGMA j : J. A j"
+ obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
+ have "inj_on u ?LEFT \<and> u `?LEFT \<le> ?RIGHT"
+ using assms unfolding u_def inj_on_def by auto
+ thus ?thesis using card_of_ordLeq by blast
+qed
+
+lemma card_of_Sigma_mono:
+assumes INJ: "inj_on f I" and IM: "f ` I \<le> J" and
+ LEQ: "\<forall>j \<in> J. |A j| \<le>o |B j|"
+shows "|SIGMA i : I. A (f i)| \<le>o |SIGMA j : J. B j|"
+proof-
+ have "\<forall>i \<in> I. |A(f i)| \<le>o |B(f i)|"
+ using IM LEQ by blast
+ hence "|SIGMA i : I. A (f i)| \<le>o |SIGMA i : I. B (f i)|"
+ using card_of_Sigma_mono1[of I] by metis
+ moreover have "|SIGMA i : I. B (f i)| \<le>o |SIGMA j : J. B j|"
+ using INJ IM card_of_Sigma_mono2 by blast
+ ultimately show ?thesis using ordLeq_transitive by blast
+qed
+
+
+lemma ordLeq_Sigma_mono1:
+assumes "\<forall>i \<in> I. p i \<le>o r i"
+shows "|SIGMA i : I. Field(p i)| \<le>o |SIGMA i : I. Field(r i)|"
+using assms by (auto simp add: card_of_Sigma_mono1)
+
+
+lemma ordLeq_Sigma_mono:
+assumes "inj_on f I" and "f ` I \<le> J" and
+ "\<forall>j \<in> J. p j \<le>o r j"
+shows "|SIGMA i : I. Field(p(f i))| \<le>o |SIGMA j : J. Field(r j)|"
+using assms card_of_mono2 card_of_Sigma_mono
+ [of f I J "\<lambda> i. Field(p i)" "\<lambda> j. Field(r j)"] by metis
+
+
+lemma card_of_Sigma_cong1:
+assumes "\<forall>i \<in> I. |A i| =o |B i|"
+shows "|SIGMA i : I. A i| =o |SIGMA i : I. B i|"
+using assms by (auto simp add: card_of_Sigma_mono1 ordIso_iff_ordLeq)
+
+
+lemma card_of_Sigma_cong2:
+assumes "bij_betw f (I::'i set) (J::'j set)"
+shows "|SIGMA i : I. (A::'j \<Rightarrow> 'a set) (f i)| =o |SIGMA j : J. A j|"
+proof-
+ let ?LEFT = "SIGMA i : I. A (f i)"
+ let ?RIGHT = "SIGMA j : J. A j"
+ obtain u where u_def: "u = (\<lambda>(i::'i,a::'a). (f i,a))" by blast
+ have "bij_betw u ?LEFT ?RIGHT"
+ using assms unfolding u_def bij_betw_def inj_on_def by auto
+ thus ?thesis using card_of_ordIso by blast
+qed
+
+lemma card_of_Sigma_cong:
+assumes BIJ: "bij_betw f I J" and
+ ISO: "\<forall>j \<in> J. |A j| =o |B j|"
+shows "|SIGMA i : I. A (f i)| =o |SIGMA j : J. B j|"
+proof-
+ have "\<forall>i \<in> I. |A(f i)| =o |B(f i)|"
+ using ISO BIJ unfolding bij_betw_def by blast
+ hence "|SIGMA i : I. A (f i)| =o |SIGMA i : I. B (f i)|"
+ using card_of_Sigma_cong1 by metis
+ moreover have "|SIGMA i : I. B (f i)| =o |SIGMA j : J. B j|"
+ using BIJ card_of_Sigma_cong2 by blast
+ ultimately show ?thesis using ordIso_transitive by blast
+qed
+
+lemma ordIso_Sigma_cong1:
+assumes "\<forall>i \<in> I. p i =o r i"
+shows "|SIGMA i : I. Field(p i)| =o |SIGMA i : I. Field(r i)|"
+using assms by (auto simp add: card_of_Sigma_cong1)
+
+lemma ordLeq_Sigma_cong:
+assumes "bij_betw f I J" and
+ "\<forall>j \<in> J. p j =o r j"
+shows "|SIGMA i : I. Field(p(f i))| =o |SIGMA j : J. Field(r j)|"
+using assms card_of_cong card_of_Sigma_cong
+ [of f I J "\<lambda> j. Field(p j)" "\<lambda> j. Field(r j)"] by blast
+
+corollary ordLeq_Sigma_Times:
+"\<forall>i \<in> I. p i \<le>o r \<Longrightarrow> |SIGMA i : I. Field (p i)| \<le>o |I \<times> (Field r)|"
+by (auto simp add: card_of_Sigma_Times)
+
+lemma card_of_UNION_Sigma2:
+assumes
+"!! i j. \<lbrakk>{i,j} <= I; i ~= j\<rbrakk> \<Longrightarrow> A i Int A j = {}"
+shows
+"|\<Union>i\<in>I. A i| =o |Sigma I A|"
+proof-
+ let ?L = "\<Union>i\<in>I. A i" let ?R = "Sigma I A"
+ have "|?L| <=o |?R|" using card_of_UNION_Sigma .
+ moreover have "|?R| <=o |?L|"
+ proof-
+ have "inj_on snd ?R"
+ unfolding inj_on_def using assms by auto
+ moreover have "snd ` ?R <= ?L" by auto
+ ultimately show ?thesis using card_of_ordLeq by blast
+ qed
+ ultimately show ?thesis by(simp add: ordIso_iff_ordLeq)
+qed
+
+corollary Plus_into_Times:
+assumes A2: "a1 \<noteq> a2 \<and> {a1,a2} \<le> A" and
+ B2: "b1 \<noteq> b2 \<and> {b1,b2} \<le> B"
+shows "\<exists>f. inj_on f (A <+> B) \<and> f ` (A <+> B) \<le> A \<times> B"
+using assms by (auto simp add: card_of_Plus_Times card_of_ordLeq)
+
+corollary Plus_into_Times_types:
+assumes A2: "(a1::'a) \<noteq> a2" and B2: "(b1::'b) \<noteq> b2"
+shows "\<exists>(f::'a + 'b \<Rightarrow> 'a * 'b). inj f"
+using assms Plus_into_Times[of a1 a2 UNIV b1 b2 UNIV]
+by auto
+
+corollary Times_same_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (A \<times> A) A"
+using assms by (auto simp add: card_of_ordIso)
+
+corollary Times_same_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)"
+shows "\<exists>(f::('a * 'a) => 'a). bij f"
+using assms Times_same_infinite_bij_betw[of "UNIV::'a set"]
+by auto
+
+corollary Times_infinite_bij_betw:
+assumes INF: "infinite A" and NE: "B \<noteq> {}" and INJ: "inj_on g B \<and> g ` B \<le> A"
+shows "(\<exists>f. bij_betw f (A \<times> B) A) \<and> (\<exists>h. bij_betw h (B \<times> A) A)"
+proof-
+ have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
+ thus ?thesis using INF NE
+ by (auto simp add: card_of_ordIso card_of_Times_infinite)
+qed
+
+corollary Times_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)" and
+ BIJ: "inj(g::'b \<Rightarrow> 'a)"
+shows "(\<exists>(f::('b * 'a) => 'a). bij f) \<and> (\<exists>(h::('a * 'b) => 'a). bij h)"
+using assms Times_infinite_bij_betw[of "UNIV::'a set" "UNIV::'b set" g]
+by auto
+
+lemma card_of_Times_ordLeq_infinite:
+"\<lbrakk>infinite C; |A| \<le>o |C|; |B| \<le>o |C|\<rbrakk>
+ \<Longrightarrow> |A <*> B| \<le>o |C|"
+by(simp add: card_of_Sigma_ordLeq_infinite)
+
+corollary Plus_infinite_bij_betw:
+assumes INF: "infinite A" and INJ: "inj_on g B \<and> g ` B \<le> A"
+shows "(\<exists>f. bij_betw f (A <+> B) A) \<and> (\<exists>h. bij_betw h (B <+> A) A)"
+proof-
+ have "|B| \<le>o |A|" using INJ card_of_ordLeq by blast
+ thus ?thesis using INF
+ by (auto simp add: card_of_ordIso)
+qed
+
+corollary Plus_infinite_bij_betw_types:
+assumes INF: "infinite(UNIV::'a set)" and
+ BIJ: "inj(g::'b \<Rightarrow> 'a)"
+shows "(\<exists>(f::('b + 'a) => 'a). bij f) \<and> (\<exists>(h::('a + 'b) => 'a). bij h)"
+using assms Plus_infinite_bij_betw[of "UNIV::'a set" g "UNIV::'b set"]
+by auto
+
+lemma card_of_Un_infinite_simps[simp]:
+"\<lbrakk>infinite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |A \<union> B| =o |A|"
+"\<lbrakk>infinite A; |B| \<le>o |A| \<rbrakk> \<Longrightarrow> |B \<union> A| =o |A|"
+using card_of_Un_infinite by auto
+
+corollary Card_order_Un_infinite:
+assumes INF: "infinite(Field r)" and CARD: "Card_order r" and
+ LEQ: "p \<le>o r"
+shows "| (Field r) \<union> (Field p) | =o r \<and> | (Field p) \<union> (Field r) | =o r"
+proof-
+ have "| Field r \<union> Field p | =o | Field r | \<and>
+ | Field p \<union> Field r | =o | Field r |"
+ using assms by (auto simp add: card_of_Un_infinite)
+ thus ?thesis
+ using assms card_of_Field_ordIso[of r]
+ ordIso_transitive[of "|Field r \<union> Field p|"]
+ ordIso_transitive[of _ "|Field r|"] by blast
+qed
+
+corollary subset_ordLeq_diff_infinite:
+assumes INF: "infinite B" and SUB: "A \<le> B" and LESS: "|A| <o |B|"
+shows "infinite (B - A)"
+using assms card_of_Un_diff_infinite card_of_ordIso_finite by blast
+
+lemma card_of_Times_ordLess_infinite[simp]:
+assumes INF: "infinite C" and
+ LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A \<times> B| <o |C|"
+proof(cases "A = {} \<or> B = {}")
+ assume Case1: "A = {} \<or> B = {}"
+ hence "A \<times> B = {}" by blast
+ moreover have "C \<noteq> {}" using
+ LESS1 card_of_empty5 by blast
+ ultimately show ?thesis by(auto simp add: card_of_empty4)
+next
+ assume Case2: "\<not>(A = {} \<or> B = {})"
+ {assume *: "|C| \<le>o |A \<times> B|"
+ hence "infinite (A \<times> B)" using INF card_of_ordLeq_finite by blast
+ hence 1: "infinite A \<or> infinite B" using finite_cartesian_product by blast
+ {assume Case21: "|A| \<le>o |B|"
+ hence "infinite B" using 1 card_of_ordLeq_finite by blast
+ hence "|A \<times> B| =o |B|" using Case2 Case21
+ by (auto simp add: card_of_Times_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 \<times> B| =o |A|" using Case2 Case22
+ by (auto simp add: card_of_Times_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 \<times> B|" "|C|"]
+ card_of_Well_order[of "A \<times> B"] card_of_Well_order[of "C"] by auto
+qed
+
+lemma card_of_Times_ordLess_infinite_Field[simp]:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A \<times> 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_Times_ordLess_infinite by blast
+ thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+lemma card_of_Un_ordLess_infinite[simp]:
+assumes INF: "infinite C" and
+ LESS1: "|A| <o |C|" and LESS2: "|B| <o |C|"
+shows "|A \<union> B| <o |C|"
+using assms card_of_Plus_ordLess_infinite card_of_Un_Plus_ordLeq
+ ordLeq_ordLess_trans by blast
+
+lemma card_of_Un_ordLess_infinite_Field[simp]:
+assumes INF: "infinite (Field r)" and r: "Card_order r" and
+ LESS1: "|A| <o r" and LESS2: "|B| <o r"
+shows "|A Un 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 Un B| <o |?C|" using INF
+ card_of_Un_ordLess_infinite by blast
+ thus ?thesis using 1 ordLess_ordIso_trans by blast
+qed
+
+lemma card_of_Un_singl_ordLess_infinite1:
+assumes "infinite B" and "|A| <o |B|"
+shows "|{a} Un A| <o |B|"
+proof-
+ have "|{a}| <o |B|" using assms by auto
+ thus ?thesis using assms card_of_Un_ordLess_infinite[of B] by fastforce
+qed
+
+lemma card_of_Un_singl_ordLess_infinite:
+assumes "infinite B"
+shows "( |A| <o |B| ) = ( |{a} Un A| <o |B| )"
+using assms card_of_Un_singl_ordLess_infinite1[of B A]
+proof(auto)
+ assume "|insert a A| <o |B|"
+ moreover have "|A| <=o |insert a A|" using card_of_mono1[of A] by blast
+ ultimately show "|A| <o |B|" using ordLeq_ordLess_trans by blast
+qed
+
+
+subsection {* Cardinals versus lists *}
+
+lemma Card_order_lists: "Card_order r \<Longrightarrow> r \<le>o |lists(Field r) |"
+using card_of_lists card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma Union_set_lists:
+"Union(set ` (lists A)) = A"
+unfolding lists_def2 proof(auto)
+ fix a assume "a \<in> A"
+ hence "set [a] \<le> A \<and> a \<in> set [a]" by auto
+ thus "\<exists>l. set l \<le> A \<and> a \<in> set l" by blast
+qed
+
+lemma inj_on_map_lists:
+assumes "inj_on f A"
+shows "inj_on (map f) (lists A)"
+using assms Union_set_lists[of A] inj_on_mapI[of f "lists A"] by auto
+
+lemma map_lists_mono:
+assumes "f ` A \<le> B"
+shows "(map f) ` (lists A) \<le> lists B"
+using assms unfolding lists_def2 by (auto, blast) (* lethal combination of methods :) *)
+
+lemma map_lists_surjective:
+assumes "f ` A = B"
+shows "(map f) ` (lists A) = lists B"
+using assms unfolding lists_def2
+proof (auto, blast)
+ fix l' assume *: "set l' \<le> f ` A"
+ have "set l' \<le> f ` A \<longrightarrow> l' \<in> map f ` {l. set l \<le> A}"
+ proof(induct l', auto)
+ fix l a
+ assume "a \<in> A" and "set l \<le> A" and
+ IH: "f ` (set l) \<le> f ` A"
+ hence "set (a # l) \<le> A" by auto
+ hence "map f (a # l) \<in> map f ` {l. set l \<le> A}" by blast
+ thus "f a # map f l \<in> map f ` {l. set l \<le> A}" by auto
+ qed
+ thus "l' \<in> map f ` {l. set l \<le> A}" using * by auto
+qed
+
+lemma bij_betw_map_lists:
+assumes "bij_betw f A B"
+shows "bij_betw (map f) (lists A) (lists B)"
+using assms unfolding bij_betw_def
+by(auto simp add: inj_on_map_lists map_lists_surjective)
+
+lemma card_of_lists_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|lists A| \<le>o |lists B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (map f) (lists A) \<and> (map f) ` (lists A) \<le> (lists B)"
+ by (auto simp add: inj_on_map_lists map_lists_mono)
+ thus ?thesis using card_of_ordLeq[of "lists A"] by metis
+qed
+
+lemma ordIso_lists_mono:
+assumes "r \<le>o r'"
+shows "|lists(Field r)| \<le>o |lists(Field r')|"
+using assms card_of_mono2 card_of_lists_mono by blast
+
+lemma card_of_lists_cong[simp]:
+assumes "|A| =o |B|"
+shows "|lists A| =o |lists B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (map f) (lists A) (lists B)"
+ by (auto simp add: bij_betw_map_lists)
+ thus ?thesis using card_of_ordIso[of "lists A"] by auto
+qed
+
+lemma ordIso_lists_cong:
+assumes "r =o r'"
+shows "|lists(Field r)| =o |lists(Field r')|"
+using assms card_of_cong card_of_lists_cong by blast
+
+corollary lists_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (lists A) A"
+using assms card_of_lists_infinite card_of_ordIso by blast
+
+corollary lists_infinite_bij_betw_types:
+assumes "infinite(UNIV :: 'a set)"
+shows "\<exists>(f::'a list \<Rightarrow> 'a). bij f"
+using assms assms lists_infinite_bij_betw[of "UNIV::'a set"]
+using lists_UNIV by auto
+
+
+subsection {* Cardinals versus the set-of-finite-sets operator *}
+
+definition Fpow :: "'a set \<Rightarrow> 'a set set"
+where "Fpow A \<equiv> {X. X \<le> A \<and> finite X}"
+
+lemma Fpow_mono: "A \<le> B \<Longrightarrow> Fpow A \<le> Fpow B"
+unfolding Fpow_def by auto
+
+lemma empty_in_Fpow: "{} \<in> Fpow A"
+unfolding Fpow_def by auto
+
+lemma Fpow_not_empty: "Fpow A \<noteq> {}"
+using empty_in_Fpow by blast
+
+lemma Fpow_subset_Pow: "Fpow A \<le> Pow A"
+unfolding Fpow_def by auto
+
+lemma card_of_Fpow[simp]: "|A| \<le>o |Fpow A|"
+proof-
+ let ?h = "\<lambda> a. {a}"
+ have "inj_on ?h A \<and> ?h ` A \<le> Fpow A"
+ unfolding inj_on_def Fpow_def by auto
+ thus ?thesis using card_of_ordLeq by metis
+qed
+
+lemma Card_order_Fpow: "Card_order r \<Longrightarrow> r \<le>o |Fpow(Field r) |"
+using card_of_Fpow card_of_Field_ordIso ordIso_ordLeq_trans ordIso_symmetric by blast
+
+lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}"
+unfolding Fpow_def Pow_def by blast
+
+lemma inj_on_image_Fpow:
+assumes "inj_on f A"
+shows "inj_on (image f) (Fpow A)"
+using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"]
+ inj_on_image_Pow by blast
+
+lemma image_Fpow_mono:
+assumes "f ` A \<le> B"
+shows "(image f) ` (Fpow A) \<le> Fpow B"
+using assms by(unfold Fpow_def, auto)
+
+lemma image_Fpow_surjective:
+assumes "f ` A = B"
+shows "(image f) ` (Fpow A) = Fpow B"
+using assms proof(unfold Fpow_def, auto)
+ fix Y assume *: "Y \<le> f ` A" and **: "finite Y"
+ hence "\<forall>b \<in> Y. \<exists>a. a \<in> A \<and> f a = b" by auto
+ with bchoice[of Y "\<lambda>b a. a \<in> A \<and> f a = b"]
+ obtain g where 1: "\<forall>b \<in> Y. g b \<in> A \<and> f(g b) = b" by blast
+ obtain X where X_def: "X = g ` Y" by blast
+ have "f ` X = Y \<and> X \<le> A \<and> finite X"
+ by(unfold X_def, force simp add: ** 1)
+ thus "Y \<in> (image f) ` {X. X \<le> A \<and> finite X}" by auto
+qed
+
+lemma bij_betw_image_Fpow:
+assumes "bij_betw f A B"
+shows "bij_betw (image f) (Fpow A) (Fpow B)"
+using assms unfolding bij_betw_def
+by (auto simp add: inj_on_image_Fpow image_Fpow_surjective)
+
+lemma card_of_Fpow_mono[simp]:
+assumes "|A| \<le>o |B|"
+shows "|Fpow A| \<le>o |Fpow B|"
+proof-
+ obtain f where "inj_on f A \<and> f ` A \<le> B"
+ using assms card_of_ordLeq[of A B] by auto
+ hence "inj_on (image f) (Fpow A) \<and> (image f) ` (Fpow A) \<le> (Fpow B)"
+ by (auto simp add: inj_on_image_Fpow image_Fpow_mono)
+ thus ?thesis using card_of_ordLeq[of "Fpow A"] by auto
+qed
+
+lemma ordIso_Fpow_mono:
+assumes "r \<le>o r'"
+shows "|Fpow(Field r)| \<le>o |Fpow(Field r')|"
+using assms card_of_mono2 card_of_Fpow_mono by blast
+
+lemma card_of_Fpow_cong[simp]:
+assumes "|A| =o |B|"
+shows "|Fpow A| =o |Fpow B|"
+proof-
+ obtain f where "bij_betw f A B"
+ using assms card_of_ordIso[of A B] by auto
+ hence "bij_betw (image f) (Fpow A) (Fpow B)"
+ by (auto simp add: bij_betw_image_Fpow)
+ thus ?thesis using card_of_ordIso[of "Fpow A"] by auto
+qed
+
+lemma ordIso_Fpow_cong:
+assumes "r =o r'"
+shows "|Fpow(Field r)| =o |Fpow(Field r')|"
+using assms card_of_cong card_of_Fpow_cong by blast
+
+lemma card_of_Fpow_lists: "|Fpow A| \<le>o |lists A|"
+proof-
+ have "set ` (lists A) = Fpow A"
+ unfolding lists_def2 Fpow_def using finite_list finite_set by blast
+ thus ?thesis using card_of_ordLeq2[of "Fpow A"] Fpow_not_empty[of A] by blast
+qed
+
+lemma card_of_Fpow_infinite[simp]:
+assumes "infinite A"
+shows "|Fpow A| =o |A|"
+using assms card_of_Fpow_lists card_of_lists_infinite card_of_Fpow
+ ordLeq_ordIso_trans ordIso_iff_ordLeq by blast
+
+corollary Fpow_infinite_bij_betw:
+assumes "infinite A"
+shows "\<exists>f. bij_betw f (Fpow A) A"
+using assms card_of_Fpow_infinite card_of_ordIso by blast
+
+
+subsection {* The cardinal $\omega$ and the finite cardinals *}
+
+subsubsection {* First as well-orders *}
+
+lemma Field_natLess: "Field natLess = (UNIV::nat set)"
+by(unfold Field_def, auto)
+
+lemma natLeq_ofilter_less: "ofilter natLeq {0 ..< n}"
+by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq, unfold rel.under_def, auto)
+
+lemma natLeq_ofilter_leq: "ofilter natLeq {0 .. n}"
+by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def,
+ simp add: Field_natLeq, unfold rel.under_def, auto)
+
+lemma natLeq_ofilter_iff:
+"ofilter natLeq A = (A = UNIV \<or> (\<exists>n. A = {0 ..< n}))"
+proof(rule iffI)
+ assume "ofilter natLeq A"
+ hence "\<forall>m n. n \<in> A \<and> m \<le> n \<longrightarrow> m \<in> A"
+ by(auto simp add: natLeq_wo_rel wo_rel.ofilter_def rel.under_def)
+ thus "A = UNIV \<or> (\<exists>n. A = {0 ..< n})" using closed_nat_set_iff by blast
+next
+ assume "A = UNIV \<or> (\<exists>n. A = {0 ..< n})"
+ thus "ofilter natLeq A"
+ by(auto simp add: natLeq_ofilter_less natLeq_UNIV_ofilter)
+qed
+
+lemma natLeq_under_leq: "under natLeq n = {0 .. n}"
+unfolding rel.under_def by auto
+
+corollary natLeq_on_ofilter:
+"ofilter(natLeq_on n) {0 ..< n}"
+by (auto simp add: natLeq_on_ofilter_less_eq)
+
+lemma natLeq_on_ofilter_less:
+"n < m \<Longrightarrow> 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_ordLess_natLeq: "natLeq_on n <o natLeq"
+using Field_natLeq Field_natLeq_on[of n] nat_infinite
+ finite_ordLess_infinite[of "natLeq_on n" natLeq]
+ natLeq_Well_order natLeq_on_Well_order[of n] by auto
+
+lemma natLeq_on_injective:
+"natLeq_on m = natLeq_on n \<Longrightarrow> m = n"
+using Field_natLeq_on[of m] Field_natLeq_on[of n]
+ atLeastLessThan_injective[of m n] by auto
+
+lemma natLeq_on_injective_ordIso:
+"(natLeq_on m =o natLeq_on n) = (m = n)"
+proof(auto simp add: natLeq_on_Well_order ordIso_reflexive)
+ assume "natLeq_on m =o natLeq_on n"
+ then obtain f where "bij_betw f {0..<m} {0..<n}"
+ using Field_natLeq_on assms unfolding ordIso_def iso_def[abs_def] by auto
+ thus "m = n" using atLeastLessThan_injective2 by blast
+qed
+
+
+subsubsection {* Then as cardinals *}
+
+lemma ordIso_natLeq_infinite1:
+"|A| =o natLeq \<Longrightarrow> infinite A"
+using ordIso_symmetric ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
+
+lemma ordIso_natLeq_infinite2:
+"natLeq =o |A| \<Longrightarrow> infinite A"
+using ordIso_imp_ordLeq infinite_iff_natLeq_ordLeq by blast
+
+lemma ordLeq_natLeq_on_imp_finite:
+assumes "|A| \<le>o natLeq_on n"
+shows "finite A"
+proof-
+ have "|A| \<le>o |{0 ..< n}|"
+ using assms card_of_less ordIso_symmetric ordLeq_ordIso_trans by blast
+ thus ?thesis by (auto simp add: card_of_ordLeq_finite)
+qed
+
+
+subsubsection {* "Backwards compatibility" with the numeric cardinal operator for finite sets *}
+
+lemma finite_card_of_iff_card:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| =o |B| ) = (card A = card B)"
+using assms card_of_ordIso[of A B] bij_betw_iff_card[of A B] by blast
+
+lemma finite_card_of_iff_card3:
+assumes FIN: "finite A" and FIN': "finite B"
+shows "( |A| <o |B| ) = (card A < card B)"
+proof-
+ have "( |A| <o |B| ) = (~ ( |B| \<le>o |A| ))" by simp
+ also have "... = (~ (card B \<le> card A))"
+ using assms by(simp add: finite_card_of_iff_card2)
+ also have "... = (card A < card B)" by auto
+ finally show ?thesis .
+qed
+
+lemma card_Field_natLeq_on:
+"card(Field(natLeq_on n)) = n"
+using Field_natLeq_on card_atLeastLessThan by auto
+
+
+subsection {* The successor of a cardinal *}
+
+lemma embed_implies_ordIso_Restr:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and EMB: "embed r' r f"
+shows "r' =o Restr r (f ` (Field r'))"
+using assms embed_implies_iso_Restr Well_order_Restr unfolding ordIso_def by blast
+
+lemma cardSuc_Well_order[simp]:
+"Card_order r \<Longrightarrow> Well_order(cardSuc r)"
+using cardSuc_Card_order unfolding card_order_on_def by blast
+
+lemma Field_cardSuc_not_empty:
+assumes "Card_order r"
+shows "Field (cardSuc r) \<noteq> {}"
+proof
+ assume "Field(cardSuc r) = {}"
+ hence "|Field(cardSuc r)| \<le>o r" using assms Card_order_empty[of r] by auto
+ hence "cardSuc r \<le>o r" using assms card_of_Field_ordIso
+ cardSuc_Card_order ordIso_symmetric ordIso_ordLeq_trans by blast
+ thus False using cardSuc_greater not_ordLess_ordLeq assms by blast
+qed
+
+lemma cardSuc_mono_ordLess[simp]:
+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 auto
+ thus ?thesis
+ using not_ordLeq_iff_ordLess not_ordLeq_iff_ordLess[of r r']
+ using cardSuc_mono_ordLeq[of r' r] assms by blast
+qed
+
+lemma card_of_Plus_ordLeq_infinite[simp]:
+assumes C: "infinite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
+shows "|A <+> B| \<le>o |C|"
+proof-
+ let ?r = "cardSuc |C|"
+ have "Card_order ?r \<and> infinite (Field ?r)" using assms by simp
+ moreover have "|A| <o ?r" and "|B| <o ?r" using A B by auto
+ ultimately have "|A <+> B| <o ?r"
+ using card_of_Plus_ordLess_infinite_Field by blast
+ thus ?thesis using C by simp
+qed
+
+lemma card_of_Un_ordLeq_infinite[simp]:
+assumes C: "infinite C" and A: "|A| \<le>o |C|" and B: "|B| \<le>o |C|"
+shows "|A Un B| \<le>o |C|"
+using assms card_of_Plus_ordLeq_infinite card_of_Un_Plus_ordLeq
+ordLeq_transitive by metis
+
+
+subsection {* Others *}
+
+lemma under_mono[simp]:
+assumes "Well_order r" and "(i,j) \<in> r"
+shows "under r i \<subseteq> under r j"
+using assms unfolding rel.under_def order_on_defs
+trans_def by blast
+
+lemma underS_under:
+assumes "i \<in> Field r"
+shows "underS r i = under r i - {i}"
+using assms unfolding rel.underS_def rel.under_def by auto
+
+lemma relChain_under:
+assumes "Well_order r"
+shows "relChain r (\<lambda> i. under r i)"
+using assms unfolding relChain_def by auto
+
+lemma infinite_card_of_diff_singl:
+assumes "infinite A"
+shows "|A - {a}| =o |A|"
+by (metis assms card_of_infinite_diff_finitte finite.emptyI finite_insert)
+
+lemma card_of_vimage:
+assumes "B \<subseteq> range f"
+shows "|B| \<le>o |f -` B|"
+apply(rule surj_imp_ordLeq[of _ f])
+using assms by (metis Int_absorb2 image_vimage_eq order_refl)
+
+lemma surj_card_of_vimage:
+assumes "surj f"
+shows "|B| \<le>o |f -` B|"
+by (metis assms card_of_vimage subset_UNIV)
+
+(* bounded powerset *)
+definition Bpow where
+"Bpow r A \<equiv> {X . X \<subseteq> A \<and> |X| \<le>o r}"
+
+lemma Bpow_empty[simp]:
+assumes "Card_order r"
+shows "Bpow r {} = {{}}"
+using assms unfolding Bpow_def by auto
+
+lemma singl_in_Bpow:
+assumes rc: "Card_order r"
+and r: "Field r \<noteq> {}" and a: "a \<in> A"
+shows "{a} \<in> Bpow r A"
+proof-
+ have "|{a}| \<le>o r" using r rc by auto
+ thus ?thesis unfolding Bpow_def using a by auto
+qed
+
+lemma ordLeq_card_Bpow:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
+shows "|A| \<le>o |Bpow r A|"
+proof-
+ have "inj_on (\<lambda> a. {a}) A" unfolding inj_on_def by auto
+ moreover have "(\<lambda> a. {a}) ` A \<subseteq> Bpow r A"
+ using singl_in_Bpow[OF assms] by auto
+ ultimately show ?thesis unfolding card_of_ordLeq[symmetric] by blast
+qed
+
+lemma infinite_Bpow:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}"
+and A: "infinite A"
+shows "infinite (Bpow r A)"
+using ordLeq_card_Bpow[OF rc r]
+by (metis A card_of_ordLeq_infinite)
+
+lemma Bpow_ordLeq_Func_Field:
+assumes rc: "Card_order r" and r: "Field r \<noteq> {}" and A: "infinite A"
+shows "|Bpow r A| \<le>o |Func (Field r) A|"
+proof-
+ let ?F = "\<lambda> f. {x | x a. f a = Some x}"
+ {fix X assume "X \<in> Bpow r A - {{}}"
+ hence XA: "X \<subseteq> A" and "|X| \<le>o r"
+ and X: "X \<noteq> {}" unfolding Bpow_def by auto
+ hence "|X| \<le>o |Field r|" by (metis Field_card_of card_of_mono2)
+ then obtain F where 1: "X = F ` (Field r)"
+ using card_of_ordLeq2[OF X] by metis
+ def f \<equiv> "\<lambda> i. if i \<in> Field r then Some (F i) else None"
+ have "\<exists> f \<in> Func (Field r) A. X = ?F f"
+ apply (intro bexI[of _ f]) using 1 XA unfolding Func_def f_def by auto
+ }
+ hence "Bpow r A - {{}} \<subseteq> ?F ` (Func (Field r) A)" by auto
+ hence "|Bpow r A - {{}}| \<le>o |Func (Field r) A|"
+ by (rule surj_imp_ordLeq)
+ moreover
+ {have 2: "infinite (Bpow r A)" using infinite_Bpow[OF rc r A] .
+ have "|Bpow r A| =o |Bpow r A - {{}}|"
+ using card_of_infinite_diff_finitte
+ by (metis Pow_empty 2 finite_Pow_iff infinite_imp_nonempty ordIso_symmetric)
+ }
+ ultimately show ?thesis by (metis ordIso_ordLeq_trans)
+qed
+
+lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
+
+lemma empty_in_Func[simp]:
+"B \<noteq> {} \<Longrightarrow> empty \<in> Func {} B"
+unfolding Func_def by auto
+
+lemma Func_mono[simp]:
+assumes "B1 \<subseteq> B2"
+shows "Func A B1 \<subseteq> Func A B2"
+using assms unfolding Func_def by force
+
+lemma Pfunc_mono[simp]:
+assumes "A1 \<subseteq> A2" and "B1 \<subseteq> B2"
+shows "Pfunc A B1 \<subseteq> Pfunc A B2"
+using assms in_mono unfolding Pfunc_def apply safe
+apply(case_tac "x a", auto)
+by (metis in_mono option.simps(5))
+
+lemma card_of_Func_UNIV_UNIV:
+"|Func (UNIV::'a set) (UNIV::'b set)| =o |UNIV::('a \<Rightarrow> 'b) set|"
+using card_of_Func_UNIV[of "UNIV::'b set"] by auto
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation_Base.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,2579 @@
+(* 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 auto
+ 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 fastforce
+ 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}"
+ using Field_natLeq_on[of m] Field_natLeq_on[of n]
+ unfolding ordLeq_def using embed_inj_on[of "natLeq_on m" "natLeq_on n"]
+ embed_Field[of "natLeq_on m" "natLeq_on n"] using natLeq_on_Well_order[of m] by fastforce
+ 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 *}
+
+(* FIXME: finitte ~> finite? *)
+lemma card_of_infinite_diff_finitte:
+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 \<equiv>
+ {f. (\<forall> a. f a \<noteq> None \<longleftrightarrow> a \<in> A) \<and> (\<forall> a \<in> A. case f a of Some b \<Rightarrow> b \<in> B |None \<Rightarrow> True)}"
+
+lemma Func_empty:
+"Func {} B = {empty}"
+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 = Some b"
+using assms unfolding Func_def by (cases "g a") force+
+
+definition curr where
+"curr A f \<equiv> \<lambda> a. if a \<in> A then Some (\<lambda> b. f (a,b)) else None"
+
+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
+ apply(cases "f1 (a,b)") apply(cases "f2 (a,b)", fastforce, fastforce)
+ apply(cases "f2 (a,b)") 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
+ apply(elim allE[of _ a]) apply simp unfolding fun_eq_iff by auto
+ 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. case g (fst ab) of None \<Rightarrow> None | Some g1 \<Rightarrow> g1 (snd ab)"
+ 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 = None" 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 = Some g1"
+ using assms using Func_elim[OF assms True] by blast
+ thus ?thesis using True unfolding curr_def by auto
+ qed
+ qed
+ show "?f \<in> Func (A <*> B) C"
+ unfolding Func_def mem_Collect_eq proof(intro conjI allI ballI)
+ fix ab show "?f ab \<noteq> None \<longleftrightarrow> ab \<in> A \<times> B"
+ proof(cases "g (fst ab)")
+ case None
+ hence "fst ab \<notin> A" using assms unfolding Func_def by force
+ thus ?thesis using None by auto
+ next
+ case (Some g1)
+ hence fst: "fst ab \<in> A" and g1: "g1 \<in> Func B C"
+ using assms unfolding Func_def[of A] by force+
+ hence "?f ab \<noteq> None \<longleftrightarrow> g1 (snd ab) \<noteq> None" using Some by auto
+ also have "... \<longleftrightarrow> snd ab \<in> B" using g1 unfolding Func_def by auto
+ also have "... \<longleftrightarrow> ab \<in> A \<times> B" using fst by (cases ab, auto)
+ finally show ?thesis .
+ qed
+ next
+ fix ab assume ab: "ab \<in> A \<times> B"
+ hence "fst ab \<in> A" and "snd ab \<in> B" by(cases ab, auto)
+ then obtain g1 where "g1 \<in> Func B C" and "g (fst ab) = Some g1"
+ using assms using Func_elim[OF assms] by blast
+ thus "case ?f ab of Some c \<Rightarrow> c \<in> C |None \<Rightarrow> True"
+ unfolding Func_def by auto
+ qed
+qed
+
+(* FIXME: betwe ~> betw? *)
+lemma bij_betwe_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_betwe_curr by blast
+
+definition Func_map where
+"Func_map B2 f1 f2 g b2 \<equiv>
+ if b2 \<in> B2 then case g (f2 b2) of None \<Rightarrow> None | Some a1 \<Rightarrow> Some (f1 a1)
+ else None"
+
+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"
+unfolding Func_def mem_Collect_eq proof(intro conjI allI ballI)
+ fix b2 show "Func_map B2 f1 f2 g b2 \<noteq> None \<longleftrightarrow> b2 \<in> B2"
+ proof(cases "b2 \<in> B2")
+ case True
+ hence "f2 b2 \<in> A2" using f2 by auto
+ then obtain a1 where "g (f2 b2) = Some a1" and "a1 \<in> A1"
+ using g unfolding Func_def by(cases "g (f2 b2)", fastforce+)
+ thus ?thesis unfolding Func_map_def using True by auto
+ qed(unfold Func_map_def, auto)
+next
+ fix b2 assume b2: "b2 \<in> B2"
+ hence "f2 b2 \<in> A2" using f2 by auto
+ then obtain a1 where "g (f2 b2) = Some a1" and "a1 \<in> A1"
+ using g unfolding Func_def by(cases "g (f2 b2)", fastforce+)
+ thus "case Func_map B2 f1 f2 g b2 of None \<Rightarrow> True | Some b1 \<Rightarrow> b1 \<in> B1"
+ unfolding Func_map_def using b2 f1 by auto
+qed
+
+lemma Func_map_empty:
+"Func_map B2 f1 f2 empty = empty"
+unfolding Func_map_def[abs_def] by (rule ext, 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 Some b else None) \<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", 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_empty)
+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")
+ case (Some b1)
+ hence b1: "b1 \<in> f1 ` A1" using True h unfolding B1 Func_def by auto
+ show ?thesis
+ using Some True A2 f_inv_into_f[OF b1]
+ unfolding g_def Func_map_def j1_def j2[OF True] by auto
+ qed(insert A2 True j2[OF True], unfold g_def Func_map_def, auto)
+ 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(force, force)
+ 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
+
+(* partial-function space: *)
+definition Pfunc where
+"Pfunc A B \<equiv>
+ {f. (\<forall>a. f a \<noteq> None \<longrightarrow> a \<in> A) \<and>
+ (\<forall>a. case f a of None \<Rightarrow> True | Some b \<Rightarrow> b \<in> B)}"
+
+lemma Func_Pfunc:
+"Func A B \<subseteq> Pfunc A B"
+unfolding Func_def Pfunc_def by auto
+
+lemma Pfunc_Func:
+"Pfunc A B = (\<Union> A' \<in> Pow A. Func A' B)"
+proof safe
+ fix f assume f: "f \<in> Pfunc A B"
+ show "f \<in> (\<Union>A'\<in>Pow A. Func A' B)"
+ proof (intro UN_I)
+ let ?A' = "{a. f a \<noteq> None}"
+ show "?A' \<in> Pow A" using f unfolding Pow_def Pfunc_def by auto
+ show "f \<in> Func ?A' B" using f unfolding Func_def Pfunc_def by auto
+ qed
+next
+ fix f A' assume "f \<in> Func A' B" and "A' \<subseteq> A"
+ thus "f \<in> Pfunc A B" unfolding Func_def Pfunc_def by 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 Some True else Some False)
+ else None"
+ 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: "A1 \<in> Pow A" and A2: "A2 \<in> Pow A" and eq: "F A1 = F A2"
+ show "A1 = A2"
+ proof-
+ {fix a
+ have "a \<in> A1 \<longleftrightarrow> F A1 a = Some True" using A1 unfolding F_def Pow_def by auto
+ also have "... \<longleftrightarrow> F A2 a = Some True" unfolding eq ..
+ also have "... \<longleftrightarrow> a \<in> A2" using A2 unfolding F_def Pow_def by auto
+ finally have "a \<in> A1 \<longleftrightarrow> a \<in> A2" .
+ }
+ thus ?thesis by auto
+ qed
+ 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 = Some True}"
+ show "f = F ?A1" unfolding F_def apply(rule ext)
+ using f unfolding Func_def mem_Collect_eq by (auto,force)
+ 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 option) a. if a \<in> A2 then (if a \<in> A1 then f1 a else Some bb)
+ else None"
+ 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 card_of_Pfunc_Pow_Func:
+assumes "B \<noteq> {}"
+shows "|Pfunc A B| \<le>o |Pow A <*> Func A B|"
+proof-
+ have "|Pfunc A B| =o |\<Union> A' \<in> Pow A. Func A' B|" (is "_ =o ?K")
+ unfolding Pfunc_Func by(rule card_of_refl)
+ also have "?K \<le>o |Sigma (Pow A) (\<lambda> A'. Func A' B)|" using card_of_UNION_Sigma .
+ also have "|Sigma (Pow A) (\<lambda> A'. Func A' B)| \<le>o |Pow A <*> Func A B|"
+ apply(rule card_of_Sigma_mono1) using card_of_Func_mono[OF _ assms] by auto
+ finally show ?thesis .
+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 Some b1 else Some b2)
+ else None"
+ 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 apply auto
+ by (metis option.simps(3))
+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)
+
+(* alternative function space avoiding the option type, with undefined instead of None *)
+definition Ffunc where
+"Ffunc A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
+
+lemma card_of_Func_Ffunc:
+"|Ffunc A B| =o |Func A B|"
+unfolding card_of_ordIso[symmetric] proof
+ let ?F = "\<lambda> f a. if a \<in> A then Some (f a) else None"
+ show "bij_betw ?F (Ffunc A B) (Func A B)"
+ unfolding bij_betw_def unfolding inj_on_def proof(intro conjI ballI impI)
+ fix f g assume f: "f \<in> Ffunc A B" and g: "g \<in> Ffunc A B" and eq: "?F f = ?F g"
+ show "f = g"
+ proof(rule ext)
+ fix a
+ show "f a = g a"
+ proof(cases "a \<in> A")
+ case True
+ have "Some (f a) = ?F f a" using True by auto
+ also have "... = ?F g a" using eq unfolding fun_eq_iff by(rule allE)
+ also have "... = Some (g a)" using True by auto
+ finally have "Some (f a) = Some (g a)" .
+ thus ?thesis by simp
+ qed(insert f g, unfold Ffunc_def, auto)
+ qed
+ next
+ show "?F ` Ffunc A B = Func A B"
+ proof safe
+ fix f assume f: "f \<in> Func A B"
+ def g \<equiv> "\<lambda> a. case f a of Some b \<Rightarrow> b | None \<Rightarrow> undefined"
+ have "g \<in> Ffunc A B"
+ using f unfolding g_def Func_def Ffunc_def by force+
+ moreover have "f = ?F g"
+ proof(rule ext)
+ fix a show "f a = ?F g a"
+ using f unfolding Func_def g_def by (cases "a \<in> A") force+
+ qed
+ ultimately show "f \<in> ?F ` (Ffunc A B)" by blast
+ qed(unfold Ffunc_def Func_def, auto)
+ qed
+qed
+
+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). Some ((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 option" assume h: "h \<in> Func UNIV B"
+ hence "\<forall> a. \<exists> b. h a = Some b" unfolding Func_def by auto
+ then obtain f where f: "\<forall> a. h a = Some (f a)" by metis
+ hence "range f \<subseteq> B" using h unfolding Func_def by auto
+ thus "h \<in> (\<lambda>f a. Some (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/Constructions_on_Wellorders.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,795 @@
+(* Title: HOL/Cardinals/Constructions_on_Wellorders.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Constructions on wellorders.
+*)
+
+header {* Constructions on Wellorders *}
+
+theory Constructions_on_Wellorders
+imports Constructions_on_Wellorders_Base Wellorder_Embedding
+begin
+
+declare
+ ordLeq_Well_order_simp[simp]
+ ordLess_Well_order_simp[simp]
+ ordIso_Well_order_simp[simp]
+ not_ordLeq_iff_ordLess[simp]
+ not_ordLess_iff_ordLeq[simp]
+
+
+subsection {* Restriction to a set *}
+
+lemma Restr_incr2:
+"r <= r' \<Longrightarrow> Restr r A <= Restr r' A"
+by blast
+
+lemma Restr_incr:
+"\<lbrakk>r \<le> r'; A \<le> A'\<rbrakk> \<Longrightarrow> Restr r A \<le> Restr r' A'"
+by blast
+
+lemma Restr_Int:
+"Restr (Restr r A) B = Restr r (A Int B)"
+by blast
+
+lemma Restr_iff: "(a,b) : Restr r A = (a : A \<and> b : A \<and> (a,b) : r)"
+by (auto simp add: Field_def)
+
+lemma Restr_subset1: "Restr r A \<le> r"
+by auto
+
+lemma Restr_subset2: "Restr r A \<le> A \<times> A"
+by auto
+
+lemma wf_Restr:
+"wf r \<Longrightarrow> wf(Restr r A)"
+using wf_subset Restr_subset by blast
+
+lemma Restr_incr1:
+"A \<le> B \<Longrightarrow> Restr r A \<le> Restr r B"
+by blast
+
+
+subsection {* Order filters versus restrictions and embeddings *}
+
+lemma ofilter_Restr:
+assumes WELL: "Well_order r" and
+ OFA: "ofilter r A" and OFB: "ofilter r B" and SUB: "A \<le> B"
+shows "ofilter (Restr r B) A"
+proof-
+ let ?rB = "Restr r B"
+ have Well: "wo_rel r" unfolding wo_rel_def using WELL .
+ hence Refl: "Refl r" by (auto 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 (auto simp add: Well_order_Restr wo_rel_def)
+ (* Main proof *)
+ show ?thesis
+ proof(auto simp add: WellB wo_rel.ofilter_def)
+ fix a assume "a \<in> A"
+ hence "a \<in> Field r \<and> a \<in> B" using assms Well
+ by (auto simp add: wo_rel.ofilter_def)
+ with Field show "a \<in> Field(Restr r B)" by auto
+ next
+ fix a b assume *: "a \<in> A" and "b \<in> under (Restr r B) a"
+ hence "b \<in> under r a"
+ using WELL OFB SUB ofilter_Restr_under[of r B a] by auto
+ thus "b \<in> A" using * Well OFA by(auto simp add: wo_rel.ofilter_def)
+ qed
+qed
+
+lemma ofilter_subset_iso:
+assumes WELL: "Well_order r" and
+ OFA: "ofilter r A" and OFB: "ofilter r B"
+shows "(A = B) = iso (Restr r A) (Restr r B) id"
+using assms
+by (auto simp add: ofilter_subset_embedS_iso)
+
+
+subsection {* Ordering the well-orders by existence of embeddings *}
+
+corollary ordLeq_refl_on: "refl_on {r. Well_order r} ordLeq"
+using ordLeq_reflexive unfolding ordLeq_def refl_on_def
+by blast
+
+corollary ordLeq_trans: "trans ordLeq"
+using trans_def[of ordLeq] ordLeq_transitive by blast
+
+corollary ordLeq_preorder_on: "preorder_on {r. Well_order r} ordLeq"
+by(auto simp add: preorder_on_def ordLeq_refl_on ordLeq_trans)
+
+corollary ordIso_refl_on: "refl_on {r. Well_order r} ordIso"
+using ordIso_reflexive unfolding refl_on_def ordIso_def
+by blast
+
+corollary ordIso_trans: "trans ordIso"
+using trans_def[of ordIso] ordIso_transitive by blast
+
+corollary ordIso_sym: "sym ordIso"
+by (auto simp add: sym_def ordIso_symmetric)
+
+corollary ordIso_equiv: "equiv {r. Well_order r} ordIso"
+by (auto simp add: equiv_def ordIso_sym ordIso_refl_on ordIso_trans)
+
+lemma ordLess_irrefl: "irrefl ordLess"
+by(unfold irrefl_def, auto simp add: ordLess_irreflexive)
+
+lemma ordLess_or_ordIso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r <o r' \<or> r' <o r \<or> r =o r'"
+unfolding ordLess_def ordIso_def
+using assms embedS_or_iso[of r r'] by auto
+
+corollary ordLeq_ordLess_Un_ordIso:
+"ordLeq = ordLess \<union> ordIso"
+by (auto simp add: ordLeq_iff_ordLess_or_ordIso)
+
+lemma not_ordLeq_ordLess:
+"r \<le>o r' \<Longrightarrow> \<not> r' <o r"
+using not_ordLess_ordLeq by blast
+
+lemma ordIso_or_ordLess:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r =o r' \<or> r <o r' \<or> r' <o r"
+using assms ordLess_or_ordLeq ordLeq_iff_ordLess_or_ordIso by blast
+
+lemmas ord_trans = ordIso_transitive ordLeq_transitive ordLess_transitive
+ ordIso_ordLeq_trans ordLeq_ordIso_trans
+ ordIso_ordLess_trans ordLess_ordIso_trans
+ ordLess_ordLeq_trans ordLeq_ordLess_trans
+
+lemma ofilter_ordLeq:
+assumes "Well_order r" and "ofilter r A"
+shows "Restr r A \<le>o r"
+proof-
+ have "A \<le> Field r" using assms by (auto simp add: wo_rel_def wo_rel.ofilter_def)
+ thus ?thesis using assms
+ by (simp add: ofilter_subset_ordLeq wo_rel.Field_ofilter
+ wo_rel_def Restr_Field)
+qed
+
+corollary under_Restr_ordLeq:
+"Well_order r \<Longrightarrow> Restr r (under r a) \<le>o r"
+by (auto simp add: ofilter_ordLeq wo_rel.under_ofilter wo_rel_def)
+
+
+subsection {* Copy via direct images *}
+
+lemma Id_dir_image: "dir_image Id f \<le> Id"
+unfolding dir_image_def by auto
+
+lemma Un_dir_image:
+"dir_image (r1 \<union> r2) f = (dir_image r1 f) \<union> (dir_image r2 f)"
+unfolding dir_image_def by auto
+
+lemma Int_dir_image:
+assumes "inj_on f (Field r1 \<union> Field r2)"
+shows "dir_image (r1 Int r2) f = (dir_image r1 f) Int (dir_image r2 f)"
+proof
+ show "dir_image (r1 Int r2) f \<le> (dir_image r1 f) Int (dir_image r2 f)"
+ using assms unfolding dir_image_def inj_on_def by auto
+next
+ show "(dir_image r1 f) Int (dir_image r2 f) \<le> dir_image (r1 Int r2) f"
+ proof(clarify)
+ fix a' b'
+ assume "(a',b') \<in> dir_image r1 f" "(a',b') \<in> dir_image r2 f"
+ then obtain a1 b1 a2 b2
+ where 1: "a' = f a1 \<and> b' = f b1 \<and> a' = f a2 \<and> b' = f b2" and
+ 2: "(a1,b1) \<in> r1 \<and> (a2,b2) \<in> r2" and
+ 3: "{a1,b1} \<le> Field r1 \<and> {a2,b2} \<le> Field r2"
+ unfolding dir_image_def Field_def by blast
+ hence "a1 = a2 \<and> b1 = b2" using assms unfolding inj_on_def by auto
+ hence "a' = f a1 \<and> b' = f b1 \<and> (a1,b1) \<in> r1 Int r2 \<and> (a2,b2) \<in> r1 Int r2"
+ using 1 2 by auto
+ thus "(a',b') \<in> dir_image (r1 \<inter> r2) f"
+ unfolding dir_image_def by blast
+ qed
+qed
+
+
+subsection {* Ordinal-like sum of two (disjoint) well-orders *}
+
+text{* This is roughly obtained by ``concatenating" the two well-orders -- thus, all elements
+of the first will be smaller than all elements of the second. This construction
+only makes sense if the fields of the two well-order relations are disjoint. *}
+
+definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60)
+where
+"r Osum r' = r \<union> r' \<union> {(a,a'). a \<in> Field r \<and> a' \<in> Field r'}"
+
+abbreviation Osum2 :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "\<union>o" 60)
+where "r \<union>o r' \<equiv> r Osum r'"
+
+lemma Field_Osum: "Field(r Osum r') = Field r \<union> Field r'"
+unfolding Osum_def Field_def by blast
+
+lemma Osum_Refl:
+assumes FLD: "Field r Int Field r' = {}" and
+ REFL: "Refl r" and REFL': "Refl r'"
+shows "Refl (r Osum r')"
+using assms (* Need first unfold Field_Osum, only then Osum_def *)
+unfolding refl_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_trans:
+assumes FLD: "Field r Int Field r' = {}" and
+ TRANS: "trans r" and TRANS': "trans r'"
+shows "trans (r Osum r')"
+proof(unfold trans_def, auto)
+ fix x y z assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, z) \<in> r \<union>o r'"
+ show "(x, z) \<in> r \<union>o r'"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case11: "(y,z) \<in> r"
+ hence "(x,z) \<in> r" using Case1 TRANS trans_def[of r] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "(y,z) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ {assume Case13: "z \<in> Field r'"
+ hence ?thesis using 1 unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case21: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ {assume Case22: "(y,z) \<in> r'"
+ hence "(x,z) \<in> r'" using Case2 TRANS' trans_def[of r'] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case23: "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume Case31: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,z) \<in> r'"
+ hence "z \<in> Field r'" unfolding Field_def by blast
+ hence ?thesis unfolding Osum_def using Case3 by auto
+ }
+ moreover
+ {assume Case33: "y \<in> Field r"
+ hence False using FLD Case3 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Preorder:
+"\<lbrakk>Field r Int Field r' = {}; Preorder r; Preorder r'\<rbrakk> \<Longrightarrow> Preorder (r Osum r')"
+unfolding preorder_on_def using Osum_Refl Osum_trans by blast
+
+lemma Osum_antisym:
+assumes FLD: "Field r Int Field r' = {}" and
+ AN: "antisym r" and AN': "antisym r'"
+shows "antisym (r Osum r')"
+proof(unfold antisym_def, auto)
+ fix x y assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, x) \<in> r \<union>o r'"
+ show "x = y"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ have "(y,x) \<in> r \<Longrightarrow> ?thesis"
+ using Case1 AN antisym_def[of r] by blast
+ moreover
+ {assume "(y,x) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ have "x \<in> Field r' \<Longrightarrow> False" using FLD 1 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ have "(y,x) \<in> r' \<Longrightarrow> ?thesis"
+ using Case2 AN' antisym_def[of r'] by blast
+ moreover
+ {assume "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,x) \<in> r'"
+ hence "x \<in> Field r'" unfolding Field_def by blast
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ have "\<not> y \<in> Field r" using FLD Case3 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Partial_order:
+"\<lbrakk>Field r Int Field r' = {}; Partial_order r; Partial_order r'\<rbrakk> \<Longrightarrow>
+ Partial_order (r Osum r')"
+unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast
+
+lemma Osum_Total:
+assumes FLD: "Field r Int Field r' = {}" and
+ TOT: "Total r" and TOT': "Total r'"
+shows "Total (r Osum r')"
+using assms
+unfolding total_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_Linear_order:
+"\<lbrakk>Field r Int Field r' = {}; Linear_order r; Linear_order r'\<rbrakk> \<Longrightarrow>
+ Linear_order (r Osum r')"
+unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast
+
+lemma Osum_wf:
+assumes FLD: "Field r Int Field r' = {}" and
+ WF: "wf r" and WF': "wf r'"
+shows "wf (r Osum r')"
+unfolding wf_eq_minimal2 unfolding Field_Osum
+proof(intro allI impI, elim conjE)
+ fix A assume *: "A \<subseteq> Field r \<union> Field r'" and **: "A \<noteq> {}"
+ obtain B where B_def: "B = A Int Field r" by blast
+ show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r \<union>o r'"
+ proof(cases "B = {}")
+ assume Case1: "B \<noteq> {}"
+ hence "B \<noteq> {} \<and> B \<le> Field r" using B_def by auto
+ then obtain a where 1: "a \<in> B" and 2: "\<forall>a1 \<in> B. (a1,a) \<notin> r"
+ using WF unfolding wf_eq_minimal2 by blast
+ hence 3: "a \<in> Field r \<and> a \<notin> Field r'" using B_def FLD by auto
+ (* *)
+ have "\<forall>a1 \<in> A. (a1,a) \<notin> r Osum r'"
+ proof(intro ballI)
+ fix a1 assume **: "a1 \<in> A"
+ {assume Case11: "a1 \<in> Field r"
+ hence "(a1,a) \<notin> r" using B_def ** 2 by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 by (auto simp add: Field_def)
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "a1 \<notin> Field r"
+ hence "(a1,a) \<notin> r" unfolding Field_def by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 unfolding Field_def by auto
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ ultimately show "(a1,a) \<notin> r Osum r'" by blast
+ qed
+ thus ?thesis using 1 B_def by auto
+ next
+ assume Case2: "B = {}"
+ hence 1: "A \<noteq> {} \<and> A \<le> Field r'" using * ** B_def by auto
+ then obtain a' where 2: "a' \<in> A" and 3: "\<forall>a1' \<in> A. (a1',a') \<notin> r'"
+ using WF' unfolding wf_eq_minimal2 by blast
+ hence 4: "a' \<in> Field r' \<and> a' \<notin> Field r" using 1 FLD by blast
+ (* *)
+ have "\<forall>a1' \<in> A. (a1',a') \<notin> r Osum r'"
+ proof(unfold Osum_def, auto simp add: 3)
+ fix a1' assume "(a1', a') \<in> r"
+ thus False using 4 unfolding Field_def by blast
+ next
+ fix a1' assume "a1' \<in> A" and "a1' \<in> Field r"
+ thus False using Case2 B_def by auto
+ qed
+ thus ?thesis using 2 by blast
+ qed
+qed
+
+lemma Osum_minus_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ NID: "\<not> (r \<le> Id)" and NID': "\<not> (r' \<le> Id)"
+shows "(r Osum r') - Id \<le> (r - Id) Osum (r' - Id)"
+proof-
+ {fix a a' assume *: "(a,a') \<in> (r Osum r')" and **: "a \<noteq> a'"
+ have "(a,a') \<in> (r - Id) Osum (r' - Id)"
+ proof-
+ {assume "(a,a') \<in> r \<or> (a,a') \<in> r'"
+ with ** have ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume "a \<in> Field r \<and> a' \<in> Field r'"
+ hence "a \<in> Field(r - Id) \<and> a' \<in> Field (r' - Id)"
+ using assms rel.Total_Id_Field by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+ }
+ thus ?thesis by(auto simp add: Osum_def)
+qed
+
+
+lemma wf_Int_Times:
+assumes "A Int B = {}"
+shows "wf(A \<times> B)"
+proof(unfold wf_def, auto)
+ fix P x
+ assume *: "\<forall>x. (\<forall>y. y \<in> A \<and> x \<in> B \<longrightarrow> P y) \<longrightarrow> P x"
+ moreover have "\<forall>y \<in> A. P y" using assms * by blast
+ ultimately show "P x" using * by (case_tac "x \<in> B", auto)
+qed
+
+lemma Osum_minus_Id1:
+assumes "r \<le> Id"
+shows "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r' - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r'"
+ with * have "(a,b) \<in> r' - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_minus_Id2:
+assumes "r' \<le> Id"
+shows "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r'"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r"
+ with * have "(a,b) \<in> r - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_wf_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ FLD: "Field r Int Field r' = {}" and
+ WF: "wf(r - Id)" and WF': "wf(r' - Id)"
+shows "wf ((r Osum r') - Id)"
+proof(cases "r \<le> Id \<or> r' \<le> Id")
+ assume Case1: "\<not>(r \<le> Id \<or> r' \<le> Id)"
+ have "Field(r - Id) Int Field(r' - Id) = {}"
+ using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r']
+ Diff_subset[of r Id] Diff_subset[of r' Id] by blast
+ thus ?thesis
+ using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"]
+ wf_subset[of "(r - Id) \<union>o (r' - Id)" "(r Osum r') - Id"] by auto
+next
+ have 1: "wf(Field r \<times> Field r')"
+ using FLD by (auto simp add: wf_Int_Times)
+ assume Case2: "r \<le> Id \<or> r' \<le> Id"
+ moreover
+ {assume Case21: "r \<le> Id"
+ hence "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id1[of r r'] by simp
+ moreover
+ {have "Domain(Field r \<times> Field r') Int Range(r' - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r' - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF' wf_Un[of "Field r \<times> Field r'" "r' - Id"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (auto simp add: wf_subset)
+ }
+ moreover
+ {assume Case22: "r' \<le> Id"
+ hence "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id2[of r' r] by simp
+ moreover
+ {have "Range(Field r \<times> Field r') Int Domain(r - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF wf_Un[of "r - Id" "Field r \<times> Field r'"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (auto simp add: wf_subset)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma Osum_Well_order:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "Well_order (r Osum r')"
+proof-
+ have "Total r \<and> Total r'" using WELL WELL'
+ by (auto simp add: order_on_defs)
+ thus ?thesis using assms unfolding well_order_on_def
+ using Osum_Linear_order Osum_wf_Id by blast
+qed
+
+lemma Osum_embed:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "embed r (r Osum r') id"
+proof-
+ have 1: "Well_order (r Osum r')"
+ using assms by (auto simp add: Osum_Well_order)
+ moreover
+ have "compat r (r Osum r') id"
+ unfolding compat_def Osum_def by auto
+ moreover
+ have "inj_on id (Field r)" by simp
+ moreover
+ have "ofilter (r Osum r') (Field r)"
+ using 1 proof(auto simp add: wo_rel_def wo_rel.ofilter_def
+ Field_Osum rel.under_def)
+ fix a b assume 2: "a \<in> Field r" and 3: "(b,a) \<in> r Osum r'"
+ moreover
+ {assume "(b,a) \<in> r'"
+ hence "a \<in> Field r'" using Field_def[of r'] by blast
+ hence False using 2 FLD by blast
+ }
+ moreover
+ {assume "a \<in> Field r'"
+ hence False using 2 FLD by blast
+ }
+ ultimately
+ show "b \<in> Field r" by (auto simp add: Osum_def Field_def)
+ qed
+ ultimately show ?thesis
+ using assms by (auto simp add: embed_iff_compat_inj_on_ofilter)
+qed
+
+corollary Osum_ordLeq:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "r \<le>o r Osum r'"
+using assms Osum_embed Osum_Well_order
+unfolding ordLeq_def by blast
+
+lemma Well_order_embed_copy:
+assumes WELL: "well_order_on A r" and
+ INJ: "inj_on f A" and SUB: "f ` A \<le> B"
+shows "\<exists>r'. well_order_on B r' \<and> r \<le>o r'"
+proof-
+ have "bij_betw f A (f ` A)"
+ using INJ inj_on_imp_bij_betw by blast
+ then obtain r'' where "well_order_on (f ` A) r''" and 1: "r =o r''"
+ using WELL Well_order_iso_copy by blast
+ hence 2: "Well_order r'' \<and> Field r'' = (f ` A)"
+ using rel.well_order_on_Well_order by blast
+ (* *)
+ let ?C = "B - (f ` A)"
+ obtain r''' where "well_order_on ?C r'''"
+ using well_order_on by blast
+ hence 3: "Well_order r''' \<and> Field r''' = ?C"
+ using rel.well_order_on_Well_order by blast
+ (* *)
+ let ?r' = "r'' Osum r'''"
+ have "Field r'' Int Field r''' = {}"
+ using 2 3 by auto
+ hence "r'' \<le>o ?r'" using Osum_ordLeq[of r'' r'''] 2 3 by blast
+ hence 4: "r \<le>o ?r'" using 1 ordIso_ordLeq_trans by blast
+ (* *)
+ hence "Well_order ?r'" unfolding ordLeq_def by auto
+ moreover
+ have "Field ?r' = B" using 2 3 SUB by (auto simp add: Field_Osum)
+ ultimately show ?thesis using 4 by blast
+qed
+
+
+subsection {* The maxim among a finite set of ordinals *}
+
+text {* The correct phrasing would be ``a maxim of ...", as @{text "\<le>o"} is only a preorder. *}
+
+definition isOmax :: "'a rel set \<Rightarrow> 'a rel \<Rightarrow> bool"
+where
+"isOmax R r == r \<in> R \<and> (ALL r' : R. r' \<le>o r)"
+
+definition omax :: "'a rel set \<Rightarrow> 'a rel"
+where
+"omax R == SOME r. isOmax R r"
+
+lemma exists_isOmax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "\<exists> r. isOmax R r"
+proof-
+ have "finite R \<Longrightarrow> R \<noteq> {} \<longrightarrow> (\<forall> r \<in> R. Well_order r) \<longrightarrow> (\<exists> r. isOmax R r)"
+ apply(erule finite_induct) apply(simp add: isOmax_def)
+ proof(clarsimp)
+ fix r :: "('a \<times> 'a) set" and R assume *: "finite R" and **: "r \<notin> R"
+ and ***: "Well_order r" and ****: "\<forall>r\<in>R. Well_order r"
+ and IH: "R \<noteq> {} \<longrightarrow> (\<exists>p. isOmax R p)"
+ let ?R' = "insert r R"
+ show "\<exists>p'. (isOmax ?R' p')"
+ proof(cases "R = {}")
+ assume Case1: "R = {}"
+ thus ?thesis unfolding isOmax_def using ***
+ by (simp add: ordLeq_reflexive)
+ next
+ assume Case2: "R \<noteq> {}"
+ then obtain p where p: "isOmax R p" using IH by auto
+ hence 1: "Well_order p" using **** unfolding isOmax_def by simp
+ {assume Case21: "r \<le>o p"
+ hence "isOmax ?R' p" using p unfolding isOmax_def by simp
+ hence ?thesis by auto
+ }
+ moreover
+ {assume Case22: "p \<le>o r"
+ {fix r' assume "r' \<in> ?R'"
+ moreover
+ {assume "r' \<in> R"
+ hence "r' \<le>o p" using p unfolding isOmax_def by simp
+ hence "r' \<le>o r" using Case22 by(rule ordLeq_transitive)
+ }
+ moreover have "r \<le>o r" using *** by(rule ordLeq_reflexive)
+ ultimately have "r' \<le>o r" by auto
+ }
+ hence "isOmax ?R' r" unfolding isOmax_def by simp
+ hence ?thesis by auto
+ }
+ moreover have "r \<le>o p \<or> p \<le>o r"
+ using 1 *** ordLeq_total by auto
+ ultimately show ?thesis by blast
+ qed
+ qed
+ thus ?thesis using assms by auto
+qed
+
+lemma omax_isOmax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "isOmax R (omax R)"
+unfolding omax_def using assms
+by(simp add: exists_isOmax someI_ex)
+
+lemma omax_in:
+assumes "finite R" and "R \<noteq> {}" and "\<forall> r \<in> R. Well_order r"
+shows "omax R \<in> R"
+using assms omax_isOmax unfolding isOmax_def by blast
+
+lemma Well_order_omax:
+assumes "finite R" and "R \<noteq> {}" and "\<forall>r\<in>R. Well_order r"
+shows "Well_order (omax R)"
+using assms apply - apply(drule omax_in) by auto
+
+lemma omax_maxim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r" and "r \<in> R"
+shows "r \<le>o omax R"
+using assms omax_isOmax unfolding isOmax_def by blast
+
+lemma omax_ordLeq:
+assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r \<le>o p"
+shows "omax R \<le>o p"
+proof-
+ have "\<forall> r \<in> R. Well_order r" using * unfolding ordLeq_def by simp
+ thus ?thesis using assms omax_in by auto
+qed
+
+lemma omax_ordLess:
+assumes "finite R" and "R \<noteq> {}" and *: "\<forall> r \<in> R. r <o p"
+shows "omax R <o p"
+proof-
+ have "\<forall> r \<in> R. Well_order r" using * unfolding ordLess_def by simp
+ thus ?thesis using assms omax_in by auto
+qed
+
+lemma omax_ordLeq_elim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "omax R \<le>o p" and "r \<in> R"
+shows "r \<le>o p"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_transitive by blast
+
+lemma omax_ordLess_elim:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "omax R <o p" and "r \<in> R"
+shows "r <o p"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_ordLess_trans by blast
+
+lemma ordLeq_omax:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "r \<in> R" and "p \<le>o r"
+shows "p \<le>o omax R"
+using assms omax_maxim[of R r] apply simp
+using ordLeq_transitive by blast
+
+lemma ordLess_omax:
+assumes "finite R" and "\<forall> r \<in> R. Well_order r"
+and "r \<in> R" and "p <o r"
+shows "p <o omax R"
+using assms omax_maxim[of R r] apply simp
+using ordLess_ordLeq_trans by blast
+
+lemma omax_ordLeq_mono:
+assumes P: "finite P" and R: "finite R"
+and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
+and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p \<le>o r"
+shows "omax P \<le>o omax R"
+proof-
+ let ?mp = "omax P" let ?mr = "omax R"
+ {fix p assume "p : P"
+ then obtain r where r: "r : R" and "p \<le>o r"
+ using LEQ by blast
+ moreover have "r <=o ?mr"
+ using r R Well_R omax_maxim by blast
+ ultimately have "p <=o ?mr"
+ using ordLeq_transitive by blast
+ }
+ thus "?mp <=o ?mr"
+ using NE_P P using omax_ordLeq by blast
+qed
+
+lemma omax_ordLess_mono:
+assumes P: "finite P" and R: "finite R"
+and NE_P: "P \<noteq> {}" and Well_R: "\<forall> r \<in> R. Well_order r"
+and LEQ: "\<forall> p \<in> P. \<exists> r \<in> R. p <o r"
+shows "omax P <o omax R"
+proof-
+ let ?mp = "omax P" let ?mr = "omax R"
+ {fix p assume "p : P"
+ then obtain r where r: "r : R" and "p <o r"
+ using LEQ by blast
+ moreover have "r <=o ?mr"
+ using r R Well_R omax_maxim by blast
+ ultimately have "p <o ?mr"
+ using ordLess_ordLeq_trans by blast
+ }
+ thus "?mp <o ?mr"
+ using NE_P P omax_ordLess by blast
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Constructions_on_Wellorders_Base.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,1633 @@
+(* 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 blast
+ }
+ 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 blast
+ 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 auto
+ }
+ 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/Fun_More.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,183 @@
+(* Title: HOL/Cardinals/Fun_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on injections, bijections and inverses.
+*)
+
+header {* More on Injections, Bijections and Inverses *}
+
+theory Fun_More
+imports Fun_More_Base
+begin
+
+
+subsection {* Purely functional properties *}
+
+(* unused *)
+(*1*)lemma notIn_Un_bij_betw2:
+assumes NIN: "b \<notin> A" and NIN': "b' \<notin> A'" and
+ BIJ: "bij_betw f A A'"
+shows "bij_betw f (A \<union> {b}) (A' \<union> {b'}) = (f b = b')"
+proof
+ assume "f b = b'"
+ thus "bij_betw f (A \<union> {b}) (A' \<union> {b'})"
+ using assms notIn_Un_bij_betw[of b A f A'] by auto
+next
+ assume *: "bij_betw f (A \<union> {b}) (A' \<union> {b'})"
+ hence "f b \<in> A' \<union> {b'}"
+ unfolding bij_betw_def by auto
+ moreover
+ {assume "f b \<in> A'"
+ then obtain b1 where 1: "b1 \<in> A" and 2: "f b1 = f b" using BIJ
+ by (auto simp add: bij_betw_def)
+ hence "b = b1" using *
+ by (auto simp add: bij_betw_def inj_on_def)
+ with 1 NIN have False by auto
+ }
+ ultimately show "f b = b'" by blast
+qed
+
+(* unused *)
+(*1*)lemma bij_betw_ball:
+assumes BIJ: "bij_betw f A B"
+shows "(\<forall>b \<in> B. phi b) = (\<forall>a \<in> A. phi(f a))"
+using assms unfolding bij_betw_def inj_on_def by blast
+
+(* unused *)
+(*1*)lemma bij_betw_diff_singl:
+assumes BIJ: "bij_betw f A A'" and IN: "a \<in> A"
+shows "bij_betw f (A - {a}) (A' - {f a})"
+proof-
+ let ?B = "A - {a}" let ?B' = "A' - {f a}"
+ have "f a \<in> A'" using IN BIJ unfolding bij_betw_def by blast
+ hence "a \<notin> ?B \<and> f a \<notin> ?B' \<and> A = ?B \<union> {a} \<and> A' = ?B' \<union> {f a}"
+ using IN by blast
+ thus ?thesis using notIn_Un_bij_betw3[of a ?B f ?B'] BIJ by simp
+qed
+
+
+subsection {* Properties involving finite and infinite sets *}
+
+(*3*)lemma inj_on_image_Pow:
+assumes "inj_on f A"
+shows "inj_on (image f) (Pow A)"
+unfolding Pow_def inj_on_def proof(clarsimp)
+ fix X Y assume *: "X \<le> A" and **: "Y \<le> A" and
+ ***: "f ` X = f ` Y"
+ show "X = Y"
+ proof(auto)
+ fix x assume ****: "x \<in> X"
+ with *** obtain y where "y \<in> Y \<and> f x = f y" by blast
+ with **** * ** assms show "x \<in> Y"
+ unfolding inj_on_def by auto
+ next
+ fix y assume ****: "y \<in> Y"
+ with *** obtain x where "x \<in> X \<and> f x = f y" by force
+ with **** * ** assms show "y \<in> X"
+ unfolding inj_on_def by auto
+ qed
+qed
+
+(*2*)lemma bij_betw_image_Pow:
+assumes "bij_betw f A B"
+shows "bij_betw (image f) (Pow A) (Pow B)"
+using assms unfolding bij_betw_def
+by (auto simp add: inj_on_image_Pow image_Pow_surj)
+
+(* unused *)
+(*1*)lemma bij_betw_inv_into_RIGHT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B' \<le> A'"
+shows "f `((inv_into A f)`B') = B'"
+using assms
+proof(auto simp add: bij_betw_inv_into_right)
+ let ?f' = "(inv_into A f)"
+ fix a' assume *: "a' \<in> B'"
+ hence "a' \<in> A'" using SUB by auto
+ hence "a' = f (?f' a')"
+ using BIJ by (auto simp add: bij_betw_inv_into_right)
+ thus "a' \<in> f ` (?f' ` B')" using * by blast
+qed
+
+(* unused *)
+(*1*)lemma bij_betw_inv_into_RIGHT_LEFT:
+assumes BIJ: "bij_betw f A A'" and SUB: "B' \<le> A'" and
+ IM: "(inv_into A f) ` B' = B"
+shows "f ` B = B'"
+proof-
+ have "f`((inv_into A f)` B') = B'"
+ using assms bij_betw_inv_into_RIGHT[of f A A' B'] by auto
+ thus ?thesis using IM by auto
+qed
+
+(* unused *)
+(*2*)lemma bij_betw_inv_into_twice:
+assumes "bij_betw f A A'"
+shows "\<forall>a \<in> A. inv_into A' (inv_into A f) a = f a"
+proof
+ let ?f' = "inv_into A f" let ?f'' = "inv_into A' ?f'"
+ have 1: "bij_betw ?f' A' A" using assms
+ by (auto simp add: bij_betw_inv_into)
+ fix a assume *: "a \<in> A"
+ then obtain a' where 2: "a' \<in> A'" and 3: "?f' a' = a"
+ using 1 unfolding bij_betw_def by force
+ hence "?f'' a = a'"
+ using * 1 3 by (auto simp add: bij_betw_inv_into_left)
+ moreover have "f a = a'" using assms 2 3
+ by (auto simp add: bij_betw_inv_into_right)
+ ultimately show "?f'' a = f a" by simp
+qed
+
+
+subsection {* Properties involving Hilbert choice *}
+
+
+subsection {* Other facts *}
+
+(*3*)lemma atLeastLessThan_injective:
+assumes "{0 ..< m::nat} = {0 ..< n}"
+shows "m = n"
+proof-
+ {assume "m < n"
+ hence "m \<in> {0 ..< n}" by auto
+ hence "{0 ..< m} < {0 ..< n}" by auto
+ hence False using assms by blast
+ }
+ moreover
+ {assume "n < m"
+ hence "n \<in> {0 ..< m}" by auto
+ hence "{0 ..< n} < {0 ..< m}" by auto
+ hence False using assms by blast
+ }
+ ultimately show ?thesis by force
+qed
+
+(*2*)lemma atLeastLessThan_injective2:
+"bij_betw f {0 ..< m::nat} {0 ..< n} \<Longrightarrow> m = n"
+using finite_atLeastLessThan[of m] finite_atLeastLessThan[of n]
+ card_atLeastLessThan[of m] card_atLeastLessThan[of n]
+ bij_betw_iff_card[of "{0 ..< m}" "{0 ..< n}"] by auto
+
+(* unused *)
+(*2*)lemma atLeastLessThan_less_eq3:
+"(\<exists>f. inj_on f {0..<(m::nat)} \<and> f ` {0..<m} \<le> {0..<n}) = (m \<le> n)"
+using atLeastLessThan_less_eq2
+proof(auto)
+ assume "m \<le> n"
+ hence "inj_on id {0..<m} \<and> id ` {0..<m} \<subseteq> {0..<n}" unfolding inj_on_def by force
+ thus "\<exists>f. inj_on f {0..<m} \<and> f ` {0..<m} \<subseteq> {0..<n}" by blast
+qed
+
+(* unused *)
+(*3*)lemma atLeastLessThan_less:
+"({0..<m} < {0..<n}) = ((m::nat) < n)"
+proof-
+ have "({0..<m} < {0..<n}) = ({0..<m} \<le> {0..<n} \<and> {0..<m} ~= {0..<n})"
+ using subset_iff_psubset_eq by blast
+ also have "\<dots> = (m \<le> n \<and> m ~= n)"
+ using atLeastLessThan_less_eq atLeastLessThan_injective by blast
+ also have "\<dots> = (m < n)" by auto
+ finally show ?thesis .
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Fun_More_Base.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,252 @@
+(* 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, auto)
+ 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/Order_Relation_More.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,582 @@
+(* Title: HOL/Cardinals/Order_Relation_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Basics on order-like relations.
+*)
+
+header {* Basics on Order-Like Relations *}
+
+theory Order_Relation_More
+imports Order_Relation_More_Base
+begin
+
+
+subsection {* The upper and lower bounds operators *}
+
+lemma (in rel) aboveS_subset_above: "aboveS a \<le> above a"
+by(auto simp add: aboveS_def above_def)
+
+lemma (in rel) AboveS_subset_Above: "AboveS A \<le> Above A"
+by(auto simp add: AboveS_def Above_def)
+
+lemma (in rel) UnderS_disjoint: "A Int (UnderS A) = {}"
+by(auto simp add: UnderS_def)
+
+lemma (in rel) aboveS_notIn: "a \<notin> aboveS a"
+by(auto simp add: aboveS_def)
+
+lemma (in rel) Refl_above_in: "\<lbrakk>Refl r; a \<in> Field r\<rbrakk> \<Longrightarrow> a \<in> above a"
+by(auto simp add: refl_on_def above_def)
+
+lemma (in rel) in_Above_under: "a \<in> Field r \<Longrightarrow> a \<in> Above (under a)"
+by(auto simp add: Above_def under_def)
+
+lemma (in rel) in_Under_above: "a \<in> Field r \<Longrightarrow> a \<in> Under (above a)"
+by(auto simp add: Under_def above_def)
+
+lemma (in rel) in_UnderS_aboveS: "a \<in> Field r \<Longrightarrow> a \<in> UnderS (aboveS a)"
+by(auto simp add: UnderS_def aboveS_def)
+
+lemma (in rel) subset_Above_Under: "B \<le> Field r \<Longrightarrow> B \<le> Above (Under B)"
+by(auto simp add: Above_def Under_def)
+
+lemma (in rel) subset_Under_Above: "B \<le> Field r \<Longrightarrow> B \<le> Under (Above B)"
+by(auto simp add: Under_def Above_def)
+
+lemma (in rel) subset_AboveS_UnderS: "B \<le> Field r \<Longrightarrow> B \<le> AboveS (UnderS B)"
+by(auto simp add: AboveS_def UnderS_def)
+
+lemma (in rel) subset_UnderS_AboveS: "B \<le> Field r \<Longrightarrow> B \<le> UnderS (AboveS B)"
+by(auto simp add: UnderS_def AboveS_def)
+
+lemma (in rel) Under_Above_Galois:
+"\<lbrakk>B \<le> Field r; C \<le> Field r\<rbrakk> \<Longrightarrow> (B \<le> Above C) = (C \<le> Under B)"
+by(unfold Above_def Under_def, blast)
+
+lemma (in rel) UnderS_AboveS_Galois:
+"\<lbrakk>B \<le> Field r; C \<le> Field r\<rbrakk> \<Longrightarrow> (B \<le> AboveS C) = (C \<le> UnderS B)"
+by(unfold AboveS_def UnderS_def, blast)
+
+lemma (in rel) Refl_above_aboveS:
+assumes REFL: "Refl r" and IN: "a \<in> Field r"
+shows "above a = aboveS a \<union> {a}"
+proof(unfold above_def aboveS_def, auto)
+ show "(a,a) \<in> r" using REFL IN refl_on_def[of _ r] by blast
+qed
+
+lemma (in rel) Linear_order_under_aboveS_Field:
+assumes LIN: "Linear_order r" and IN: "a \<in> Field r"
+shows "Field r = under a \<union> aboveS a"
+proof(unfold under_def aboveS_def, auto)
+ assume "a \<in> Field r" "(a, a) \<notin> r"
+ with LIN IN order_on_defs[of _ r] refl_on_def[of _ r]
+ show False by auto
+next
+ fix b assume "b \<in> Field r" "(b, a) \<notin> r"
+ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r]
+ have "(a,b) \<in> r \<or> a = b" by blast
+ thus "(a,b) \<in> r"
+ using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto
+next
+ fix b assume "(b, a) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast
+next
+ fix b assume "b \<noteq> a" "(a, b) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast
+qed
+
+lemma (in rel) Linear_order_underS_above_Field:
+assumes LIN: "Linear_order r" and IN: "a \<in> Field r"
+shows "Field r = underS a \<union> above a"
+proof(unfold underS_def above_def, auto)
+ assume "a \<in> Field r" "(a, a) \<notin> r"
+ with LIN IN order_on_defs[of _ r] refl_on_def[of _ r]
+ show False by auto
+next
+ fix b assume "b \<in> Field r" "(a, b) \<notin> r"
+ with LIN IN order_on_defs[of "Field r" r] total_on_def[of "Field r" r]
+ have "(b,a) \<in> r \<or> b = a" by blast
+ thus "(b,a) \<in> r"
+ using LIN IN order_on_defs[of _ r] refl_on_def[of _ r] by auto
+next
+ fix b assume "b \<noteq> a" "(b, a) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of _ r] refl_on_def[of _ r] by blast
+next
+ fix b assume "(a, b) \<in> r"
+ thus "b \<in> Field r"
+ using LIN order_on_defs[of "Field r" r] refl_on_def[of "Field r" r] by blast
+qed
+
+lemma (in rel) under_empty: "a \<notin> Field r \<Longrightarrow> under a = {}"
+unfolding Field_def under_def by auto
+
+lemma (in rel) above_Field: "above a \<le> Field r"
+by(unfold above_def Field_def, auto)
+
+lemma (in rel) aboveS_Field: "aboveS a \<le> Field r"
+by(unfold aboveS_def Field_def, auto)
+
+lemma (in rel) Above_Field: "Above A \<le> Field r"
+by(unfold Above_def Field_def, auto)
+
+lemma (in rel) Refl_under_Under:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "Under A = (\<Inter> a \<in> A. under a)"
+proof
+ show "Under A \<subseteq> (\<Inter> a \<in> A. under a)"
+ by(unfold Under_def under_def, auto)
+next
+ show "(\<Inter> a \<in> A. under a) \<subseteq> Under A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> under xa"
+ hence "\<forall>xa \<in> A. (x,xa) \<in> r"
+ by (simp add: under_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> under a" by simp
+ hence "x \<in> Field r"
+ using under_Field[of a] by auto
+ }
+ ultimately show "x \<in> Under A"
+ unfolding Under_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_underS_UnderS:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "UnderS A = (\<Inter> a \<in> A. underS a)"
+proof
+ show "UnderS A \<subseteq> (\<Inter> a \<in> A. underS a)"
+ by(unfold UnderS_def underS_def, auto)
+next
+ show "(\<Inter> a \<in> A. underS a) \<subseteq> UnderS A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> underS xa"
+ hence "\<forall>xa \<in> A. x \<noteq> xa \<and> (x,xa) \<in> r"
+ by (auto simp add: underS_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> underS a" by simp
+ hence "x \<in> Field r"
+ using underS_Field[of a] by auto
+ }
+ ultimately show "x \<in> UnderS A"
+ unfolding UnderS_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_above_Above:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "Above A = (\<Inter> a \<in> A. above a)"
+proof
+ show "Above A \<subseteq> (\<Inter> a \<in> A. above a)"
+ by(unfold Above_def above_def, auto)
+next
+ show "(\<Inter> a \<in> A. above a) \<subseteq> Above A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> above xa"
+ hence "\<forall>xa \<in> A. (xa,x) \<in> r"
+ by (simp add: above_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> above a" by simp
+ hence "x \<in> Field r"
+ using above_Field[of a] by auto
+ }
+ ultimately show "x \<in> Above A"
+ unfolding Above_def by auto
+ qed
+qed
+
+lemma (in rel) Refl_aboveS_AboveS:
+assumes REFL: "Refl r" and NE: "A \<noteq> {}"
+shows "AboveS A = (\<Inter> a \<in> A. aboveS a)"
+proof
+ show "AboveS A \<subseteq> (\<Inter> a \<in> A. aboveS a)"
+ by(unfold AboveS_def aboveS_def, auto)
+next
+ show "(\<Inter> a \<in> A. aboveS a) \<subseteq> AboveS A"
+ proof(auto)
+ fix x
+ assume *: "\<forall>xa \<in> A. x \<in> aboveS xa"
+ hence "\<forall>xa \<in> A. xa \<noteq> x \<and> (xa,x) \<in> r"
+ by (auto simp add: aboveS_def)
+ moreover
+ {from NE obtain a where "a \<in> A" by blast
+ with * have "x \<in> aboveS a" by simp
+ hence "x \<in> Field r"
+ using aboveS_Field[of a] by auto
+ }
+ ultimately show "x \<in> AboveS A"
+ unfolding AboveS_def by auto
+ qed
+qed
+
+lemma (in rel) under_Under_singl: "under a = Under {a}"
+by(unfold Under_def under_def, auto simp add: Field_def)
+
+lemma (in rel) underS_UnderS_singl: "underS a = UnderS {a}"
+by(unfold UnderS_def underS_def, auto simp add: Field_def)
+
+lemma (in rel) above_Above_singl: "above a = Above {a}"
+by(unfold Above_def above_def, auto simp add: Field_def)
+
+lemma (in rel) aboveS_AboveS_singl: "aboveS a = AboveS {a}"
+by(unfold AboveS_def aboveS_def, auto simp add: Field_def)
+
+lemma (in rel) Under_decr: "A \<le> B \<Longrightarrow> Under B \<le> Under A"
+by(unfold Under_def, auto)
+
+lemma (in rel) UnderS_decr: "A \<le> B \<Longrightarrow> UnderS B \<le> UnderS A"
+by(unfold UnderS_def, auto)
+
+lemma (in rel) Above_decr: "A \<le> B \<Longrightarrow> Above B \<le> Above A"
+by(unfold Above_def, auto)
+
+lemma (in rel) AboveS_decr: "A \<le> B \<Longrightarrow> AboveS B \<le> AboveS A"
+by(unfold AboveS_def, auto)
+
+lemma (in rel) under_incl_iff:
+assumes TRANS: "trans r" and REFL: "Refl r" and IN: "a \<in> Field r"
+shows "(under a \<le> under b) = ((a,b) \<in> r)"
+proof
+ assume "(a,b) \<in> r"
+ thus "under a \<le> under b" using TRANS
+ by (auto simp add: under_incr)
+next
+ assume "under a \<le> under b"
+ moreover
+ have "a \<in> under a" using REFL IN
+ by (auto simp add: Refl_under_in)
+ ultimately show "(a,b) \<in> r"
+ by (auto simp add: under_def)
+qed
+
+lemma (in rel) above_decr:
+assumes TRANS: "trans r" and REL: "(a,b) \<in> r"
+shows "above b \<le> above a"
+proof(unfold above_def, auto)
+ fix x assume "(b,x) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(a,x) \<in> r" by blast
+qed
+
+lemma (in rel) aboveS_decr:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ REL: "(a,b) \<in> r"
+shows "aboveS b \<le> aboveS a"
+proof(unfold aboveS_def, auto)
+ assume *: "a \<noteq> b" and **: "(b,a) \<in> r"
+ with ANTISYM antisym_def[of r] REL
+ show False by auto
+next
+ fix x assume "x \<noteq> b" "(b,x) \<in> r"
+ with REL TRANS trans_def[of r]
+ show "(a,x) \<in> r" by blast
+qed
+
+lemma (in rel) 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> (b,c) \<in> r"
+ using IN1 IN2 under_def by auto
+ hence "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ thus ?thesis unfolding under_def by simp
+qed
+
+lemma (in rel) 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-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 under_def underS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "b \<noteq> c" using IN2 underS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding underS_def by simp
+qed
+
+lemma (in rel) underS_under_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> under c"
+shows "a \<in> underS c"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 under_def underS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "a \<noteq> b" using IN1 underS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding underS_def by simp
+qed
+
+lemma (in rel) underS_underS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> underS c"
+shows "a \<in> underS c"
+proof-
+ have "a \<in> under b"
+ using IN1 underS_subset_under by auto
+ with assms under_underS_trans show ?thesis by auto
+qed
+
+lemma (in rel) above_trans:
+assumes TRANS: "trans r" and
+ IN1: "b \<in> above a" and IN2: "c \<in> above b"
+shows "c \<in> above a"
+proof-
+ have "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def by auto
+ hence "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ thus ?thesis unfolding above_def by simp
+qed
+
+lemma (in rel) above_aboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> above a" and IN2: "c \<in> aboveS b"
+shows "c \<in> aboveS a"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def aboveS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "b \<noteq> c" using IN2 aboveS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding aboveS_def by simp
+qed
+
+lemma (in rel) aboveS_above_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> aboveS a" and IN2: "c \<in> above b"
+shows "c \<in> aboveS a"
+proof-
+ have 0: "(a,b) \<in> r \<and> (b,c) \<in> r"
+ using IN1 IN2 above_def aboveS_def by auto
+ hence 1: "(a,c) \<in> r"
+ using TRANS trans_def[of r] by blast
+ have 2: "a \<noteq> b" using IN1 aboveS_def by auto
+ have 3: "a \<noteq> c"
+ proof
+ assume "a = c" with 0 2 ANTISYM antisym_def[of r]
+ show False by auto
+ qed
+ from 1 3 show ?thesis unfolding aboveS_def by simp
+qed
+
+lemma (in rel) aboveS_aboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "b \<in> aboveS a" and IN2: "c \<in> aboveS b"
+shows "c \<in> aboveS a"
+proof-
+ have "b \<in> above a"
+ using IN1 aboveS_subset_above by auto
+ with assms above_aboveS_trans show ?thesis by auto
+qed
+
+lemma (in rel) underS_Under_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS b" and IN2: "b \<in> Under C"
+shows "a \<in> UnderS C"
+proof-
+ from IN1 have "a \<in> under b"
+ using underS_subset_under[of b] by blast
+ with assms under_Under_trans
+ have "a \<in> Under C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "b \<noteq> a \<and> (a,b) \<in> r"
+ using IN1 underS_def[of b] by auto
+ have "\<forall>c \<in> C. (b,c) \<in> r"
+ using IN2 Under_def[of C] by auto
+ with * have "(b,a) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding UnderS_def
+ using Under_def by auto
+qed
+
+lemma (in rel) underS_UnderS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> underS 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 underS_Under_trans assms
+ show ?thesis by auto
+qed
+
+lemma (in rel) above_Above_trans:
+assumes TRANS: "trans r" and
+ IN1: "a \<in> above b" and IN2: "b \<in> Above C"
+shows "a \<in> Above C"
+proof-
+ have "(b,a) \<in> r \<and> (\<forall>c \<in> C. (c,b) \<in> r)"
+ using IN1 IN2 above_def Above_def by auto
+ hence "\<forall>c \<in> C. (c,a) \<in> r"
+ using TRANS trans_def[of r] by blast
+ moreover
+ have "a \<in> Field r" using IN1 Field_def above_def by force
+ ultimately
+ show ?thesis unfolding Above_def by auto
+qed
+
+lemma (in rel) aboveS_Above_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> aboveS b" and IN2: "b \<in> Above C"
+shows "a \<in> AboveS C"
+proof-
+ from IN1 have "a \<in> above b"
+ using aboveS_subset_above[of b] by blast
+ with assms above_Above_trans
+ have "a \<in> Above C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "b \<noteq> a \<and> (b,a) \<in> r"
+ using IN1 aboveS_def[of b] by auto
+ have "\<forall>c \<in> C. (c,b) \<in> r"
+ using IN2 Above_def[of C] by auto
+ with * have "(a,b) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding AboveS_def
+ using Above_def by auto
+qed
+
+lemma (in rel) above_AboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> above b" and IN2: "b \<in> AboveS C"
+shows "a \<in> AboveS C"
+proof-
+ from IN2 have "b \<in> Above C"
+ using AboveS_subset_Above[of C] by blast
+ with assms above_Above_trans
+ have "a \<in> Above C" by auto
+ (* *)
+ moreover
+ have "a \<notin> C"
+ proof
+ assume *: "a \<in> C"
+ have 1: "(b,a) \<in> r"
+ using IN1 above_def[of b] by auto
+ have "\<forall>c \<in> C. b \<noteq> c \<and> (c,b) \<in> r"
+ using IN2 AboveS_def[of C] by auto
+ with * have "b \<noteq> a \<and> (a,b) \<in> r" by simp
+ with 1 ANTISYM antisym_def[of r]
+ show False by blast
+ qed
+ (* *)
+ ultimately
+ show ?thesis unfolding AboveS_def
+ using Above_def by auto
+qed
+
+lemma (in rel) aboveS_AboveS_trans:
+assumes TRANS: "trans r" and ANTISYM: "antisym r" and
+ IN1: "a \<in> aboveS b" and IN2: "b \<in> AboveS C"
+shows "a \<in> AboveS C"
+proof-
+ from IN2 have "b \<in> Above C"
+ using AboveS_subset_Above[of C] by blast
+ with aboveS_Above_trans assms
+ show ?thesis by auto
+qed
+
+
+subsection {* Properties depending on more than one relation *}
+
+abbreviation "under \<equiv> rel.under"
+abbreviation "underS \<equiv> rel.underS"
+abbreviation "Under \<equiv> rel.Under"
+abbreviation "UnderS \<equiv> rel.UnderS"
+abbreviation "above \<equiv> rel.above"
+abbreviation "aboveS \<equiv> rel.aboveS"
+abbreviation "Above \<equiv> rel.Above"
+abbreviation "AboveS \<equiv> rel.AboveS"
+
+lemma under_incr2:
+"r \<le> r' \<Longrightarrow> under r a \<le> under r' a"
+unfolding rel.under_def by blast
+
+lemma underS_incr2:
+"r \<le> r' \<Longrightarrow> underS r a \<le> underS r' a"
+unfolding rel.underS_def by blast
+
+lemma Under_incr:
+"r \<le> r' \<Longrightarrow> Under r A \<le> Under r A"
+unfolding rel.Under_def by blast
+
+lemma UnderS_incr:
+"r \<le> r' \<Longrightarrow> UnderS r A \<le> UnderS r A"
+unfolding rel.UnderS_def by blast
+
+lemma Under_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> Under r A \<le> Under r A'"
+unfolding rel.Under_def by blast
+
+lemma UnderS_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> UnderS r A \<le> UnderS r A'"
+unfolding rel.UnderS_def by blast
+
+lemma above_incr2:
+"r \<le> r' \<Longrightarrow> above r a \<le> above r' a"
+unfolding rel.above_def by blast
+
+lemma aboveS_incr2:
+"r \<le> r' \<Longrightarrow> aboveS r a \<le> aboveS r' a"
+unfolding rel.aboveS_def by blast
+
+lemma Above_incr:
+"r \<le> r' \<Longrightarrow> Above r A \<le> Above r A"
+unfolding rel.Above_def by blast
+
+lemma AboveS_incr:
+"r \<le> r' \<Longrightarrow> AboveS r A \<le> AboveS r A"
+unfolding rel.AboveS_def by blast
+
+lemma Above_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> Above r A \<le> Above r A'"
+unfolding rel.Above_def by blast
+
+lemma AboveS_incr_decr:
+"\<lbrakk>r \<le> r'; A' \<le> A\<rbrakk> \<Longrightarrow> AboveS r A \<le> AboveS r A'"
+unfolding rel.AboveS_def by blast
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Order_Relation_More_Base.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,304 @@
+(* 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_Id_Field:
+assumes TOT: "Total r" and NID: "\<not> (r <= Id)"
+shows "Field r = Field(r - Id)"
+using mono_Field[of "r - Id" r] Diff_subset[of r Id]
+proof(auto)
+ have "r \<noteq> {}" using NID by fast
+ then obtain b and c where "b \<noteq> c \<and> (b,c) \<in> r" using NID by fast
+ hence 1: "b \<noteq> c \<and> {b,c} \<le> Field r" by (auto simp: Field_def)
+ (* *)
+ fix a assume *: "a \<in> Field r"
+ obtain d where 2: "d \<in> Field r" and 3: "d \<noteq> a"
+ using * 1 by blast
+ hence "(a,d) \<in> r \<or> (d,a) \<in> r" using * TOT
+ by (simp add: total_on_def)
+ thus "a \<in> Field(r - Id)" using 3 unfolding Field_def by blast
+qed
+
+
+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/README.txt Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,218 @@
+Authors: Andrei Popescu & Dmitriy Traytel
+
+
+PDF Documentation:
+-----------------
+
+See "document/root.pdf".
+
+
+
+Short description of the theories' main content:
+-----------------------------------------------
+
+The "minor" theories Fun_More, Wellfounded_More and Order_Relation_More are
+extensions of the existing theories Fun, Wellfounded and
+Order_Relation:
+-- Fun_More states more facts (mainly) concerning injections, bijections,
+inverses, and (numeric) cardinals of finite sets.
+-- Wellfounded_More states variations of well-founded
+recursion and well-founded recursion.
+-- Order_Relation_More fixes a relation, defines upper and lower bounds
+operators and proves many basic properties for these
+(depending on assumptions such as reflexivity or transitivity).
+
+The "major" theories are:
+-- Wellorder_Relation: Here one fixes a well-order relation, and then:
+----- 1) Defines the concepts of maximum (of two elements), minimum (of a set), supremum,
+ successor (of a set), and order filter (i.e., downwards closed set, a.k.a.
+ initial segment).
+-- Wellorder_Embedding:
+----- 2) For two well-order relations,
+ defines *well-order embeddings* as injective functions copying
+ the source into an order filter of the target and *compatible functions*
+ as those preserving the order. Also, *isomorphisms*
+ and *strict embeddings* are defined to be embeddings that are, and respectively
+ are not, bijections.
+
+-- Constructions_on_Wellorders:
+----- 1) Defines direct images, restrictions, disjoint unions and
+ bounded squares of well-orders.
+----- 2) Defines the relations "ordLeq", "ordLess" and "ordIso"
+ between well-order relations (concrete syntax: r <=o r', r <o r' and
+ r =o r', respectively), defined by the existence of an embedding,
+ strict embedding and isomorphism, respectively between the two members.
+ Among the properties proved for these relations:
+--------- ordLeq is total;
+--------- ordLess (on a fixed type) is well-founded.
+
+-- Cardinal_Order_Relation:
+---- 1) Defines a *cardinal order* to be a well-order minim w.r.t. "ordLeq"
+ (or, equivalently, minimal w.r.t. "ordLess").
+ ordLess being well-founded together with the well-ordering theorem (from theory Zorn.thy)
+ ensures the existence of a cardinal relation on any given set. In addition,
+ a cardinal relation on a set is unique up to order isomorphism.
+---- 2) Defines the cardinal of a set A, |A|, to be SOME cardinal
+ order on it (unique up to =o, according to the above).
+---- 3) Proves properties of cardinals, including their
+ interactions with sums, products, unions, lists,
+ powersets, sets of finite sets. Among them, nontrivial
+ facts concerning the invariance of infinite cardinals
+ under some of these constructs -- e.g., if A is infinite,
+ than the cardinal of lists/words over A is the same (up to
+ the "cardinal equality" =o) as that of A.
+---- 5) Studies the connection between the introduced (order-based) notion
+ of cardinal and the numeric one previously defined for
+ finite sets (operator "card"). On the way, one introduces the cardinal omega
+ (natLeq) and the finite cardinals (natLeq_on n).
+---- 6) Defines and proves the existence of successor cardinals.
+
+-- Cardinal_Arithmetic
+
+
+Here is a list of names of proved facts concerning cardinalities that are
+expressed independently of notions of order, and of potential interest
+for "working mathematicians":
+--- one_set_greater, one_type_greater (their proofs use the
+ fact that ordinals are totally ordered)
+--- Plus_into_Times, Plus_into_Times_types,
+ Plus_infinite_bij_betw, Plus_infinite_bij_betw_types,
+ Times_same_infinite_bij_betw, Times_same_infinite_bij_betw_types,
+ Times_infinite_bij_betw, Times_infinite_bij_betw_types
+ inj_on_UNION_infinite, List_infinite_bij_betw, List_infinite_bij_betw_types
+ Fpow_infinite_bij_betw
+ (their proofs employ cardinals)
+
+
+
+
+Minor technicalities and naming issues:
+---------------------------------------
+
+1. Most of the definitions and theorems are proved in files suffixed with
+"_Base". Bootstrapping considerations (for the (co)datatype package) made this
+division desirable.
+
+
+2. Even though we would have preferred to use "initial segment" instead of
+"order filter", we chose the latter to avoid terminological clash with
+the operator "init_seg_of" from Zorn.thy. The latter expresses a related, but different
+concept -- it considers a relation, rather than a set, as initial segment of a relation.
+
+
+3. We prefer to define the upper-bound operations under, underS,
+etc., as opposed to working with combinations of relation image,
+converse and diagonal, because the former seem more intuitive
+notations when we think of orderings (but of course we cannot
+define them as abbreviations, as this would have a global
+effect, also affecting cases where one does not think of relations
+as orders). Moreover, in my locales the relation parameter r for
+under, underS etc. is fixed, hence these operations can keep r
+implicit. To get a concrete glimpse at my aesthetic reason for
+introducing these operations: otherwise, instead of "underS a",
+we would have to write "(r - Id)^-1 `` {a}" or "r^-1 `` {a} - Id".
+
+
+4. Even though the main focus of this development are
+the well-order relations, we prove the basic results on order relations
+and bounds as generally as possible.
+To the contrary, the results concerning minima, suprema and successors
+are stated for well-order relations, not maximally generally.
+
+
+5. "Refl_on A r" requires in particular that "r <= A <*> A",
+and therefore whenever "Refl_on A r", we have that necessarily
+"A = Field r". This means that in a theory of orders the domain
+A would be redundant -- we decided not to include it explicitly
+for most of the tehory.
+
+
+6. An infinite ordinal/cardinal is one for which the field is infinite.
+We always prefer the slightly more verbose "finite (Field r)" to the more
+compact but less standard equivalent condition "finite r".
+
+
+7. After we proved lots of facts about injections and
+bijections, we discovered that a couple of
+fancier (set-restricted) version of some of them are proved in
+the theory FuncSet. However, we did not need here restricted
+abstraction and such, and felt we should not import the whole
+theory for just a couple of minor facts.
+
+
+
+
+
+
+
+Notes for anyone who would like to enrich these theories in the future
+--------------------------------------------------------------------------------------
+
+Theory Fun_More (and Fun_More_Base):
+- Careful: "inj" is an abbreviation for "inj_on UNIV", while
+ "bij" is not an abreviation for "bij_betw UNIV UNIV", but
+ a defined constant; there is no "surj_betw", but only "surj".
+ "inv" is an abbreviation for "inv_into UNIV"
+- In subsection "Purely functional properties":
+-- Recall lemma "comp_inj_on".
+-- A lemma for inj_on corresponding to "bij_betw_same_card" already exists, and is called "card_inj_on_le".
+- In subsection "Properties involving Hilbert choice":
+-- One should refrain from trying to prove "intuitive" properties of f conditioned
+ by properties of (inv_into f A), such as "bij_betw A' A (inv_into f A) ==> bij_betw A A' f".
+ They usually do not hold, since one cannot usually infer the well-definedness of "inv_into f A".
+- A lemma "bij_betw_inv_into_LEFT" -- why didn't
+"proof(auto simp add: bij_betw_inv_into_left)" finish the proof?
+-- Recall lemma "bij_betw_inv_into".
+- In subsection "Other facts":
+-- Does the lemma "atLeastLessThan_injective" already exist anywhere?
+
+Theory Order_Relation_More (and Order_Relation_More_Base):
+- In subsection "Auxiliaries":
+-- Recall the lemmas "order_on_defs", "Field_def", "Domain_def", "Range_def", "converse_def".
+-- Recall that "refl_on r A" forces r to not be defined outside A.
+ This is why "partial_order_def"
+ can afford to use non-parameterized versions of antisym and trans.
+-- Recall the ASCII notation for "converse r": "r ^-1".
+-- Recall the abbreviations:
+ abbreviation "Refl r ≡ refl_on (Field r) r"
+ abbreviation "Preorder r ≡ preorder_on (Field r) r"
+ abbreviation "Partial_order r ≡ partial_order_on (Field r) r"
+ abbreviation "Total r ≡ total_on (Field r) r"
+ abbreviation "Linear_order r ≡ linear_order_on (Field r) r"
+ abbreviation "Well_order r ≡ well_order_on (Field r) r"
+
+Theory Wellorder_Relation (and Wellorder_Relation_Base):
+- In subsection "Auxiliaries": recall lemmas "order_on_defs"
+- In subsection "The notions of maximum, minimum, supremum, successor and order filter":
+ Should we define all constants from "wo_rel" in "rel" instead,
+ so that their outside definition not be conditional in "wo_rel r"?
+
+Theory Wellfounded_More (and Wellfounded_More_Base):
+ Recall the lemmas "wfrec" and "wf_induct".
+
+Theory Wellorder_Embedding (and Wellorder_Embedding_Base):
+- Recall "inj_on_def" and "bij_betw_def".
+- Line 5 in the proof of lemma embed_in_Field: we have to figure out for this and many other
+ situations: Why did it work without annotations to Refl_under_in?
+- At the proof of theorem "wellorders_totally_ordered" (and, similarly, elsewhere):
+ Had we used metavariables instead of local definitions for H, f, g and test, the
+ goals (in the goal window) would have become unreadable,
+ making impossible to debug theorem instantiations.
+- At lemma "embed_unique": If we add the attribute "rule format" at lemma, we get an error at qed.
+
+Theory Constructions_on_Wellorders (and Constructions_on_Wellorders_Base):
+- Some of the lemmas in this section are about more general kinds of relations than
+ well-orders, but it is not clear whether they are useful in such more general contexts.
+- Recall that "equiv" does not have the "equiv_on" and "Equiv" versions,
+ like the order relation. "equiv" corresponds, for instance, to "well_order_on".
+- The lemmas "ord_trans" are not clearly useful together, as their employment within blast or auto
+tends to diverge.
+
+Theory Cardinal_Order_Relation (and Cardinal_Order_Relation_Base):
+- Careful: if "|..|" meets an outer parehthesis, an extra space needs to be inserted, as in
+ "( |A| )".
+- At lemma like ordLeq_Sigma_mono1: Not worth stating something like ordLeq_Sigma_mono2 --
+ would be a mere instance of card_of_Sigma_mono2.
+- At lemma ordLeq_Sigma_cong2: Again, no reason for stating something like ordLeq_Sigma_cong2.
+- At lemma Fpow_Pow_finite: why wouldn't a version of this lemma with "... Int finite"
+also be proved by blast?
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/TODO.txt Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,59 @@
+-1. At the new Isabelle release (after 2009-2), replace the operator List
+from Cardinal_Order_Relation with "lists" from List.thy in standard library.
+A lemma for "lists" is the actual definition of "List".
+
+0. Add:
+
+lemma finite_iff_cardOf_nat:
+"finite A = ( |A| <o |UNIV :: nat set| )"
+using infinite_iff_card_of_nat[of A]
+by (auto simp add: card_of_Well_order ordLess_iff_not_ordLeq)
+
+
+
+
+1. Remember that the version from my computer has a few more theorems than the one
+from the Isabelle site.
+
+2. Many basic cardinal arithmetic facts can be listed as simps!
+The the proofs can be simplified.
+
+3. Add:
+
+lemma card_of_Plus_ordLeq_infinite:
+assumes "infinite C" and "|A| );£o |C|" and "|B| £o |C|"-A
+shows "|A <+> B| );£o |C|"-A
+proof-
+ let ?phi = "(EX a1 a2. a1 ~= a2 );Ù {a1,a2} <= A) Ù -A
+ (EX b1 b2. b1 ~= b2 );Ù {b1,b2} <= B)"-A
+ show ?thesis
+ proof(cases ?phi, auto)
+ fix a1 b1 a2 b2
+ assume "a1 );¹ a2" "b1 ¹ b2" "a1 Î A" "a2 Î A" "b1 Î B" "b2 Î B"-A
+ hence "|A <+> B| <=o |A <*> B|"
+ apply - apply(rule card_of_Plus_Times) by auto
+ moreover have "|A <*> B| );£o |C|"-A
+ using assms by (auto simp add: card_of_Times_ordLeq_infinite)
+ ultimately show ?thesis using ordLeq_transitive by blast
+ next
+ assume "-=Õa1. a1 );Î A (<_(B (-=Õa2. a1 = a2 );Ú a2 Ï A)"-A
+ then obtain a where "A <= {a}" by blast
+ {assume "A = {}"
+ hence "|A <+> B| =o |B|"
+ using ordIso_symmetric card_of_Plus_empty2 by blast
+ hence ?thesis using assms ordIso_ordLeq_trans by blast
+ }
+ moreover
+ {assume "A = {a}"
+
+ }
+ qed
+qed
+
+
+lemma card_of_Plus_ordLess_infinite:
+assumes "infinite C" and "|A| <o |C|" and "|B| <o |C|"
+shows "|A <+> B| <o |C|"
+
+
+
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Wellfounded_More.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,50 @@
+(* Title: HOL/Cardinals/Wellfounded_More.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on well-founded relations.
+*)
+
+header {* More on Well-Founded Relations *}
+
+theory Wellfounded_More
+imports Wellfounded_More_Base Order_Relation_More
+begin
+
+
+subsection {* Well-founded recursion via genuine fixpoints *}
+
+(*2*)lemma adm_wf_unique_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" and
+ f :: "'a \<Rightarrow> 'b" and g :: "'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H" and fFP: "f = H f" and gFP: "g = H g"
+shows "f = g"
+proof-
+ {fix x
+ have "f x = g x"
+ proof(rule wf_induct[of r "(\<lambda>x. f x = g x)"],
+ auto simp add: WF)
+ fix x assume "\<forall>y. (y, x) \<in> r \<longrightarrow> f y = g y"
+ hence "H f x = H g x" using ADM adm_wf_def[of r H] by auto
+ thus "f x = g x" using fFP and gFP by simp
+ qed
+ }
+ thus ?thesis by (simp add: ext)
+qed
+
+(*2*)lemma wfrec_unique_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" and
+ f :: "'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H" and
+ fp: "f = H f"
+shows "f = wfrec r H"
+proof-
+ have "H (wfrec r H) = wfrec r H"
+ using assms wfrec_fixpoint[of r H] by simp
+ thus ?thesis
+ using assms adm_wf_unique_fixpoint[of r H "wfrec r H"] by simp
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Wellfounded_More_Base.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,194 @@
+(* Title: HOL/Cardinals/Wellfounded_More_Base.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+More on well-founded relations (base).
+*)
+
+header {* More on Well-Founded Relations (Base) *}
+
+theory Wellfounded_More_Base
+imports Wellfounded Order_Relation_More_Base "~~/src/HOL/Library/Wfrec"
+begin
+
+
+text {* This section contains some variations of results in the theory
+@{text "Wellfounded.thy"}:
+\begin{itemize}
+\item means for slightly more direct definitions by well-founded recursion;
+\item variations of well-founded induction;
+\item means for proving a linear order to be a well-order.
+\end{itemize} *}
+
+
+subsection {* Well-founded recursion via genuine fixpoints *}
+
+
+(*2*)lemma wfrec_fixpoint:
+fixes r :: "('a * 'a) set" and
+ H :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+assumes WF: "wf r" and ADM: "adm_wf r H"
+shows "wfrec r H = H (wfrec r H)"
+proof(rule ext)
+ fix x
+ have "wfrec r H x = H (cut (wfrec r H) r x) x"
+ using wfrec[of r H] WF by simp
+ also
+ {have "\<And> y. (y,x) : r \<Longrightarrow> (cut (wfrec r H) r x) y = (wfrec r H) y"
+ by (auto simp add: cut_apply)
+ hence "H (cut (wfrec r H) r x) x = H (wfrec r H) x"
+ using ADM adm_wf_def[of r H] by auto
+ }
+ finally show "wfrec r H x = H (wfrec r H) x" .
+qed
+
+
+
+subsection {* Characterizations of well-founded-ness *}
+
+
+text {* A transitive relation is well-founded iff it is ``locally" well-founded,
+i.e., iff its restriction to the lower bounds of of any element is well-founded. *}
+
+(*3*)lemma trans_wf_iff:
+assumes "trans r"
+shows "wf r = (\<forall>a. wf(r Int (r^-1``{a} \<times> r^-1``{a})))"
+proof-
+ obtain R where R_def: "R = (\<lambda> a. r Int (r^-1``{a} \<times> r^-1``{a}))" by blast
+ {assume *: "wf r"
+ {fix a
+ have "wf(R a)"
+ using * R_def wf_subset[of r "R a"] by auto
+ }
+ }
+ (* *)
+ moreover
+ {assume *: "\<forall>a. wf(R a)"
+ have "wf r"
+ proof(unfold wf_def, clarify)
+ fix phi a
+ assume **: "\<forall>a. (\<forall>b. (b,a) \<in> r \<longrightarrow> phi b) \<longrightarrow> phi a"
+ obtain chi where chi_def: "chi = (\<lambda>b. (b,a) \<in> r \<longrightarrow> phi b)" by blast
+ with * have "wf (R a)" by auto
+ hence "(\<forall>b. (\<forall>c. (c,b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b) \<longrightarrow> (\<forall>b. chi b)"
+ unfolding wf_def by blast
+ moreover
+ have "\<forall>b. (\<forall>c. (c,b) \<in> R a \<longrightarrow> chi c) \<longrightarrow> chi b"
+ proof(auto simp add: chi_def R_def)
+ fix b
+ assume 1: "(b,a) \<in> r" and 2: "\<forall>c. (c, b) \<in> r \<and> (c, a) \<in> r \<longrightarrow> phi c"
+ hence "\<forall>c. (c, b) \<in> r \<longrightarrow> phi c"
+ using assms trans_def[of r] by blast
+ thus "phi b" using ** by blast
+ qed
+ ultimately have "\<forall>b. chi b" by (rule mp)
+ with ** chi_def show "phi a" by blast
+ qed
+ }
+ ultimately show ?thesis using R_def by blast
+qed
+
+
+text {* The next lemma is a variation of @{text "wf_eq_minimal"} from Wellfounded,
+allowing one to assume the set included in the field. *}
+
+(*2*)lemma wf_eq_minimal2:
+"wf r = (\<forall>A. A <= Field r \<and> A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. \<not> (a',a) \<in> r))"
+proof-
+ let ?phi = "\<lambda> A. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. \<not> (a',a) \<in> r)"
+ have "wf r = (\<forall>A. ?phi A)"
+ by (auto simp: ex_in_conv [THEN sym], erule wfE_min, assumption, blast)
+ (rule wfI_min, metis)
+ (* *)
+ also have "(\<forall>A. ?phi A) = (\<forall>B \<le> Field r. ?phi B)"
+ proof
+ assume "\<forall>A. ?phi A"
+ thus "\<forall>B \<le> Field r. ?phi B" by simp
+ next
+ assume *: "\<forall>B \<le> Field r. ?phi B"
+ show "\<forall>A. ?phi A"
+ proof(clarify)
+ fix A::"'a set" assume **: "A \<noteq> {}"
+ obtain B where B_def: "B = A Int (Field r)" by blast
+ show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r"
+ proof(cases "B = {}")
+ assume Case1: "B = {}"
+ obtain a where 1: "a \<in> A \<and> a \<notin> Field r"
+ using ** Case1 unfolding B_def by blast
+ hence "\<forall>a' \<in> A. (a',a) \<notin> r" using 1 unfolding Field_def by blast
+ thus ?thesis using 1 by blast
+ next
+ assume Case2: "B \<noteq> {}" have 1: "B \<le> Field r" unfolding B_def by blast
+ obtain a where 2: "a \<in> B \<and> (\<forall>a' \<in> B. (a',a) \<notin> r)"
+ using Case2 1 * by blast
+ have "\<forall>a' \<in> A. (a',a) \<notin> r"
+ proof(clarify)
+ fix a' assume "a' \<in> A" and **: "(a',a) \<in> r"
+ hence "a' \<in> B" unfolding B_def Field_def by blast
+ thus False using 2 ** by blast
+ qed
+ thus ?thesis using 2 unfolding B_def by blast
+ qed
+ qed
+ qed
+ finally show ?thesis by blast
+qed
+
+subsection {* Characterizations of well-founded-ness *}
+
+text {* The next lemma and its corollary enable one to prove that
+a linear order is a well-order in a way which is more standard than
+via well-founded-ness of the strict version of the relation. *}
+
+(*3*)
+lemma Linear_order_wf_diff_Id:
+assumes LI: "Linear_order r"
+shows "wf(r - Id) = (\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r))"
+proof(cases "r \<le> Id")
+ assume Case1: "r \<le> Id"
+ hence temp: "r - Id = {}" by blast
+ hence "wf(r - Id)" by (simp add: temp)
+ moreover
+ {fix A assume *: "A \<le> Field r" and **: "A \<noteq> {}"
+ obtain a where 1: "r = {} \<or> r = {(a,a)}" using LI
+ unfolding order_on_defs using Case1 rel.Total_subset_Id by metis
+ hence "A = {a} \<and> r = {(a,a)}" using * ** unfolding Field_def by blast
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r" using 1 by blast
+ }
+ ultimately show ?thesis by blast
+next
+ assume Case2: "\<not> r \<le> Id"
+ hence 1: "Field r = Field(r - Id)" using rel.Total_Id_Field LI
+ unfolding order_on_defs by blast
+ show ?thesis
+ proof
+ assume *: "wf(r - Id)"
+ show "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+ proof(clarify)
+ fix A assume **: "A \<le> Field r" and ***: "A \<noteq> {}"
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id"
+ using 1 * unfolding wf_eq_minimal2 by simp
+ moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. ((a,a') \<in> r) = ((a',a) \<notin> r - Id)"
+ using rel.Linear_order_in_diff_Id[of r] ** LI by blast
+ ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r" by blast
+ qed
+ next
+ assume *: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+ show "wf(r - Id)"
+ proof(unfold wf_eq_minimal2, clarify)
+ fix A assume **: "A \<le> Field(r - Id)" and ***: "A \<noteq> {}"
+ hence "\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r"
+ using 1 * by simp
+ moreover have "\<forall>a \<in> A. \<forall>a' \<in> A. ((a,a') \<in> r) = ((a',a) \<notin> r - Id)"
+ using rel.Linear_order_in_diff_Id[of r] ** LI mono_Field[of "r - Id" r] by blast
+ ultimately show "\<exists>a \<in> A. \<forall>a' \<in> A. (a',a) \<notin> r - Id" by blast
+ qed
+ qed
+qed
+
+(*3*)corollary Linear_order_Well_order_iff:
+assumes "Linear_order r"
+shows "Well_order r = (\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r))"
+using assms unfolding well_order_on_def using Linear_order_wf_diff_Id[of r] by blast
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Wellorder_Embedding.thy Wed Sep 12 05:29:21 2012 +0200
@@ -0,0 +1,175 @@
+(* Title: HOL/Cardinals/Wellorder_Embedding.thy
+ Author: Andrei Popescu, TU Muenchen
+ Copyright 2012
+
+Well-order embeddings.
+*)
+
+header {* Well-Order Embeddings *}
+
+theory Wellorder_Embedding
+imports Wellorder_Embedding_Base Fun_More Wellorder_Relation
+begin
+
+
+subsection {* Auxiliaries *}
+
+lemma UNION_bij_betw_ofilter:
+assumes WELL: "Well_order r" and
+ OF: "\<And> i. i \<in> I \<Longrightarrow> ofilter r (A i)" and
+ BIJ: "\<And> i. i \<in> I \<Longrightarrow> bij_betw f (A i) (A' i)"
+shows "bij_betw f (\<Union> i \<in> I. A i) (\<Union> i \<in> I. A' i)"
+proof-
+ have "wo_rel r" using WELL by (simp add: wo_rel_def)
+ hence "\<And> i j. \<lbrakk>i \<in> I; j \<in> I\<rbrakk> \<Longrightarrow> A i \<le> A j \<or> A j \<le> A i"
+ using wo_rel.ofilter_linord[of r] OF by blast
+ with WELL BIJ show ?thesis
+ by (auto simp add: bij_betw_UNION_chain)
+qed
+
+
+subsection {* (Well-order) embeddings, strict embeddings, isomorphisms and order-compatible
+functions *}
+
+lemma embed_halfcong:
+assumes EQ: "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a" and
+ EMB: "embed r r' f"
+shows "embed r r' g"
+proof(unfold embed_def, auto)
+ fix a assume *: "a \<in> Field r"
+ hence "bij_betw f (under r a) (under r' (f a))"
+ using EMB unfolding embed_def by simp
+ moreover
+ {have "under r a \<le> Field r"
+ by (auto simp add: rel.under_Field)
+ hence "\<And> b. b \<in> under r a \<Longrightarrow> f b = g b"
+ using EQ by blast
+ }
+ moreover have "f a = g a" using * EQ by auto
+ ultimately show "bij_betw g (under r a) (under r' (g a))"
+ using bij_betw_cong[of "under r a" f g "under r' (f a)"] by auto
+qed
+
+lemma embed_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "embed r r' f = embed r r' g"
+using assms embed_halfcong[of r f g r']
+ embed_halfcong[of r g f r'] by auto
+
+lemma embedS_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "embedS r r' f = embedS r r' g"
+unfolding embedS_def using assms
+embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast
+
+lemma iso_cong[fundef_cong]:
+assumes "\<And> a. a \<in> Field r \<Longrightarrow> f a = g a"
+shows "iso r r' f = iso r r' g"
+unfolding iso_def using assms
+embed_cong[of r f g r'] bij_betw_cong[of "Field r" f g "Field r'"] by blast
+
+lemma id_compat: "compat r r id"
+by(auto simp add: id_def compat_def)
+
+lemma comp_compat:
+"\<lbrakk>compat r r' f; compat r' r'' f'\<rbrakk> \<Longrightarrow> compat r r'' (f' o f)"
+by(auto simp add: comp_def compat_def)
+
+corollary one_set_greater:
+"(\<exists>f::'a \<Rightarrow> 'a'. f ` A \<le> A' \<and> inj_on f A) \<or> (\<exists>g::'a' \<Rightarrow> 'a. g ` A' \<le> A \<and> inj_on g A')"
+proof-
+ obtain r where "well_order_on A r" by (fastforce simp add: well_order_on)
+ hence 1: "A = Field r \<and> Well_order r"
+ using rel.well_order_on_Well_order by auto
+ obtain r' where 2: "well_order_on A' r'" by (fastforce simp add: well_order_on)
+ hence 2: "A' = Field r' \<and> Well_order r'"
+ using rel.well_order_on_Well_order by auto
+ hence "(\<exists>f. embed r r' f) \<or> (\<exists>g. embed r' r g)"
+ using 1 2 by (auto simp add: wellorders_totally_ordered)
+ moreover
+ {fix f assume "embed r r' f"
+ hence "f`A \<le> A' \<and> inj_on f A"
+ using 1 2 by (auto simp add: embed_Field embed_inj_on)
+ }
+ moreover
+ {fix g assume "embed r' r g"
+ hence "g`A' \<le> A \<and> inj_on g A'"
+ using 1 2 by (auto simp add: embed_Field embed_inj_on)
+ }
+ ultimately show ?thesis by blast
+qed
+
+corollary one_type_greater:
+"(\<exists>f::'a \<Rightarrow> 'a'. inj f) \<or> (\<exists>g::'a' \<Rightarrow> 'a. inj g)"
+using one_set_greater[of UNIV UNIV] by auto
+
+
+subsection {* Uniqueness of embeddings *}
+
+lemma comp_embedS:
+assumes WELL: "Well_order r" and WELL': "Well_order r'" and WELL'': "Well_order r''"
+ and EMB: "embedS r r' f" and EMB': "embedS r' r'' f'"
+shows "embedS r r'' (f' o f)"
+proof-
+ have "embed r' r'' f'" using EMB' unfolding embedS_def by simp
+ thus ?thesis using assms by (auto simp add: embedS_comp_embed)
+qed
+
+lemma iso_iff4:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "iso r r' f = (embed r r' f \<and> embed r' r (inv_into (Field r) f))"
+using assms embed_bothWays_iso
+by(unfold iso_def, auto simp add: inv_into_Field_embed_bij_betw)
+
+lemma embed_embedS_iso:
+"embed r r' f = (embedS r r' f \<or> iso r r' f)"
+unfolding embedS_def iso_def by blast
+
+lemma not_embedS_iso:
+"\<not> (embedS r r' f \<and> iso r r' f)"
+unfolding embedS_def iso_def by blast
+
+lemma embed_embedS_iff_not_iso:
+assumes "embed r r' f"
+shows "embedS r r' f = (\<not> iso r r' f)"
+using assms unfolding embedS_def iso_def by blast
+
+lemma iso_inv_into:
+assumes WELL: "Well_order r" and ISO: "iso r r' f"
+shows "iso r' r (inv_into (Field r) f)"
+using assms unfolding iso_def
+using bij_betw_inv_into inv_into_Field_embed_bij_betw by blast
+
+lemma embedS_or_iso:
+assumes WELL: "Well_order r" and WELL': "Well_order r'"
+shows "(\<exists>g. embedS r r' g) \<or> (\<exists>h. embedS r' r h) \<or> (\<exists>f. iso r r' f)"
+proof-
+ {fix f assume *: "embed r r' f"
+ {assume "bij_betw f (Field r) (Field r')"