renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
authorblanchet
Mon, 01 Sep 2014 16:34:39 +0200
changeset 58127 b7cab82f488e
parent 58126 3831312eb476
child 58128 43a1ba26a8cb
renamed '(BNF_)Constructions_on_Wellorders' to '(BNF_)Wellorder_Constructions'
src/HOL/BNF_Cardinal_Order_Relation.thy
src/HOL/BNF_Constructions_on_Wellorders.thy
src/HOL/BNF_Wellorder_Constructions.thy
src/HOL/Cardinals/Cardinal_Arithmetic.thy
src/HOL/Cardinals/Cardinal_Order_Relation.thy
src/HOL/Cardinals/Constructions_on_Wellorders.thy
src/HOL/Cardinals/Order_Union.thy
src/HOL/Cardinals/Ordinal_Arithmetic.thy
src/HOL/Cardinals/README.txt
src/HOL/Cardinals/Wellorder_Constructions.thy
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -8,7 +8,7 @@
 header {* Cardinal-Order Relations as Needed by Bounded Natural Functors *}
 
 theory BNF_Cardinal_Order_Relation
-imports Zorn BNF_Constructions_on_Wellorders
+imports Zorn BNF_Wellorder_Constructions
 begin
 
 text{* In this section, we define cardinal-order relations to be minim well-orders
--- a/src/HOL/BNF_Constructions_on_Wellorders.thy	Mon Sep 01 16:34:38 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1656 +0,0 @@
-(*  Title:      HOL/BNF_Constructions_on_Wellorders.thy
-    Author:     Andrei Popescu, TU Muenchen
-    Copyright   2012
-
-Constructions on wellorders as needed by bounded natural functors.
-*)
-
-header {* Constructions on Wellorders as Needed by Bounded Natural Functors *}
-
-theory BNF_Constructions_on_Wellorders
-imports BNF_Wellorder_Embedding
-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"
-unfolding refl_on_def Field_def by blast
-
-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 "under (Restr r A) a = under r a"
-using assms wo_rel_def
-proof(auto simp add: wo_rel.ofilter_def under_def)
-  fix b assume *: "a \<in> A" and "(b,a) \<in> r"
-  hence "b \<in> under r a \<and> a \<in> Field r"
-  unfolding 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 (under (Restr r A) a) (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 (under (Restr r A) a) (under r a)"
-   using * unfolding embed_def by auto
-   hence "under r a \<le> under (Restr r A) a"
-   unfolding bij_betw_def by auto
-   also have "\<dots> \<le> Field(Restr r A)" by (simp add: under_Field)
-   also have "\<dots> \<le> A" by (simp add: Field_Restr_subset)
-   finally have "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 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 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 = underS r a \<and> B = 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: 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
-
-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 (underS r a) <o r"
-proof-
-  have "underS r a < Field r" using assms
-  by (simp add: 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 (underS r a))"
-proof(auto)
-  fix a assume *: "a \<in> Field r" and **: "r' =o Restr r (underS r a)"
-  hence "Restr r (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: "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 (underS r a)" using 4 by auto
-  thus "\<exists>a \<in> Field r. r' =o Restr r (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 (underS r a)"
-  using ordLess_iff_ordIso_Restr by blast
-  let ?p = "Restr r (underS r a)"
-  have "wo_rel.ofilter r (underS r a)" using 0
-  by (simp add: wo_rel_def wo_rel.underS_ofilter)
-  hence "Field ?p = underS r a" using 0 Field_Restr_ofilter by blast
-  hence "Field ?p < Field r" using underS_Field2 1 by fast
-  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 (underS r a) <o r')"
-proof(auto)
-  assume *: "r \<le>o r'"
-  fix a assume "a \<in> Field r"
-  hence "Restr r (underS r a) <o r"
-  using WELL underS_Restr_ordLess[of r] by blast
-  thus "Restr r (underS r a) <o r'"
-  using * ordLess_ordLeq_trans by blast
-next
-  assume *: "\<forall>a \<in> Field r. Restr r (underS r a) <o r'"
-  {assume "r' <o r"
-   then obtain a where "a \<in> Field r \<and> r' =o Restr r (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: "\<not>finite(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 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 * ** 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) = f ` (Field r)"
-unfolding dir_image_def Field_def Range_def Domain_def by fast
-
-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'"
-    unfolding 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 SUB NE by (auto simp: dir_image_Field)
-  then obtain a where 1: "a \<in> A \<and> (\<forall>b \<in> A. (b,a) \<notin> r)"
-  using spec[OF WF[unfolded wf_eq_minimal2], of A] 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_bij_betw:
-"\<lbrakk>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_Field 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 well_order_on_Well_order by blast
-   hence 2: "iso r ?r' f"
-   using dir_image_iso using BIJ unfolding bij_betw_def by auto
-   hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast
-   hence "Field ?r' = A'"
-   using 1 BIJ unfolding bij_betw_def by auto
-   moreover have "Well_order ?r'"
-   using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast
-   ultimately show ?thesis unfolding ordIso_def using 1 2 by blast
-qed
-
-
-subsection {* Bounded square *}
-
-text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
-order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
-following criteria (in this order):
-\begin{itemize}
-\item compare the maximums;
-\item compare the first components;
-\item compare the second components.
-\end{itemize}
-%
-The only application of this construction that we are aware of is
-at proving that the square of an infinite set has the same cardinal
-as that set. The essential property required there (and which is ensured by this
-construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
-in a product of proper filters on the original relation (assumed to be a well-order). *}
-
-definition bsqr :: "'a rel => ('a * 'a)rel"
-where
-"bsqr r = {((a1,a2),(b1,b2)).
-           {a1,a2,b1,b2} \<le> Field r \<and>
-           (a1 = b1 \<and> a2 = b2 \<or>
-            (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1  \<and> (a2,b2) \<in> r - Id
-           )}"
-
-lemma Field_bsqr:
-"Field (bsqr r) = Field r \<times> Field r"
-proof
-  show "Field (bsqr r) \<le> Field r \<times> Field r"
-  proof-
-    {fix a1 a2 assume "(a1,a2) \<in> Field (bsqr r)"
-     moreover
-     have "\<And> b1 b2. ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r \<Longrightarrow>
-                      a1 \<in> Field r \<and> a2 \<in> Field r" unfolding bsqr_def by auto
-     ultimately have "a1 \<in> Field r \<and> a2 \<in> Field r" unfolding Field_def by auto
-    }
-    thus ?thesis unfolding Field_def by force
-  qed
-next
-  show "Field r \<times> Field r \<le> Field (bsqr r)"
-  proof(auto)
-    fix a1 a2 assume "a1 \<in> Field r" and "a2 \<in> Field r"
-    hence "((a1,a2),(a1,a2)) \<in> bsqr r" unfolding bsqr_def by blast
-    thus "(a1,a2) \<in> Field (bsqr r)" unfolding Field_def by auto
-  qed
-qed
-
-lemma bsqr_Refl: "Refl(bsqr r)"
-by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def)
-
-lemma bsqr_Trans:
-assumes "Well_order r"
-shows "trans (bsqr r)"
-proof(unfold trans_def, auto)
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Trans: "trans r" using wo_rel.TRANS by auto
-  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
-  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
-  (* Main proof *)
-  fix a1 a2 b1 b2 c1 c2
-  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(c1,c2)) \<in> bsqr r"
-  hence 0: "{a1,a2,b1,b2,c1,c2} \<le> Field r" unfolding bsqr_def by auto
-  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-  using * unfolding bsqr_def by auto
-  have 2: "b1 = c1 \<and> b2 = c2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
-  using ** unfolding bsqr_def by auto
-  show "((a1,a2),(c1,c2)) \<in> bsqr r"
-  proof-
-    {assume Case1: "a1 = b1 \<and> a2 = b2"
-     hence ?thesis using ** by simp
-    }
-    moreover
-    {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
-     {assume Case21: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      using Case2 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2"
-      hence ?thesis using Case2 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
-     {assume Case31: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence ?thesis using Case3 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
-      hence "(a1,c1) \<in> r - Id"
-      using Case3 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1"
-      hence ?thesis using Case3 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-     {assume Case41: "b1 = c1 \<and> b2 = c2"
-      hence ?thesis using * by simp
-     }
-     moreover
-     {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
-      hence ?thesis using Case4 0 unfolding bsqr_def by force
-     }
-     moreover
-     {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
-      hence ?thesis using Case4 0 unfolding bsqr_def by auto
-     }
-     moreover
-     {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
-      hence "(a2,c2) \<in> r - Id"
-      using Case4 TransS trans_def[of "r - Id"] by blast
-      hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    ultimately show ?thesis using 0 1 by auto
-  qed
-qed
-
-lemma bsqr_antisym:
-assumes "Well_order r"
-shows "antisym (bsqr r)"
-proof(unfold antisym_def, clarify)
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Trans: "trans r" using wo_rel.TRANS by auto
-  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
-  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
-  hence IrrS: "\<forall>a b. \<not>((a,b) \<in> r - Id \<and> (b,a) \<in> r - Id)"
-  using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast
-  (* Main proof *)
-  fix a1 a2 b1 b2
-  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(a1,a2)) \<in> bsqr r"
-  hence 0: "{a1,a2,b1,b2} \<le> Field r" unfolding bsqr_def by auto
-  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
-           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-  using * unfolding bsqr_def by auto
-  have 2: "b1 = a1 \<and> b2 = a2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> (b1,a1) \<in> r - Id \<or>
-           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> b1 = a1 \<and> (b2,a2) \<in> r - Id"
-  using ** unfolding bsqr_def by auto
-  show "a1 = b1 \<and> a2 = b2"
-  proof-
-    {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
-     {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-      hence False using Case1 IrrS by blast
-     }
-     moreover
-     {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2"
-      hence False using Case1 by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
-     {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-       hence False using Case2 by auto
-     }
-     moreover
-     {assume Case22: "(b1,a1) \<in> r - Id"
-      hence False using Case2 IrrS by blast
-     }
-     moreover
-     {assume Case23: "b1 = a1"
-      hence False using Case2 by auto
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    moreover
-    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
-     moreover
-     {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
-      hence False using Case3 by auto
-     }
-     moreover
-     {assume Case32: "(b1,a1) \<in> r - Id"
-      hence False using Case3 by auto
-     }
-     moreover
-     {assume Case33: "(b2,a2) \<in> r - Id"
-      hence False using Case3 IrrS by blast
-     }
-     ultimately have ?thesis using 0 2 by auto
-    }
-    ultimately show ?thesis using 0 1 by blast
-  qed
-qed
-
-lemma bsqr_Total:
-assumes "Well_order r"
-shows "Total(bsqr r)"
-proof-
-  (* Preliminary facts *)
-  have Well: "wo_rel r" using assms wo_rel_def by auto
-  hence Total: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
-  using wo_rel.TOTALS by auto
-  (* Main proof *)
-  {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \<le> Field(bsqr r)"
-   hence 0: "a1 \<in> Field r \<and> a2 \<in> Field r \<and> b1 \<in> Field r \<and> b2 \<in> Field r"
-   using Field_bsqr by blast
-   have "((a1,a2) = (b1,b2) \<or> ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r)"
-   proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0)
-       (* Why didn't clarsimp simp add: Well 0 do the same job? *)
-     assume Case1: "(a1,a2) \<in> r"
-     hence 1: "wo_rel.max2 r a1 a2 = a2"
-     using Well 0 by (simp add: wo_rel.max2_equals2)
-     show ?thesis
-     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
-       assume Case11: "(b1,b2) \<in> r"
-       hence 2: "wo_rel.max2 r b1 b2 = b2"
-       using Well 0 by (simp add: wo_rel.max2_equals2)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-         assume Case111: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
-       next
-         assume Case112: "a2 = b2"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case1121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto
-         next
-           assume Case1122: "a1 = b1"
-           thus ?thesis using Case112 by auto
-         qed
-       qed
-     next
-       assume Case12: "(b2,b1) \<in> r"
-       hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0)
-         assume Case121: "(a2,b1) \<in> r - Id \<or> (b1,a2) \<in> r - Id"
-         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
-       next
-         assume Case122: "a2 = b1"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case1221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto
-         next
-           assume Case1222: "a1 = b1"
-           show ?thesis
-           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-             assume Case12221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-             thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto
-           next
-             assume Case12222: "a2 = b2"
-             thus ?thesis using Case122 Case1222 by auto
-           qed
-         qed
-       qed
-     qed
-   next
-     assume Case2: "(a2,a1) \<in> r"
-     hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1)
-     show ?thesis
-     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
-       assume Case21: "(b1,b2) \<in> r"
-       hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0)
-         assume Case211: "(a1,b2) \<in> r - Id \<or> (b2,a1) \<in> r - Id"
-         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
-       next
-         assume Case212: "a1 = b2"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-           assume Case2121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-           thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto
-         next
-           assume Case2122: "a1 = b1"
-           show ?thesis
-           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-             assume Case21221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-             thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto
-           next
-             assume Case21222: "a2 = b2"
-             thus ?thesis using Case2122 Case212 by auto
-           qed
-         qed
-       qed
-     next
-       assume Case22: "(b2,b1) \<in> r"
-       hence 3: "wo_rel.max2 r b1 b2 = b1"  using Well 0 by (simp add: wo_rel.max2_equals1)
-       show ?thesis
-       proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
-         assume Case221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
-         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
-       next
-         assume Case222: "a1 = b1"
-         show ?thesis
-         proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
-           assume Case2221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
-           thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto
-         next
-           assume Case2222: "a2 = b2"
-           thus ?thesis using Case222 by auto
-         qed
-       qed
-     qed
-   qed
-  }
-  thus ?thesis unfolding total_on_def by fast
-qed
-
-lemma bsqr_Linear_order:
-assumes "Well_order r"
-shows "Linear_order(bsqr r)"
-unfolding order_on_defs
-using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast
-
-lemma bsqr_Well_order:
-assumes "Well_order r"
-shows "Well_order(bsqr r)"
-using assms
-proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI)
-  have 0: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
-  using assms well_order_on_def Linear_order_Well_order_iff by blast
-  fix D assume *: "D \<le> Field (bsqr r)" and **: "D \<noteq> {}"
-  hence 1: "D \<le> Field r \<times> Field r" unfolding Field_bsqr by simp
-  (*  *)
-  obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \<in> D}" by blast
-  have "M \<noteq> {}" using 1 M_def ** by auto
-  moreover
-  have "M \<le> Field r" unfolding M_def
-  using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
-  ultimately obtain m where m_min: "m \<in> M \<and> (\<forall>a \<in> M. (m,a) \<in> r)"
-  using 0 by blast
-  (*  *)
-  obtain A1 where A1_def: "A1 = {a1. \<exists>a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
-  have "A1 \<le> Field r" unfolding A1_def using 1 by auto
-  moreover have "A1 \<noteq> {}" unfolding A1_def using m_min unfolding M_def by blast
-  ultimately obtain a1 where a1_min: "a1 \<in> A1 \<and> (\<forall>a \<in> A1. (a1,a) \<in> r)"
-  using 0 by blast
-  (*  *)
-  obtain A2 where A2_def: "A2 = {a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
-  have "A2 \<le> Field r" unfolding A2_def using 1 by auto
-  moreover have "A2 \<noteq> {}" unfolding A2_def
-  using m_min a1_min unfolding A1_def M_def by blast
-  ultimately obtain a2 where a2_min: "a2 \<in> A2 \<and> (\<forall>a \<in> A2. (a2,a) \<in> r)"
-  using 0 by blast
-  (*   *)
-  have 2: "wo_rel.max2 r a1 a2 = m"
-  using a1_min a2_min unfolding A1_def A2_def by auto
-  have 3: "(a1,a2) \<in> D" using a2_min unfolding A2_def by auto
-  (*  *)
-  moreover
-  {fix b1 b2 assume ***: "(b1,b2) \<in> D"
-   hence 4: "{a1,a2,b1,b2} \<le> Field r" using 1 3 by blast
-   have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
-   using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto
-   have "((a1,a2),(b1,b2)) \<in> bsqr r"
-   proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2")
-     assume Case1: "wo_rel.max2 r a1 a2 \<noteq> wo_rel.max2 r b1 b2"
-     thus ?thesis unfolding bsqr_def using 4 5 by auto
-   next
-     assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
-     hence "b1 \<in> A1" unfolding A1_def using 2 *** by auto
-     hence 6: "(a1,b1) \<in> r" using a1_min by auto
-     show ?thesis
-     proof(cases "a1 = b1")
-       assume Case21: "a1 \<noteq> b1"
-       thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto
-     next
-       assume Case22: "a1 = b1"
-       hence "b2 \<in> A2" unfolding A2_def using 2 *** Case2 by auto
-       hence 7: "(a2,b2) \<in> r" using a2_min by auto
-       thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto
-     qed
-   qed
-  }
-  (*  *)
-  ultimately show "\<exists>d \<in> D. \<forall>d' \<in> D. (d,d') \<in> bsqr r" by fastforce
-qed
-
-lemma bsqr_max2:
-assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \<in> bsqr r"
-shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
-proof-
-  have "{(a1,a2),(b1,b2)} \<le> Field(bsqr r)"
-  using LEQ unfolding Field_def by auto
-  hence "{a1,a2,b1,b2} \<le> Field r" unfolding Field_bsqr by auto
-  hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \<le> Field r"
-  using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
-  moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r \<or> wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
-  using LEQ unfolding bsqr_def by auto
-  ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto
-qed
-
-lemma bsqr_ofilter:
-assumes WELL: "Well_order r" and
-        OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \<times> Field r" and
-        NE: "\<not> (\<exists>a. Field r = 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 = 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> (under r ?m) \<times> (under r ?m)"
-  proof(unfold 1)
-    {fix b1 b2
-     let ?n = "wo_rel.max2 r b1 b2"
-     assume "(b1,b2) \<in> underS ?r' (a1,a2)"
-     hence 3: "((b1,b2),(a1,a2)) \<in> ?r'"
-     unfolding 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> (under r ?m) \<times> (under r ?m)" unfolding under_def by simp}
-     thus "underS ?r' (a1,a2) \<le> (under r ?m) \<times> (under r ?m)" by auto
-  qed
-  moreover have "wo_rel.ofilter r (under r ?m)"
-  using Well by (simp add: wo_rel.under_ofilter)
-  moreover have "under r ?m < Field r"
-  using NE under_Field[of r ?m] by blast
-  ultimately show ?thesis by blast
-qed
-
-definition Func where
-"Func A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
-
-lemma Func_empty:
-"Func {} B = {\<lambda>x. undefined}"
-unfolding Func_def by auto
-
-lemma Func_elim:
-assumes "g \<in> Func A B" and "a \<in> A"
-shows "\<exists> b. b \<in> B \<and> g a = b"
-using assms unfolding Func_def by (cases "g a = undefined") auto
-
-definition curr where
-"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
-
-lemma curr_in:
-assumes f: "f \<in> Func (A <*> B) C"
-shows "curr A f \<in> Func A (Func B C)"
-using assms unfolding curr_def Func_def by auto
-
-lemma curr_inj:
-assumes "f1 \<in> Func (A <*> B) C" and "f2 \<in> Func (A <*> B) C"
-shows "curr A f1 = curr A f2 \<longleftrightarrow> f1 = f2"
-proof safe
-  assume c: "curr A f1 = curr A f2"
-  show "f1 = f2"
-  proof (rule ext, clarify)
-    fix a b show "f1 (a, b) = f2 (a, b)"
-    proof (cases "(a,b) \<in> A <*> B")
-      case False
-      thus ?thesis using assms unfolding Func_def by auto
-    next
-      case True hence a: "a \<in> A" and b: "b \<in> B" by auto
-      thus ?thesis
-      using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp
-    qed
-  qed
-qed
-
-lemma curr_surj:
-assumes "g \<in> Func A (Func B C)"
-shows "\<exists> f \<in> Func (A <*> B) C. curr A f = g"
-proof
-  let ?f = "\<lambda> ab. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
-  show "curr A ?f = g"
-  proof (rule ext)
-    fix a show "curr A ?f a = g a"
-    proof (cases "a \<in> A")
-      case False
-      hence "g a = undefined" using assms unfolding Func_def by auto
-      thus ?thesis unfolding curr_def using False by simp
-    next
-      case True
-      obtain g1 where "g1 \<in> Func B C" and "g a = g1"
-      using assms using Func_elim[OF assms True] by blast
-      thus ?thesis using True unfolding Func_def curr_def by auto
-    qed
-  qed
-  show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
-qed
-
-lemma bij_betw_curr:
-"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
-unfolding bij_betw_def inj_on_def image_def
-apply (intro impI conjI ballI)
-apply (erule curr_inj[THEN iffD1], assumption+)
-apply auto
-apply (erule curr_in)
-using curr_surj by blast
-
-definition Func_map where
-"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
-
-lemma Func_map:
-assumes g: "g \<in> Func A2 A1" and f1: "f1 ` A1 \<subseteq> B1" and f2: "f2 ` B2 \<subseteq> A2"
-shows "Func_map B2 f1 f2 g \<in> Func B2 B1"
-using assms unfolding Func_def Func_map_def mem_Collect_eq by auto
-
-lemma Func_non_emp:
-assumes "B \<noteq> {}"
-shows "Func A B \<noteq> {}"
-proof-
-  obtain b where b: "b \<in> B" using assms by auto
-  hence "(\<lambda> a. if a \<in> A then b else undefined) \<in> Func A B" unfolding Func_def by auto
-  thus ?thesis by blast
-qed
-
-lemma Func_is_emp:
-"Func A B = {} \<longleftrightarrow> A \<noteq> {} \<and> B = {}" (is "?L \<longleftrightarrow> ?R")
-proof
-  assume L: ?L
-  moreover {assume "A = {}" hence False using L Func_empty by auto}
-  moreover {assume "B \<noteq> {}" hence False using L Func_non_emp[of B A] by simp }
-  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 blast
-   with R have False by blast
-  }
-  thus ?L by blast
-qed
-
-lemma Func_map_surj:
-assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \<subseteq> A2"
-and B2A2: "B2 = {} \<Longrightarrow> A2 = {}"
-shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1"
-proof(cases "B2 = {}")
-  case True
-  thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def)
-next
-  case False note B2 = False
-  show ?thesis
-  proof safe
-    fix h assume h: "h \<in> Func B2 B1"
-    def j1 \<equiv> "inv_into A1 f1"
-    have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
-    then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> f2 (k a2) = a2"
-      by atomize_elim (rule bchoice)
-    {fix b2 assume b2: "b2 \<in> B2"
-     hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto
-     moreover have "k (f2 b2) \<in> B2" using b2 A2(2) k by auto
-     ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast
-    } note kk = this
-    obtain b22 where b22: "b22 \<in> B2" using B2 by auto
-    def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
-    have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
-    have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
-    using kk unfolding j2_def by auto
-    def g \<equiv> "Func_map A2 j1 j2 h"
-    have "Func_map B2 f1 f2 g = h"
-    proof (rule ext)
-      fix b2 show "Func_map B2 f1 f2 g b2 = h b2"
-      proof(cases "b2 \<in> B2")
-        case True
-        show ?thesis
-        proof (cases "h b2 = undefined")
-          case True
-          hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
-          show ?thesis using A2 f_inv_into_f[OF b1]
-            unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
-        qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
-          auto intro: f_inv_into_f)
-      qed(insert h, unfold Func_def Func_map_def, auto)
-    qed
-    moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
-    using j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+
-    ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
-    unfolding Func_map_def[abs_def] by auto
-  qed(insert B1 Func_map[OF _ _ A2(2)], auto)
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/BNF_Wellorder_Constructions.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -0,0 +1,1656 @@
+(*  Title:      HOL/BNF_Wellorder_Constructions.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Constructions on wellorders as needed by bounded natural functors.
+*)
+
+header {* Constructions on Wellorders as Needed by Bounded Natural Functors *}
+
+theory BNF_Wellorder_Constructions
+imports BNF_Wellorder_Embedding
+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"
+unfolding refl_on_def Field_def by blast
+
+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 "under (Restr r A) a = under r a"
+using assms wo_rel_def
+proof(auto simp add: wo_rel.ofilter_def under_def)
+  fix b assume *: "a \<in> A" and "(b,a) \<in> r"
+  hence "b \<in> under r a \<and> a \<in> Field r"
+  unfolding 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 (under (Restr r A) a) (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 (under (Restr r A) a) (under r a)"
+   using * unfolding embed_def by auto
+   hence "under r a \<le> under (Restr r A) a"
+   unfolding bij_betw_def by auto
+   also have "\<dots> \<le> Field(Restr r A)" by (simp add: under_Field)
+   also have "\<dots> \<le> A" by (simp add: Field_Restr_subset)
+   finally have "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 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 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 = underS r a \<and> B = 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: 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
+
+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 (underS r a) <o r"
+proof-
+  have "underS r a < Field r" using assms
+  by (simp add: 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 (underS r a))"
+proof(auto)
+  fix a assume *: "a \<in> Field r" and **: "r' =o Restr r (underS r a)"
+  hence "Restr r (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: "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 (underS r a)" using 4 by auto
+  thus "\<exists>a \<in> Field r. r' =o Restr r (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 (underS r a)"
+  using ordLess_iff_ordIso_Restr by blast
+  let ?p = "Restr r (underS r a)"
+  have "wo_rel.ofilter r (underS r a)" using 0
+  by (simp add: wo_rel_def wo_rel.underS_ofilter)
+  hence "Field ?p = underS r a" using 0 Field_Restr_ofilter by blast
+  hence "Field ?p < Field r" using underS_Field2 1 by fast
+  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 (underS r a) <o r')"
+proof(auto)
+  assume *: "r \<le>o r'"
+  fix a assume "a \<in> Field r"
+  hence "Restr r (underS r a) <o r"
+  using WELL underS_Restr_ordLess[of r] by blast
+  thus "Restr r (underS r a) <o r'"
+  using * ordLess_ordLeq_trans by blast
+next
+  assume *: "\<forall>a \<in> Field r. Restr r (underS r a) <o r'"
+  {assume "r' <o r"
+   then obtain a where "a \<in> Field r \<and> r' =o Restr r (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: "\<not>finite(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 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 * ** 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) = f ` (Field r)"
+unfolding dir_image_def Field_def Range_def Domain_def by fast
+
+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'"
+    unfolding 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 SUB NE by (auto simp: dir_image_Field)
+  then obtain a where 1: "a \<in> A \<and> (\<forall>b \<in> A. (b,a) \<notin> r)"
+  using spec[OF WF[unfolded wf_eq_minimal2], of A] 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_bij_betw:
+"\<lbrakk>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_Field 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 well_order_on_Well_order by blast
+   hence 2: "iso r ?r' f"
+   using dir_image_iso using BIJ unfolding bij_betw_def by auto
+   hence "f ` (Field r) = Field ?r'" using 1 iso_iff[of r ?r'] by blast
+   hence "Field ?r' = A'"
+   using 1 BIJ unfolding bij_betw_def by auto
+   moreover have "Well_order ?r'"
+   using 1 Well_order_dir_image BIJ unfolding bij_betw_def by blast
+   ultimately show ?thesis unfolding ordIso_def using 1 2 by blast
+qed
+
+
+subsection {* Bounded square *}
+
+text{* This construction essentially defines, for an order relation @{text "r"}, a lexicographic
+order @{text "bsqr r"} on @{text "(Field r) \<times> (Field r)"}, applying the
+following criteria (in this order):
+\begin{itemize}
+\item compare the maximums;
+\item compare the first components;
+\item compare the second components.
+\end{itemize}
+%
+The only application of this construction that we are aware of is
+at proving that the square of an infinite set has the same cardinal
+as that set. The essential property required there (and which is ensured by this
+construction) is that any proper order filter of the product order is included in a rectangle, i.e.,
+in a product of proper filters on the original relation (assumed to be a well-order). *}
+
+definition bsqr :: "'a rel => ('a * 'a)rel"
+where
+"bsqr r = {((a1,a2),(b1,b2)).
+           {a1,a2,b1,b2} \<le> Field r \<and>
+           (a1 = b1 \<and> a2 = b2 \<or>
+            (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+            wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1  \<and> (a2,b2) \<in> r - Id
+           )}"
+
+lemma Field_bsqr:
+"Field (bsqr r) = Field r \<times> Field r"
+proof
+  show "Field (bsqr r) \<le> Field r \<times> Field r"
+  proof-
+    {fix a1 a2 assume "(a1,a2) \<in> Field (bsqr r)"
+     moreover
+     have "\<And> b1 b2. ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r \<Longrightarrow>
+                      a1 \<in> Field r \<and> a2 \<in> Field r" unfolding bsqr_def by auto
+     ultimately have "a1 \<in> Field r \<and> a2 \<in> Field r" unfolding Field_def by auto
+    }
+    thus ?thesis unfolding Field_def by force
+  qed
+next
+  show "Field r \<times> Field r \<le> Field (bsqr r)"
+  proof(auto)
+    fix a1 a2 assume "a1 \<in> Field r" and "a2 \<in> Field r"
+    hence "((a1,a2),(a1,a2)) \<in> bsqr r" unfolding bsqr_def by blast
+    thus "(a1,a2) \<in> Field (bsqr r)" unfolding Field_def by auto
+  qed
+qed
+
+lemma bsqr_Refl: "Refl(bsqr r)"
+by(unfold refl_on_def Field_bsqr, auto simp add: bsqr_def)
+
+lemma bsqr_Trans:
+assumes "Well_order r"
+shows "trans (bsqr r)"
+proof(unfold trans_def, auto)
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Trans: "trans r" using wo_rel.TRANS by auto
+  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+  (* Main proof *)
+  fix a1 a2 b1 b2 c1 c2
+  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(c1,c2)) \<in> bsqr r"
+  hence 0: "{a1,a2,b1,b2,c1,c2} \<le> Field r" unfolding bsqr_def by auto
+  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+  using * unfolding bsqr_def by auto
+  have 2: "b1 = c1 \<and> b2 = c2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+  using ** unfolding bsqr_def by auto
+  show "((a1,a2),(c1,c2)) \<in> bsqr r"
+  proof-
+    {assume Case1: "a1 = b1 \<and> a2 = b2"
+     hence ?thesis using ** by simp
+    }
+    moreover
+    {assume Case2: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+     {assume Case21: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case22: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence "(wo_rel.max2 r a1 a2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      using Case2 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case23_4: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2"
+      hence ?thesis using Case2 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+     {assume Case31: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case32: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence ?thesis using Case3 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+      hence "(a1,c1) \<in> r - Id"
+      using Case3 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using Case3 Case33 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case33: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1"
+      hence ?thesis using Case3 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case4: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+     {assume Case41: "b1 = c1 \<and> b2 = c2"
+      hence ?thesis using * by simp
+     }
+     moreover
+     {assume Case42: "(wo_rel.max2 r b1 b2, wo_rel.max2 r c1 c2) \<in> r - Id"
+      hence ?thesis using Case4 0 unfolding bsqr_def by force
+     }
+     moreover
+     {assume Case43: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> (b1,c1) \<in> r - Id"
+      hence ?thesis using Case4 0 unfolding bsqr_def by auto
+     }
+     moreover
+     {assume Case44: "wo_rel.max2 r b1 b2 = wo_rel.max2 r c1 c2 \<and> b1 = c1 \<and> (b2,c2) \<in> r - Id"
+      hence "(a2,c2) \<in> r - Id"
+      using Case4 TransS trans_def[of "r - Id"] by blast
+      hence ?thesis using Case4 Case44 0 unfolding bsqr_def by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    ultimately show ?thesis using 0 1 by auto
+  qed
+qed
+
+lemma bsqr_antisym:
+assumes "Well_order r"
+shows "antisym (bsqr r)"
+proof(unfold antisym_def, clarify)
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Trans: "trans r" using wo_rel.TRANS by auto
+  have Anti: "antisym r" using wo_rel.ANTISYM Well by auto
+  hence TransS: "trans(r - Id)" using Trans by (simp add: trans_diff_Id)
+  hence IrrS: "\<forall>a b. \<not>((a,b) \<in> r - Id \<and> (b,a) \<in> r - Id)"
+  using Anti trans_def[of "r - Id"] antisym_def[of "r - Id"] by blast
+  (* Main proof *)
+  fix a1 a2 b1 b2
+  assume *: "((a1,a2),(b1,b2)) \<in> bsqr r" and **: "((b1,b2),(a1,a2)) \<in> bsqr r"
+  hence 0: "{a1,a2,b1,b2} \<le> Field r" unfolding bsqr_def by auto
+  have 1: "a1 = b1 \<and> a2 = b2 \<or> (wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id \<or>
+           wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+  using * unfolding bsqr_def by auto
+  have 2: "b1 = a1 \<and> b2 = a2 \<or> (wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> (b1,a1) \<in> r - Id \<or>
+           wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2 \<and> b1 = a1 \<and> (b2,a2) \<in> r - Id"
+  using ** unfolding bsqr_def by auto
+  show "a1 = b1 \<and> a2 = b2"
+  proof-
+    {assume Case1: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r - Id"
+     {assume Case11: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+      hence False using Case1 IrrS by blast
+     }
+     moreover
+     {assume Case12_3: "wo_rel.max2 r b1 b2 = wo_rel.max2 r a1 a2"
+      hence False using Case1 by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> (a1,b1) \<in> r - Id"
+     {assume Case21: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+       hence False using Case2 by auto
+     }
+     moreover
+     {assume Case22: "(b1,a1) \<in> r - Id"
+      hence False using Case2 IrrS by blast
+     }
+     moreover
+     {assume Case23: "b1 = a1"
+      hence False using Case2 by auto
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    moreover
+    {assume Case3: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2 \<and> a1 = b1 \<and> (a2,b2) \<in> r - Id"
+     moreover
+     {assume Case31: "(wo_rel.max2 r b1 b2, wo_rel.max2 r a1 a2) \<in> r - Id"
+      hence False using Case3 by auto
+     }
+     moreover
+     {assume Case32: "(b1,a1) \<in> r - Id"
+      hence False using Case3 by auto
+     }
+     moreover
+     {assume Case33: "(b2,a2) \<in> r - Id"
+      hence False using Case3 IrrS by blast
+     }
+     ultimately have ?thesis using 0 2 by auto
+    }
+    ultimately show ?thesis using 0 1 by blast
+  qed
+qed
+
+lemma bsqr_Total:
+assumes "Well_order r"
+shows "Total(bsqr r)"
+proof-
+  (* Preliminary facts *)
+  have Well: "wo_rel r" using assms wo_rel_def by auto
+  hence Total: "\<forall>a \<in> Field r. \<forall>b \<in> Field r. (a,b) \<in> r \<or> (b,a) \<in> r"
+  using wo_rel.TOTALS by auto
+  (* Main proof *)
+  {fix a1 a2 b1 b2 assume "{(a1,a2), (b1,b2)} \<le> Field(bsqr r)"
+   hence 0: "a1 \<in> Field r \<and> a2 \<in> Field r \<and> b1 \<in> Field r \<and> b2 \<in> Field r"
+   using Field_bsqr by blast
+   have "((a1,a2) = (b1,b2) \<or> ((a1,a2),(b1,b2)) \<in> bsqr r \<or> ((b1,b2),(a1,a2)) \<in> bsqr r)"
+   proof(rule wo_rel.cases_Total[of r a1 a2], clarsimp simp add: Well, simp add: 0)
+       (* Why didn't clarsimp simp add: Well 0 do the same job? *)
+     assume Case1: "(a1,a2) \<in> r"
+     hence 1: "wo_rel.max2 r a1 a2 = a2"
+     using Well 0 by (simp add: wo_rel.max2_equals2)
+     show ?thesis
+     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+       assume Case11: "(b1,b2) \<in> r"
+       hence 2: "wo_rel.max2 r b1 b2 = b2"
+       using Well 0 by (simp add: wo_rel.max2_equals2)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+         assume Case111: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+       next
+         assume Case112: "a2 = b2"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case1121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 2 Case112 unfolding bsqr_def by auto
+         next
+           assume Case1122: "a1 = b1"
+           thus ?thesis using Case112 by auto
+         qed
+       qed
+     next
+       assume Case12: "(b2,b1) \<in> r"
+       hence 3: "wo_rel.max2 r b1 b2 = b1" using Well 0 by (simp add: wo_rel.max2_equals1)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a2 b1], clarsimp simp add: Well, simp add: 0)
+         assume Case121: "(a2,b1) \<in> r - Id \<or> (b1,a2) \<in> r - Id"
+         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+       next
+         assume Case122: "a2 = b1"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case1221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 3 Case122 unfolding bsqr_def by auto
+         next
+           assume Case1222: "a1 = b1"
+           show ?thesis
+           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+             assume Case12221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+             thus ?thesis using 0 1 3 Case122 Case1222 unfolding bsqr_def by auto
+           next
+             assume Case12222: "a2 = b2"
+             thus ?thesis using Case122 Case1222 by auto
+           qed
+         qed
+       qed
+     qed
+   next
+     assume Case2: "(a2,a1) \<in> r"
+     hence 1: "wo_rel.max2 r a1 a2 = a1" using Well 0 by (simp add: wo_rel.max2_equals1)
+     show ?thesis
+     proof(rule wo_rel.cases_Total[of r b1 b2], clarsimp simp add: Well, simp add: 0)
+       assume Case21: "(b1,b2) \<in> r"
+       hence 2: "wo_rel.max2 r b1 b2 = b2" using Well 0 by (simp add: wo_rel.max2_equals2)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a1 b2], clarsimp simp add: Well, simp add: 0)
+         assume Case211: "(a1,b2) \<in> r - Id \<or> (b2,a1) \<in> r - Id"
+         thus ?thesis using 0 1 2 unfolding bsqr_def by auto
+       next
+         assume Case212: "a1 = b2"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+           assume Case2121: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+           thus ?thesis using 0 1 2 Case212 unfolding bsqr_def by auto
+         next
+           assume Case2122: "a1 = b1"
+           show ?thesis
+           proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+             assume Case21221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+             thus ?thesis using 0 1 2 Case212 Case2122 unfolding bsqr_def by auto
+           next
+             assume Case21222: "a2 = b2"
+             thus ?thesis using Case2122 Case212 by auto
+           qed
+         qed
+       qed
+     next
+       assume Case22: "(b2,b1) \<in> r"
+       hence 3: "wo_rel.max2 r b1 b2 = b1"  using Well 0 by (simp add: wo_rel.max2_equals1)
+       show ?thesis
+       proof(rule wo_rel.cases_Total3[of r a1 b1], clarsimp simp add: Well, simp add: 0)
+         assume Case221: "(a1,b1) \<in> r - Id \<or> (b1,a1) \<in> r - Id"
+         thus ?thesis using 0 1 3 unfolding bsqr_def by auto
+       next
+         assume Case222: "a1 = b1"
+         show ?thesis
+         proof(rule wo_rel.cases_Total3[of r a2 b2], clarsimp simp add: Well, simp add: 0)
+           assume Case2221: "(a2,b2) \<in> r - Id \<or> (b2,a2) \<in> r - Id"
+           thus ?thesis using 0 1 3 Case222 unfolding bsqr_def by auto
+         next
+           assume Case2222: "a2 = b2"
+           thus ?thesis using Case222 by auto
+         qed
+       qed
+     qed
+   qed
+  }
+  thus ?thesis unfolding total_on_def by fast
+qed
+
+lemma bsqr_Linear_order:
+assumes "Well_order r"
+shows "Linear_order(bsqr r)"
+unfolding order_on_defs
+using assms bsqr_Refl bsqr_Trans bsqr_antisym bsqr_Total by blast
+
+lemma bsqr_Well_order:
+assumes "Well_order r"
+shows "Well_order(bsqr r)"
+using assms
+proof(simp add: bsqr_Linear_order Linear_order_Well_order_iff, intro allI impI)
+  have 0: "\<forall>A \<le> Field r. A \<noteq> {} \<longrightarrow> (\<exists>a \<in> A. \<forall>a' \<in> A. (a,a') \<in> r)"
+  using assms well_order_on_def Linear_order_Well_order_iff by blast
+  fix D assume *: "D \<le> Field (bsqr r)" and **: "D \<noteq> {}"
+  hence 1: "D \<le> Field r \<times> Field r" unfolding Field_bsqr by simp
+  (*  *)
+  obtain M where M_def: "M = {wo_rel.max2 r a1 a2| a1 a2. (a1,a2) \<in> D}" by blast
+  have "M \<noteq> {}" using 1 M_def ** by auto
+  moreover
+  have "M \<le> Field r" unfolding M_def
+  using 1 assms wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+  ultimately obtain m where m_min: "m \<in> M \<and> (\<forall>a \<in> M. (m,a) \<in> r)"
+  using 0 by blast
+  (*  *)
+  obtain A1 where A1_def: "A1 = {a1. \<exists>a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+  have "A1 \<le> Field r" unfolding A1_def using 1 by auto
+  moreover have "A1 \<noteq> {}" unfolding A1_def using m_min unfolding M_def by blast
+  ultimately obtain a1 where a1_min: "a1 \<in> A1 \<and> (\<forall>a \<in> A1. (a1,a) \<in> r)"
+  using 0 by blast
+  (*  *)
+  obtain A2 where A2_def: "A2 = {a2. (a1,a2) \<in> D \<and> wo_rel.max2 r a1 a2 = m}" by blast
+  have "A2 \<le> Field r" unfolding A2_def using 1 by auto
+  moreover have "A2 \<noteq> {}" unfolding A2_def
+  using m_min a1_min unfolding A1_def M_def by blast
+  ultimately obtain a2 where a2_min: "a2 \<in> A2 \<and> (\<forall>a \<in> A2. (a2,a) \<in> r)"
+  using 0 by blast
+  (*   *)
+  have 2: "wo_rel.max2 r a1 a2 = m"
+  using a1_min a2_min unfolding A1_def A2_def by auto
+  have 3: "(a1,a2) \<in> D" using a2_min unfolding A2_def by auto
+  (*  *)
+  moreover
+  {fix b1 b2 assume ***: "(b1,b2) \<in> D"
+   hence 4: "{a1,a2,b1,b2} \<le> Field r" using 1 3 by blast
+   have 5: "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+   using *** a1_min a2_min m_min unfolding A1_def A2_def M_def by auto
+   have "((a1,a2),(b1,b2)) \<in> bsqr r"
+   proof(cases "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2")
+     assume Case1: "wo_rel.max2 r a1 a2 \<noteq> wo_rel.max2 r b1 b2"
+     thus ?thesis unfolding bsqr_def using 4 5 by auto
+   next
+     assume Case2: "wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+     hence "b1 \<in> A1" unfolding A1_def using 2 *** by auto
+     hence 6: "(a1,b1) \<in> r" using a1_min by auto
+     show ?thesis
+     proof(cases "a1 = b1")
+       assume Case21: "a1 \<noteq> b1"
+       thus ?thesis unfolding bsqr_def using 4 Case2 6 by auto
+     next
+       assume Case22: "a1 = b1"
+       hence "b2 \<in> A2" unfolding A2_def using 2 *** Case2 by auto
+       hence 7: "(a2,b2) \<in> r" using a2_min by auto
+       thus ?thesis unfolding bsqr_def using 4 7 Case2 Case22 by auto
+     qed
+   qed
+  }
+  (*  *)
+  ultimately show "\<exists>d \<in> D. \<forall>d' \<in> D. (d,d') \<in> bsqr r" by fastforce
+qed
+
+lemma bsqr_max2:
+assumes WELL: "Well_order r" and LEQ: "((a1,a2),(b1,b2)) \<in> bsqr r"
+shows "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r"
+proof-
+  have "{(a1,a2),(b1,b2)} \<le> Field(bsqr r)"
+  using LEQ unfolding Field_def by auto
+  hence "{a1,a2,b1,b2} \<le> Field r" unfolding Field_bsqr by auto
+  hence "{wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2} \<le> Field r"
+  using WELL wo_rel_def[of r] wo_rel.max2_among[of r] by fastforce
+  moreover have "(wo_rel.max2 r a1 a2, wo_rel.max2 r b1 b2) \<in> r \<or> wo_rel.max2 r a1 a2 = wo_rel.max2 r b1 b2"
+  using LEQ unfolding bsqr_def by auto
+  ultimately show ?thesis using WELL unfolding order_on_defs refl_on_def by auto
+qed
+
+lemma bsqr_ofilter:
+assumes WELL: "Well_order r" and
+        OF: "wo_rel.ofilter (bsqr r) D" and SUB: "D < Field r \<times> Field r" and
+        NE: "\<not> (\<exists>a. Field r = 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 = 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> (under r ?m) \<times> (under r ?m)"
+  proof(unfold 1)
+    {fix b1 b2
+     let ?n = "wo_rel.max2 r b1 b2"
+     assume "(b1,b2) \<in> underS ?r' (a1,a2)"
+     hence 3: "((b1,b2),(a1,a2)) \<in> ?r'"
+     unfolding 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> (under r ?m) \<times> (under r ?m)" unfolding under_def by simp}
+     thus "underS ?r' (a1,a2) \<le> (under r ?m) \<times> (under r ?m)" by auto
+  qed
+  moreover have "wo_rel.ofilter r (under r ?m)"
+  using Well by (simp add: wo_rel.under_ofilter)
+  moreover have "under r ?m < Field r"
+  using NE under_Field[of r ?m] by blast
+  ultimately show ?thesis by blast
+qed
+
+definition Func where
+"Func A B = {f . (\<forall> a \<in> A. f a \<in> B) \<and> (\<forall> a. a \<notin> A \<longrightarrow> f a = undefined)}"
+
+lemma Func_empty:
+"Func {} B = {\<lambda>x. undefined}"
+unfolding Func_def by auto
+
+lemma Func_elim:
+assumes "g \<in> Func A B" and "a \<in> A"
+shows "\<exists> b. b \<in> B \<and> g a = b"
+using assms unfolding Func_def by (cases "g a = undefined") auto
+
+definition curr where
+"curr A f \<equiv> \<lambda> a. if a \<in> A then \<lambda>b. f (a,b) else undefined"
+
+lemma curr_in:
+assumes f: "f \<in> Func (A <*> B) C"
+shows "curr A f \<in> Func A (Func B C)"
+using assms unfolding curr_def Func_def by auto
+
+lemma curr_inj:
+assumes "f1 \<in> Func (A <*> B) C" and "f2 \<in> Func (A <*> B) C"
+shows "curr A f1 = curr A f2 \<longleftrightarrow> f1 = f2"
+proof safe
+  assume c: "curr A f1 = curr A f2"
+  show "f1 = f2"
+  proof (rule ext, clarify)
+    fix a b show "f1 (a, b) = f2 (a, b)"
+    proof (cases "(a,b) \<in> A <*> B")
+      case False
+      thus ?thesis using assms unfolding Func_def by auto
+    next
+      case True hence a: "a \<in> A" and b: "b \<in> B" by auto
+      thus ?thesis
+      using c unfolding curr_def fun_eq_iff by(elim allE[of _ a]) simp
+    qed
+  qed
+qed
+
+lemma curr_surj:
+assumes "g \<in> Func A (Func B C)"
+shows "\<exists> f \<in> Func (A <*> B) C. curr A f = g"
+proof
+  let ?f = "\<lambda> ab. if fst ab \<in> A \<and> snd ab \<in> B then g (fst ab) (snd ab) else undefined"
+  show "curr A ?f = g"
+  proof (rule ext)
+    fix a show "curr A ?f a = g a"
+    proof (cases "a \<in> A")
+      case False
+      hence "g a = undefined" using assms unfolding Func_def by auto
+      thus ?thesis unfolding curr_def using False by simp
+    next
+      case True
+      obtain g1 where "g1 \<in> Func B C" and "g a = g1"
+      using assms using Func_elim[OF assms True] by blast
+      thus ?thesis using True unfolding Func_def curr_def by auto
+    qed
+  qed
+  show "?f \<in> Func (A <*> B) C" using assms unfolding Func_def mem_Collect_eq by auto
+qed
+
+lemma bij_betw_curr:
+"bij_betw (curr A) (Func (A <*> B) C) (Func A (Func B C))"
+unfolding bij_betw_def inj_on_def image_def
+apply (intro impI conjI ballI)
+apply (erule curr_inj[THEN iffD1], assumption+)
+apply auto
+apply (erule curr_in)
+using curr_surj by blast
+
+definition Func_map where
+"Func_map B2 f1 f2 g b2 \<equiv> if b2 \<in> B2 then f1 (g (f2 b2)) else undefined"
+
+lemma Func_map:
+assumes g: "g \<in> Func A2 A1" and f1: "f1 ` A1 \<subseteq> B1" and f2: "f2 ` B2 \<subseteq> A2"
+shows "Func_map B2 f1 f2 g \<in> Func B2 B1"
+using assms unfolding Func_def Func_map_def mem_Collect_eq by auto
+
+lemma Func_non_emp:
+assumes "B \<noteq> {}"
+shows "Func A B \<noteq> {}"
+proof-
+  obtain b where b: "b \<in> B" using assms by auto
+  hence "(\<lambda> a. if a \<in> A then b else undefined) \<in> Func A B" unfolding Func_def by auto
+  thus ?thesis by blast
+qed
+
+lemma Func_is_emp:
+"Func A B = {} \<longleftrightarrow> A \<noteq> {} \<and> B = {}" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume L: ?L
+  moreover {assume "A = {}" hence False using L Func_empty by auto}
+  moreover {assume "B \<noteq> {}" hence False using L Func_non_emp[of B A] by simp }
+  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 blast
+   with R have False by blast
+  }
+  thus ?L by blast
+qed
+
+lemma Func_map_surj:
+assumes B1: "f1 ` A1 = B1" and A2: "inj_on f2 B2" "f2 ` B2 \<subseteq> A2"
+and B2A2: "B2 = {} \<Longrightarrow> A2 = {}"
+shows "Func B2 B1 = Func_map B2 f1 f2 ` Func A2 A1"
+proof(cases "B2 = {}")
+  case True
+  thus ?thesis using B2A2 by (auto simp: Func_empty Func_map_def)
+next
+  case False note B2 = False
+  show ?thesis
+  proof safe
+    fix h assume h: "h \<in> Func B2 B1"
+    def j1 \<equiv> "inv_into A1 f1"
+    have "\<forall> a2 \<in> f2 ` B2. \<exists> b2. b2 \<in> B2 \<and> f2 b2 = a2" by blast
+    then obtain k where k: "\<forall> a2 \<in> f2 ` B2. k a2 \<in> B2 \<and> f2 (k a2) = a2"
+      by atomize_elim (rule bchoice)
+    {fix b2 assume b2: "b2 \<in> B2"
+     hence "f2 (k (f2 b2)) = f2 b2" using k A2(2) by auto
+     moreover have "k (f2 b2) \<in> B2" using b2 A2(2) k by auto
+     ultimately have "k (f2 b2) = b2" using b2 A2(1) unfolding inj_on_def by blast
+    } note kk = this
+    obtain b22 where b22: "b22 \<in> B2" using B2 by auto
+    def j2 \<equiv> "\<lambda> a2. if a2 \<in> f2 ` B2 then k a2 else b22"
+    have j2A2: "j2 ` A2 \<subseteq> B2" unfolding j2_def using k b22 by auto
+    have j2: "\<And> b2. b2 \<in> B2 \<Longrightarrow> j2 (f2 b2) = b2"
+    using kk unfolding j2_def by auto
+    def g \<equiv> "Func_map A2 j1 j2 h"
+    have "Func_map B2 f1 f2 g = h"
+    proof (rule ext)
+      fix b2 show "Func_map B2 f1 f2 g b2 = h b2"
+      proof(cases "b2 \<in> B2")
+        case True
+        show ?thesis
+        proof (cases "h b2 = undefined")
+          case True
+          hence b1: "h b2 \<in> f1 ` A1" using h `b2 \<in> B2` unfolding B1 Func_def by auto
+          show ?thesis using A2 f_inv_into_f[OF b1]
+            unfolding True g_def Func_map_def j1_def j2[OF `b2 \<in> B2`] by auto
+        qed(insert A2 True j2[OF True] h B1, unfold j1_def g_def Func_def Func_map_def,
+          auto intro: f_inv_into_f)
+      qed(insert h, unfold Func_def Func_map_def, auto)
+    qed
+    moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
+    using j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+
+    ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
+    unfolding Func_map_def[abs_def] by auto
+  qed(insert B1 Func_map[OF _ _ A2(2)], auto)
+qed
+
+end
--- a/src/HOL/Cardinals/Cardinal_Arithmetic.thy	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/Cardinals/Cardinal_Arithmetic.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -219,7 +219,7 @@
 lemma czero_cexp: "Cnotzero r \<Longrightarrow> czero ^c r =o czero"
   by (drule Cnotzero_imp_not_empty) (simp add: cexp_def czero_def card_of_empty_ordIso)
 
-lemma Func_singleton: 
+lemma Func_singleton:
 fixes x :: 'b and A :: "'a set"
 shows "|Func A {x}| =o |{x}|"
 proof (rule ordIso_symmetric)
--- a/src/HOL/Cardinals/Cardinal_Order_Relation.thy	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/Cardinals/Cardinal_Order_Relation.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -8,7 +8,7 @@
 header {* Cardinal-Order Relations *}
 
 theory Cardinal_Order_Relation
-imports BNF_Cardinal_Order_Relation Constructions_on_Wellorders
+imports BNF_Cardinal_Order_Relation Wellorder_Constructions
 begin
 
 declare
--- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy	Mon Sep 01 16:34:38 2014 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1077 +0,0 @@
-(*  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
-  BNF_Constructions_on_Wellorders Wellorder_Embedding Order_Union
-  "../Library/Cardinal_Notations"
-begin
-
-declare
-  ordLeq_Well_order_simp[simp]
-  not_ordLeq_iff_ordLess[simp]
-  not_ordLess_iff_ordLeq[simp]
-  Func_empty[simp]
-  Func_is_emp[simp]
-
-lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
-
-
-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 Restr_subset by (elim wf_subset) simp
-
-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_Well_order_simp[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[simp]:
-assumes "r =o r'"
-shows "Well_order r \<and> Well_order r'"
-using assms unfolding ordIso_def by simp
-
-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
-
-(* More facts on ordinal sum: *)
-
-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 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 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 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
-
-
-subsection {* Limit and succesor ordinals *}
-
-lemma embed_underS2:
-assumes r: "Well_order r" and s: "Well_order s"  and g: "embed r s g" and a: "a \<in> Field r"
-shows "g ` underS r a = underS s (g a)"
-using embed_underS[OF assms] unfolding bij_betw_def by auto
-
-lemma bij_betw_insert:
-assumes "b \<notin> A" and "f b \<notin> A'" and "bij_betw f A A'"
-shows "bij_betw f (insert b A) (insert (f b) A')"
-using notIn_Un_bij_betw[OF assms] by auto
-
-context wo_rel
-begin
-
-lemma underS_induct:
-  assumes "\<And>a. (\<And> a1. a1 \<in> underS a \<Longrightarrow> P a1) \<Longrightarrow> P a"
-  shows "P a"
-  by (induct rule: well_order_induct) (rule assms, simp add: underS_def)
-
-lemma suc_underS:
-assumes B: "B \<subseteq> Field r" and A: "AboveS B \<noteq> {}" and b: "b \<in> B"
-shows "b \<in> underS (suc B)"
-using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto
-
-lemma underS_supr:
-assumes bA: "b \<in> underS (supr A)" and A: "A \<subseteq> Field r"
-shows "\<exists> a \<in> A. b \<in> underS a"
-proof(rule ccontr, auto)
-  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
-  assume "\<forall>a\<in>A.  b \<notin> underS a"
-  hence 0: "\<forall>a \<in> A. (a,b) \<in> r" using A bA unfolding underS_def
-  by simp (metis REFL in_mono max2_def max2_greater refl_on_domain)
-  have "(supr A, b) \<in> r" apply(rule supr_least[OF A bb]) using 0 by auto
-  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
-qed
-
-lemma underS_suc:
-assumes bA: "b \<in> underS (suc A)" and A: "A \<subseteq> Field r"
-shows "\<exists> a \<in> A. b \<in> under a"
-proof(rule ccontr, auto)
-  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
-  assume "\<forall>a\<in>A.  b \<notin> under a"
-  hence 0: "\<forall>a \<in> A. a \<in> underS b" using A bA unfolding underS_def
-  by simp (metis (lifting) bb max2_def max2_greater mem_Collect_eq under_def set_rev_mp)
-  have "(suc A, b) \<in> r" apply(rule suc_least[OF A bb]) using 0 unfolding underS_def by auto
-  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
-qed
-
-lemma (in wo_rel) in_underS_supr:
-assumes j: "j \<in> underS i" and i: "i \<in> A" and A: "A \<subseteq> Field r" and AA: "Above A \<noteq> {}"
-shows "j \<in> underS (supr A)"
-proof-
-  have "(i,supr A) \<in> r" using supr_greater[OF A AA i] .
-  thus ?thesis using j unfolding underS_def
-  by simp (metis REFL TRANS max2_def max2_equals1 refl_on_domain transD)
-qed
-
-lemma inj_on_Field:
-assumes A: "A \<subseteq> Field r" and f: "\<And> a b. \<lbrakk>a \<in> A; b \<in> A; a \<in> underS b\<rbrakk> \<Longrightarrow> f a \<noteq> f b"
-shows "inj_on f A"
-unfolding inj_on_def proof safe
-  fix a b assume a: "a \<in> A" and b: "b \<in> A" and fab: "f a = f b"
-  {assume "a \<in> underS b"
-   hence False using f[OF a b] fab by auto
-  }
-  moreover
-  {assume "b \<in> underS a"
-   hence False using f[OF b a] fab by auto
-  }
-  ultimately show "a = b" using TOTALS A a b unfolding underS_def by auto
-qed
-
-lemma in_notinI:
-assumes "(j,i) \<notin> r \<or> j = i" and "i \<in> Field r" and "j \<in> Field r"
-shows "(i,j) \<in> r" by (metis assms max2_def max2_greater_among)
-
-lemma ofilter_init_seg_of:
-assumes "ofilter F"
-shows "Restr r F initial_segment_of r"
-using assms unfolding ofilter_def init_seg_of_def under_def by auto
-
-lemma underS_init_seg_of_Collect:
-assumes "\<And>j1 j2. \<lbrakk>j2 \<in> underS i; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
-shows "{R j |j. j \<in> underS i} \<in> Chains init_seg_of"
-unfolding Chains_def proof safe
-  fix j ja assume jS: "j \<in> underS i" and jaS: "ja \<in> underS i"
-  and init: "(R ja, R j) \<notin> init_seg_of"
-  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
-  and jjai: "(j,i) \<in> r" "(ja,i) \<in> r"
-  and i: "i \<notin> {j,ja}" unfolding Field_def underS_def by auto
-  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
-  show "R j initial_segment_of R ja"
-  using jja init jjai i
-  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: underS_def)
-qed
-
-lemma (in wo_rel) Field_init_seg_of_Collect:
-assumes "\<And>j1 j2. \<lbrakk>j2 \<in> Field r; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
-shows "{R j |j. j \<in> Field r} \<in> Chains init_seg_of"
-unfolding Chains_def proof safe
-  fix j ja assume jS: "j \<in> Field r" and jaS: "ja \<in> Field r"
-  and init: "(R ja, R j) \<notin> init_seg_of"
-  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
-  unfolding Field_def underS_def by auto
-  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
-  show "R j initial_segment_of R ja"
-  using jja init
-  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def)
-qed
-
-subsubsection {* Successor and limit elements of an ordinal *}
-
-definition "succ i \<equiv> suc {i}"
-
-definition "isSucc i \<equiv> \<exists> j. aboveS j \<noteq> {} \<and> i = succ j"
-
-definition "zero = minim (Field r)"
-
-definition "isLim i \<equiv> \<not> isSucc i"
-
-lemma zero_smallest[simp]:
-assumes "j \<in> Field r" shows "(zero, j) \<in> r"
-unfolding zero_def
-by (metis AboveS_Field assms subset_AboveS_UnderS subset_antisym subset_refl suc_def suc_least_AboveS)
-
-lemma zero_in_Field: assumes "Field r \<noteq> {}"  shows "zero \<in> Field r"
-using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def)
-
-lemma leq_zero_imp[simp]:
-"(x, zero) \<in> r \<Longrightarrow> x = zero"
-by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest)
-
-lemma leq_zero[simp]:
-assumes "Field r \<noteq> {}"  shows "(x, zero) \<in> r \<longleftrightarrow> x = zero"
-using zero_in_Field[OF assms] in_notinI[of x zero] by auto
-
-lemma under_zero[simp]:
-assumes "Field r \<noteq> {}" shows "under zero = {zero}"
-using assms unfolding under_def by auto
-
-lemma underS_zero[simp,intro]: "underS zero = {}"
-unfolding underS_def by auto
-
-lemma isSucc_succ: "aboveS i \<noteq> {} \<Longrightarrow> isSucc (succ i)"
-unfolding isSucc_def succ_def by auto
-
-lemma succ_in_diff:
-assumes "aboveS i \<noteq> {}"  shows "(i,succ i) \<in> r \<and> succ i \<noteq> i"
-using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto
-
-lemmas succ_in[simp] = succ_in_diff[THEN conjunct1]
-lemmas succ_diff[simp] = succ_in_diff[THEN conjunct2]
-
-lemma succ_in_Field[simp]:
-assumes "aboveS i \<noteq> {}"  shows "succ i \<in> Field r"
-using succ_in[OF assms] unfolding Field_def by auto
-
-lemma succ_not_in:
-assumes "aboveS i \<noteq> {}" shows "(succ i, i) \<notin> r"
-proof
-  assume 1: "(succ i, i) \<in> r"
-  hence "succ i \<in> Field r \<and> i \<in> Field r" unfolding Field_def by auto
-  hence "(i, succ i) \<in> r \<and> succ i \<noteq> i" using assms by auto
-  thus False using 1 by (metis ANTISYM antisymD)
-qed
-
-lemma not_isSucc_zero: "\<not> isSucc zero"
-proof
-  assume "isSucc zero"
-  moreover
-  then obtain i where "aboveS i \<noteq> {}" and 1: "minim (Field r) = succ i"
-  unfolding isSucc_def zero_def by auto
-  hence "succ i \<in> Field r" by auto
-  ultimately show False by (metis REFL isSucc_def minim_least refl_on_domain
-    subset_refl succ_in succ_not_in zero_def)
-qed
-
-lemma isLim_zero[simp]: "isLim zero"
-  by (metis isLim_def not_isSucc_zero)
-
-lemma succ_smallest:
-assumes "(i,j) \<in> r" and "i \<noteq> j"
-shows "(succ i, j) \<in> r"
-unfolding succ_def apply(rule suc_least)
-using assms unfolding Field_def by auto
-
-lemma isLim_supr:
-assumes f: "i \<in> Field r" and l: "isLim i"
-shows "i = supr (underS i)"
-proof(rule equals_supr)
-  fix j assume j: "j \<in> Field r" and 1: "\<And> j'. j' \<in> underS i \<Longrightarrow> (j',j) \<in> r"
-  show "(i,j) \<in> r" proof(intro in_notinI[OF _ f j], safe)
-    assume ji: "(j,i) \<in> r" "j \<noteq> i"
-    hence a: "aboveS j \<noteq> {}" unfolding aboveS_def by auto
-    hence "i \<noteq> succ j" using l unfolding isLim_def isSucc_def by auto
-    moreover have "(succ j, i) \<in> r" using succ_smallest[OF ji] by auto
-    ultimately have "succ j \<in> underS i" unfolding underS_def by auto
-    hence "(succ j, j) \<in> r" using 1 by auto
-    thus False using succ_not_in[OF a] by simp
-  qed
-qed(insert f, unfold underS_def Field_def, auto)
-
-definition "pred i \<equiv> SOME j. j \<in> Field r \<and> aboveS j \<noteq> {} \<and> succ j = i"
-
-lemma pred_Field_succ:
-assumes "isSucc i" shows "pred i \<in> Field r \<and> aboveS (pred i) \<noteq> {} \<and> succ (pred i) = i"
-proof-
-  obtain j where a: "aboveS j \<noteq> {}" and i: "i = succ j" using assms unfolding isSucc_def by auto
-  have 1: "j \<in> Field r" "j \<noteq> i" unfolding Field_def i
-  using succ_diff[OF a] a unfolding aboveS_def by auto
-  show ?thesis unfolding pred_def apply(rule someI_ex) using 1 i a by auto
-qed
-
-lemmas pred_Field[simp] = pred_Field_succ[THEN conjunct1]
-lemmas aboveS_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct1]
-lemmas succ_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct2]
-
-lemma isSucc_pred_in:
-assumes "isSucc i"  shows "(pred i, i) \<in> r"
-proof-
-  def j \<equiv> "pred i"
-  have i: "i = succ j" using assms unfolding j_def by simp
-  have a: "aboveS j \<noteq> {}" unfolding j_def using assms by auto
-  show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] .
-qed
-
-lemma isSucc_pred_diff:
-assumes "isSucc i"  shows "pred i \<noteq> i"
-by (metis aboveS_pred assms succ_diff succ_pred)
-
-(* todo: pred maximal, pred injective? *)
-
-lemma succ_inj[simp]:
-assumes "aboveS i \<noteq> {}" and "aboveS j \<noteq> {}"
-shows "succ i = succ j \<longleftrightarrow> i = j"
-proof safe
-  assume s: "succ i = succ j"
-  {assume "i \<noteq> j" and "(i,j) \<in> r"
-   hence "(succ i, j) \<in> r" using assms by (metis succ_smallest)
-   hence False using s assms by (metis succ_not_in)
-  }
-  moreover
-  {assume "i \<noteq> j" and "(j,i) \<in> r"
-   hence "(succ j, i) \<in> r" using assms by (metis succ_smallest)
-   hence False using s assms by (metis succ_not_in)
-  }
-  ultimately show "i = j" by (metis TOTALS WELL assms(1) assms(2) succ_in_diff well_order_on_domain)
-qed
-
-lemma pred_succ[simp]:
-assumes "aboveS j \<noteq> {}"  shows "pred (succ j) = j"
-unfolding pred_def apply(rule some_equality)
-using assms apply(force simp: Field_def aboveS_def)
-by (metis assms succ_inj)
-
-lemma less_succ[simp]:
-assumes "aboveS i \<noteq> {}"
-shows "(j, succ i) \<in> r \<longleftrightarrow> (j,i) \<in> r \<or> j = succ i"
-apply safe
-  apply (metis WELL assms in_notinI well_order_on_domain suc_singl_pred succ_def succ_in_diff)
-  apply (metis (hide_lams, full_types) REFL TRANS assms max2_def max2_equals1 refl_on_domain succ_in_Field succ_not_in transD)
-  apply (metis assms in_notinI succ_in_Field)
-done
-
-lemma underS_succ[simp]:
-assumes "aboveS i \<noteq> {}"
-shows "underS (succ i) = under i"
-unfolding underS_def under_def by (auto simp: assms succ_not_in)
-
-lemma succ_mono:
-assumes "aboveS j \<noteq> {}" and "(i,j) \<in> r"
-shows "(succ i, succ j) \<in> r"
-by (metis (full_types) assms less_succ succ_smallest)
-
-lemma under_succ[simp]:
-assumes "aboveS i \<noteq> {}"
-shows "under (succ i) = insert (succ i) (under i)"
-using less_succ[OF assms] unfolding under_def by auto
-
-definition mergeSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
-where
-"mergeSL S L f i \<equiv>
- if isSucc i then S (pred i) (f (pred i))
- else L f i"
-
-
-subsubsection {* Well-order recursion with (zero), succesor, and limit *}
-
-definition worecSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
-where "worecSL S L \<equiv> worec (mergeSL S L)"
-
-definition "adm_woL L \<equiv> \<forall>f g i. isLim i \<and> (\<forall>j\<in>underS i. f j = g j) \<longrightarrow> L f i = L g i"
-
-lemma mergeSL:
-assumes "adm_woL L"  shows "adm_wo (mergeSL S L)"
-unfolding adm_wo_def proof safe
-  fix f g :: "'a => 'b" and i :: 'a
-  assume 1: "\<forall>j\<in>underS i. f j = g j"
-  show "mergeSL S L f i = mergeSL S L g i"
-  proof(cases "isSucc i")
-    case True
-    hence "pred i \<in> underS i" unfolding underS_def using isSucc_pred_in isSucc_pred_diff by auto
-    thus ?thesis using True 1 unfolding mergeSL_def by auto
-  next
-    case False hence "isLim i" unfolding isLim_def by auto
-    thus ?thesis using assms False 1 unfolding mergeSL_def adm_woL_def by auto
-  qed
-qed
-
-lemma worec_fixpoint1: "adm_wo H \<Longrightarrow> worec H i = H (worec H) i"
-by (metis worec_fixpoint)
-
-lemma worecSL_isSucc:
-assumes a: "adm_woL L" and i: "isSucc i"
-shows "worecSL S L i = S (pred i) (worecSL S L (pred i))"
-proof-
-  let ?H = "mergeSL S L"
-  have "worecSL S L i = ?H (worec ?H) i"
-  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
-  also have "... = S (pred i) (worecSL S L (pred i))"
-  unfolding worecSL_def mergeSL_def using i by simp
-  finally show ?thesis .
-qed
-
-lemma worecSL_succ:
-assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
-shows "worecSL S L (succ j) = S j (worecSL S L j)"
-proof-
-  def i \<equiv> "succ j"
-  have i: "isSucc i" by (metis i i_def isSucc_succ)
-  have ij: "j = pred i" unfolding i_def using assms by simp
-  have 0: "succ (pred i) = i" using i by simp
-  show ?thesis unfolding ij using worecSL_isSucc[OF a i] unfolding 0 .
-qed
-
-lemma worecSL_isLim:
-assumes a: "adm_woL L" and i: "isLim i"
-shows "worecSL S L i = L (worecSL S L) i"
-proof-
-  let ?H = "mergeSL S L"
-  have "worecSL S L i = ?H (worec ?H) i"
-  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
-  also have "... = L (worecSL S L) i"
-  using i unfolding worecSL_def mergeSL_def isLim_def by simp
-  finally show ?thesis .
-qed
-
-definition worecZSL :: "'b \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
-where "worecZSL Z S L \<equiv> worecSL S (\<lambda> f a. if a = zero then Z else L f a)"
-
-lemma worecZSL_zero:
-assumes a: "adm_woL L"
-shows "worecZSL Z S L zero = Z"
-proof-
-  let ?L = "\<lambda> f a. if a = zero then Z else L f a"
-  have "worecZSL Z S L zero = ?L (worecZSL Z S L) zero"
-  unfolding worecZSL_def apply(rule worecSL_isLim)
-  using assms unfolding adm_woL_def by auto
-  also have "... = Z" by simp
-  finally show ?thesis .
-qed
-
-lemma worecZSL_succ:
-assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
-shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)"
-unfolding worecZSL_def apply(rule  worecSL_succ)
-using assms unfolding adm_woL_def by auto
-
-lemma worecZSL_isLim:
-assumes a: "adm_woL L" and "isLim i" and "i \<noteq> zero"
-shows "worecZSL Z S L i = L (worecZSL Z S L) i"
-proof-
-  let ?L = "\<lambda> f a. if a = zero then Z else L f a"
-  have "worecZSL Z S L i = ?L (worecZSL Z S L) i"
-  unfolding worecZSL_def apply(rule worecSL_isLim)
-  using assms unfolding adm_woL_def by auto
-  also have "... = L (worecZSL Z S L) i" using assms by simp
-  finally show ?thesis .
-qed
-
-
-subsubsection {* Well-order succ-lim induction *}
-
-lemma ord_cases:
-obtains j where "i = succ j" and "aboveS j \<noteq> {}"  | "isLim i"
-by (metis isLim_def isSucc_def)
-
-lemma well_order_inductSL[case_names Suc Lim]:
-assumes SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
-LIM: "\<And>i. \<lbrakk>isLim i; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
-shows "P i"
-proof(induction rule: well_order_induct)
-  fix i assume 0:  "\<forall>j. j \<noteq> i \<and> (j, i) \<in> r \<longrightarrow> P j"
-  show "P i" proof(cases i rule: ord_cases)
-    fix j assume i: "i = succ j" and j: "aboveS j \<noteq> {}"
-    hence "j \<noteq> i \<and> (j, i) \<in> r" by (metis  succ_diff succ_in)
-    hence 1: "P j" using 0 by simp
-    show "P i" unfolding i apply(rule SUCC) using 1 j by auto
-  qed(insert 0 LIM, unfold underS_def, auto)
-qed
-
-lemma well_order_inductZSL[case_names Zero Suc Lim]:
-assumes ZERO: "P zero"
-and SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
-LIM: "\<And>i. \<lbrakk>isLim i; i \<noteq> zero; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
-shows "P i"
-apply(induction rule: well_order_inductSL) using assms by auto
-
-(* Succesor and limit ordinals *)
-definition "isSuccOrd \<equiv> \<exists> j \<in> Field r. \<forall> i \<in> Field r. (i,j) \<in> r"
-definition "isLimOrd \<equiv> \<not> isSuccOrd"
-
-lemma isLimOrd_succ:
-assumes isLimOrd and "i \<in> Field r"
-shows "succ i \<in> Field r"
-using assms unfolding isLimOrd_def isSuccOrd_def
-by (metis REFL in_notinI refl_on_domain succ_smallest)
-
-lemma isLimOrd_aboveS:
-assumes l: isLimOrd and i: "i \<in> Field r"
-shows "aboveS i \<noteq> {}"
-proof-
-  obtain j where "j \<in> Field r" and "(j,i) \<notin> r"
-  using assms unfolding isLimOrd_def isSuccOrd_def by auto
-  hence "(i,j) \<in> r \<and> j \<noteq> i" by (metis i max2_def max2_greater)
-  thus ?thesis unfolding aboveS_def by auto
-qed
-
-lemma succ_aboveS_isLimOrd:
-assumes "\<forall> i \<in> Field r. aboveS i \<noteq> {} \<and> succ i \<in> Field r"
-shows isLimOrd
-unfolding isLimOrd_def isSuccOrd_def proof safe
-  fix j assume j: "j \<in> Field r" "\<forall>i\<in>Field r. (i, j) \<in> r"
-  hence "(succ j, j) \<in> r" using assms by auto
-  moreover have "aboveS j \<noteq> {}" using assms j unfolding aboveS_def by auto
-  ultimately show False by (metis succ_not_in)
-qed
-
-lemma isLim_iff:
-assumes l: "isLim i" and j: "j \<in> underS i"
-shows "\<exists> k. k \<in> underS i \<and> j \<in> underS k"
-proof-
-  have a: "aboveS j \<noteq> {}" using j unfolding underS_def aboveS_def by auto
-  show ?thesis apply(rule exI[of _ "succ j"]) apply safe
-  using assms a unfolding underS_def isLim_def
-  apply (metis (lifting, full_types) isSucc_def mem_Collect_eq succ_smallest)
-  by (metis (lifting, full_types) a mem_Collect_eq succ_diff succ_in)
-qed
-
-end (* context wo_rel *)
-
-abbreviation "zero \<equiv> wo_rel.zero"
-abbreviation "succ \<equiv> wo_rel.succ"
-abbreviation "pred \<equiv> wo_rel.pred"
-abbreviation "isSucc \<equiv> wo_rel.isSucc"
-abbreviation "isLim \<equiv> wo_rel.isLim"
-abbreviation "isLimOrd \<equiv> wo_rel.isLimOrd"
-abbreviation "isSuccOrd \<equiv> wo_rel.isSuccOrd"
-abbreviation "adm_woL \<equiv> wo_rel.adm_woL"
-abbreviation "worecSL \<equiv> wo_rel.worecSL"
-abbreviation "worecZSL \<equiv> wo_rel.worecZSL"
-
-
-subsection {* Projections of wellorders *}
- 
-definition "oproj r s f \<equiv> Field s \<subseteq> f ` (Field r) \<and> compat r s f"
-
-lemma oproj_in: 
-assumes "oproj r s f" and "(a,a') \<in> r"
-shows "(f a, f a') \<in> s"
-using assms unfolding oproj_def compat_def by auto
-
-lemma oproj_Field:
-assumes f: "oproj r s f" and a: "a \<in> Field r"
-shows "f a \<in> Field s"
-using oproj_in[OF f] a unfolding Field_def by auto
-
-lemma oproj_Field2:
-assumes f: "oproj r s f" and a: "b \<in> Field s"
-shows "\<exists> a \<in> Field r. f a = b"
-using assms unfolding oproj_def by auto
-
-lemma oproj_under: 
-assumes f:  "oproj r s f" and a: "a \<in> under r a'"
-shows "f a \<in> under s (f a')"
-using oproj_in[OF f] a unfolding under_def by auto
-
-(* An ordinal is embedded in another whenever it is embedded as an order 
-(not necessarily as initial segment):*)
-theorem embedI:
-assumes r: "Well_order r" and s: "Well_order s" 
-and f: "\<And> a. a \<in> Field r \<Longrightarrow> f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a)"
-shows "\<exists> g. embed r s g"
-proof-  
-  interpret r!: wo_rel r by unfold_locales (rule r)
-  interpret s!: wo_rel s by unfold_locales (rule s)
-  let ?G = "\<lambda> g a. suc s (g ` underS r a)"
-  def g \<equiv> "worec r ?G"
-  have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto
-  (*  *)
-  {fix a assume "a \<in> Field r"
-   hence "bij_betw g (under r a) (under s (g a)) \<and> 
-          g a \<in> under s (f a)"
-   proof(induction a rule: r.underS_induct)
-     case (1 a)
-     hence a: "a \<in> Field r"
-     and IH1a: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> inj_on g (under r a1)"
-     and IH1b: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g ` under r a1 = under s (g a1)"
-     and IH2: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> under s (f a1)"
-     unfolding underS_def Field_def bij_betw_def by auto
-     have fa: "f a \<in> Field s" using f[OF a] by auto
-     have g: "g a = suc s (g ` underS r a)" 
-     using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by simp
-     have A0: "g ` underS r a \<subseteq> Field s" 
-     using IH1b by (metis IH2 image_subsetI in_mono under_Field)
-     {fix a1 assume a1: "a1 \<in> underS r a"
-      from IH2[OF this] have "g a1 \<in> under s (f a1)" .
-      moreover have "f a1 \<in> underS s (f a)" using f[OF a] a1 by auto
-      ultimately have "g a1 \<in> underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans)
-     }
-     hence "f a \<in> AboveS s (g ` underS r a)" unfolding AboveS_def 
-     using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def)
-     hence A: "AboveS s (g ` underS r a) \<noteq> {}" by auto
-     have B: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (g a)"
-     unfolding g apply(rule s.suc_underS[OF A0 A]) by auto
-     {fix a1 a2 assume a2: "a2 \<in> underS r a" and 1: "a1 \<in> underS r a2"
-      hence a12: "{a1,a2} \<subseteq> under r a2" and "a1 \<noteq> a2" using r.REFL a
-      unfolding underS_def under_def refl_on_def Field_def by auto 
-      hence "g a1 \<noteq> g a2" using IH1a[OF a2] unfolding inj_on_def by auto
-      hence "g a1 \<in> underS s (g a2)" using IH1b[OF a2] a12 
-      unfolding underS_def under_def by auto
-     } note C = this
-     have ga: "g a \<in> Field s" unfolding g using s.suc_inField[OF A0 A] .
-     have aa: "a \<in> under r a" 
-     using a r.REFL unfolding under_def underS_def refl_on_def by auto
-     show ?case proof safe
-       show "bij_betw g (under r a) (under s (g a))" unfolding bij_betw_def proof safe
-         show "inj_on g (under r a)" proof(rule r.inj_on_Field)
-           fix a1 a2 assume "a1 \<in> under r a" and a2: "a2 \<in> under r a" and a1: "a1 \<in> underS r a2"
-           hence a22: "a2 \<in> under r a2" and a12: "a1 \<in> under r a2" "a1 \<noteq> a2"
-           using a r.REFL unfolding under_def underS_def refl_on_def by auto
-           show "g a1 \<noteq> g a2"
-           proof(cases "a2 = a")
-             case False hence "a2 \<in> underS r a" 
-             using a2 unfolding underS_def under_def by auto 
-             from IH1a[OF this] show ?thesis using a12 a22 unfolding inj_on_def by auto
-           qed(insert B a1, unfold underS_def, auto)
-         qed(unfold under_def Field_def, auto)
-       next
-         fix a1 assume a1: "a1 \<in> under r a" 
-         show "g a1 \<in> under s (g a)"
-         proof(cases "a1 = a")
-           case True thus ?thesis 
-           using ga s.REFL unfolding refl_on_def under_def by auto
-         next
-           case False
-           hence a1: "a1 \<in> underS r a" using a1 unfolding underS_def under_def by auto 
-           thus ?thesis using B unfolding underS_def under_def by auto
-         qed
-       next
-         fix b1 assume b1: "b1 \<in> under s (g a)"
-         show "b1 \<in> g ` under r a"
-         proof(cases "b1 = g a")
-           case True thus ?thesis using aa by auto
-         next
-           case False 
-           hence "b1 \<in> underS s (g a)" using b1 unfolding underS_def under_def by auto
-           from s.underS_suc[OF this[unfolded g] A0]
-           obtain a1 where a1: "a1 \<in> underS r a" and b1: "b1 \<in> under s (g a1)" by auto
-           obtain a2 where "a2 \<in> under r a1" and b1: "b1 = g a2" using IH1b[OF a1] b1 by auto
-           hence "a2 \<in> under r a" using a1 
-           by (metis r.ANTISYM r.TRANS in_mono underS_subset_under under_underS_trans)
-           thus ?thesis using b1 by auto
-         qed
-       qed
-     next
-       have "(g a, f a) \<in> s" unfolding g proof(rule s.suc_least[OF A0])
-         fix b1 assume "b1 \<in> g ` underS r a" 
-         then obtain a1 where a1: "b1 = g a1" and a1: "a1 \<in> underS r a" by auto
-         hence "b1 \<in> underS s (f a)" 
-         using a by (metis `\<And>a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (f a)`)
-         thus "f a \<noteq> b1 \<and> (b1, f a) \<in> s" unfolding underS_def by auto
-       qed(insert fa, auto)
-       thus "g a \<in> under s (f a)" unfolding under_def by auto
-     qed
-   qed
-  }
-  thus ?thesis unfolding embed_def by auto
-qed
-
-corollary ordLeq_def2:
-  "r \<le>o s \<longleftrightarrow> Well_order r \<and> Well_order s \<and>
-               (\<exists> f. \<forall> a \<in> Field r. f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a))"
-using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s]
-unfolding ordLeq_def by fast
-
-lemma iso_oproj:
-  assumes r: "Well_order r" and s: "Well_order s" and f: "iso r s f"
-  shows "oproj r s f"
-using assms unfolding oproj_def
-by (subst (asm) iso_iff3) (auto simp: bij_betw_def)
-
-theorem oproj_embed:
-assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
-shows "\<exists> g. embed s r g"
-proof (rule embedI[OF s r, of "inv_into (Field r) f"], unfold underS_def, safe)
-  fix b assume "b \<in> Field s"
-  thus "inv_into (Field r) f b \<in> Field r" using oproj_Field2[OF f] by (metis imageI inv_into_into)
-next
-  fix a b assume "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
-    "inv_into (Field r) f a = inv_into (Field r) f b"
-  with f show False by (auto dest!: inv_into_injective simp: Field_def oproj_def)
-next
-  fix a b assume *: "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
-  { assume "(inv_into (Field r) f a, inv_into (Field r) f b) \<notin> r"
-    moreover
-    from *(3) have "a \<in> Field s" unfolding Field_def by auto
-    with *(1,2) have
-      "inv_into (Field r) f a \<in> Field r" "inv_into (Field r) f b \<in> Field r"
-      "inv_into (Field r) f a \<noteq> inv_into (Field r) f b"
-      by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into)
-    ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \<in> r"
-      using r by (auto simp: well_order_on_def linear_order_on_def total_on_def)
-    with f[unfolded oproj_def compat_def] *(1) `a \<in> Field s`
-      f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"]
-      have "(b, a) \<in> s" by (metis in_mono)
-    with *(2,3) s have False
-      by (auto simp: well_order_on_def linear_order_on_def partial_order_on_def antisym_def)
-  } thus "(inv_into (Field r) f a, inv_into (Field r) f b) \<in> r" by blast
-qed
-
-corollary oproj_ordLeq:
-assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
-shows "s \<le>o r"
-using oproj_embed[OF assms] r s unfolding ordLeq_def by auto
-
-end
--- a/src/HOL/Cardinals/Order_Union.thy	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/Cardinals/Order_Union.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -78,7 +78,7 @@
 assumes FLD: "Field r Int Field r' = {}" and
         REFL: "Refl r" and REFL': "Refl r'"
 shows "Refl (r Osum r')"
-using assms 
+using assms
 unfolding refl_on_def Field_Osum unfolding Osum_def by blast
 
 lemma Osum_trans:
--- a/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/Cardinals/Ordinal_Arithmetic.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -8,12 +8,12 @@
 header {* Ordinal Arithmetic *}
 
 theory Ordinal_Arithmetic
-imports Constructions_on_Wellorders
+imports Wellorder_Constructions
 begin
 
 definition osum :: "'a rel \<Rightarrow> 'b rel \<Rightarrow> ('a + 'b) rel"  (infixr "+o" 70)
 where
-  "r +o r' = map_prod Inl Inl ` r \<union> map_prod Inr Inr ` r' \<union> 
+  "r +o r' = map_prod Inl Inl ` r \<union> map_prod Inr Inr ` r' \<union>
      {(Inl a, Inr a') | a a' . a \<in> Field r \<and> a' \<in> Field r'}"
 
 lemma Field_osum: "Field(r +o r') = Inl ` Field r \<union> Inr ` Field r'"
@@ -175,7 +175,7 @@
 lemma oprod_Refl:"\<lbrakk>Refl r; Refl r'\<rbrakk> \<Longrightarrow> Refl (r *o r')"
   unfolding refl_on_def Field_oprod unfolding oprod_def by auto
 
-lemma oprod_trans: 
+lemma oprod_trans:
   assumes "trans r" "trans r'" "antisym r" "antisym r'"
   shows "trans (r *o r')"
 proof(unfold trans_def, safe)
@@ -309,7 +309,7 @@
   from assms(3) have r': "Field r' \<noteq> {}" unfolding Field_def by auto
   have minim[simp]: "minim r' (Field r') \<in> Field r'"
     using wo_rel.minim_inField[unfolded wo_rel_def, OF WELL' _ r'] by auto
-  { fix b 
+  { fix b
     assume "(b, minim r' (Field r')) \<in> r'"
     moreover hence "b \<in> Field r'" unfolding Field_def by auto
     hence "(minim r' (Field r'), b) \<in> r'"
@@ -390,7 +390,7 @@
   unfolding isMaxim_def using antisymD[OF ANTISYM, of x y] by auto
 
 lemma maxim_isMaxim: "\<lbrakk>finite A; A \<noteq> {}; A \<subseteq> Field r\<rbrakk> \<Longrightarrow> isMaxim A (maxim A)"
-unfolding maxim_def 
+unfolding maxim_def
 proof (rule theI', rule ex_ex1I[OF _ isMaxim_unique, rotated], assumption+,
   induct A rule: finite_induct)
   case (insert x A)
@@ -494,7 +494,7 @@
 
 locale wo_rel2 =
   fixes r s
-  assumes rWELL: "Well_order r" 
+  assumes rWELL: "Well_order r"
   and     sWELL: "Well_order s"
 begin
 
@@ -539,7 +539,7 @@
 
 lemma max_fun_diff_max2:
   assumes ineq: "s.max_fun_diff f g = s.max_fun_diff g h \<longrightarrow>
-    f (s.max_fun_diff f g) \<noteq> h (s.max_fun_diff g h)" and 
+    f (s.max_fun_diff f g) \<noteq> h (s.max_fun_diff g h)" and
     fg: "f \<noteq> g" and gh: "g \<noteq> h" and fh: "f \<noteq> h" and
     f: "f \<in> FINFUNC" and g: "g \<in> FINFUNC" and h: "h \<in> FINFUNC"
   shows "s.max_fun_diff f h = s.max2 (s.max_fun_diff f g) (s.max_fun_diff g h)"
@@ -708,7 +708,7 @@
   by (auto intro: finite_subset[OF support_upd_subset])
 
 lemma fun_upd_same_oexp:
-  assumes "(f, g) \<in> oexp" "f x = g x" "x \<in> Field s" "y \<in> Field r" 
+  assumes "(f, g) \<in> oexp" "f x = g x" "x \<in> Field s" "y \<in> Field r"
   shows   "(f(x := y), g(x := y)) \<in> oexp"
 proof -
   from assms(1) fun_upd_FINFUNC[OF _ assms(3,4)] have fg: "f(x := y) \<in> FINFUNC" "g(x := y) \<in> FINFUNC"
@@ -829,7 +829,7 @@
                 thus ?thesis
                 proof (cases "s.maxim (SUPP f) = z \<and> f z = x0")
                   case True
-                  with f have "f(z := r.zero) \<in> G" unfolding G_def by blast 
+                  with f have "f(z := r.zero) \<in> G" unfolding G_def by blast
                   with g0(2) f0z have "(f0(z := r.zero), f(z := r.zero)) \<in> oexp" by auto
                   hence "(f0(z := r.zero, z := x0), f(z := r.zero, z := x0)) \<in> oexp"
                     by (elim fun_upd_same_oexp[OF _ _ zField x0Field]) simp
@@ -922,10 +922,10 @@
   unfolding ozero_def by simp
 
 lemma iso_ozero_empty[simp]: "r =o ozero = (r = {})"
-  unfolding ozero_def ordIso_def iso_def[abs_def] embed_def bij_betw_def 
+  unfolding ozero_def ordIso_def iso_def[abs_def] embed_def bij_betw_def
   by (auto dest: well_order_on_domain)
 
-lemma ozero_ordLeq: 
+lemma ozero_ordLeq:
 assumes "Well_order r"  shows "ozero \<le>o r"
 using assms unfolding ozero_def ordLeq_def embed_def[abs_def] under_def by auto
 
@@ -959,7 +959,7 @@
   with f have "bij_betw ?f (Field ?L) (Field ?R)"
     unfolding Field_osum iso_def bij_betw_def image_image image_Un by auto
   moreover from f have "compat ?L ?R ?f"
-    unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def 
+    unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def
     by (auto simp: map_prod_imageI)
   ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: osum_Well_order r s t)
   thus ?thesis unfolding ordIso_def by (auto intro: osum_Well_order r s t)
@@ -977,7 +977,7 @@
   with f have "bij_betw ?f (Field ?L) (Field ?R)"
     unfolding Field_osum iso_def bij_betw_def image_image image_Un by auto
   moreover from f have "compat ?L ?R ?f"
-    unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def 
+    unfolding osum_def iso_iff3[OF r s] compat_def bij_betw_def
     by (auto simp: map_prod_imageI)
   ultimately have "iso ?L ?R ?f" by (subst iso_iff3) (auto intro: osum_Well_order r s t)
   thus ?thesis unfolding ordIso_def by (auto intro: osum_Well_order r s t)
@@ -1269,7 +1269,7 @@
 lemma ozero_oexp: "\<not> (s =o ozero) \<Longrightarrow> ozero ^o s =o ozero"
   unfolding oexp_def[OF ozero_Well_order s] FinFunc_def
   by simp (metis Func_emp2 bot.extremum_uniqueI emptyE well_order_on_domain s subrelI)
-      
+
 lemma oone_oexp: "oone ^o s =o oone" (is "?L =o ?R")
   by (rule oone_ordIso_oexp[OF ordIso_reflexive[OF oone_Well_order] s])
 
@@ -1490,7 +1490,7 @@
   thus ?thesis unfolding ordLeq_def2 by (fast intro: oexp_Well_order r s)
 qed
 
-lemma FinFunc_osum: 
+lemma FinFunc_osum:
   "fg \<in> FinFunc r (s +o t) = (fg o Inl \<in> FinFunc r s \<and> fg o Inr \<in> FinFunc r t)"
   (is "?L = (?R1 \<and> ?R2)")
 proof safe
@@ -1696,7 +1696,7 @@
         unfolding inj_on_def FinFunc_def Func_def Field_oprod rs.Field_oexp rev_curr_def[abs_def]
         by (auto simp: fun_eq_iff) metis
       show "rev_curr ` (FinFunc r (s *o t)) = FinFunc (r ^o s) t" by (rule rev_curr_FinFunc[OF Field])
-    qed               
+    qed
     moreover
     have "compat ?L ?R rev_curr"
     unfolding compat_def proof safe
--- a/src/HOL/Cardinals/README.txt	Mon Sep 01 16:34:38 2014 +0200
+++ b/src/HOL/Cardinals/README.txt	Mon Sep 01 16:34:39 2014 +0200
@@ -35,7 +35,7 @@
       and *strict embeddings* are defined to be embeddings that are, and respectively
       are not, bijections.
 
--- Constructions_on_Wellorders:
+-- Wellorder_Constructions:
 ----- 1) Defines direct images, restrictions, disjoint unions and 
       bounded squares of well-orders.
 ----- 2) Defines the relations "ordLeq", "ordLess" and "ordIso" 
@@ -203,7 +203,7 @@
   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 BNF_Constructions_on_Wellorders):
+Theory Wellorder_Constructions (and BNF_Wellorder_Constructions):
 - 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, 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Wellorder_Constructions.thy	Mon Sep 01 16:34:39 2014 +0200
@@ -0,0 +1,1077 @@
+(*  Title:      HOL/Cardinals/Wellorder_Constructions.thy
+    Author:     Andrei Popescu, TU Muenchen
+    Copyright   2012
+
+Constructions on wellorders.
+*)
+
+header {* Constructions on Wellorders *}
+
+theory Wellorder_Constructions
+imports
+  BNF_Wellorder_Constructions Wellorder_Embedding Order_Union
+  "../Library/Cardinal_Notations"
+begin
+
+declare
+  ordLeq_Well_order_simp[simp]
+  not_ordLeq_iff_ordLess[simp]
+  not_ordLess_iff_ordLeq[simp]
+  Func_empty[simp]
+  Func_is_emp[simp]
+
+lemma Func_emp2[simp]: "A \<noteq> {} \<Longrightarrow> Func A {} = {}" by auto
+
+
+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 Restr_subset by (elim wf_subset) simp
+
+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_Well_order_simp[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[simp]:
+assumes "r =o r'"
+shows "Well_order r \<and> Well_order r'"
+using assms unfolding ordIso_def by simp
+
+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
+
+(* More facts on ordinal sum: *)
+
+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 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 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 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
+
+
+subsection {* Limit and succesor ordinals *}
+
+lemma embed_underS2:
+assumes r: "Well_order r" and s: "Well_order s"  and g: "embed r s g" and a: "a \<in> Field r"
+shows "g ` underS r a = underS s (g a)"
+using embed_underS[OF assms] unfolding bij_betw_def by auto
+
+lemma bij_betw_insert:
+assumes "b \<notin> A" and "f b \<notin> A'" and "bij_betw f A A'"
+shows "bij_betw f (insert b A) (insert (f b) A')"
+using notIn_Un_bij_betw[OF assms] by auto
+
+context wo_rel
+begin
+
+lemma underS_induct:
+  assumes "\<And>a. (\<And> a1. a1 \<in> underS a \<Longrightarrow> P a1) \<Longrightarrow> P a"
+  shows "P a"
+  by (induct rule: well_order_induct) (rule assms, simp add: underS_def)
+
+lemma suc_underS:
+assumes B: "B \<subseteq> Field r" and A: "AboveS B \<noteq> {}" and b: "b \<in> B"
+shows "b \<in> underS (suc B)"
+using suc_AboveS[OF B A] b unfolding underS_def AboveS_def by auto
+
+lemma underS_supr:
+assumes bA: "b \<in> underS (supr A)" and A: "A \<subseteq> Field r"
+shows "\<exists> a \<in> A. b \<in> underS a"
+proof(rule ccontr, auto)
+  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
+  assume "\<forall>a\<in>A.  b \<notin> underS a"
+  hence 0: "\<forall>a \<in> A. (a,b) \<in> r" using A bA unfolding underS_def
+  by simp (metis REFL in_mono max2_def max2_greater refl_on_domain)
+  have "(supr A, b) \<in> r" apply(rule supr_least[OF A bb]) using 0 by auto
+  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
+qed
+
+lemma underS_suc:
+assumes bA: "b \<in> underS (suc A)" and A: "A \<subseteq> Field r"
+shows "\<exists> a \<in> A. b \<in> under a"
+proof(rule ccontr, auto)
+  have bb: "b \<in> Field r" using bA unfolding underS_def Field_def by auto
+  assume "\<forall>a\<in>A.  b \<notin> under a"
+  hence 0: "\<forall>a \<in> A. a \<in> underS b" using A bA unfolding underS_def
+  by simp (metis (lifting) bb max2_def max2_greater mem_Collect_eq under_def set_rev_mp)
+  have "(suc A, b) \<in> r" apply(rule suc_least[OF A bb]) using 0 unfolding underS_def by auto
+  thus False using bA unfolding underS_def by simp (metis ANTISYM antisymD)
+qed
+
+lemma (in wo_rel) in_underS_supr:
+assumes j: "j \<in> underS i" and i: "i \<in> A" and A: "A \<subseteq> Field r" and AA: "Above A \<noteq> {}"
+shows "j \<in> underS (supr A)"
+proof-
+  have "(i,supr A) \<in> r" using supr_greater[OF A AA i] .
+  thus ?thesis using j unfolding underS_def
+  by simp (metis REFL TRANS max2_def max2_equals1 refl_on_domain transD)
+qed
+
+lemma inj_on_Field:
+assumes A: "A \<subseteq> Field r" and f: "\<And> a b. \<lbrakk>a \<in> A; b \<in> A; a \<in> underS b\<rbrakk> \<Longrightarrow> f a \<noteq> f b"
+shows "inj_on f A"
+unfolding inj_on_def proof safe
+  fix a b assume a: "a \<in> A" and b: "b \<in> A" and fab: "f a = f b"
+  {assume "a \<in> underS b"
+   hence False using f[OF a b] fab by auto
+  }
+  moreover
+  {assume "b \<in> underS a"
+   hence False using f[OF b a] fab by auto
+  }
+  ultimately show "a = b" using TOTALS A a b unfolding underS_def by auto
+qed
+
+lemma in_notinI:
+assumes "(j,i) \<notin> r \<or> j = i" and "i \<in> Field r" and "j \<in> Field r"
+shows "(i,j) \<in> r" by (metis assms max2_def max2_greater_among)
+
+lemma ofilter_init_seg_of:
+assumes "ofilter F"
+shows "Restr r F initial_segment_of r"
+using assms unfolding ofilter_def init_seg_of_def under_def by auto
+
+lemma underS_init_seg_of_Collect:
+assumes "\<And>j1 j2. \<lbrakk>j2 \<in> underS i; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
+shows "{R j |j. j \<in> underS i} \<in> Chains init_seg_of"
+unfolding Chains_def proof safe
+  fix j ja assume jS: "j \<in> underS i" and jaS: "ja \<in> underS i"
+  and init: "(R ja, R j) \<notin> init_seg_of"
+  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
+  and jjai: "(j,i) \<in> r" "(ja,i) \<in> r"
+  and i: "i \<notin> {j,ja}" unfolding Field_def underS_def by auto
+  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
+  show "R j initial_segment_of R ja"
+  using jja init jjai i
+  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: underS_def)
+qed
+
+lemma (in wo_rel) Field_init_seg_of_Collect:
+assumes "\<And>j1 j2. \<lbrakk>j2 \<in> Field r; (j1, j2) \<in> r\<rbrakk> \<Longrightarrow> R j1 initial_segment_of R j2"
+shows "{R j |j. j \<in> Field r} \<in> Chains init_seg_of"
+unfolding Chains_def proof safe
+  fix j ja assume jS: "j \<in> Field r" and jaS: "ja \<in> Field r"
+  and init: "(R ja, R j) \<notin> init_seg_of"
+  hence jja: "{j,ja} \<subseteq> Field r" and j: "j \<in> Field r" and ja: "ja \<in> Field r"
+  unfolding Field_def underS_def by auto
+  have jj: "(j,j) \<in> r" and jaja: "(ja,ja) \<in> r" using j ja by (metis in_notinI)+
+  show "R j initial_segment_of R ja"
+  using jja init
+  by (elim cases_Total3 disjE) (auto elim: cases_Total3 intro!: assms simp: Field_def)
+qed
+
+subsubsection {* Successor and limit elements of an ordinal *}
+
+definition "succ i \<equiv> suc {i}"
+
+definition "isSucc i \<equiv> \<exists> j. aboveS j \<noteq> {} \<and> i = succ j"
+
+definition "zero = minim (Field r)"
+
+definition "isLim i \<equiv> \<not> isSucc i"
+
+lemma zero_smallest[simp]:
+assumes "j \<in> Field r" shows "(zero, j) \<in> r"
+unfolding zero_def
+by (metis AboveS_Field assms subset_AboveS_UnderS subset_antisym subset_refl suc_def suc_least_AboveS)
+
+lemma zero_in_Field: assumes "Field r \<noteq> {}"  shows "zero \<in> Field r"
+using assms unfolding zero_def by (metis Field_ofilter minim_in ofilter_def)
+
+lemma leq_zero_imp[simp]:
+"(x, zero) \<in> r \<Longrightarrow> x = zero"
+by (metis ANTISYM WELL antisymD well_order_on_domain zero_smallest)
+
+lemma leq_zero[simp]:
+assumes "Field r \<noteq> {}"  shows "(x, zero) \<in> r \<longleftrightarrow> x = zero"
+using zero_in_Field[OF assms] in_notinI[of x zero] by auto
+
+lemma under_zero[simp]:
+assumes "Field r \<noteq> {}" shows "under zero = {zero}"
+using assms unfolding under_def by auto
+
+lemma underS_zero[simp,intro]: "underS zero = {}"
+unfolding underS_def by auto
+
+lemma isSucc_succ: "aboveS i \<noteq> {} \<Longrightarrow> isSucc (succ i)"
+unfolding isSucc_def succ_def by auto
+
+lemma succ_in_diff:
+assumes "aboveS i \<noteq> {}"  shows "(i,succ i) \<in> r \<and> succ i \<noteq> i"
+using assms suc_greater[of "{i}"] unfolding succ_def AboveS_def aboveS_def Field_def by auto
+
+lemmas succ_in[simp] = succ_in_diff[THEN conjunct1]
+lemmas succ_diff[simp] = succ_in_diff[THEN conjunct2]
+
+lemma succ_in_Field[simp]:
+assumes "aboveS i \<noteq> {}"  shows "succ i \<in> Field r"
+using succ_in[OF assms] unfolding Field_def by auto
+
+lemma succ_not_in:
+assumes "aboveS i \<noteq> {}" shows "(succ i, i) \<notin> r"
+proof
+  assume 1: "(succ i, i) \<in> r"
+  hence "succ i \<in> Field r \<and> i \<in> Field r" unfolding Field_def by auto
+  hence "(i, succ i) \<in> r \<and> succ i \<noteq> i" using assms by auto
+  thus False using 1 by (metis ANTISYM antisymD)
+qed
+
+lemma not_isSucc_zero: "\<not> isSucc zero"
+proof
+  assume "isSucc zero"
+  moreover
+  then obtain i where "aboveS i \<noteq> {}" and 1: "minim (Field r) = succ i"
+  unfolding isSucc_def zero_def by auto
+  hence "succ i \<in> Field r" by auto
+  ultimately show False by (metis REFL isSucc_def minim_least refl_on_domain
+    subset_refl succ_in succ_not_in zero_def)
+qed
+
+lemma isLim_zero[simp]: "isLim zero"
+  by (metis isLim_def not_isSucc_zero)
+
+lemma succ_smallest:
+assumes "(i,j) \<in> r" and "i \<noteq> j"
+shows "(succ i, j) \<in> r"
+unfolding succ_def apply(rule suc_least)
+using assms unfolding Field_def by auto
+
+lemma isLim_supr:
+assumes f: "i \<in> Field r" and l: "isLim i"
+shows "i = supr (underS i)"
+proof(rule equals_supr)
+  fix j assume j: "j \<in> Field r" and 1: "\<And> j'. j' \<in> underS i \<Longrightarrow> (j',j) \<in> r"
+  show "(i,j) \<in> r" proof(intro in_notinI[OF _ f j], safe)
+    assume ji: "(j,i) \<in> r" "j \<noteq> i"
+    hence a: "aboveS j \<noteq> {}" unfolding aboveS_def by auto
+    hence "i \<noteq> succ j" using l unfolding isLim_def isSucc_def by auto
+    moreover have "(succ j, i) \<in> r" using succ_smallest[OF ji] by auto
+    ultimately have "succ j \<in> underS i" unfolding underS_def by auto
+    hence "(succ j, j) \<in> r" using 1 by auto
+    thus False using succ_not_in[OF a] by simp
+  qed
+qed(insert f, unfold underS_def Field_def, auto)
+
+definition "pred i \<equiv> SOME j. j \<in> Field r \<and> aboveS j \<noteq> {} \<and> succ j = i"
+
+lemma pred_Field_succ:
+assumes "isSucc i" shows "pred i \<in> Field r \<and> aboveS (pred i) \<noteq> {} \<and> succ (pred i) = i"
+proof-
+  obtain j where a: "aboveS j \<noteq> {}" and i: "i = succ j" using assms unfolding isSucc_def by auto
+  have 1: "j \<in> Field r" "j \<noteq> i" unfolding Field_def i
+  using succ_diff[OF a] a unfolding aboveS_def by auto
+  show ?thesis unfolding pred_def apply(rule someI_ex) using 1 i a by auto
+qed
+
+lemmas pred_Field[simp] = pred_Field_succ[THEN conjunct1]
+lemmas aboveS_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct1]
+lemmas succ_pred[simp] = pred_Field_succ[THEN conjunct2, THEN conjunct2]
+
+lemma isSucc_pred_in:
+assumes "isSucc i"  shows "(pred i, i) \<in> r"
+proof-
+  def j \<equiv> "pred i"
+  have i: "i = succ j" using assms unfolding j_def by simp
+  have a: "aboveS j \<noteq> {}" unfolding j_def using assms by auto
+  show ?thesis unfolding j_def[symmetric] unfolding i using succ_in[OF a] .
+qed
+
+lemma isSucc_pred_diff:
+assumes "isSucc i"  shows "pred i \<noteq> i"
+by (metis aboveS_pred assms succ_diff succ_pred)
+
+(* todo: pred maximal, pred injective? *)
+
+lemma succ_inj[simp]:
+assumes "aboveS i \<noteq> {}" and "aboveS j \<noteq> {}"
+shows "succ i = succ j \<longleftrightarrow> i = j"
+proof safe
+  assume s: "succ i = succ j"
+  {assume "i \<noteq> j" and "(i,j) \<in> r"
+   hence "(succ i, j) \<in> r" using assms by (metis succ_smallest)
+   hence False using s assms by (metis succ_not_in)
+  }
+  moreover
+  {assume "i \<noteq> j" and "(j,i) \<in> r"
+   hence "(succ j, i) \<in> r" using assms by (metis succ_smallest)
+   hence False using s assms by (metis succ_not_in)
+  }
+  ultimately show "i = j" by (metis TOTALS WELL assms(1) assms(2) succ_in_diff well_order_on_domain)
+qed
+
+lemma pred_succ[simp]:
+assumes "aboveS j \<noteq> {}"  shows "pred (succ j) = j"
+unfolding pred_def apply(rule some_equality)
+using assms apply(force simp: Field_def aboveS_def)
+by (metis assms succ_inj)
+
+lemma less_succ[simp]:
+assumes "aboveS i \<noteq> {}"
+shows "(j, succ i) \<in> r \<longleftrightarrow> (j,i) \<in> r \<or> j = succ i"
+apply safe
+  apply (metis WELL assms in_notinI well_order_on_domain suc_singl_pred succ_def succ_in_diff)
+  apply (metis (hide_lams, full_types) REFL TRANS assms max2_def max2_equals1 refl_on_domain succ_in_Field succ_not_in transD)
+  apply (metis assms in_notinI succ_in_Field)
+done
+
+lemma underS_succ[simp]:
+assumes "aboveS i \<noteq> {}"
+shows "underS (succ i) = under i"
+unfolding underS_def under_def by (auto simp: assms succ_not_in)
+
+lemma succ_mono:
+assumes "aboveS j \<noteq> {}" and "(i,j) \<in> r"
+shows "(succ i, succ j) \<in> r"
+by (metis (full_types) assms less_succ succ_smallest)
+
+lemma under_succ[simp]:
+assumes "aboveS i \<noteq> {}"
+shows "under (succ i) = insert (succ i) (under i)"
+using less_succ[OF assms] unfolding under_def by auto
+
+definition mergeSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+where
+"mergeSL S L f i \<equiv>
+ if isSucc i then S (pred i) (f (pred i))
+ else L f i"
+
+
+subsubsection {* Well-order recursion with (zero), succesor, and limit *}
+
+definition worecSL :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+where "worecSL S L \<equiv> worec (mergeSL S L)"
+
+definition "adm_woL L \<equiv> \<forall>f g i. isLim i \<and> (\<forall>j\<in>underS i. f j = g j) \<longrightarrow> L f i = L g i"
+
+lemma mergeSL:
+assumes "adm_woL L"  shows "adm_wo (mergeSL S L)"
+unfolding adm_wo_def proof safe
+  fix f g :: "'a => 'b" and i :: 'a
+  assume 1: "\<forall>j\<in>underS i. f j = g j"
+  show "mergeSL S L f i = mergeSL S L g i"
+  proof(cases "isSucc i")
+    case True
+    hence "pred i \<in> underS i" unfolding underS_def using isSucc_pred_in isSucc_pred_diff by auto
+    thus ?thesis using True 1 unfolding mergeSL_def by auto
+  next
+    case False hence "isLim i" unfolding isLim_def by auto
+    thus ?thesis using assms False 1 unfolding mergeSL_def adm_woL_def by auto
+  qed
+qed
+
+lemma worec_fixpoint1: "adm_wo H \<Longrightarrow> worec H i = H (worec H) i"
+by (metis worec_fixpoint)
+
+lemma worecSL_isSucc:
+assumes a: "adm_woL L" and i: "isSucc i"
+shows "worecSL S L i = S (pred i) (worecSL S L (pred i))"
+proof-
+  let ?H = "mergeSL S L"
+  have "worecSL S L i = ?H (worec ?H) i"
+  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
+  also have "... = S (pred i) (worecSL S L (pred i))"
+  unfolding worecSL_def mergeSL_def using i by simp
+  finally show ?thesis .
+qed
+
+lemma worecSL_succ:
+assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
+shows "worecSL S L (succ j) = S j (worecSL S L j)"
+proof-
+  def i \<equiv> "succ j"
+  have i: "isSucc i" by (metis i i_def isSucc_succ)
+  have ij: "j = pred i" unfolding i_def using assms by simp
+  have 0: "succ (pred i) = i" using i by simp
+  show ?thesis unfolding ij using worecSL_isSucc[OF a i] unfolding 0 .
+qed
+
+lemma worecSL_isLim:
+assumes a: "adm_woL L" and i: "isLim i"
+shows "worecSL S L i = L (worecSL S L) i"
+proof-
+  let ?H = "mergeSL S L"
+  have "worecSL S L i = ?H (worec ?H) i"
+  unfolding worecSL_def using worec_fixpoint1[OF mergeSL[OF a]] .
+  also have "... = L (worecSL S L) i"
+  using i unfolding worecSL_def mergeSL_def isLim_def by simp
+  finally show ?thesis .
+qed
+
+definition worecZSL :: "'b \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> (('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+where "worecZSL Z S L \<equiv> worecSL S (\<lambda> f a. if a = zero then Z else L f a)"
+
+lemma worecZSL_zero:
+assumes a: "adm_woL L"
+shows "worecZSL Z S L zero = Z"
+proof-
+  let ?L = "\<lambda> f a. if a = zero then Z else L f a"
+  have "worecZSL Z S L zero = ?L (worecZSL Z S L) zero"
+  unfolding worecZSL_def apply(rule worecSL_isLim)
+  using assms unfolding adm_woL_def by auto
+  also have "... = Z" by simp
+  finally show ?thesis .
+qed
+
+lemma worecZSL_succ:
+assumes a: "adm_woL L" and i: "aboveS j \<noteq> {}"
+shows "worecZSL Z S L (succ j) = S j (worecZSL Z S L j)"
+unfolding worecZSL_def apply(rule  worecSL_succ)
+using assms unfolding adm_woL_def by auto
+
+lemma worecZSL_isLim:
+assumes a: "adm_woL L" and "isLim i" and "i \<noteq> zero"
+shows "worecZSL Z S L i = L (worecZSL Z S L) i"
+proof-
+  let ?L = "\<lambda> f a. if a = zero then Z else L f a"
+  have "worecZSL Z S L i = ?L (worecZSL Z S L) i"
+  unfolding worecZSL_def apply(rule worecSL_isLim)
+  using assms unfolding adm_woL_def by auto
+  also have "... = L (worecZSL Z S L) i" using assms by simp
+  finally show ?thesis .
+qed
+
+
+subsubsection {* Well-order succ-lim induction *}
+
+lemma ord_cases:
+obtains j where "i = succ j" and "aboveS j \<noteq> {}"  | "isLim i"
+by (metis isLim_def isSucc_def)
+
+lemma well_order_inductSL[case_names Suc Lim]:
+assumes SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
+LIM: "\<And>i. \<lbrakk>isLim i; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
+shows "P i"
+proof(induction rule: well_order_induct)
+  fix i assume 0:  "\<forall>j. j \<noteq> i \<and> (j, i) \<in> r \<longrightarrow> P j"
+  show "P i" proof(cases i rule: ord_cases)
+    fix j assume i: "i = succ j" and j: "aboveS j \<noteq> {}"
+    hence "j \<noteq> i \<and> (j, i) \<in> r" by (metis  succ_diff succ_in)
+    hence 1: "P j" using 0 by simp
+    show "P i" unfolding i apply(rule SUCC) using 1 j by auto
+  qed(insert 0 LIM, unfold underS_def, auto)
+qed
+
+lemma well_order_inductZSL[case_names Zero Suc Lim]:
+assumes ZERO: "P zero"
+and SUCC: "\<And>i. \<lbrakk>aboveS i \<noteq> {}; P i\<rbrakk> \<Longrightarrow> P (succ i)" and
+LIM: "\<And>i. \<lbrakk>isLim i; i \<noteq> zero; \<And>j. j \<in> underS i \<Longrightarrow> P j\<rbrakk> \<Longrightarrow> P i"
+shows "P i"
+apply(induction rule: well_order_inductSL) using assms by auto
+
+(* Succesor and limit ordinals *)
+definition "isSuccOrd \<equiv> \<exists> j \<in> Field r. \<forall> i \<in> Field r. (i,j) \<in> r"
+definition "isLimOrd \<equiv> \<not> isSuccOrd"
+
+lemma isLimOrd_succ:
+assumes isLimOrd and "i \<in> Field r"
+shows "succ i \<in> Field r"
+using assms unfolding isLimOrd_def isSuccOrd_def
+by (metis REFL in_notinI refl_on_domain succ_smallest)
+
+lemma isLimOrd_aboveS:
+assumes l: isLimOrd and i: "i \<in> Field r"
+shows "aboveS i \<noteq> {}"
+proof-
+  obtain j where "j \<in> Field r" and "(j,i) \<notin> r"
+  using assms unfolding isLimOrd_def isSuccOrd_def by auto
+  hence "(i,j) \<in> r \<and> j \<noteq> i" by (metis i max2_def max2_greater)
+  thus ?thesis unfolding aboveS_def by auto
+qed
+
+lemma succ_aboveS_isLimOrd:
+assumes "\<forall> i \<in> Field r. aboveS i \<noteq> {} \<and> succ i \<in> Field r"
+shows isLimOrd
+unfolding isLimOrd_def isSuccOrd_def proof safe
+  fix j assume j: "j \<in> Field r" "\<forall>i\<in>Field r. (i, j) \<in> r"
+  hence "(succ j, j) \<in> r" using assms by auto
+  moreover have "aboveS j \<noteq> {}" using assms j unfolding aboveS_def by auto
+  ultimately show False by (metis succ_not_in)
+qed
+
+lemma isLim_iff:
+assumes l: "isLim i" and j: "j \<in> underS i"
+shows "\<exists> k. k \<in> underS i \<and> j \<in> underS k"
+proof-
+  have a: "aboveS j \<noteq> {}" using j unfolding underS_def aboveS_def by auto
+  show ?thesis apply(rule exI[of _ "succ j"]) apply safe
+  using assms a unfolding underS_def isLim_def
+  apply (metis (lifting, full_types) isSucc_def mem_Collect_eq succ_smallest)
+  by (metis (lifting, full_types) a mem_Collect_eq succ_diff succ_in)
+qed
+
+end (* context wo_rel *)
+
+abbreviation "zero \<equiv> wo_rel.zero"
+abbreviation "succ \<equiv> wo_rel.succ"
+abbreviation "pred \<equiv> wo_rel.pred"
+abbreviation "isSucc \<equiv> wo_rel.isSucc"
+abbreviation "isLim \<equiv> wo_rel.isLim"
+abbreviation "isLimOrd \<equiv> wo_rel.isLimOrd"
+abbreviation "isSuccOrd \<equiv> wo_rel.isSuccOrd"
+abbreviation "adm_woL \<equiv> wo_rel.adm_woL"
+abbreviation "worecSL \<equiv> wo_rel.worecSL"
+abbreviation "worecZSL \<equiv> wo_rel.worecZSL"
+
+
+subsection {* Projections of wellorders *}
+
+definition "oproj r s f \<equiv> Field s \<subseteq> f ` (Field r) \<and> compat r s f"
+
+lemma oproj_in:
+assumes "oproj r s f" and "(a,a') \<in> r"
+shows "(f a, f a') \<in> s"
+using assms unfolding oproj_def compat_def by auto
+
+lemma oproj_Field:
+assumes f: "oproj r s f" and a: "a \<in> Field r"
+shows "f a \<in> Field s"
+using oproj_in[OF f] a unfolding Field_def by auto
+
+lemma oproj_Field2:
+assumes f: "oproj r s f" and a: "b \<in> Field s"
+shows "\<exists> a \<in> Field r. f a = b"
+using assms unfolding oproj_def by auto
+
+lemma oproj_under:
+assumes f:  "oproj r s f" and a: "a \<in> under r a'"
+shows "f a \<in> under s (f a')"
+using oproj_in[OF f] a unfolding under_def by auto
+
+(* An ordinal is embedded in another whenever it is embedded as an order
+(not necessarily as initial segment):*)
+theorem embedI:
+assumes r: "Well_order r" and s: "Well_order s"
+and f: "\<And> a. a \<in> Field r \<Longrightarrow> f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a)"
+shows "\<exists> g. embed r s g"
+proof-
+  interpret r!: wo_rel r by unfold_locales (rule r)
+  interpret s!: wo_rel s by unfold_locales (rule s)
+  let ?G = "\<lambda> g a. suc s (g ` underS r a)"
+  def g \<equiv> "worec r ?G"
+  have adm: "adm_wo r ?G" unfolding r.adm_wo_def image_def by auto
+  (*  *)
+  {fix a assume "a \<in> Field r"
+   hence "bij_betw g (under r a) (under s (g a)) \<and>
+          g a \<in> under s (f a)"
+   proof(induction a rule: r.underS_induct)
+     case (1 a)
+     hence a: "a \<in> Field r"
+     and IH1a: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> inj_on g (under r a1)"
+     and IH1b: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g ` under r a1 = under s (g a1)"
+     and IH2: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> under s (f a1)"
+     unfolding underS_def Field_def bij_betw_def by auto
+     have fa: "f a \<in> Field s" using f[OF a] by auto
+     have g: "g a = suc s (g ` underS r a)"
+     using r.worec_fixpoint[OF adm] unfolding g_def fun_eq_iff by simp
+     have A0: "g ` underS r a \<subseteq> Field s"
+     using IH1b by (metis IH2 image_subsetI in_mono under_Field)
+     {fix a1 assume a1: "a1 \<in> underS r a"
+      from IH2[OF this] have "g a1 \<in> under s (f a1)" .
+      moreover have "f a1 \<in> underS s (f a)" using f[OF a] a1 by auto
+      ultimately have "g a1 \<in> underS s (f a)" by (metis s.ANTISYM s.TRANS under_underS_trans)
+     }
+     hence "f a \<in> AboveS s (g ` underS r a)" unfolding AboveS_def
+     using fa by simp (metis (lifting, full_types) mem_Collect_eq underS_def)
+     hence A: "AboveS s (g ` underS r a) \<noteq> {}" by auto
+     have B: "\<And> a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (g a)"
+     unfolding g apply(rule s.suc_underS[OF A0 A]) by auto
+     {fix a1 a2 assume a2: "a2 \<in> underS r a" and 1: "a1 \<in> underS r a2"
+      hence a12: "{a1,a2} \<subseteq> under r a2" and "a1 \<noteq> a2" using r.REFL a
+      unfolding underS_def under_def refl_on_def Field_def by auto
+      hence "g a1 \<noteq> g a2" using IH1a[OF a2] unfolding inj_on_def by auto
+      hence "g a1 \<in> underS s (g a2)" using IH1b[OF a2] a12
+      unfolding underS_def under_def by auto
+     } note C = this
+     have ga: "g a \<in> Field s" unfolding g using s.suc_inField[OF A0 A] .
+     have aa: "a \<in> under r a"
+     using a r.REFL unfolding under_def underS_def refl_on_def by auto
+     show ?case proof safe
+       show "bij_betw g (under r a) (under s (g a))" unfolding bij_betw_def proof safe
+         show "inj_on g (under r a)" proof(rule r.inj_on_Field)
+           fix a1 a2 assume "a1 \<in> under r a" and a2: "a2 \<in> under r a" and a1: "a1 \<in> underS r a2"
+           hence a22: "a2 \<in> under r a2" and a12: "a1 \<in> under r a2" "a1 \<noteq> a2"
+           using a r.REFL unfolding under_def underS_def refl_on_def by auto
+           show "g a1 \<noteq> g a2"
+           proof(cases "a2 = a")
+             case False hence "a2 \<in> underS r a"
+             using a2 unfolding underS_def under_def by auto
+             from IH1a[OF this] show ?thesis using a12 a22 unfolding inj_on_def by auto
+           qed(insert B a1, unfold underS_def, auto)
+         qed(unfold under_def Field_def, auto)
+       next
+         fix a1 assume a1: "a1 \<in> under r a"
+         show "g a1 \<in> under s (g a)"
+         proof(cases "a1 = a")
+           case True thus ?thesis
+           using ga s.REFL unfolding refl_on_def under_def by auto
+         next
+           case False
+           hence a1: "a1 \<in> underS r a" using a1 unfolding underS_def under_def by auto
+           thus ?thesis using B unfolding underS_def under_def by auto
+         qed
+       next
+         fix b1 assume b1: "b1 \<in> under s (g a)"
+         show "b1 \<in> g ` under r a"
+         proof(cases "b1 = g a")
+           case True thus ?thesis using aa by auto
+         next
+           case False
+           hence "b1 \<in> underS s (g a)" using b1 unfolding underS_def under_def by auto
+           from s.underS_suc[OF this[unfolded g] A0]
+           obtain a1 where a1: "a1 \<in> underS r a" and b1: "b1 \<in> under s (g a1)" by auto
+           obtain a2 where "a2 \<in> under r a1" and b1: "b1 = g a2" using IH1b[OF a1] b1 by auto
+           hence "a2 \<in> under r a" using a1
+           by (metis r.ANTISYM r.TRANS in_mono underS_subset_under under_underS_trans)
+           thus ?thesis using b1 by auto
+         qed
+       qed
+     next
+       have "(g a, f a) \<in> s" unfolding g proof(rule s.suc_least[OF A0])
+         fix b1 assume "b1 \<in> g ` underS r a"
+         then obtain a1 where a1: "b1 = g a1" and a1: "a1 \<in> underS r a" by auto
+         hence "b1 \<in> underS s (f a)"
+         using a by (metis `\<And>a1. a1 \<in> underS r a \<Longrightarrow> g a1 \<in> underS s (f a)`)
+         thus "f a \<noteq> b1 \<and> (b1, f a) \<in> s" unfolding underS_def by auto
+       qed(insert fa, auto)
+       thus "g a \<in> under s (f a)" unfolding under_def by auto
+     qed
+   qed
+  }
+  thus ?thesis unfolding embed_def by auto
+qed
+
+corollary ordLeq_def2:
+  "r \<le>o s \<longleftrightarrow> Well_order r \<and> Well_order s \<and>
+               (\<exists> f. \<forall> a \<in> Field r. f a \<in> Field s \<and> f ` underS r a \<subseteq> underS s (f a))"
+using embed_in_Field[of r s] embed_underS2[of r s] embedI[of r s]
+unfolding ordLeq_def by fast
+
+lemma iso_oproj:
+  assumes r: "Well_order r" and s: "Well_order s" and f: "iso r s f"
+  shows "oproj r s f"
+using assms unfolding oproj_def
+by (subst (asm) iso_iff3) (auto simp: bij_betw_def)
+
+theorem oproj_embed:
+assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
+shows "\<exists> g. embed s r g"
+proof (rule embedI[OF s r, of "inv_into (Field r) f"], unfold underS_def, safe)
+  fix b assume "b \<in> Field s"
+  thus "inv_into (Field r) f b \<in> Field r" using oproj_Field2[OF f] by (metis imageI inv_into_into)
+next
+  fix a b assume "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
+    "inv_into (Field r) f a = inv_into (Field r) f b"
+  with f show False by (auto dest!: inv_into_injective simp: Field_def oproj_def)
+next
+  fix a b assume *: "b \<in> Field s" "a \<noteq> b" "(a, b) \<in> s"
+  { assume "(inv_into (Field r) f a, inv_into (Field r) f b) \<notin> r"
+    moreover
+    from *(3) have "a \<in> Field s" unfolding Field_def by auto
+    with *(1,2) have
+      "inv_into (Field r) f a \<in> Field r" "inv_into (Field r) f b \<in> Field r"
+      "inv_into (Field r) f a \<noteq> inv_into (Field r) f b"
+      by (auto dest!: oproj_Field2[OF f] inv_into_injective intro!: inv_into_into)
+    ultimately have "(inv_into (Field r) f b, inv_into (Field r) f a) \<in> r"
+      using r by (auto simp: well_order_on_def linear_order_on_def total_on_def)
+    with f[unfolded oproj_def compat_def] *(1) `a \<in> Field s`
+      f_inv_into_f[of b f "Field r"] f_inv_into_f[of a f "Field r"]
+      have "(b, a) \<in> s" by (metis in_mono)
+    with *(2,3) s have False
+      by (auto simp: well_order_on_def linear_order_on_def partial_order_on_def antisym_def)
+  } thus "(inv_into (Field r) f a, inv_into (Field r) f b) \<in> r" by blast
+qed
+
+corollary oproj_ordLeq:
+assumes r: "Well_order r" and s: "Well_order s" and f: "oproj r s f"
+shows "s \<le>o r"
+using oproj_embed[OF assms] r s unfolding ordLeq_def by auto
+
+end