Generalized Zorn and added well-ordering theorem
authornipkow
Sun, 02 Mar 2008 15:02:06 +0100
changeset 26191 ae537f315b34
parent 26190 cf51a23c0cd0
child 26192 52617dca8386
Generalized Zorn and added well-ordering theorem
src/HOL/Library/Zorn.thy
--- a/src/HOL/Library/Zorn.thy	Sat Mar 01 15:01:03 2008 +0100
+++ b/src/HOL/Library/Zorn.thy	Sun Mar 02 15:02:06 2008 +0100
@@ -1,7 +1,8 @@
 (*  Title       : HOL/Library/Zorn.thy
     ID          : $Id$
-    Author      : Jacques D. Fleuriot
-    Description : Zorn's Lemma -- see Larry Paulson's Zorn.thy in ZF
+    Author      : Jacques D. Fleuriot, Tobias Nipkow
+    Description : Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF)
+                  The well-ordering theorem
 *)
 
 header {* Zorn's Lemma *}
@@ -152,7 +153,7 @@
 lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S"
   by (unfold maxchain_def) blast
 
-lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> ? d. d \<in> super S c"
+lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> EX d. d \<in> super S c"
   by (unfold super_def maxchain_def) auto
 
 lemma select_super:
@@ -261,4 +262,353 @@
 lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S"
   by (unfold chain_def) blast
 
+
+(* FIXME into Relation.thy *)
+
+lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
+by(auto simp:Field_def Domain_def Range_def)
+
+lemma Field_empty[simp]: "Field {} = {}"
+by(auto simp:Field_def)
+
+lemma Field_insert[simp]: "Field (insert (a,b) r) = {a,b} \<union> Field r"
+by(auto simp:Field_def)
+
+lemma Field_Un[simp]: "Field (r \<union> s) = Field r \<union> Field s"
+by(auto simp:Field_def)
+
+lemma Field_Union[simp]: "Field (\<Union>R) = \<Union>(Field ` R)"
+by(auto simp:Field_def)
+
+lemma Domain_converse[simp]: "Domain(r^-1) = Range r"
+by blast
+
+lemma Range_converse[simp]: "Range(r^-1) = Domain r"
+by blast
+
+lemma Field_converse[simp]: "Field(r^-1) = Field r"
+by(auto simp:Field_def)
+
+lemma reflexive_reflcl[simp]: "reflexive(r^=)"
+by(simp add:refl_def)
+
+lemma antisym_reflcl[simp]: "antisym(r^=) = antisym r"
+by(simp add:antisym_def)
+
+lemma trans_reflclI[simp]: "trans r \<Longrightarrow> trans(r^=)"
+unfolding trans_def by blast
+
+(*********************************************************)
+
+(* Define globally? In Set.thy?
+   Use in def of chain at the beginning *)
+definition "subset_chain C \<equiv> \<forall>A\<in>C.\<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A"
+
+(* Define globally? In Relation.thy? *)
+definition Chain :: "('a*'a)set \<Rightarrow> 'a set set" where
+"Chain r \<equiv> {A. \<forall>a\<in>A.\<forall>b\<in>A. (a,b) : r \<or> (b,a) \<in> r}"
+
+lemma mono_Chain: "r \<subseteq> s \<Longrightarrow> Chain r \<subseteq> Chain s"
+unfolding Chain_def by blast
+
+(* Are the following definitions the "right" ones?
+
+Key point: should the set appear as an explicit argument,
+(as currently in "refl A r") or should it remain implicitly the Field
+(as in Refl below)? I use refl/Refl merely to illusrate the point.
+
+The notation "refl A r" is closer to the usual (A,<=) in the literature
+whereas "Refl r" is shorter and avoids naming the set.
+Note that "refl A r \<Longrightarrow> A = Field r & Refl r" and "Refl r \<Longrightarrow> refl (Field r) r"
+This makes the A look redundant.
+
+A slight advantage of having the A around is that one can write "a:A"
+rather than "a:Field r". A disavantage is the multiple occurrences of
+"refl (Field r) r" (etc) in the proof of the well-ordering thm.
+
+I propose to move the definitions into Main, either as they are or
+with an additional A argument.
+
+Naming: The capital letters were chosen to distinguish them from
+versions on the whole type we have (eg reflexive) or may want to have
+(eg preorder). In case of an additional A argument one could append
+"_on" to distinguish the relativized versions.
+*)
+
+definition "Refl r \<equiv> \<forall>x \<in> Field r. (x,x) \<in> r"
+definition "Preorder r \<equiv> Refl r \<and> trans r"
+definition "Partial_order r \<equiv> Preorder r \<and> antisym r"
+definition "Total r \<equiv> \<forall>x\<in>Field r.\<forall>y\<in>Field r. x\<noteq>y \<longrightarrow> (x,y)\<in>r \<or> (y,x)\<in>r"
+definition "Linear_order r \<equiv> Partial_order r \<and> Total r"
+definition "Well_order r \<equiv> Linear_order r \<and> wf(r - Id)"
+
+lemmas Order_defs =
+  Preorder_def Partial_order_def Linear_order_def Well_order_def
+
+lemma Refl_empty[simp]: "Refl {}"
+by(simp add:Refl_def)
+lemma Preorder_empty[simp]: "Preorder {}"
+by(simp add:Preorder_def trans_def)
+lemma Partial_order_empty[simp]: "Partial_order {}"
+by(simp add:Partial_order_def)
+lemma Total_empty[simp]: "Total {}"
+by(simp add:Total_def)
+lemma Linear_order_empty[simp]: "Linear_order {}"
+by(simp add:Linear_order_def)
+lemma Well_order_empty[simp]: "Well_order {}"
+by(simp add:Well_order_def)
+
+lemma Refl_converse[simp]: "Refl(r^-1) = Refl r"
+by(simp add:Refl_def)
+
+lemma Preorder_converse[simp]: "Preorder (r^-1) = Preorder r"
+by (simp add:Preorder_def)
+
+lemma Partial_order_converse[simp]:
+  "Partial_order (r^-1) = Partial_order r"
+by (simp add: Partial_order_def)
+
+lemma subset_Image_Image_iff:
+  "\<lbrakk> Preorder r; A \<subseteq> Field r; B \<subseteq> Field r\<rbrakk> \<Longrightarrow>
+   r `` A \<subseteq> r `` B \<longleftrightarrow> (\<forall>a\<in>A.\<exists>b\<in>B. (b,a):r)"
+apply(auto simp add:subset_def Preorder_def Refl_def Image_def)
+apply metis
+by(metis trans_def)
+
+lemma subset_Image1_Image1_iff:
+  "\<lbrakk> Preorder r; a : Field r; b : Field r\<rbrakk> \<Longrightarrow> r `` {a} \<subseteq> r `` {b} \<longleftrightarrow> (b,a):r"
+by(simp add:subset_Image_Image_iff)
+
+lemma Refl_antisym_eq_Image1_Image1_iff:
+  "\<lbrakk>Refl r; antisym r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b"
+by(simp add:Preorder_def expand_set_eq Partial_order_def antisym_def Refl_def)
+  metis
+
+lemma Partial_order_eq_Image1_Image1_iff:
+  "\<lbrakk>Partial_order r; a:Field r; b:Field r\<rbrakk> \<Longrightarrow> r `` {a} = r `` {b} \<longleftrightarrow> a=b"
+by(auto simp:Preorder_def Partial_order_def Refl_antisym_eq_Image1_Image1_iff)
+
+text{* Zorn's lemma for partial orders: *}
+
+lemma Zorns_po_lemma:
+assumes po: "Partial_order r" and u: "\<forall>C\<in>Chain r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a,u):r"
+shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m,a):r \<longrightarrow> a=m"
+proof-
+  have "Preorder r" using po by(simp add:Partial_order_def)
+--{* Mirror r in the set of subsets below (wrt r) elements of A*}
+  let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r"
+  have "\<forall>C \<in> chain ?S. EX U:?S. ALL A:C. A\<subseteq>U"
+  proof (auto simp:chain_def)
+    fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C.\<forall>B\<in>C. A\<subseteq>B | B\<subseteq>A"
+    let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}"
+    have "C = ?B ` ?A" using 1 by(auto simp: image_def)
+    have "?A\<in>Chain r"
+    proof (simp add:Chain_def, intro allI impI, elim conjE)
+      fix a b
+      assume "a \<in> Field r" "?B a \<in> C" "b \<in> Field r" "?B b \<in> C"
+      hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by auto
+      thus "(a, b) \<in> r \<or> (b, a) \<in> r" using `Preorder r` `a:Field r` `b:Field r`
+	by(simp add:subset_Image1_Image1_iff)
+    qed
+    then obtain u where uA: "u:Field r" "\<forall>a\<in>?A. (a,u) : r" using u by auto
+    have "\<forall>A\<in>C. A \<subseteq> r^-1 `` {u}" (is "?P u")
+    proof auto
+      fix a B assume aB: "B:C" "a:B"
+      with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto
+      thus "(a,u) : r" using uA aB `Preorder r`
+	by (auto simp add: Preorder_def Refl_def) (metis transD)
+    qed
+    thus "EX u:Field r. ?P u" using `u:Field r` by blast
+  qed
+  from Zorn_Lemma2[OF this]
+  obtain m B where "m:Field r" "B = r^-1 `` {m}"
+    "\<forall>x\<in>Field r. B \<subseteq> r^-1 `` {x} \<longrightarrow> B = r^-1 `` {x}"
+    by(auto simp:image_def) blast
+  hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m" using po `Preorder r` `m:Field r`
+    by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff)
+  thus ?thesis using `m:Field r` by blast
+qed
+
+(* The initial segment of a relation appears generally useful.
+   Move to Relation.thy?
+   Definition correct/most general?
+   Naming?
+*)
+definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where
+"init_seg_of == {(r,s). r \<subseteq> s \<and> (\<forall>a b c. (a,b):s \<and> (b,c):r \<longrightarrow> (a,b):r)}"
+
+abbreviation initialSegmentOf :: "('a*'a)set \<Rightarrow> ('a*'a)set \<Rightarrow> bool"
+             (infix "initial'_segment'_of" 55) where
+"r initial_segment_of s == (r,s):init_seg_of"
+
+lemma refl_init_seg_of[simp]: "r initial_segment_of r"
+by(simp add:init_seg_of_def)
+
+lemma trans_init_seg_of:
+  "r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t"
+by(simp (no_asm_use) add: init_seg_of_def)
+  (metis Domain_iff UnCI Un_absorb2 subset_trans)
+
+lemma antisym_init_seg_of:
+  "r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r=s"
+by(auto simp:init_seg_of_def)
+
+lemma Chain_init_seg_of_Union:
+  "R \<in> Chain init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R"
+by(auto simp add:init_seg_of_def Chain_def Ball_def) blast
+
+lemma subset_chain_trans_Union:
+  "subset_chain R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans(\<Union>R)"
+apply(auto simp add:subset_chain_def)
+apply(simp (no_asm_use) add:trans_def)
+apply (metis subsetD)
+done
+
+lemma subset_chain_antisym_Union:
+  "subset_chain R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym(\<Union>R)"
+apply(auto simp add:subset_chain_def antisym_def)
+apply (metis subsetD)
+done
+
+lemma subset_chain_Total_Union:
+assumes "subset_chain R" "\<forall>r\<in>R. Total r"
+shows "Total (\<Union>R)"
+proof (simp add: Total_def Ball_def, auto del:disjCI)
+  fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a\<noteq>b"
+  from `subset_chain R` `r:R` `s:R` have "r\<subseteq>s \<or> s\<subseteq>r"
+    by(simp add:subset_chain_def)
+  thus "(\<exists>r\<in>R. (a,b) \<in> r) \<or> (\<exists>r\<in>R. (b,a) \<in> r)"
+  proof
+    assume "r\<subseteq>s" hence "(a,b):s \<or> (b,a):s" using assms(2) A
+      by(simp add:Total_def)(metis mono_Field subsetD)
+    thus ?thesis using `s:R` by blast
+  next
+    assume "s\<subseteq>r" hence "(a,b):r \<or> (b,a):r" using assms(2) A
+      by(simp add:Total_def)(metis mono_Field subsetD)
+    thus ?thesis using `r:R` by blast
+  qed
+qed
+
+lemma wf_Union_wf_init_segs:
+assumes "R \<in> Chain init_seg_of" and "\<forall>r\<in>R. wf r" shows "wf(\<Union>R)"
+proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto)
+  fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f(Suc i), f i) \<in> r"
+  then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto
+  { fix i have "(f(Suc i), f i) \<in> r"
+    proof(induct i)
+      case 0 show ?case by fact
+    next
+      case (Suc i)
+      moreover obtain s where "s\<in>R" and "(f(Suc(Suc i)), f(Suc i)) \<in> s"
+	using 1 by auto
+      moreover hence "s initial_segment_of r \<or> r initial_segment_of s"
+	using assms(1) `r:R` by(simp add: Chain_def)
+      ultimately show ?case by(simp add:init_seg_of_def) blast
+    qed
+  }
+  thus False using assms(2) `r:R`
+    by(simp add:wf_iff_no_infinite_down_chain) blast
+qed
+
+lemma Chain_inits_DiffI:
+  "R \<in> Chain init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chain init_seg_of"
+apply(auto simp:Chain_def init_seg_of_def)
+apply (metis subsetD)
+apply (metis subsetD)
+done
+
+theorem well_ordering: "\<exists>r::('a*'a)set. Well_order r"
+proof-
+-- {*The initial segment relation on well-orders: *}
+  let ?WO = "{r::('a*'a)set. Well_order r}"
+  def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO"
+  have I_init: "I \<subseteq> init_seg_of" by(auto simp:I_def)
+  hence subch: "!!R. R : Chain I \<Longrightarrow> subset_chain R"
+    by(auto simp:init_seg_of_def subset_chain_def Chain_def)
+  have Chain_wo: "!!R r. R \<in> Chain I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r"
+    by(simp add:Chain_def I_def) blast
+  have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def)
+  hence 0: "Partial_order I"
+    by(auto simp add: Partial_order_def Preorder_def antisym_def antisym_init_seg_of Refl_def trans_def I_def)(metis trans_init_seg_of)
+-- {*I-chains have upper bounds in ?WO wrt I: their Union*}
+  { fix R assume "R \<in> Chain I"
+    hence Ris: "R \<in> Chain init_seg_of" using mono_Chain[OF I_init] by blast
+    have subch: "subset_chain R" using `R : Chain I` I_init
+      by(auto simp:init_seg_of_def subset_chain_def Chain_def)
+    have "\<forall>r\<in>R. Refl r" "\<forall>r\<in>R. trans r" "\<forall>r\<in>R. antisym r" "\<forall>r\<in>R. Total r"
+         "\<forall>r\<in>R. wf(r-Id)"
+      using Chain_wo[OF `R \<in> Chain I`] by(simp_all add:Order_defs)
+    have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by(auto simp:Refl_def)
+    moreover have "trans (\<Union>R)"
+      by(rule subset_chain_trans_Union[OF subch `\<forall>r\<in>R. trans r`])
+    moreover have "antisym(\<Union>R)"
+      by(rule subset_chain_antisym_Union[OF subch `\<forall>r\<in>R. antisym r`])
+    moreover have "Total (\<Union>R)"
+      by(rule subset_chain_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 Chain_inits_DiffI[OF Ris]]
+      show ?thesis by (simp (no_asm_simp)) blast
+    qed
+    ultimately have "Well_order (\<Union>R)" by(simp add:Order_defs)
+    moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris
+      by(simp add: Chain_init_seg_of_Union)
+    ultimately have "\<Union>R : ?WO \<and> (\<forall>r\<in>R. (r,\<Union>R) : I)"
+      using mono_Chain[OF I_init] `R \<in> Chain I`
+      by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo subsetD)
+  }
+  hence 1: "\<forall>R \<in> Chain I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r,u) : I" by (subst FI) blast
+--{*Zorn's Lemma yields a maximal well-order m:*}
+  then obtain m::"('a*'a)set" where "Well_order m" and
+    max: "\<forall>r. Well_order r \<and> (m,r):I \<longrightarrow> r=m"
+    using Zorns_po_lemma[OF 0 1] by (auto simp:FI)
+--{*Now show by contradiction that m covers the whole type:*}
+  { fix x::'a assume "x \<notin> Field m"
+--{*We assume that x is not covered and extend m at the top with x*}
+    have "m \<noteq> {}"
+    proof
+      assume "m={}"
+      moreover have "Well_order {(x,x)}"
+	by(simp add:Order_defs Refl_def trans_def antisym_def Total_def Field_def Domain_def Range_def)
+      ultimately show False using max
+	by (auto simp:I_def init_seg_of_def simp del:Field_insert)
+    qed
+    hence "Field m \<noteq> {}" by(auto simp:Field_def)
+    moreover have "wf(m-Id)" using `Well_order m` by(simp add:Well_order_def)
+--{*The extension of m by x:*}
+    let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s"
+    have Fm: "Field ?m = insert x (Field m)"
+      apply(simp add:Field_insert Field_Un)
+      unfolding Field_def by auto
+    have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)"
+      using `Well_order m` by(simp_all add:Order_defs)
+--{*We show that the extension is a well-order*}
+    have "Refl ?m" using `Refl m` Fm by(auto simp:Refl_def)
+    moreover have "trans ?m" using `trans m` `x \<notin> Field m`
+      unfolding trans_def Field_def Domain_def Range_def by blast
+    moreover have "antisym ?m" using `antisym m` `x \<notin> Field m`
+      unfolding antisym_def Field_def Domain_def Range_def by blast
+    moreover have "Total ?m" using `Total m` Fm by(auto simp: Total_def)
+    moreover have "wf(?m-Id)"
+    proof-
+      have "wf ?s" using `x \<notin> Field m`
+	by(auto simp add:wf_eq_minimal Field_def Domain_def Range_def) metis
+      thus ?thesis using `wf(m-Id)` `x \<notin> Field m`
+	wf_subset[OF `wf ?s` Diff_subset]
+	by (fastsimp intro!: wf_Un simp add: Un_Diff Field_def)
+    qed
+    ultimately have "Well_order ?m" by(simp add:Order_defs)
+--{*We show that the extension is above m*}
+    moreover hence "(m,?m) : I" using `Well_order m` `x \<notin> Field m`
+      by(fastsimp simp:I_def init_seg_of_def Field_def Domain_def Range_def)
+    ultimately
+--{*This contradicts maximality of m:*}
+    have False using max `x \<notin> Field m` unfolding Field_def by blast
+  }
+  hence "Field m = UNIV" by auto
+  with `Well_order m` have "Well_order m" by simp
+  thus ?thesis ..
+qed
+
 end