effectively reverted d25fc4c0ff62, to avoid quasi-cyclic dependencies with HOL-Cardinals and minimize BNF dependencies
--- a/src/HOL/Cardinals/Cardinals.thy Wed Nov 20 21:28:58 2013 +0100
+++ b/src/HOL/Cardinals/Cardinals.thy Wed Nov 20 23:14:06 2013 +0100
@@ -9,7 +9,7 @@
header {* Theory of Ordinals and Cardinals *}
theory Cardinals
-imports Cardinal_Order_Relation Cardinal_Arithmetic
+imports Cardinal_Order_Relation Cardinal_Arithmetic Wellorder_Extension
begin
end
--- a/src/HOL/Cardinals/Constructions_on_Wellorders.thy Wed Nov 20 21:28:58 2013 +0100
+++ b/src/HOL/Cardinals/Constructions_on_Wellorders.thy Wed Nov 20 23:14:06 2013 +0100
@@ -8,7 +8,7 @@
header {* Constructions on Wellorders *}
theory Constructions_on_Wellorders
-imports Constructions_on_Wellorders_FP Wellorder_Embedding
+imports Constructions_on_Wellorders_FP Wellorder_Embedding Order_Union
begin
declare
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Order_Union.thy Wed Nov 20 23:14:06 2013 +0100
@@ -0,0 +1,370 @@
+(* Title: HOL/Cardinals/Order_Union.thy
+ Author: Andrei Popescu, TU Muenchen
+
+The ordinal-like sum of two orders with disjoint fields
+*)
+
+header {* Order Union *}
+
+theory Order_Union
+imports Wellfounded_More_FP
+begin
+
+definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60) where
+ "r Osum r' = r \<union> r' \<union> {(a, a'). a \<in> Field r \<and> a' \<in> Field r'}"
+
+notation Osum (infix "\<union>o" 60)
+
+lemma Field_Osum: "Field (r \<union>o r') = Field r \<union> Field r'"
+ unfolding Osum_def Field_def by blast
+
+lemma Osum_wf:
+assumes FLD: "Field r Int Field r' = {}" and
+ WF: "wf r" and WF': "wf r'"
+shows "wf (r Osum r')"
+unfolding wf_eq_minimal2 unfolding Field_Osum
+proof(intro allI impI, elim conjE)
+ fix A assume *: "A \<subseteq> Field r \<union> Field r'" and **: "A \<noteq> {}"
+ obtain B where B_def: "B = A Int Field r" by blast
+ show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r \<union>o r'"
+ proof(cases "B = {}")
+ assume Case1: "B \<noteq> {}"
+ hence "B \<noteq> {} \<and> B \<le> Field r" using B_def by auto
+ then obtain a where 1: "a \<in> B" and 2: "\<forall>a1 \<in> B. (a1,a) \<notin> r"
+ using WF unfolding wf_eq_minimal2 by metis
+ hence 3: "a \<in> Field r \<and> a \<notin> Field r'" using B_def FLD by auto
+ (* *)
+ have "\<forall>a1 \<in> A. (a1,a) \<notin> r Osum r'"
+ proof(intro ballI)
+ fix a1 assume **: "a1 \<in> A"
+ {assume Case11: "a1 \<in> Field r"
+ hence "(a1,a) \<notin> r" using B_def ** 2 by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 by (auto simp add: Field_def)
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "a1 \<notin> Field r"
+ hence "(a1,a) \<notin> r" unfolding Field_def by auto
+ moreover
+ have "(a1,a) \<notin> r'" using 3 unfolding Field_def by auto
+ ultimately have "(a1,a) \<notin> r Osum r'"
+ using 3 unfolding Osum_def by auto
+ }
+ ultimately show "(a1,a) \<notin> r Osum r'" by blast
+ qed
+ thus ?thesis using 1 B_def by auto
+ next
+ assume Case2: "B = {}"
+ hence 1: "A \<noteq> {} \<and> A \<le> Field r'" using * ** B_def by auto
+ then obtain a' where 2: "a' \<in> A" and 3: "\<forall>a1' \<in> A. (a1',a') \<notin> r'"
+ using WF' unfolding wf_eq_minimal2 by metis
+ hence 4: "a' \<in> Field r' \<and> a' \<notin> Field r" using 1 FLD by blast
+ (* *)
+ have "\<forall>a1' \<in> A. (a1',a') \<notin> r Osum r'"
+ proof(unfold Osum_def, auto simp add: 3)
+ fix a1' assume "(a1', a') \<in> r"
+ thus False using 4 unfolding Field_def by blast
+ next
+ fix a1' assume "a1' \<in> A" and "a1' \<in> Field r"
+ thus False using Case2 B_def by auto
+ qed
+ thus ?thesis using 2 by blast
+ qed
+qed
+
+lemma Osum_Refl:
+assumes FLD: "Field r Int Field r' = {}" and
+ REFL: "Refl r" and REFL': "Refl r'"
+shows "Refl (r Osum r')"
+using assms
+unfolding refl_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_trans:
+assumes FLD: "Field r Int Field r' = {}" and
+ TRANS: "trans r" and TRANS': "trans r'"
+shows "trans (r Osum r')"
+proof(unfold trans_def, auto)
+ fix x y z assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, z) \<in> r \<union>o r'"
+ show "(x, z) \<in> r \<union>o r'"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case11: "(y,z) \<in> r"
+ hence "(x,z) \<in> r" using Case1 TRANS trans_def[of r] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case12: "(y,z) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ {assume Case13: "z \<in> Field r'"
+ hence ?thesis using 1 unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume Case21: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ {assume Case22: "(y,z) \<in> r'"
+ hence "(x,z) \<in> r'" using Case2 TRANS' trans_def[of r'] by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume Case23: "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume Case31: "(y,z) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,z) \<in> r'"
+ hence "z \<in> Field r'" unfolding Field_def by blast
+ hence ?thesis unfolding Osum_def using Case3 by auto
+ }
+ moreover
+ {assume Case33: "y \<in> Field r"
+ hence False using FLD Case3 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Preorder:
+"\<lbrakk>Field r Int Field r' = {}; Preorder r; Preorder r'\<rbrakk> \<Longrightarrow> Preorder (r Osum r')"
+unfolding preorder_on_def using Osum_Refl Osum_trans by blast
+
+lemma Osum_antisym:
+assumes FLD: "Field r Int Field r' = {}" and
+ AN: "antisym r" and AN': "antisym r'"
+shows "antisym (r Osum r')"
+proof(unfold antisym_def, auto)
+ fix x y assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, x) \<in> r \<union>o r'"
+ show "x = y"
+ proof-
+ {assume Case1: "(x,y) \<in> r"
+ hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ have "(y,x) \<in> r \<Longrightarrow> ?thesis"
+ using Case1 AN antisym_def[of r] by blast
+ moreover
+ {assume "(y,x) \<in> r'"
+ hence "y \<in> Field r'" unfolding Field_def by auto
+ hence False using FLD 1 by auto
+ }
+ moreover
+ have "x \<in> Field r' \<Longrightarrow> False" using FLD 1 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case2: "(x,y) \<in> r'"
+ hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD 2 by auto
+ }
+ moreover
+ have "(y,x) \<in> r' \<Longrightarrow> ?thesis"
+ using Case2 AN' antisym_def[of r'] by blast
+ moreover
+ {assume "y \<in> Field r"
+ hence False using FLD 2 by auto
+ }
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ moreover
+ {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
+ have ?thesis
+ proof-
+ {assume "(y,x) \<in> r"
+ hence "y \<in> Field r" unfolding Field_def by auto
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ {assume Case32: "(y,x) \<in> r'"
+ hence "x \<in> Field r'" unfolding Field_def by blast
+ hence False using FLD Case3 by auto
+ }
+ moreover
+ have "\<not> y \<in> Field r" using FLD Case3 by auto
+ ultimately show ?thesis using ** unfolding Osum_def by blast
+ qed
+ }
+ ultimately show ?thesis using * unfolding Osum_def by blast
+ qed
+qed
+
+lemma Osum_Partial_order:
+"\<lbrakk>Field r Int Field r' = {}; Partial_order r; Partial_order r'\<rbrakk> \<Longrightarrow>
+ Partial_order (r Osum r')"
+unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast
+
+lemma Osum_Total:
+assumes FLD: "Field r Int Field r' = {}" and
+ TOT: "Total r" and TOT': "Total r'"
+shows "Total (r Osum r')"
+using assms
+unfolding total_on_def Field_Osum unfolding Osum_def by blast
+
+lemma Osum_Linear_order:
+"\<lbrakk>Field r Int Field r' = {}; Linear_order r; Linear_order r'\<rbrakk> \<Longrightarrow>
+ Linear_order (r Osum r')"
+unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast
+
+lemma Osum_minus_Id1:
+assumes "r \<le> Id"
+shows "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r' - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r'"
+ with * have "(a,b) \<in> r' - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_minus_Id2:
+assumes "r' \<le> Id"
+shows "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+proof-
+ let ?Left = "(r Osum r') - Id"
+ let ?Right = "(r - Id) \<union> (Field r \<times> Field r')"
+ {fix a::'a and b assume *: "(a,b) \<notin> Id"
+ {assume "(a,b) \<in> r'"
+ with * have False using assms by auto
+ }
+ moreover
+ {assume "(a,b) \<in> r"
+ with * have "(a,b) \<in> r - Id" by auto
+ }
+ ultimately
+ have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
+ unfolding Osum_def by auto
+ }
+ thus ?thesis by auto
+qed
+
+lemma Osum_minus_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ NID: "\<not> (r \<le> Id)" and NID': "\<not> (r' \<le> Id)"
+shows "(r Osum r') - Id \<le> (r - Id) Osum (r' - Id)"
+proof-
+ {fix a a' assume *: "(a,a') \<in> (r Osum r')" and **: "a \<noteq> a'"
+ have "(a,a') \<in> (r - Id) Osum (r' - Id)"
+ proof-
+ {assume "(a,a') \<in> r \<or> (a,a') \<in> r'"
+ with ** have ?thesis unfolding Osum_def by auto
+ }
+ moreover
+ {assume "a \<in> Field r \<and> a' \<in> Field r'"
+ hence "a \<in> Field(r - Id) \<and> a' \<in> Field (r' - Id)"
+ using assms Total_Id_Field by blast
+ hence ?thesis unfolding Osum_def by auto
+ }
+ ultimately show ?thesis using * unfolding Osum_def by fast
+ qed
+ }
+ thus ?thesis by(auto simp add: Osum_def)
+qed
+
+lemma wf_Int_Times:
+assumes "A Int B = {}"
+shows "wf(A \<times> B)"
+unfolding wf_def using assms by blast
+
+lemma Osum_wf_Id:
+assumes TOT: "Total r" and TOT': "Total r'" and
+ FLD: "Field r Int Field r' = {}" and
+ WF: "wf(r - Id)" and WF': "wf(r' - Id)"
+shows "wf ((r Osum r') - Id)"
+proof(cases "r \<le> Id \<or> r' \<le> Id")
+ assume Case1: "\<not>(r \<le> Id \<or> r' \<le> Id)"
+ have "Field(r - Id) Int Field(r' - Id) = {}"
+ using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r']
+ Diff_subset[of r Id] Diff_subset[of r' Id] by blast
+ thus ?thesis
+ using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"]
+ wf_subset[of "(r - Id) \<union>o (r' - Id)" "(r Osum r') - Id"] by auto
+next
+ have 1: "wf(Field r \<times> Field r')"
+ using FLD by (auto simp add: wf_Int_Times)
+ assume Case2: "r \<le> Id \<or> r' \<le> Id"
+ moreover
+ {assume Case21: "r \<le> Id"
+ hence "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id1[of r r'] by simp
+ moreover
+ {have "Domain(Field r \<times> Field r') Int Range(r' - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r' - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF' wf_Un[of "Field r \<times> Field r'" "r' - Id"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (metis wf_subset)
+ }
+ moreover
+ {assume Case22: "r' \<le> Id"
+ hence "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
+ using Osum_minus_Id2[of r' r] by simp
+ moreover
+ {have "Range(Field r \<times> Field r') Int Domain(r - Id) = {}"
+ using FLD unfolding Field_def by blast
+ hence "wf((r - Id) \<union> (Field r \<times> Field r'))"
+ using 1 WF wf_Un[of "r - Id" "Field r \<times> Field r'"]
+ by (auto simp add: Un_commute)
+ }
+ ultimately have ?thesis by (metis wf_subset)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma Osum_Well_order:
+assumes FLD: "Field r Int Field r' = {}" and
+ WELL: "Well_order r" and WELL': "Well_order r'"
+shows "Well_order (r Osum r')"
+proof-
+ have "Total r \<and> Total r'" using WELL WELL'
+ by (auto simp add: order_on_defs)
+ thus ?thesis using assms unfolding well_order_on_def
+ using Osum_Linear_order Osum_wf_Id by blast
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Cardinals/Wellorder_Extension.thy Wed Nov 20 23:14:06 2013 +0100
@@ -0,0 +1,213 @@
+(* Title: HOL/Cardinals/Wellorder_Extension.thy
+ Author: Christian Sternagel, JAIST
+*)
+
+header {* Extending Well-founded Relations to Wellorders *}
+
+theory Wellorder_Extension
+imports "~~/src/HOL/Library/Zorn" 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 = (\<forall>x y. (x, y) \<in> r \<and> y \<in> A \<longrightarrow> x \<in> A)"
+
+(*
+text {*Connection to order filters of the @{theory Cardinals} theory.*}
+lemma (in wo_rel) ofilter_downset_on_conv:
+ "ofilter A \<longleftrightarrow> downset_on A r \<and> A \<subseteq> Field r"
+ by (auto simp: downset_on_def ofilter_def under_def)
+*)
+
+lemma downset_onI:
+ "(\<And>x y. (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A) \<Longrightarrow> downset_on A r"
+ by (auto simp: downset_on_def)
+
+lemma downset_onD:
+ "downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> 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 = (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> s \<longrightarrow> (x, y) \<in> r)"
+
+lemma extension_onI:
+ "(\<And>x y. \<lbrakk>x \<in> A; y \<in> A; (x, y) \<in> s\<rbrakk> \<Longrightarrow> (x, y) \<in> r) \<Longrightarrow> extension_on A r s"
+ by (auto simp: extension_on_def)
+
+lemma extension_onD:
+ "extension_on A r s \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> (x, y) \<in> s \<Longrightarrow> (x, y) \<in> r"
+ by (auto simp: extension_on_def)
+
+lemma downset_on_Union:
+ assumes "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p"
+ shows "downset_on (Field (\<Union>R)) p"
+ using assms by (auto intro: downset_onI dest: downset_onD)
+
+lemma chain_subset_extension_on_Union:
+ assumes "chain\<^sub>\<subseteq> R" and "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
+ shows "extension_on (Field (\<Union>R)) (\<Union>R) p"
+ using assms
+ by (simp add: chain_subset_def extension_on_def)
+ (metis (no_types) mono_Field set_mp)
+
+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 "\<exists>w. p \<subseteq> w \<and> Well_order w"
+proof -
+ let ?K = "{r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p}"
+ def I \<equiv> "init_seg_of \<inter> ?K \<times> ?K"
+ have I_init: "I \<subseteq> init_seg_of" by (simp add: I_def)
+ then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
+ by (auto simp: init_seg_of_def chain_subset_def Chains_def)
+ have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow>
+ Well_order r \<and> downset_on (Field r) p \<and> 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)
+ { fix R assume "R \<in> Chains I"
+ then have Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
+ have subch: "chain\<^sub>\<subseteq> R" using `R \<in> Chains I` I_init
+ by (auto simp: init_seg_of_def chain_subset_def Chains_def)
+ have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r" and
+ "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)" and
+ "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and
+ "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
+ using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
+ have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
+ moreover have "trans (\<Union>R)"
+ by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
+ moreover have "antisym (\<Union>R)"
+ by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
+ moreover have "Total (\<Union>R)"
+ by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
+ moreover have "wf ((\<Union>R) - Id)"
+ proof -
+ have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
+ with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
+ show ?thesis by fastforce
+ qed
+ ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs)
+ moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris
+ by (simp add: Chains_init_seg_of_Union)
+ moreover have "downset_on (Field (\<Union>R)) p"
+ by (rule downset_on_Union [OF `\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p`])
+ moreover have "extension_on (Field (\<Union>R)) (\<Union>R) p"
+ by (rule chain_subset_extension_on_Union [OF subch `\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p`])
+ ultimately have "\<Union>R \<in> ?K \<and> (\<forall>r\<in>R. (r,\<Union>R) \<in> I)"
+ using mono_Chains [OF I_init] and `R \<in> Chains I`
+ by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
+ }
+ then have 1: "\<forall>R\<in>Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
+ txt {*Zorn's Lemma yields a maximal wellorder m.*}
+ from Zorns_po_lemma [OF 0 1] obtain m :: "('a \<times> 'a) set"
+ where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and
+ max: "\<forall>r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p \<and>
+ (m, r) \<in> I \<longrightarrow> r = m"
+ by (auto simp: FI)
+ have "Field p \<subseteq> Field m"
+ proof (rule ccontr)
+ let ?Q = "Field p - Field m"
+ assume "\<not> (Field p \<subseteq> Field m)"
+ with assms [unfolded wf_eq_minimal, THEN spec, of ?Q]
+ obtain x where "x \<in> Field p" and "x \<notin> Field m" and
+ min: "\<forall>y. (y, x) \<in> p \<longrightarrow> y \<notin> ?Q" by blast
+ txt {*Add @{term x} as topmost element to @{term m}.*}
+ let ?s = "{(y, x) | y. y \<in> Field m}"
+ let ?m = "insert (x, x) m \<union> ?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 \<notin> Field m`
+ unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
+ moreover have "antisym ?m" using `antisym m` `x \<notin> 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 \<notin> Field m`
+ by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
+ thus ?thesis using `wf (m - Id)` `x \<notin> 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) \<in> 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 \<notin> 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 \<notin> Field m` unfolding Field_def by blast
+ qed
+ have "p \<subseteq> m"
+ using `Field p \<subseteq> 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 "\<exists>w. p \<subseteq> w \<and> Well_order w \<and> Field w = UNIV"
+proof -
+ from well_order_extension [OF assms] obtain w
+ where "p \<subseteq> 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 rel.well_order_on_Well_order [OF wo'] by simp
+ have *: "Field w \<inter> Field w' = {}" by simp
+ let ?w = "w \<union>o w'"
+ have "p \<subseteq> ?w" using `p \<subseteq> 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 \<subseteq> A"
+ shows "\<exists>w. p \<subseteq> w \<and> well_order_on A w"
+proof -
+ from total_well_order_extension [OF `wf p`] obtain r
+ where "p \<subseteq> r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast
+ let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
+ from `p \<subseteq> r` have "p \<subseteq> ?r" using `Field p \<subseteq> A` by (auto simp: Field_def)
+ have 1: "Field ?r = A" using wo univ
+ by (fastforce simp: Field_def order_on_defs refl_on_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 \<subseteq> ?r` show ?thesis by blast
+qed
+
+end
--- a/src/HOL/Library/Library.thy Wed Nov 20 21:28:58 2013 +0100
+++ b/src/HOL/Library/Library.thy Wed Nov 20 23:14:06 2013 +0100
@@ -41,7 +41,6 @@
Numeral_Type
OptionalSugar
Option_ord
- Order_Union
Parallel
Permutation
Permutations
@@ -66,7 +65,6 @@
Transitive_Closure_Table
Wfrec
While_Combinator
- Zorn
begin
end
(*>*)
--- a/src/HOL/Library/Order_Union.thy Wed Nov 20 21:28:58 2013 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,371 +0,0 @@
-(* Title: HOL/Library/Order_Union.thy
- Author: Andrei Popescu, TU Muenchen
-
-The ordinal-like sum of two orders with disjoint fields
-*)
-
-header {* Order Union *}
-
-theory Order_Union
-imports "~~/src/HOL/Cardinals/Wellfounded_More_FP"
-begin
-
-definition Osum :: "'a rel \<Rightarrow> 'a rel \<Rightarrow> 'a rel" (infix "Osum" 60) where
- "r Osum r' = r \<union> r' \<union> {(a, a'). a \<in> Field r \<and> a' \<in> Field r'}"
-
-notation Osum (infix "\<union>o" 60)
-
-lemma Field_Osum: "Field (r \<union>o r') = Field r \<union> Field r'"
- unfolding Osum_def Field_def by blast
-
-lemma Osum_wf:
-assumes FLD: "Field r Int Field r' = {}" and
- WF: "wf r" and WF': "wf r'"
-shows "wf (r Osum r')"
-unfolding wf_eq_minimal2 unfolding Field_Osum
-proof(intro allI impI, elim conjE)
- fix A assume *: "A \<subseteq> Field r \<union> Field r'" and **: "A \<noteq> {}"
- obtain B where B_def: "B = A Int Field r" by blast
- show "\<exists>a\<in>A. \<forall>a'\<in>A. (a', a) \<notin> r \<union>o r'"
- proof(cases "B = {}")
- assume Case1: "B \<noteq> {}"
- hence "B \<noteq> {} \<and> B \<le> Field r" using B_def by auto
- then obtain a where 1: "a \<in> B" and 2: "\<forall>a1 \<in> B. (a1,a) \<notin> r"
- using WF unfolding wf_eq_minimal2 by metis
- hence 3: "a \<in> Field r \<and> a \<notin> Field r'" using B_def FLD by auto
- (* *)
- have "\<forall>a1 \<in> A. (a1,a) \<notin> r Osum r'"
- proof(intro ballI)
- fix a1 assume **: "a1 \<in> A"
- {assume Case11: "a1 \<in> Field r"
- hence "(a1,a) \<notin> r" using B_def ** 2 by auto
- moreover
- have "(a1,a) \<notin> r'" using 3 by (auto simp add: Field_def)
- ultimately have "(a1,a) \<notin> r Osum r'"
- using 3 unfolding Osum_def by auto
- }
- moreover
- {assume Case12: "a1 \<notin> Field r"
- hence "(a1,a) \<notin> r" unfolding Field_def by auto
- moreover
- have "(a1,a) \<notin> r'" using 3 unfolding Field_def by auto
- ultimately have "(a1,a) \<notin> r Osum r'"
- using 3 unfolding Osum_def by auto
- }
- ultimately show "(a1,a) \<notin> r Osum r'" by blast
- qed
- thus ?thesis using 1 B_def by auto
- next
- assume Case2: "B = {}"
- hence 1: "A \<noteq> {} \<and> A \<le> Field r'" using * ** B_def by auto
- then obtain a' where 2: "a' \<in> A" and 3: "\<forall>a1' \<in> A. (a1',a') \<notin> r'"
- using WF' unfolding wf_eq_minimal2 by metis
- hence 4: "a' \<in> Field r' \<and> a' \<notin> Field r" using 1 FLD by blast
- (* *)
- have "\<forall>a1' \<in> A. (a1',a') \<notin> r Osum r'"
- proof(unfold Osum_def, auto simp add: 3)
- fix a1' assume "(a1', a') \<in> r"
- thus False using 4 unfolding Field_def by blast
- next
- fix a1' assume "a1' \<in> A" and "a1' \<in> Field r"
- thus False using Case2 B_def by auto
- qed
- thus ?thesis using 2 by blast
- qed
-qed
-
-lemma Osum_Refl:
-assumes FLD: "Field r Int Field r' = {}" and
- REFL: "Refl r" and REFL': "Refl r'"
-shows "Refl (r Osum r')"
-using assms
-unfolding refl_on_def Field_Osum unfolding Osum_def by blast
-
-lemma Osum_trans:
-assumes FLD: "Field r Int Field r' = {}" and
- TRANS: "trans r" and TRANS': "trans r'"
-shows "trans (r Osum r')"
-proof(unfold trans_def, auto)
- fix x y z assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, z) \<in> r \<union>o r'"
- show "(x, z) \<in> r \<union>o r'"
- proof-
- {assume Case1: "(x,y) \<in> r"
- hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
- have ?thesis
- proof-
- {assume Case11: "(y,z) \<in> r"
- hence "(x,z) \<in> r" using Case1 TRANS trans_def[of r] by blast
- hence ?thesis unfolding Osum_def by auto
- }
- moreover
- {assume Case12: "(y,z) \<in> r'"
- hence "y \<in> Field r'" unfolding Field_def by auto
- hence False using FLD 1 by auto
- }
- moreover
- {assume Case13: "z \<in> Field r'"
- hence ?thesis using 1 unfolding Osum_def by auto
- }
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- moreover
- {assume Case2: "(x,y) \<in> r'"
- hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
- have ?thesis
- proof-
- {assume Case21: "(y,z) \<in> r"
- hence "y \<in> Field r" unfolding Field_def by auto
- hence False using FLD 2 by auto
- }
- moreover
- {assume Case22: "(y,z) \<in> r'"
- hence "(x,z) \<in> r'" using Case2 TRANS' trans_def[of r'] by blast
- hence ?thesis unfolding Osum_def by auto
- }
- moreover
- {assume Case23: "y \<in> Field r"
- hence False using FLD 2 by auto
- }
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- moreover
- {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
- have ?thesis
- proof-
- {assume Case31: "(y,z) \<in> r"
- hence "y \<in> Field r" unfolding Field_def by auto
- hence False using FLD Case3 by auto
- }
- moreover
- {assume Case32: "(y,z) \<in> r'"
- hence "z \<in> Field r'" unfolding Field_def by blast
- hence ?thesis unfolding Osum_def using Case3 by auto
- }
- moreover
- {assume Case33: "y \<in> Field r"
- hence False using FLD Case3 by auto
- }
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- ultimately show ?thesis using * unfolding Osum_def by blast
- qed
-qed
-
-lemma Osum_Preorder:
-"\<lbrakk>Field r Int Field r' = {}; Preorder r; Preorder r'\<rbrakk> \<Longrightarrow> Preorder (r Osum r')"
-unfolding preorder_on_def using Osum_Refl Osum_trans by blast
-
-lemma Osum_antisym:
-assumes FLD: "Field r Int Field r' = {}" and
- AN: "antisym r" and AN': "antisym r'"
-shows "antisym (r Osum r')"
-proof(unfold antisym_def, auto)
- fix x y assume *: "(x, y) \<in> r \<union>o r'" and **: "(y, x) \<in> r \<union>o r'"
- show "x = y"
- proof-
- {assume Case1: "(x,y) \<in> r"
- hence 1: "x \<in> Field r \<and> y \<in> Field r" unfolding Field_def by auto
- have ?thesis
- proof-
- have "(y,x) \<in> r \<Longrightarrow> ?thesis"
- using Case1 AN antisym_def[of r] by blast
- moreover
- {assume "(y,x) \<in> r'"
- hence "y \<in> Field r'" unfolding Field_def by auto
- hence False using FLD 1 by auto
- }
- moreover
- have "x \<in> Field r' \<Longrightarrow> False" using FLD 1 by auto
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- moreover
- {assume Case2: "(x,y) \<in> r'"
- hence 2: "x \<in> Field r' \<and> y \<in> Field r'" unfolding Field_def by auto
- have ?thesis
- proof-
- {assume "(y,x) \<in> r"
- hence "y \<in> Field r" unfolding Field_def by auto
- hence False using FLD 2 by auto
- }
- moreover
- have "(y,x) \<in> r' \<Longrightarrow> ?thesis"
- using Case2 AN' antisym_def[of r'] by blast
- moreover
- {assume "y \<in> Field r"
- hence False using FLD 2 by auto
- }
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- moreover
- {assume Case3: "x \<in> Field r \<and> y \<in> Field r'"
- have ?thesis
- proof-
- {assume "(y,x) \<in> r"
- hence "y \<in> Field r" unfolding Field_def by auto
- hence False using FLD Case3 by auto
- }
- moreover
- {assume Case32: "(y,x) \<in> r'"
- hence "x \<in> Field r'" unfolding Field_def by blast
- hence False using FLD Case3 by auto
- }
- moreover
- have "\<not> y \<in> Field r" using FLD Case3 by auto
- ultimately show ?thesis using ** unfolding Osum_def by blast
- qed
- }
- ultimately show ?thesis using * unfolding Osum_def by blast
- qed
-qed
-
-lemma Osum_Partial_order:
-"\<lbrakk>Field r Int Field r' = {}; Partial_order r; Partial_order r'\<rbrakk> \<Longrightarrow>
- Partial_order (r Osum r')"
-unfolding partial_order_on_def using Osum_Preorder Osum_antisym by blast
-
-lemma Osum_Total:
-assumes FLD: "Field r Int Field r' = {}" and
- TOT: "Total r" and TOT': "Total r'"
-shows "Total (r Osum r')"
-using assms
-unfolding total_on_def Field_Osum unfolding Osum_def by blast
-
-lemma Osum_Linear_order:
-"\<lbrakk>Field r Int Field r' = {}; Linear_order r; Linear_order r'\<rbrakk> \<Longrightarrow>
- Linear_order (r Osum r')"
-unfolding linear_order_on_def using Osum_Partial_order Osum_Total by blast
-
-lemma Osum_minus_Id1:
-assumes "r \<le> Id"
-shows "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
-proof-
- let ?Left = "(r Osum r') - Id"
- let ?Right = "(r' - Id) \<union> (Field r \<times> Field r')"
- {fix a::'a and b assume *: "(a,b) \<notin> Id"
- {assume "(a,b) \<in> r"
- with * have False using assms by auto
- }
- moreover
- {assume "(a,b) \<in> r'"
- with * have "(a,b) \<in> r' - Id" by auto
- }
- ultimately
- have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
- unfolding Osum_def by auto
- }
- thus ?thesis by auto
-qed
-
-lemma Osum_minus_Id2:
-assumes "r' \<le> Id"
-shows "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
-proof-
- let ?Left = "(r Osum r') - Id"
- let ?Right = "(r - Id) \<union> (Field r \<times> Field r')"
- {fix a::'a and b assume *: "(a,b) \<notin> Id"
- {assume "(a,b) \<in> r'"
- with * have False using assms by auto
- }
- moreover
- {assume "(a,b) \<in> r"
- with * have "(a,b) \<in> r - Id" by auto
- }
- ultimately
- have "(a,b) \<in> ?Left \<Longrightarrow> (a,b) \<in> ?Right"
- unfolding Osum_def by auto
- }
- thus ?thesis by auto
-qed
-
-lemma Osum_minus_Id:
-assumes TOT: "Total r" and TOT': "Total r'" and
- NID: "\<not> (r \<le> Id)" and NID': "\<not> (r' \<le> Id)"
-shows "(r Osum r') - Id \<le> (r - Id) Osum (r' - Id)"
-proof-
- {fix a a' assume *: "(a,a') \<in> (r Osum r')" and **: "a \<noteq> a'"
- have "(a,a') \<in> (r - Id) Osum (r' - Id)"
- proof-
- {assume "(a,a') \<in> r \<or> (a,a') \<in> r'"
- with ** have ?thesis unfolding Osum_def by auto
- }
- moreover
- {assume "a \<in> Field r \<and> a' \<in> Field r'"
- hence "a \<in> Field(r - Id) \<and> a' \<in> Field (r' - Id)"
- using assms Total_Id_Field by blast
- hence ?thesis unfolding Osum_def by auto
- }
- ultimately show ?thesis using * unfolding Osum_def by fast
- qed
- }
- thus ?thesis by(auto simp add: Osum_def)
-qed
-
-lemma wf_Int_Times:
-assumes "A Int B = {}"
-shows "wf(A \<times> B)"
-unfolding wf_def using assms by blast
-
-lemma Osum_wf_Id:
-assumes TOT: "Total r" and TOT': "Total r'" and
- FLD: "Field r Int Field r' = {}" and
- WF: "wf(r - Id)" and WF': "wf(r' - Id)"
-shows "wf ((r Osum r') - Id)"
-proof(cases "r \<le> Id \<or> r' \<le> Id")
- assume Case1: "\<not>(r \<le> Id \<or> r' \<le> Id)"
- have "Field(r - Id) Int Field(r' - Id) = {}"
- using FLD mono_Field[of "r - Id" r] mono_Field[of "r' - Id" r']
- Diff_subset[of r Id] Diff_subset[of r' Id] by blast
- thus ?thesis
- using Case1 Osum_minus_Id[of r r'] assms Osum_wf[of "r - Id" "r' - Id"]
- wf_subset[of "(r - Id) \<union>o (r' - Id)" "(r Osum r') - Id"] by auto
-next
- have 1: "wf(Field r \<times> Field r')"
- using FLD by (auto simp add: wf_Int_Times)
- assume Case2: "r \<le> Id \<or> r' \<le> Id"
- moreover
- {assume Case21: "r \<le> Id"
- hence "(r Osum r') - Id \<le> (r' - Id) \<union> (Field r \<times> Field r')"
- using Osum_minus_Id1[of r r'] by simp
- moreover
- {have "Domain(Field r \<times> Field r') Int Range(r' - Id) = {}"
- using FLD unfolding Field_def by blast
- hence "wf((r' - Id) \<union> (Field r \<times> Field r'))"
- using 1 WF' wf_Un[of "Field r \<times> Field r'" "r' - Id"]
- by (auto simp add: Un_commute)
- }
- ultimately have ?thesis by (metis wf_subset)
- }
- moreover
- {assume Case22: "r' \<le> Id"
- hence "(r Osum r') - Id \<le> (r - Id) \<union> (Field r \<times> Field r')"
- using Osum_minus_Id2[of r' r] by simp
- moreover
- {have "Range(Field r \<times> Field r') Int Domain(r - Id) = {}"
- using FLD unfolding Field_def by blast
- hence "wf((r - Id) \<union> (Field r \<times> Field r'))"
- using 1 WF wf_Un[of "r - Id" "Field r \<times> Field r'"]
- by (auto simp add: Un_commute)
- }
- ultimately have ?thesis by (metis wf_subset)
- }
- ultimately show ?thesis by blast
-qed
-
-lemma Osum_Well_order:
-assumes FLD: "Field r Int Field r' = {}" and
- WELL: "Well_order r" and WELL': "Well_order r'"
-shows "Well_order (r Osum r')"
-proof-
- have "Total r \<and> Total r'" using WELL WELL'
- by (auto simp add: order_on_defs)
- thus ?thesis using assms unfolding well_order_on_def
- using Osum_Linear_order Osum_wf_Id by blast
-qed
-
-end
-
--- a/src/HOL/Library/Zorn.thy Wed Nov 20 21:28:58 2013 +0100
+++ b/src/HOL/Library/Zorn.thy Wed Nov 20 23:14:06 2013 +0100
@@ -5,13 +5,12 @@
Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF).
The well-ordering theorem.
-The extension of any well-founded relation to a well-order.
*)
header {* Zorn's Lemma *}
theory Zorn
-imports Order_Union
+imports Order_Relation
begin
subsection {* Zorn's Lemma for the Subset Relation *}
@@ -710,207 +709,4 @@
with 1 show ?thesis by auto
qed
-subsection {* Extending Well-founded Relations to Well-Orders *}
-
-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 = (\<forall>x y. (x, y) \<in> r \<and> y \<in> A \<longrightarrow> x \<in> A)"
-
-(*
-text {*Connection to order filters of the @{theory Cardinals} theory.*}
-lemma (in wo_rel) ofilter_downset_on_conv:
- "ofilter A \<longleftrightarrow> downset_on A r \<and> A \<subseteq> Field r"
- by (auto simp: downset_on_def ofilter_def under_def)
-*)
-
-lemma downset_onI:
- "(\<And>x y. (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A) \<Longrightarrow> downset_on A r"
- by (auto simp: downset_on_def)
-
-lemma downset_onD:
- "downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> 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 = (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> s \<longrightarrow> (x, y) \<in> r)"
-
-lemma extension_onI:
- "(\<And>x y. \<lbrakk>x \<in> A; y \<in> A; (x, y) \<in> s\<rbrakk> \<Longrightarrow> (x, y) \<in> r) \<Longrightarrow> extension_on A r s"
- by (auto simp: extension_on_def)
-
-lemma extension_onD:
- "extension_on A r s \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> (x, y) \<in> s \<Longrightarrow> (x, y) \<in> r"
- by (auto simp: extension_on_def)
-
-lemma downset_on_Union:
- assumes "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p"
- shows "downset_on (Field (\<Union>R)) p"
- using assms by (auto intro: downset_onI dest: downset_onD)
-
-lemma chain_subset_extension_on_Union:
- assumes "chain\<^sub>\<subseteq> R" and "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
- shows "extension_on (Field (\<Union>R)) (\<Union>R) p"
- using assms
- by (simp add: chain_subset_def extension_on_def)
- (metis (no_types) mono_Field set_mp)
-
-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 well-order.*}
-theorem well_order_extension:
- assumes "wf p"
- shows "\<exists>w. p \<subseteq> w \<and> Well_order w"
-proof -
- let ?K = "{r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p}"
- def I \<equiv> "init_seg_of \<inter> ?K \<times> ?K"
- have I_init: "I \<subseteq> init_seg_of" by (simp add: I_def)
- then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R"
- by (auto simp: init_seg_of_def chain_subset_def Chains_def)
- have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow>
- Well_order r \<and> downset_on (Field r) p \<and> 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)
- { fix R assume "R \<in> Chains I"
- then have Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast
- have subch: "chain\<^sub>\<subseteq> R" using `R \<in> Chains I` I_init
- by (auto simp: init_seg_of_def chain_subset_def Chains_def)
- have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r" and
- "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)" and
- "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and
- "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p"
- using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs)
- have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` unfolding refl_on_def by fastforce
- moreover have "trans (\<Union>R)"
- by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`])
- moreover have "antisym (\<Union>R)"
- by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`])
- moreover have "Total (\<Union>R)"
- by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`])
- moreover have "wf ((\<Union>R) - Id)"
- proof -
- have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast
- with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]]
- show ?thesis by fastforce
- qed
- ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs)
- moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris
- by (simp add: Chains_init_seg_of_Union)
- moreover have "downset_on (Field (\<Union>R)) p"
- by (rule downset_on_Union [OF `\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p`])
- moreover have "extension_on (Field (\<Union>R)) (\<Union>R) p"
- by (rule chain_subset_extension_on_Union [OF subch `\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p`])
- ultimately have "\<Union>R \<in> ?K \<and> (\<forall>r\<in>R. (r,\<Union>R) \<in> I)"
- using mono_Chains [OF I_init] and `R \<in> Chains I`
- by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo)
- }
- then have 1: "\<forall>R\<in>Chains I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r, u) \<in> I" by (subst FI) blast
- txt {*Zorn's Lemma yields a maximal well-order m.*}
- from Zorns_po_lemma [OF 0 1] obtain m :: "('a \<times> 'a) set"
- where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and
- max: "\<forall>r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p \<and>
- (m, r) \<in> I \<longrightarrow> r = m"
- by (auto simp: FI)
- have "Field p \<subseteq> Field m"
- proof (rule ccontr)
- let ?Q = "Field p - Field m"
- assume "\<not> (Field p \<subseteq> Field m)"
- with assms [unfolded wf_eq_minimal, THEN spec, of ?Q]
- obtain x where "x \<in> Field p" and "x \<notin> Field m" and
- min: "\<forall>y. (y, x) \<in> p \<longrightarrow> y \<notin> ?Q" by blast
- txt {*Add @{term x} as topmost element to @{term m}.*}
- let ?s = "{(y, x) | y. y \<in> Field m}"
- let ?m = "insert (x, x) m \<union> ?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 well-order.*}
- have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def)
- moreover have "trans ?m" using `trans m` `x \<notin> Field m`
- unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast
- moreover have "antisym ?m" using `antisym m` `x \<notin> 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 \<notin> Field m`
- by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis
- thus ?thesis using `wf (m - Id)` `x \<notin> 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) \<in> 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 \<notin> 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 \<notin> Field m` unfolding Field_def by blast
- qed
- have "p \<subseteq> m"
- using `Field p \<subseteq> 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 well-order.*}
-corollary total_well_order_extension:
- assumes "wf p"
- shows "\<exists>w. p \<subseteq> w \<and> Well_order w \<and> Field w = UNIV"
-proof -
- from well_order_extension [OF assms] obtain w
- where "p \<subseteq> 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 rel.well_order_on_Well_order [OF wo'] by simp
- have *: "Field w \<inter> Field w' = {}" by simp
- let ?w = "w \<union>o w'"
- have "p \<subseteq> ?w" using `p \<subseteq> 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 \<subseteq> A"
- shows "\<exists>w. p \<subseteq> w \<and> well_order_on A w"
-proof -
- from total_well_order_extension [OF `wf p`] obtain r
- where "p \<subseteq> r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast
- let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}"
- from `p \<subseteq> r` have "p \<subseteq> ?r" using `Field p \<subseteq> A` by (auto simp: Field_def)
- have 1: "Field ?r = A" using wo univ
- by (fastforce simp: Field_def order_on_defs refl_on_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 \<subseteq> ?r` show ?thesis by blast
-qed
-
end
-