Theory Szpilrajn.Szpilrajn

(*<*)
theory Szpilrajn 
  imports Main
begin
  (*>*)

text ‹
  We formalize a more general version of Szpilrajn's extension theorem~cite"Szpilrajn:1930",
  employing the terminology of Bossert and Suzumura~cite"Bossert:2010". We also formalize 
  Theorem 2.7 of their book. Our extension theorem states that any preorder can be extended to a
  total preorder while maintaining its structure. The proof of the extension theorem follows the
  proof presented in the Wikipedia article~citeWiki.
›

section ‹Definitions›

subsection ‹Symmetric and asymmetric factor of a relation›

text ‹
  According to Bossert and Suzumura, every relation can be partitioned into its symmetric
  and asymmetric factor. The symmetric factor of a relation termr contains all pairs
  term(x, y)  r where term(y, x)  r. Conversely, the asymmetric factor contains all pairs
   where this is not the case. In terms of an order term(≤), the asymmetric factor contains all
  term(x, y)  {(x, y) |x y. x  y} where termx < y.
›
definition sym_factor :: "'a rel  'a rel"
  where "sym_factor r  {(x, y)  r. (y, x)  r}"

lemma sym_factor_def': "sym_factor r = r  r¯"
  unfolding sym_factor_def by fast

definition asym_factor :: "'a rel  'a rel"
  where "asym_factor r = {(x, y)  r. (y, x)  r}"


subsubsection ‹Properties of the symmetric factor›

lemma sym_factorI[intro]: "(x, y)  r  (y, x)  r  (x, y)  sym_factor r"
  unfolding sym_factor_def by blast

lemma sym_factorE[elim?]:
  assumes "(x, y)  sym_factor r" obtains "(x, y)  r" "(y, x)  r"
  using assms[unfolded sym_factor_def] by blast

lemma sym_sym_factor[simp]: "sym (sym_factor r)"
  unfolding sym_factor_def
  by (auto intro!: symI) 

lemma trans_sym_factor[simp]: "trans r  trans (sym_factor r)"
  unfolding sym_factor_def' using trans_Int by force

lemma refl_on_sym_factor[simp]: "refl_on A r  refl_on A (sym_factor r)"
  unfolding sym_factor_def
  by (auto intro!: refl_onI dest: refl_onD refl_onD1)

lemma sym_factor_absorb_if_sym[simp]: "sym r  sym_factor r = r"
  unfolding sym_factor_def'
  by (simp add: sym_conv_converse_eq)

lemma sym_factor_idem[simp]: "sym_factor (sym_factor r) = sym_factor r"
  using sym_factor_absorb_if_sym[OF sym_sym_factor] .

lemma sym_factor_reflc[simp]: "sym_factor (r=) = (sym_factor r)="
  unfolding sym_factor_def by auto

lemma sym_factor_Restr[simp]: "sym_factor (Restr r A) = Restr (sym_factor r) A"
  unfolding sym_factor_def by blast

text ‹
  In contrast to termasym_factor, the termsym_factor is monotone.
›
lemma sym_factor_mono: "r  s  sym_factor r  sym_factor s"
  unfolding sym_factor_def by auto


subsubsection ‹Properties of the asymmetric factor›

lemma asym_factorI[intro]: "(x, y)  r  (y, x)  r  (x, y)  asym_factor r"
  unfolding asym_factor_def by blast

lemma asym_factorE[elim?]:
  assumes "(x, y)  asym_factor r" obtains "(x, y)  r"
  using assms unfolding asym_factor_def by blast

lemma refl_not_in_asym_factor[simp]: "(x, x)  asym_factor r"
  unfolding asym_factor_def by blast

lemma irrefl_asym_factor[simp]: "irrefl (asym_factor r)"
  unfolding asym_factor_def irrefl_def by fast

lemma asym_asym_factor[simp]: "asym (asym_factor r)"
  using irrefl_asym_factor
  by (auto intro!: asymI simp: asym_factor_def)

lemma trans_asym_factor[simp]: "trans r  trans (asym_factor r)"
  unfolding asym_factor_def trans_def by fast

lemma asym_if_irrefl_trans: "irrefl r  trans r  asym r"
  by (intro asymI) (auto simp: irrefl_def trans_def)

lemma antisym_if_irrefl_trans: "irrefl r  trans r  antisym r"
  using antisym_def asym_if_irrefl_trans by (auto dest: asymD)
    
lemma asym_factor_asym_rel[simp]: "asym r  asym_factor r = r"
  unfolding asym_factor_def
  by (auto dest: asymD)

lemma irrefl_trans_asym_factor_id[simp]: "irrefl r  trans r  asym_factor r = r"
  using asym_factor_asym_rel[OF asym_if_irrefl_trans] .

lemma asym_factor_id[simp]: "asym_factor (asym_factor r) = asym_factor r"
  using asym_factor_asym_rel[OF asym_asym_factor] .

lemma asym_factor_rtrancl: "asym_factor (r*) = asym_factor (r+)"
  unfolding asym_factor_def
  by (auto simp add: rtrancl_eq_or_trancl)

lemma asym_factor_Restr[simp]: "asym_factor (Restr r A) = Restr (asym_factor r) A"
  unfolding asym_factor_def by blast

lemma acyclic_asym_factor[simp]: "acyclic r  acyclic (asym_factor r)"
  unfolding asym_factor_def by (auto intro: acyclic_subset)


subsubsection ‹Relations between symmetric and asymmetric factor›

text ‹
  We prove that termsym_factor and termasym_factor partition the input relation.
›
lemma sym_asym_factor_Un: "sym_factor r  asym_factor r = r"
  unfolding sym_factor_def asym_factor_def by blast

lemma disjnt_sym_asym_factor[simp]: "disjnt (sym_factor r) (asym_factor r)"
  unfolding disjnt_def
  unfolding sym_factor_def asym_factor_def by blast

lemma Field_sym_asym_factor_Un:
  "Field (sym_factor r)  Field (asym_factor r) = Field r"
  using sym_asym_factor_Un Field_Un by metis

lemma asym_factor_tranclE:
  assumes "(a, b)  (asym_factor r)+" shows "(a, b)  r+"
  using assms sym_asym_factor_Un
  by (metis UnCI subsetI trancl_mono)


subsection ‹Extension of Orders›

text ‹
  We use the definition of Bossert and Suzumura for extends›. The requirement termr  R is
  obvious. The second requirement termasym_factor r  asym_factor R enforces that the 
  extension termR maintains all strict preferences of termr (viewing termr as a 
  preference relation).
›
                    
definition extends :: "'a rel  'a rel  bool"
  where "extends R r  r  R  asym_factor r  asym_factor R"

text ‹
  We define a stronger notion of termextends where we also demand that
  termsym_factor R  (sym_factor r)=. This enforces that the extension does not introduce
  preference cycles between previously unrelated pairs term(x, y)  R - r.
›

definition strict_extends :: "'a rel  'a rel  bool"
  where "strict_extends R r  extends R r  sym_factor R  (sym_factor r)="

lemma extendsI[intro]: "r  R  asym_factor r  asym_factor R  extends R r"
  unfolding extends_def by (intro conjI)

lemma extendsE:
  assumes "extends R r"
  obtains "r  R" "asym_factor r  asym_factor R"
  using assms unfolding extends_def by blast

lemma trancl_subs_extends_if_trans: "extends r_ext r  trans r_ext  r+  r_ext"
  unfolding extends_def asym_factor_def
  by (metis subrelI trancl_id trancl_mono)

lemma extends_if_strict_extends: "strict_extends r_ext ext  extends r_ext ext"
  unfolding strict_extends_def by blast

lemma strict_extendsI[intro]:
  assumes "r  R" "asym_factor r  asym_factor R" "sym_factor R  (sym_factor r)="
  shows "strict_extends R r"
  unfolding strict_extends_def using assms by (intro conjI extendsI)

lemma strict_extendsE:
  assumes "strict_extends R r"
  obtains "r  R" "asym_factor r  asym_factor R" "sym_factor R  (sym_factor r)="
  using assms extendsE unfolding strict_extends_def by blast

lemma strict_extends_antisym_Restr:
  assumes "strict_extends R r"
  assumes "antisym (Restr r A)"
  shows "antisym ((R - r)  Restr r A)"
proof(rule antisymI, rule ccontr)
  fix x y assume "(x, y)  (R - r)  Restr r A" "(y, x)  (R - r)  Restr r A" "x  y"
  with strict_extends R r have "(x, y)  sym_factor R"
    unfolding sym_factor_def by (auto elim!: strict_extendsE)
  with assms x  y have "(x, y)  sym_factor r"
    by (auto elim!: strict_extendsE)
  then have "(x, y)  r" "(y, x)  r"
    unfolding sym_factor_def by simp_all
  with antisym (Restr r A) x  y (y, x)  R - r  Restr r A show False
    using antisymD by fastforce
qed

text ‹Here we prove that we have no preference cycles between previously unrelated pairs.›
lemma antisym_Diff_if_strict_extends:
  assumes "strict_extends R r"
  shows "antisym (R - r)"
  using strict_extends_antisym_Restr[OF assms, where ?A="{}"] by simp

lemma strict_extends_antisym:
  assumes "strict_extends R r"
  assumes "antisym r"
  shows "antisym R"
  using assms strict_extends_antisym_Restr[OF assms(1), where ?A=UNIV]
  by (auto elim!: strict_extendsE simp: antisym_def) 

lemma strict_extends_if_strict_extends_reflc:
  assumes "strict_extends r_ext (r=)"
  shows "strict_extends r_ext r"
proof(intro strict_extendsI)
  from assms show "r  r_ext"
    by (auto elim: strict_extendsE)

  from assms r  r_ext show "asym_factor r  asym_factor r_ext"
    unfolding strict_extends_def
    by (auto simp: asym_factor_def sym_factor_def)

  from assms show "sym_factor r_ext  (sym_factor r)="
    by (auto simp: sym_factor_def strict_extends_def)
qed

lemma strict_extends_diff_Id:
  assumes "irrefl r" "trans r"
  assumes "strict_extends r_ext (r=)"
  shows "strict_extends (r_ext - Id) r"
proof(intro strict_extendsI)
  from assms show "r  r_ext - Id"
    by (auto elim: strict_extendsE simp: irrefl_def)

  note antisym_r = antisym_if_irrefl_trans[OF assms(1,2)]
  with assms strict_extends_if_strict_extends_reflc show "asym_factor r  asym_factor (r_ext - Id)"
    unfolding asym_factor_def
    by (auto intro: strict_extends_antisym[THEN antisymD] elim: strict_extendsE transE)

  from assms antisym_r show "sym_factor (r_ext - Id)  (sym_factor r)="
    unfolding sym_factor_def
    by (auto intro: strict_extends_antisym[THEN antisymD])
qed

text ‹
  Both termextends and termstrict_extends form a partial order since they
  are reflexive, transitive, and antisymmetric.
›
lemma shows
    reflp_extends: "reflp extends" and
    transp_extends: "transp extends" and
    antisymp_extends: "antisymp extends"
  unfolding extends_def reflp_def transp_def antisymp_def
  by auto

lemma shows
    reflp_strict_extends: "reflp strict_extends" and
    transp_strict_extends: "transp strict_extends" and
    antisymp_strict_extends: "antisymp strict_extends"
  using reflp_extends transp_extends antisymp_extends
  unfolding strict_extends_def reflp_def transp_def antisymp_def
  by auto

subsection ‹Missing order definitions›

lemma preorder_onD[dest?]:
  assumes "preorder_on A r"
  shows "refl_on A r" "trans r"
  using assms unfolding preorder_on_def by blast+

lemma preorder_onI[intro]: "refl_on A r  trans r  preorder_on A r"
  unfolding preorder_on_def by (intro conjI)

abbreviation "preorder  preorder_on UNIV"

lemma preorder_rtrancl: "preorder (r*)"
  by (intro preorder_onI refl_rtrancl trans_rtrancl)

definition "total_preorder_on A r  preorder_on A r  total_on A r"

abbreviation "total_preorder r  total_preorder_on UNIV r"

lemma total_preorder_onI[intro]:
  "refl_on A r  trans r  total_on A r  total_preorder_on A r"
  unfolding total_preorder_on_def by (intro conjI preorder_onI)

lemma total_preorder_onD[dest?]:
  assumes "total_preorder_on A r"
  shows "refl_on A r" "trans r" "total_on A r"
  using assms unfolding total_preorder_on_def preorder_on_def by blast+

definition "strict_partial_order r  trans r  irrefl r"

lemma strict_partial_orderI[intro]:
  "trans r  irrefl r  strict_partial_order r"
  unfolding strict_partial_order_def by blast

lemma strict_partial_orderD[dest?]:
  assumes "strict_partial_order r"
  shows "trans r" "irrefl r"
  using assms unfolding strict_partial_order_def by blast+

lemma strict_partial_order_acyclic:
  assumes "strict_partial_order r"
  shows "acyclic r"
  by (metis acyclic_irrefl assms strict_partial_order_def trancl_id)


abbreviation "partial_order  partial_order_on UNIV"

lemma partial_order_onI[intro]:
  "refl_on A r  trans r  antisym r  partial_order_on A r"
  using partial_order_on_def by blast

lemma linear_order_onI[intro]:
  "refl_on A r  trans r  antisym r  total_on A r  linear_order_on A r"
  using linear_order_on_def by blast

lemma linear_order_onD[dest?]:
  assumes "linear_order_on A r"
  shows "refl_on A r" "trans r" "antisym r" "total_on A r"
  using assms[unfolded linear_order_on_def] partial_order_onD by blast+

text ‹A typical example is term(⊂) on sets:›

lemma strict_partial_order_subset:
  "strict_partial_order {(x,y). x  y}"
proof
  show "trans {(x,y). x  y}"
    by (auto simp add: trans_def)
  show "irrefl {(x, y). x  y}"
    by (simp add: irrefl_def)
qed

text ‹We already have a definition of a strict linear order in termstrict_linear_order.›

section ‹Extending preorders to total preorders›

text ‹
  We start by proving that a preorder with two incomparable elements termx and termy can be
  strictly extended to a preorder where termx < y.
›

lemma can_extend_preorder: 
  assumes "preorder_on A r"
    and "y  A" "x  A" "(y, x)  r"
  shows
    "preorder_on A ((insert (x, y) r)+)" "strict_extends ((insert (x, y) r)+) r"
proof -
  note preorder_onD[OF preorder_on A r]
  then have "insert (x, y) r  A × A"
    using y  A x  A refl_on_domain by fast
  with refl_on A r show "preorder_on A ((insert (x, y) r)+)"
    by (intro preorder_onI refl_onI trans_trancl)
       (auto simp: trancl_subset_Sigma intro!: r_into_trancl' dest: refl_onD)

  show "strict_extends ((insert (x, y) r)+) r"
  proof(intro strict_extendsI)
    from preorder_onD(2)[OF preorder_on A r] (y, x)  r
    show "asym_factor r  asym_factor ((insert (x, y) r)+)"
       unfolding asym_factor_def trancl_insert
       using rtranclD rtrancl_into_trancl1 r_r_into_trancl
       by fastforce

     from assms have "(y, x)  (insert (x, y) r)+"
       unfolding preorder_on_def trancl_insert
       using refl_onD rtranclD by fastforce
     with trans r show "sym_factor ((insert (x, y) r)+)  (sym_factor r)="
       unfolding trancl_insert sym_factor_def by (fastforce intro: rtrancl_trans)
  qed auto
qed


text ‹
  With this, we can start the proof of our main extension theorem.
  For this we will use a variant of Zorns Lemma, which only considers nonempty chains:
›
lemma Zorns_po_lemma_nonempty:
  assumes po: "Partial_order r"
    and u: "C. C  Chains r; C{}  uField r. aC. (a, u)  r"
    and "r  {}"
  shows "mField r. aField r. (m, a)  r  a = m"
proof -
  from r  {} obtain x where "x  Field r"
    using FieldI2 by fastforce
  with assms show ?thesis
    using Zorns_po_lemma by (metis empty_iff)  
qed


theorem strict_extends_preorder_on:
  assumes "preorder_on A base_r"
  shows "r. total_preorder_on A r  strict_extends r base_r" 
proof -

  text ‹
    We define an order on the set of strict extensions of the base relation termbase_r, 
    where termr  s iff termstrict_extends r base_r and termstrict_extends s r:
  ›

  define order_of_orders :: "('a rel) rel" where "order_of_orders =
    Restr {(r, s). strict_extends r base_r  strict_extends s r} {r. preorder_on A r}"

  text ‹
    We show that this order consists of those relations that are preorders and that strictly extend
    the base relation termbase_r

  have Field_order_of_orders: "Field order_of_orders =
    {r. preorder_on A r  strict_extends r base_r}"
    using transp_strict_extends
  proof(safe)
    fix r assume "preorder_on A r" "strict_extends r base_r"
    with reflp_strict_extends have
      "(r, r)  {(r, s). strict_extends r base_r  strict_extends s r}"
      by (auto elim!: reflpE)
    with preorder_on A r show "r  Field order_of_orders"
      unfolding order_of_orders_def by (auto simp: Field_def)
  qed (auto simp: order_of_orders_def Field_def elim: transpE)

  text ‹
    We now show that this set has a maximum and that any maximum of this set is a total preorder
    and as thus is one of the extensions we are looking for.
    We begin by showing the existence of a maximal element using Zorn's lemma.
  ›

  have "m  Field order_of_orders.
      a  Field order_of_orders. (m, a)  order_of_orders  a = m"
  proof (rule Zorns_po_lemma_nonempty)

    text ‹
      Zorn's Lemma requires us to prove that our termorder_of_orders is a nonempty partial order
      and that every nonempty chain has an upper bound. 
      The partial order property is trivial, since we used termstrict_extends for the relation, 
      which is a partial order as shown above.
    ›

    from reflp_strict_extends transp_strict_extends
    have "Refl {(r, s). strict_extends r base_r  strict_extends s r}"
      unfolding refl_on_def Field_def by (auto elim: transpE reflpE)
    moreover have "trans {(r, s). strict_extends r base_r  strict_extends s r}"
      using transp_strict_extends  by (auto elim: transpE intro: transI)
    moreover have "antisym {(r, s). strict_extends r base_r  strict_extends s r}"
      using antisymp_strict_extends by (fastforce dest: antisympD intro: antisymI)

    ultimately show "Partial_order order_of_orders"
      unfolding order_of_orders_def order_on_defs
      using Field_order_of_orders Refl_Restr trans_Restr antisym_Restr
      by blast

    text ‹Also, our order is obviously not empty since it contains term(base_r, base_r):›

    have "(base_r, base_r)  order_of_orders"
      unfolding order_of_orders_def
      using assms reflp_strict_extends by (auto dest: reflpD)
    thus "order_of_orders  {}" by force


    text ‹
      Next we show that each chain has an upper bound.
      For the upper bound we take the union of all relations in the chain.
    ›

    show "u  Field order_of_orders. a  C. (a, u)  order_of_orders" 
      if C_def: "C  Chains order_of_orders" and C_nonempty: "C  {}"
      for C
    proof (rule bexI[where x="C"])

      text ‹
        Obviously each element in the chain is a strict extension of termbase_r by definition
        and as such it is also a preorder.
      ›

      have preorder_r: "preorder_on A r" and extends_r: "strict_extends r base_r" if "r  C" for r
        using that C_def[unfolded order_of_orders_def Chains_def] by blast+

      text ‹
        Because a chain is partially ordered, the union of the chain is reflexive and transitive.
      ›

      have total_subs_C: "r  s  s  r" if "r  C" and "s  C" for r s
        using C_def that
        unfolding Chains_def order_of_orders_def strict_extends_def extends_def
        by blast

      have preorder_UnC: "preorder_on A (C)"
      proof(intro preorder_onI)
        show "refl_on A (C)"
          using preorder_onD(1)[OF preorder_r] C_nonempty
          unfolding refl_on_def by auto

        from total_subs_C show "trans (C)"
          using chain_subset_trans_Union[unfolded chain_subset_def]
          by (metis preorder_onD(2)[OF preorder_r])
      qed

      text ‹We show that termC strictly extends the base relation.›
    
      have strict_extends_UnC: "strict_extends (C) base_r"
      proof(intro strict_extendsI)
        note extends_r_unfolded = extends_r[unfolded extends_def strict_extends_def]

        show "base_r  (C)"
          using C_nonempty extends_r_unfolded
          by blast

        then show "asym_factor base_r  asym_factor (C)"
          using extends_r_unfolded
          unfolding asym_factor_def by auto

        show "sym_factor (C)  (sym_factor base_r)="
        proof(safe)
          fix x y assume "(x, y)  sym_factor (C)" "(x, y)  sym_factor base_r"
          then have "(x, y)  C" "(y, x)  C"
            unfolding sym_factor_def by blast+

          with extends_r obtain c where "c  C" "(x, y)  c" "(y, x)  c"
            "strict_extends c base_r"
            using total_subs_C by blast
          then have "(x, y)  sym_factor c"
            unfolding sym_factor_def by blast
          with strict_extends c base_r (x, y)  sym_factor base_r
          show "x = y"
            unfolding strict_extends_def by blast
        qed
      qed

      from preorder_UnC strict_extends_UnC show "(C)  Field order_of_orders"
        unfolding Field_order_of_orders by simp

      text ‹
        Lastly, we prove by contradiction that termC is an upper bound for the chain.
      ›

      show "a  C. (a, C)  order_of_orders"
      proof(rule ccontr)
        presume "a  C. (a, C)  order_of_orders"
        then obtain m where m: "m  C" "(m, C)  order_of_orders"
          by blast

        hence strict_extends_m: "strict_extends m base_r" "preorder_on A m"
          using extends_r preorder_r by blast+
        with m have "¬ strict_extends (C) m"
          using preorder_UnC unfolding order_of_orders_def by blast

        from m have "m  C"
          by blast
        moreover
        have "sym_factor (C)  (sym_factor m)="
        proof(safe)
          fix a b
          assume "(a, b)  sym_factor ( C)" "(a, b)  sym_factor m"
          then have "(a, b)  sym_factor base_r  (a, b)  Id"
            using strict_extends_UnC[unfolded strict_extends_def] by blast
          with (a, b)  sym_factor m strict_extends_m(1) show "a = b"
            by (auto elim: strict_extendsE simp: sym_factor_mono[THEN in_mono])
        qed
        ultimately
        have "¬ asym_factor m  asym_factor (C)"
          using ¬ strict_extends (C) m unfolding strict_extends_def extends_def by blast

        then obtain x y where
          "(x, y)  m" "(y, x)  m" "(x, y)  asym_factor m" "(x, y)  asym_factor (C)"
          unfolding asym_factor_def by blast
    
        then obtain w where "w  C" "(y, x)  w"
          unfolding asym_factor_def using m  C by auto

        with (y, x)  m have "¬ extends m w"
          unfolding extends_def by auto
        moreover
        from (x, y)  m have "¬ extends w m"
        proof(cases "(x, y)  w")
          case True
          with (y, x)  w have "(x, y)  asym_factor w"
            unfolding asym_factor_def by simp
          with (x, y)  asym_factor m show "¬ extends w m"
            unfolding extends_def by auto
        qed (auto simp: extends_def)

        ultimately show False
          using m  C w  C
          using C_def[unfolded Chains_def order_of_orders_def strict_extends_def]
          by auto
      qed blast
    qed
  qed

  text ‹Let our maximal element be named termmax:›

  from this obtain max 
    where max_field: "max  Field order_of_orders"
      and is_max: 
        "aField order_of_orders. (max, a)  order_of_orders  a = max"
    by auto

  from max_field have max_extends_base: "preorder_on A max" "strict_extends max base_r"
    using Field_order_of_orders by blast+

  text ‹
    We still have to show, that termmax is a strict linear order,
    meaning that it is also a total order:
  ›

  have "total_on A max"
  proof
    fix x y :: 'a
    assume "x  y" "x  A" "y  A"

    show "(x, y)  max  (y, x)  max"
    proof (rule ccontr)

      text ‹
        Assume that termmax is not total, and termx and termy are incomparable.
        Then we can extend termmax by setting $x < y$:
      ›

      presume "(x, y)  max" and "(y, x)  max"
      let ?max' = "(insert (x, y) max)+"

      note max'_extends_max = can_extend_preorder[OF
          preorder_on A max y  A x  A (y, x)  max]

      hence max'_extends_base: "strict_extends ?max' base_r"
        using strict_extends max base_r transp_strict_extends by (auto elim: transpE)


      text ‹The extended relation is greater than termmax, which is a contradiction.›

      have "(max, ?max')  order_of_orders"
        using max'_extends_base max'_extends_max max_extends_base
        unfolding order_of_orders_def by simp
      thus False
        using FieldI2 (x, y)  max is_max by fastforce
    qed simp_all
  qed

  with preorder_on A max have "total_preorder_on A max"
    unfolding total_preorder_on_def by simp

  with strict_extends max base_r show "?thesis" by blast
qed

text ‹
  With this extension theorem, we can easily prove Szpilrajn's theorem and its equivalent for
  partial orders.
›

corollary partial_order_extension:
  assumes "partial_order_on A r"
  shows "r_ext. linear_order_on A r_ext  r  r_ext"
proof -
  from assms strict_extends_preorder_on obtain r_ext where r_ext:
    "total_preorder_on A r_ext" "strict_extends r_ext r"
    unfolding partial_order_on_def by blast

  with assms have "antisym r_ext"
    unfolding partial_order_on_def using strict_extends_antisym by blast

  with assms r_ext have "linear_order_on A r_ext  r  r_ext"
    unfolding total_preorder_on_def order_on_defs strict_extends_def extends_def
    by blast
  then show ?thesis ..
qed

corollary Szpilrajn:
  assumes "strict_partial_order r"
  shows "r_ext. strict_linear_order r_ext  r  r_ext"
proof -
  from assms have "partial_order (r=)"
    by (auto simp: antisym_if_irrefl_trans strict_partial_order_def)
  from partial_order_extension[OF this] obtain r_ext where "linear_order r_ext" "(r=)  r_ext"
    by blast
  with assms have "r  r_ext - Id" "strict_linear_order (r_ext - Id)"
    by (auto simp: irrefl_def strict_linear_order_on_diff_Id dest: strict_partial_orderD(2))
  then show ?thesis by blast
qed

corollary acyclic_order_extension:
  assumes "acyclic r"
  shows "r_ext. strict_linear_order r_ext  r  r_ext"
proof -
  from assms have "strict_partial_order (r+)"
    unfolding strict_partial_order_def using acyclic_irrefl trans_trancl by blast
  thus ?thesis
    by (meson Szpilrajn r_into_trancl' subset_iff)
qed

section ‹Consistency›

text ‹
  As a weakening of transitivity, Suzumura introduces the notion of consistency which rules out
  all preference cycles that contain at least one strict preference.
  Consistency characterises those order relations which can be extended (in terms of termextends)
  to a total order relation. 
›

definition consistent :: "'a rel  bool"
  where "consistent r = ((x, y)  r+. (y, x)  asym_factor r)"

lemma consistentI: "(x y. (x, y)  r+  (y, x)  asym_factor r)  consistent r"
  unfolding consistent_def by blast

lemma consistent_if_preorder_on[simp]:
  "preorder_on A r  consistent r"
  unfolding preorder_on_def consistent_def asym_factor_def by auto

lemma consistent_asym_factor[simp]: "consistent r  consistent (asym_factor r)"
  unfolding consistent_def
  using asym_factor_tranclE by fastforce

lemma acyclic_asym_factor_if_consistent[simp]: "consistent r  acyclic (asym_factor r)"
  unfolding consistent_def acyclic_def
  using asym_factor_tranclE by (metis case_prodD trancl.simps)

lemma consistent_Restr[simp]: "consistent r  consistent (Restr r A)"
  unfolding consistent_def asym_factor_def
  using trancl_mono by fastforce

text ‹
  This corresponds to Theorem 2.2~cite"Bossert:2010".
›
theorem trans_if_refl_total_consistent:
  assumes "refl r" "total r" and "consistent r"
  shows "trans r"
proof
  fix x y z assume "(x, y)  r" "(y, z)  r"
  
  from (x, y)  r (y, z)  r have "(x, z)  r+"
    by simp
  hence "(z, x)  asym_factor r"
    using consistent r unfolding consistent_def by blast
  hence "x  z  (x, z)  r"
    unfolding asym_factor_def using total r
    by (auto simp: total_on_def)
  then show "(x, z)  r"
    apply(cases "x = z")
    using refl_onD[OF refl r] by blast+ 
qed


lemma order_extension_if_consistent:
  assumes "consistent r"
  obtains r_ext where "extends r_ext r" "total_preorder r_ext"  
proof -
  from assms have extends: "extends (r*) r"
    unfolding extends_def consistent_def asym_factor_def
    using rtranclD by (fastforce simp: Field_def)
  have preorder: "preorder (r*)"
    unfolding preorder_on_def using refl_on_def trans_def by fastforce

  from strict_extends_preorder_on[OF preorder] extends obtain r_ext where
    "total_preorder r_ext" "extends r_ext r"
    using transpE[OF transp_extends] unfolding strict_extends_def by blast
  then show thesis using that by blast
qed

lemma consistent_if_extends_trans:
  assumes "extends r_ext r" "trans r_ext"
  shows "consistent r"
proof(rule consistentI, standard)
  fix x y assume *: "(x, y)  r+" "(y, x)  asym_factor r"
  with assms have "(x, y)  r_ext"
    using trancl_subs_extends_if_trans[OF assms] by blast
  moreover from * assms have "(x, y)  r_ext"
    unfolding extends_def asym_factor_def by auto
  ultimately show False by blast
qed

text ‹
  With Theorem 2.6~cite"Bossert:2010", we show that termconsistent characterises the existence
  of order extensions.
›
corollary order_extension_iff_consistent:
  "(r_ext. extends r_ext r  total_preorder r_ext)  consistent r"
  using order_extension_if_consistent consistent_if_extends_trans
  by (metis total_preorder_onD(2))


text ‹
  The following theorem corresponds to Theorem 2.7~cite"Bossert:2010".
  Bossert and Suzumura claim that this theorem generalises Szpilrajn's theorem; however, we cannot
  use the theorem to strictly extend a given order termQ. Therefore, it is not strong enough to
  extend a strict partial order to a strict linear order. It works for total preorders (called 
  orderings by Bossert and Suzumura). Unfortunately, we were not able to generalise the theorem
  to allow for strict extensions.
›

theorem general_order_extension_iff_consistent:
  assumes "x y.  x  S; y  S; x  y   (x, y)  Q+"
  assumes "total_preorder_on S Ord"
  shows "(Ext. extends Ext Q  total_preorder Ext  Restr Ext S = Ord)
      consistent Q" (is "?ExExt  _")
proof
  assume "?ExExt"
  then obtain Ext where
    "extends Ext Q"
    "refl Ext" "trans Ext" "total Ext"
    "Restr Ext S = Restr Ord S"
    using total_preorder_onD by fast
  show "consistent Q"
  proof(rule consistentI)
    fix x y assume "(x, y)  Q+"
    with extends Ext Q trans Ext have "(x, y)  Ext"
      unfolding extends_def by (metis trancl_id trancl_mono)
    then have "(y, x)  asym_factor Ext"
      unfolding asym_factor_def by blast
    with extends Ext Q show "(y, x)  asym_factor Q"
      unfolding extends_def asym_factor_def by blast
  qed
next
  assume "consistent Q"

  define Q' where "Q'  Q*  Ord  Ord O Q*  Q* O Ord  (Q* O Ord) O Q*"

  have "refl (Q*)" "trans (Q*)" "refl_on S Ord" "trans Ord" "total_on S Ord"
    using refl_rtrancl trans_rtrancl total_preorder_onD[OF total_preorder_on S Ord]
    by - assumption

  have preorder_Q': "preorder Q'"
  proof
    show "refl Q'"
      unfolding Q'_def refl_on_def by auto

    from trans (Q*) refl_on S Ord trans Ord show "trans Q'"
      unfolding Q'_def[simplified]
      apply(safe intro!: transI)
      unfolding relcomp.simps
      by (metis assms(1) refl_on_domain rtranclD transD)+
  qed

  have "consistent Q'"
    using consistent_if_preorder_on preorder_Q' by blast

  have "extends Q' Q"
  proof(rule extendsI)
    have "Q  Restr (Q*) (Field Q)"
      by (auto intro: FieldI1 FieldI2)
    then show "Q  Q'"
      unfolding Q'_def by blast

    from consistent Q have consistentD: "(x, y)  Q+  (y, x)  Q  (x, y)  Q" for x y
      unfolding consistent_def asym_factor_def using rtranclD by fastforce
    have refl_on_domainE: " (x, y)  Ord; x  S  y  S  P   P" for x y P
      using refl_on_domain[OF refl_on S Ord] by blast

    show "asym_factor Q  asym_factor Q'"
      unfolding Q'_def asym_factor_def Field_def
      apply(safe)
      using assms(1) consistentD refl_on_domainE
      by (metis r_into_rtrancl rtranclD rtrancl_trancl_trancl)+
  qed

  with strict_extends_preorder_on[OF preorder Q']
  obtain Ext where Ext: "extends Ext Q'" "extends Ext Q" "total_preorder Ext"
    unfolding strict_extends_def
    by (metis transpE transp_extends)

  have not_in_Q': "x  S  y  S  (x, y)  Ord  (x, y)  Q'" for x y
    using assms(1) unfolding Q'_def
    apply(safe)
    by (metis refl_on S Ord refl_on_def refl_on_domain rtranclD)+

  have "Restr Ext S = Ord"
  proof
    from extends Ext Q' have "Ord  Ext"
      unfolding Q'_def extends_def by auto
    with refl_on S Ord show "Ord  Restr Ext S"
      using refl_on_domain by fast
  next
    have "(x, y)  Ord" if "x  S" and "y  S" and "(x, y)  Ext" for x y
    proof(rule ccontr)
      assume "(x, y)  Ord"
      with that not_in_Q' have "(x, y)  Q'"
        by blast
      with refl_on S Ord total_on S Ord x  S y  S (x, y)  Ord
      have "(y, x)  Ord"
        unfolding refl_on_def total_on_def by fast
      hence "(y, x)  Q'"
        unfolding Q'_def by blast
      with (x, y)  Q' (y, x)  Q' extends Ext Q'
      have "(x, y)  Ext"
        unfolding extends_def asym_factor_def by auto
      with (x, y)  Ext show False by blast
    qed
    then show "Restr Ext S  Ord"
      by blast
  qed

  with Ext show "?ExExt" by blast
qed

section ‹Strong consistency›

text ‹
  We define a stronger version of termconsistent which requires that the relation does not
  contain hidden preference cycles, i.e. if there is a preference cycle then all the elements
  in the cycle should already be related (in both directions).
  In contrast to consistency which characterises relations that can be extended, strong consistency
  characterises relations that can be extended strictly (cf. termstrict_extends).
›

definition "strongly_consistent r  sym_factor (r+)  sym_factor (r=)"

lemma consistent_if_strongly_consistent: "strongly_consistent r  consistent r"
  unfolding strongly_consistent_def consistent_def
  by (auto simp: sym_factor_def asym_factor_def) 

lemma strongly_consistentI: "sym_factor (r+)  sym_factor (r=)  strongly_consistent r"
  unfolding strongly_consistent_def by blast

lemma strongly_consistent_if_trans_strict_extension:
  assumes "strict_extends r_ext r"
  assumes "trans r_ext"
  shows   "strongly_consistent r"
proof(unfold strongly_consistent_def, standard)
  fix x assume "x  sym_factor (r+)"
  then show "x  sym_factor (r=)"
    using assms trancl_subs_extends_if_trans[OF extends_if_strict_extends]
    by (metis sym_factor_mono strict_extendsE subsetD sym_factor_reflc)
qed

lemma strict_order_extension_if_consistent:
  assumes "strongly_consistent r"
  obtains r_ext where "strict_extends r_ext r" "total_preorder r_ext" 
proof -
  from assms have "strict_extends (r+) r"
    unfolding strongly_consistent_def strict_extends_def extends_def asym_factor_def sym_factor_def
    by (auto simp: Field_def dest: tranclD)
  moreover have "strict_extends (r*) (r+)"
    unfolding strict_extends_def extends_def
    by (auto simp: asym_factor_rtrancl sym_factor_def dest: rtranclD)
  ultimately have extends: "strict_extends (r*) r"
    using transpE[OF transp_strict_extends] by blast

  have "preorder (r*)"
    unfolding preorder_on_def using refl_on_def trans_def by fastforce
  from strict_extends_preorder_on[OF this] extends obtain r_ext where
    "total_preorder r_ext" "strict_extends r_ext r"
    using transpE[OF transp_strict_extends] by blast
  then show thesis using that by blast
qed


experiment begin

text ‹We can instantiate the above theorem to get Szpilrajn's theorem.›
lemma
  assumes "strict_partial_order r"
  shows "r_ext. strict_linear_order r_ext  r  r_ext"
proof -                  
  from assms[unfolded strict_partial_order_def] have "strongly_consistent r" "antisym r"
    unfolding strongly_consistent_def by (simp_all add: antisym_if_irrefl_trans)
  from strict_order_extension_if_consistent[OF this(1)] obtain r_ext
    where "strict_extends r_ext r" "total_preorder r_ext" 
    by blast
  with assms[unfolded strict_partial_order_def] 
  have "trans (r_ext - Id)" "irrefl (r_ext - Id)" "total (r_ext - Id)" "r  (r_ext - Id)"
    using strict_extends_antisym[OF _ antisym r]
    by (auto simp: irrefl_def elim: strict_extendsE intro: trans_diff_Id dest: total_preorder_onD)
  then show ?thesis
    unfolding strict_linear_order_on_def by blast
qed

end

 
(*<*)
end
(*>*)