Theory Wellorder_Extension

(*  Title:      HOL/Cardinals/Wellorder_Extension.thy
    Author:     Christian Sternagel, JAIST
*)

section ‹Extending Well-founded Relations to Wellorders›

theory Wellorder_Extension
imports Main Order_Union
begin

subsection ‹Extending Well-founded Relations to Wellorders›

text ‹A \emph{downset} (also lower set, decreasing set, initial segment, or
downward closed set) is closed w.r.t.\ smaller elements.›
definition downset_on where
  "downset_on A r = (x y. (x, y)  r  y  A  x  A)"

(*
text {*Connection to order filters of the @{theory Cardinals} theory.*}
lemma (in wo_rel) ofilter_downset_on_conv:
  "ofilter A ⟷ downset_on A r ∧ A ⊆ Field r"
  by (auto simp: downset_on_def ofilter_def under_def)
*)

lemma downset_onI:
  "(x y. (x, y)  r  y  A  x  A)  downset_on A r"
  by (auto simp: downset_on_def)

lemma downset_onD:
  "downset_on A r  (x, y)  r  y  A  x  A"
  unfolding downset_on_def by blast

text ‹Extensions of relations w.r.t.\ a given set.›
definition extension_on where
  "extension_on A r s = (xA. yA. (x, y)  s  (x, y)  r)"

lemma extension_onI:
  "(x y. x  A; y  A; (x, y)  s  (x, y)  r)  extension_on A r s"
  by (auto simp: extension_on_def)

lemma extension_onD:
  "extension_on A r s  x  A  y  A  (x, y)  s  (x, y)  r"
  by (auto simp: extension_on_def)

lemma downset_on_Union:
  assumes "r. r  R  downset_on (Field r) p"
  shows "downset_on (Field (R)) p"
  using assms by (auto intro: downset_onI dest: downset_onD)

lemma chain_subset_extension_on_Union:
  assumes "chain R" and "r. r  R  extension_on (Field r) r p"
  shows "extension_on (Field (R)) (R) p"
  using assms
  by (simp add: chain_subset_def extension_on_def)
     (metis (no_types) mono_Field subsetD)

lemma downset_on_empty [simp]: "downset_on {} p"
  by (auto simp: downset_on_def)

lemma extension_on_empty [simp]: "extension_on {} p q"
  by (auto simp: extension_on_def)

text ‹Every well-founded relation can be extended to a wellorder.›
theorem well_order_extension:
  assumes "wf p"
  shows "w. p  w  Well_order w"
proof -
  let ?K = "{r. Well_order r  downset_on (Field r) p  extension_on (Field r) r p}"
  define I where "I = init_seg_of  ?K × ?K"
  have I_init: "I  init_seg_of" by (simp add: I_def)
  then have subch: "R. R  Chains I  chain R"
    by (auto simp: init_seg_of_def chain_subset_def Chains_def)
  have Chains_wo: "R r. R  Chains I  r  R 
      Well_order r  downset_on (Field r) p  extension_on (Field r) r p"
    by (simp add: Chains_def I_def) blast
  have FI: "Field I = ?K" by (auto simp: I_def init_seg_of_def Field_def)
  then have 0: "Partial_order I"
    by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def
      trans_def I_def elim: trans_init_seg_of)
  have "R  ?K  (rR. (r,R)  I)" if "R  Chains I" for R
  proof -
    from that have Ris: "R  Chains init_seg_of" using mono_Chains [OF I_init] by blast
    have subch: "chain R" using R  Chains I I_init
      by (auto simp: init_seg_of_def chain_subset_def Chains_def)
    have "rR. Refl r" and "rR. trans r" and "rR. antisym r" and
      "rR. Total r" and "rR. wf (r - Id)" and
      "r. r  R  downset_on (Field r) p" and
      "r. r  R  extension_on (Field r) r p"
      using Chains_wo [OF R  Chains I] by (simp_all add: order_on_defs)
    have "Refl (R)" using rR. Refl r  unfolding refl_on_def by fastforce
    moreover have "trans (R)"
      by (rule chain_subset_trans_Union [OF subch rR. trans r])
    moreover have "antisym (R)"
      by (rule chain_subset_antisym_Union [OF subch rR. antisym r])
    moreover have "Total (R)"
      by (rule chain_subset_Total_Union [OF subch rR. Total r])
    moreover have "wf ((R) - Id)"
    proof -
      have "(R) - Id = {r - Id | r. r  R}" by blast
      with rR. wf (r - Id) wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
      show ?thesis by fastforce
    qed
    ultimately have "Well_order (R)" by (simp add: order_on_defs)
    moreover have "rR. r initial_segment_of R" using Ris
      by (simp add: Chains_init_seg_of_Union)
    moreover have "downset_on (Field (R)) p"
      by (rule downset_on_Union [OF r. r  R  downset_on (Field r) p])
    moreover have "extension_on (Field (R)) (R) p"
      by (rule chain_subset_extension_on_Union [OF subch r. r  R  extension_on (Field r) r p])
    ultimately show ?thesis
      using mono_Chains [OF I_init] and R  Chains I
      by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
  qed
  then have 1: "uField I. rR. (r, u)  I" if "RChains I" for R 
    using that by (subst FI) blast
  txt ‹Zorn's Lemma yields a maximal wellorder m.›
  from Zorns_po_lemma [OF 0 1] obtain m :: "('a × 'a) set"
    where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and
    max: "r. Well_order r  downset_on (Field r) p  extension_on (Field r) r p 
      (m, r)  I  r = m"
    by (auto simp: FI)
  have "Field p  Field m"
  proof (rule ccontr)
    let ?Q = "Field p - Field m"
    assume "¬ (Field p  Field m)"
    with assms [unfolded wf_eq_minimal, THEN spec, of ?Q]
      obtain x where "x  Field p" and "x  Field m" and
      min: "y. (y, x)  p  y  ?Q" by blast
    txt ‹Add termx as topmost element to termm.›
    let ?s = "{(y, x) | y. y  Field m}"
    let ?m = "insert (x, x) m  ?s"
    have Fm: "Field ?m = insert x (Field m)" by (auto simp: Field_def)
    have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)"
      using Well_order m by (simp_all add: order_on_defs)
    txt ‹We show that the extension is a wellorder.›
    have "Refl ?m" using Refl m Fm by (auto simp: refl_on_def)
    moreover have "trans ?m" using trans m x  Field m
      unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
    moreover have "antisym ?m" using antisym m x  Field m
      unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast
    moreover have "Total ?m" using Total m Fm by (auto simp: Relation.total_on_def)
    moreover have "wf (?m - Id)"
    proof -
      have "wf ?s" using x  Field m
        by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
      thus ?thesis using wf (m - Id) x  Field m
        wf_subset [OF wf ?s Diff_subset]
        by (fastforce intro!: wf_Un simp add: Un_Diff Field_def)
    qed
    ultimately have "Well_order ?m" by (simp add: order_on_defs)
    moreover have "extension_on (Field ?m) ?m p"
      using extension_on (Field m) m p downset_on (Field m) p
      by (subst Fm) (auto simp: extension_on_def dest: downset_onD)
    moreover have "downset_on (Field ?m) p"
      apply (subst Fm)
      using downset_on (Field m) p and min
      unfolding downset_on_def Field_def by blast
    moreover have "(m, ?m)  I"
      using Well_order m and Well_order ?m and
      downset_on (Field m) p and downset_on (Field ?m) p and
      extension_on (Field m) m p and extension_on (Field ?m) ?m p and
      Refl m and x  Field m
      by (auto simp: I_def init_seg_of_def refl_on_def)
    ultimately
    ― ‹This contradicts maximality of m:›
    show False using max and x  Field m unfolding Field_def by blast
  qed
  have "p  m"
    using Field p  Field m and extension_on (Field m) m p
    unfolding Field_def extension_on_def by auto fast
  with Well_order m show ?thesis by blast
qed

text ‹Every well-founded relation can be extended to a total wellorder.›
corollary total_well_order_extension:
  assumes "wf p"
  shows "w. p  w  Well_order w  Field w = UNIV"
proof -
  from well_order_extension [OF assms] obtain w
    where "p  w" and wo: "Well_order w" by blast
  let ?A = "UNIV - Field w"
  from well_order_on [of ?A] obtain w' where wo': "well_order_on ?A w'" ..
  have [simp]: "Field w' = ?A" using well_order_on_Well_order [OF wo'] by simp
  have *: "Field w  Field w' = {}" by simp
  let ?w = "w ∪o w'"
  have "p  ?w" using p  w by (auto simp: Osum_def)
  moreover have "Well_order ?w" using Osum_Well_order [OF * wo] and wo' by simp
  moreover have "Field ?w = UNIV" by (simp add: Field_Osum)
  ultimately show ?thesis by blast
qed

corollary well_order_on_extension:
  assumes "wf p" and "Field p  A"
  shows "w. p  w  well_order_on A w"
proof -
  from total_well_order_extension [OF wf p] obtain r
    where "p  r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast
  let ?r = "{(x, y). x  A  y  A  (x, y)  r}"
  from p  r have "p  ?r" using Field p  A by (auto simp: Field_def)
  have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)"
    using Well_order r by (simp_all add: order_on_defs)
  have "refl_on A ?r" using Refl r by (auto simp: refl_on_def univ)
  moreover have "trans ?r" using trans r
    unfolding trans_def by blast
  moreover have "antisym ?r" using antisym r
    unfolding antisym_def by blast
  moreover have "total_on A ?r" using Total r by (simp add: total_on_def univ)
  moreover have "wf (?r - Id)" by (rule wf_subset [OF wf(r - Id)]) blast
  ultimately have "well_order_on A ?r" by (simp add: order_on_defs)
  with p  ?r show ?thesis by blast
qed

end