wenzelm@32960: (* Title: HOL/Library/Zorn.thy wenzelm@32960: Author: Jacques D. Fleuriot, Tobias Nipkow wenzelm@32960: wenzelm@32960: Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF). wenzelm@32960: The well-ordering theorem. wenzelm@14706: *) paulson@13551: wenzelm@14706: header {* Zorn's Lemma *} paulson@13551: nipkow@15131: theory Zorn blanchet@48750: imports Order_Relation nipkow@15131: begin paulson@13551: nipkow@26272: (* Define globally? In Set.thy? *) haftmann@46980: definition chain_subset :: "'a set set \ bool" ("chain\<^bsub>\\<^esub>") haftmann@46980: where haftmann@46980: "chain\<^bsub>\\<^esub> C \ \A\C.\B\C. A \ B \ B \ A" nipkow@26272: wenzelm@14706: text{* wenzelm@14706: The lemma and section numbers refer to an unpublished article wenzelm@14706: \cite{Abrial-Laffitte}. wenzelm@14706: *} paulson@13551: haftmann@46980: definition chain :: "'a set set \ 'a set set set" haftmann@46980: where haftmann@46980: "chain S = {F. F \ S \ chain\<^bsub>\\<^esub> F}" paulson@13551: haftmann@46980: definition super :: "'a set set \ 'a set set \ 'a set set set" haftmann@46980: where haftmann@46980: "super S c = {d. d \ chain S \ c \ d}" paulson@13551: haftmann@46980: definition maxchain :: "'a set set \ 'a set set set" haftmann@46980: where haftmann@46980: "maxchain S = {c. c \ chain S \ super S c = {}}" paulson@13551: haftmann@46980: definition succ :: "'a set set \ 'a set set \ 'a set set" haftmann@46980: where haftmann@46980: "succ S c = (if c \ chain S \ c \ maxchain S then c else SOME c'. c' \ super S c)" haftmann@46980: haftmann@46980: inductive_set TFin :: "'a set set \ 'a set set set" haftmann@46980: for S :: "'a set set" haftmann@46980: where haftmann@46980: succI: "x \ TFin S \ succ S x \ TFin S" haftmann@46980: | Pow_UnionI: "Y \ Pow (TFin S) \ \Y \ TFin S" paulson@13551: paulson@13551: paulson@13551: subsection{*Mathematical Preamble*} paulson@13551: wenzelm@17200: lemma Union_lemma0: paulson@18143: "(\x \ C. x \ A | B \ x) ==> Union(C) \ A | B \ Union(C)" wenzelm@17200: by blast paulson@13551: paulson@13551: paulson@13551: text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*} wenzelm@17200: paulson@13551: lemma Abrial_axiom1: "x \ succ S x" berghofe@26806: apply (auto simp add: succ_def super_def maxchain_def) wenzelm@18585: apply (rule contrapos_np, assumption) wenzelm@46008: apply (rule someI2) wenzelm@46008: apply blast+ wenzelm@17200: done paulson@13551: paulson@13551: lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] paulson@13551: wenzelm@14706: lemma TFin_induct: wenzelm@46008: assumes H: "n \ TFin S" and wenzelm@46008: I: "!!x. x \ TFin S ==> P x ==> P (succ S x)" wenzelm@46008: "!!Y. Y \ TFin S ==> Ball Y P ==> P (Union Y)" wenzelm@46008: shows "P n" wenzelm@46008: using H by induct (blast intro: I)+ paulson@13551: paulson@13551: lemma succ_trans: "x \ y ==> x \ succ S y" wenzelm@17200: apply (erule subset_trans) wenzelm@17200: apply (rule Abrial_axiom1) wenzelm@17200: done paulson@13551: paulson@13551: text{*Lemma 1 of section 3.1*} paulson@13551: lemma TFin_linear_lemma1: wenzelm@14706: "[| n \ TFin S; m \ TFin S; wenzelm@14706: \x \ TFin S. x \ m --> x = m | succ S x \ m paulson@13551: |] ==> n \ m | succ S m \ n" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (erule_tac [2] Union_lemma0) wenzelm@17200: apply (blast del: subsetI intro: succ_trans) wenzelm@17200: done paulson@13551: paulson@13551: text{* Lemma 2 of section 3.2 *} paulson@13551: lemma TFin_linear_lemma2: paulson@13551: "m \ TFin S ==> \n \ TFin S. n \ m --> n=m | succ S n \ m" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (rule impI [THEN ballI]) wenzelm@17200: txt{*case split using @{text TFin_linear_lemma1}*} wenzelm@17200: apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], wenzelm@17200: assumption+) wenzelm@17200: apply (drule_tac x = n in bspec, assumption) wenzelm@17200: apply (blast del: subsetI intro: succ_trans, blast) wenzelm@17200: txt{*second induction step*} wenzelm@17200: apply (rule impI [THEN ballI]) wenzelm@17200: apply (rule Union_lemma0 [THEN disjE]) wenzelm@17200: apply (rule_tac [3] disjI2) wenzelm@17200: prefer 2 apply blast wenzelm@17200: apply (rule ballI) wenzelm@17200: apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], wenzelm@17200: assumption+, auto) wenzelm@17200: apply (blast intro!: Abrial_axiom1 [THEN subsetD]) wenzelm@17200: done paulson@13551: paulson@13551: text{*Re-ordering the premises of Lemma 2*} paulson@13551: lemma TFin_subsetD: paulson@13551: "[| n \ m; m \ TFin S; n \ TFin S |] ==> n=m | succ S n \ m" wenzelm@17200: by (rule TFin_linear_lemma2 [rule_format]) paulson@13551: paulson@13551: text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*} paulson@13551: lemma TFin_subset_linear: "[| m \ TFin S; n \ TFin S|] ==> n \ m | m \ n" wenzelm@17200: apply (rule disjE) wenzelm@17200: apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) wenzelm@17200: apply (assumption+, erule disjI2) wenzelm@17200: apply (blast del: subsetI wenzelm@17200: intro: subsetI Abrial_axiom1 [THEN subset_trans]) wenzelm@17200: done paulson@13551: paulson@13551: text{*Lemma 3 of section 3.3*} paulson@13551: lemma eq_succ_upper: "[| n \ TFin S; m \ TFin S; m = succ S m |] ==> n \ m" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (drule TFin_subsetD) wenzelm@17200: apply (assumption+, force, blast) wenzelm@17200: done paulson@13551: paulson@13551: text{*Property 3.3 of section 3.3*} paulson@13551: lemma equal_succ_Union: "m \ TFin S ==> (m = succ S m) = (m = Union(TFin S))" wenzelm@17200: apply (rule iffI) wenzelm@17200: apply (rule Union_upper [THEN equalityI]) paulson@18143: apply assumption paulson@18143: apply (rule eq_succ_upper [THEN Union_least], assumption+) wenzelm@17200: apply (erule ssubst) wenzelm@17200: apply (rule Abrial_axiom1 [THEN equalityI]) wenzelm@17200: apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI) wenzelm@17200: done paulson@13551: paulson@13551: subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*} paulson@13551: wenzelm@14706: text{*NB: We assume the partial ordering is @{text "\"}, paulson@13551: the subset relation!*} paulson@13551: paulson@13551: lemma empty_set_mem_chain: "({} :: 'a set set) \ chain S" blanchet@48750: by (unfold chain_def chain_subset_def) simp paulson@13551: paulson@13551: lemma super_subset_chain: "super S c \ chain S" wenzelm@17200: by (unfold super_def) blast paulson@13551: paulson@13551: lemma maxchain_subset_chain: "maxchain S \ chain S" wenzelm@17200: by (unfold maxchain_def) blast paulson@13551: nipkow@26191: lemma mem_super_Ex: "c \ chain S - maxchain S ==> EX d. d \ super S c" blanchet@48750: by (unfold super_def maxchain_def) simp paulson@13551: paulson@18143: lemma select_super: paulson@18143: "c \ chain S - maxchain S ==> (\c'. c': super S c): super S c" wenzelm@17200: apply (erule mem_super_Ex [THEN exE]) wenzelm@46008: apply (rule someI2) blanchet@48750: apply simp+ wenzelm@17200: done paulson@13551: paulson@18143: lemma select_not_equals: paulson@18143: "c \ chain S - maxchain S ==> (\c'. c': super S c) \ c" wenzelm@17200: apply (rule notI) wenzelm@17200: apply (drule select_super) berghofe@26806: apply (simp add: super_def less_le) wenzelm@17200: done paulson@13551: wenzelm@17200: lemma succI3: "c \ chain S - maxchain S ==> succ S c = (\c'. c': super S c)" wenzelm@17200: by (unfold succ_def) (blast intro!: if_not_P) paulson@13551: paulson@13551: lemma succ_not_equals: "c \ chain S - maxchain S ==> succ S c \ c" wenzelm@17200: apply (frule succI3) wenzelm@17200: apply (simp (no_asm_simp)) wenzelm@17200: apply (rule select_not_equals, assumption) wenzelm@17200: done paulson@13551: paulson@13551: lemma TFin_chain_lemma4: "c \ TFin S ==> (c :: 'a set set): chain S" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) nipkow@26272: apply (unfold chain_def chain_subset_def) wenzelm@17200: apply (rule CollectI, safe) wenzelm@17200: apply (drule bspec, assumption) wenzelm@46008: apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE]) wenzelm@46008: apply blast+ wenzelm@17200: done wenzelm@14706: paulson@13551: theorem Hausdorff: "\c. (c :: 'a set set): maxchain S" paulson@18143: apply (rule_tac x = "Union (TFin S)" in exI) wenzelm@17200: apply (rule classical) wenzelm@17200: apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") wenzelm@17200: prefer 2 wenzelm@17200: apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) wenzelm@17200: apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) wenzelm@17200: apply (drule DiffI [THEN succ_not_equals], blast+) wenzelm@17200: done paulson@13551: paulson@13551: wenzelm@14706: subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then paulson@13551: There Is a Maximal Element*} paulson@13551: wenzelm@14706: lemma chain_extend: nipkow@26272: "[| c \ chain S; z \ S; \x \ c. x \ (z:: 'a set) |] ==> {z} Un c \ chain S" nipkow@26272: by (unfold chain_def chain_subset_def) blast paulson@13551: paulson@13551: lemma chain_Union_upper: "[| c \ chain S; x \ c |] ==> x \ Union(c)" nipkow@26272: by auto paulson@13551: paulson@13551: lemma chain_ball_Union_upper: "c \ chain S ==> \x \ c. x \ Union(c)" nipkow@26272: by auto paulson@13551: paulson@13551: lemma maxchain_Zorn: nipkow@26272: "[| c \ maxchain S; u \ S; Union(c) \ u |] ==> Union(c) = u" nipkow@26272: apply (rule ccontr) nipkow@26272: apply (simp add: maxchain_def) nipkow@26272: apply (erule conjE) nipkow@26272: apply (subgoal_tac "({u} Un c) \ super S c") nipkow@26272: apply simp berghofe@26806: apply (unfold super_def less_le) nipkow@26272: apply (blast intro: chain_extend dest: chain_Union_upper) nipkow@26272: done paulson@13551: paulson@13551: theorem Zorn_Lemma: nipkow@26272: "\c \ chain S. Union(c): S ==> \y \ S. \z \ S. y \ z --> y = z" nipkow@26272: apply (cut_tac Hausdorff maxchain_subset_chain) nipkow@26272: apply (erule exE) nipkow@26272: apply (drule subsetD, assumption) nipkow@26272: apply (drule bspec, assumption) nipkow@26272: apply (rule_tac x = "Union(c)" in bexI) nipkow@26272: apply (rule ballI, rule impI) nipkow@26272: apply (blast dest!: maxchain_Zorn, assumption) nipkow@26272: done paulson@13551: paulson@13551: subsection{*Alternative version of Zorn's Lemma*} paulson@13551: paulson@13551: lemma Zorn_Lemma2: wenzelm@17200: "\c \ chain S. \y \ S. \x \ c. x \ y wenzelm@17200: ==> \y \ S. \x \ S. (y :: 'a set) \ x --> y = x" nipkow@26272: apply (cut_tac Hausdorff maxchain_subset_chain) nipkow@26272: apply (erule exE) nipkow@26272: apply (drule subsetD, assumption) nipkow@26272: apply (drule bspec, assumption, erule bexE) nipkow@26272: apply (rule_tac x = y in bexI) nipkow@26272: prefer 2 apply assumption nipkow@26272: apply clarify nipkow@26272: apply (rule ccontr) nipkow@26272: apply (frule_tac z = x in chain_extend) nipkow@26272: apply (assumption, blast) berghofe@26806: apply (unfold maxchain_def super_def less_le) nipkow@26272: apply (blast elim!: equalityCE) nipkow@26272: done paulson@13551: paulson@13551: text{*Various other lemmas*} paulson@13551: paulson@13551: lemma chainD: "[| c \ chain S; x \ c; y \ c |] ==> x \ y | y \ x" nipkow@26272: by (unfold chain_def chain_subset_def) blast paulson@13551: paulson@13551: lemma chainD2: "!!(c :: 'a set set). c \ chain S ==> c \ S" nipkow@26272: by (unfold chain_def) blast nipkow@26191: nipkow@26191: nipkow@26191: (* Define globally? In Relation.thy? *) nipkow@26191: definition Chain :: "('a*'a)set \ 'a set set" where nipkow@26191: "Chain r \ {A. \a\A.\b\A. (a,b) : r \ (b,a) \ r}" nipkow@26191: nipkow@26191: lemma mono_Chain: "r \ s \ Chain r \ Chain s" nipkow@26191: unfolding Chain_def by blast nipkow@26191: nipkow@26191: text{* Zorn's lemma for partial orders: *} nipkow@26191: nipkow@26191: lemma Zorns_po_lemma: nipkow@26191: assumes po: "Partial_order r" and u: "\C\Chain r. \u\Field r. \a\C. (a,u):r" nipkow@26191: shows "\m\Field r. \a\Field r. (m,a):r \ a=m" nipkow@26191: proof- nipkow@26295: have "Preorder r" using po by(simp add:partial_order_on_def) nipkow@26191: --{* Mirror r in the set of subsets below (wrt r) elements of A*} nipkow@26191: let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r" nipkow@26191: have "\C \ chain ?S. EX U:?S. ALL A:C. A\U" nipkow@26272: proof (auto simp:chain_def chain_subset_def) nipkow@26191: fix C assume 1: "C \ ?S" and 2: "\A\C.\B\C. A\B | B\A" nipkow@26191: let ?A = "{x\Field r. \M\C. M = ?B x}" nipkow@26191: have "C = ?B ` ?A" using 1 by(auto simp: image_def) nipkow@26191: have "?A\Chain r" nipkow@26191: proof (simp add:Chain_def, intro allI impI, elim conjE) nipkow@26191: fix a b nipkow@26191: assume "a \ Field r" "?B a \ C" "b \ Field r" "?B b \ C" blanchet@48750: hence "?B a \ ?B b \ ?B b \ ?B a" using 2 by simp nipkow@26191: thus "(a, b) \ r \ (b, a) \ r" using `Preorder r` `a:Field r` `b:Field r` wenzelm@32960: by (simp add:subset_Image1_Image1_iff) nipkow@26191: qed nipkow@26191: then obtain u where uA: "u:Field r" "\a\?A. (a,u) : r" using u by auto nipkow@26191: have "\A\C. A \ r^-1 `` {u}" (is "?P u") nipkow@26191: proof auto nipkow@26191: fix a B assume aB: "B:C" "a:B" nipkow@26191: with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto nipkow@26191: thus "(a,u) : r" using uA aB `Preorder r` blanchet@48750: by (simp add: preorder_on_def refl_on_def) (rule transD, blast+) nipkow@26191: qed nipkow@26191: thus "EX u:Field r. ?P u" using `u:Field r` by blast nipkow@26191: qed nipkow@26191: from Zorn_Lemma2[OF this] nipkow@26191: obtain m B where "m:Field r" "B = r^-1 `` {m}" nipkow@26191: "\x\Field r. B \ r^-1 `` {x} \ B = r^-1 `` {x}" ballarin@27064: by auto nipkow@26191: hence "\a\Field r. (m, a) \ r \ a = m" using po `Preorder r` `m:Field r` nipkow@26191: by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff) nipkow@26191: thus ?thesis using `m:Field r` by blast nipkow@26191: qed nipkow@26191: nipkow@26191: (* The initial segment of a relation appears generally useful. nipkow@26191: Move to Relation.thy? nipkow@26191: Definition correct/most general? nipkow@26191: Naming? nipkow@26191: *) nipkow@26191: definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where nipkow@26191: "init_seg_of == {(r,s). r \ s \ (\a b c. (a,b):s \ (b,c):r \ (a,b):r)}" nipkow@26191: nipkow@26191: abbreviation initialSegmentOf :: "('a*'a)set \ ('a*'a)set \ bool" nipkow@26191: (infix "initial'_segment'_of" 55) where nipkow@26191: "r initial_segment_of s == (r,s):init_seg_of" nipkow@26191: nipkow@30198: lemma refl_on_init_seg_of[simp]: "r initial_segment_of r" nipkow@26191: by(simp add:init_seg_of_def) nipkow@26191: nipkow@26191: lemma trans_init_seg_of: nipkow@26191: "r initial_segment_of s \ s initial_segment_of t \ r initial_segment_of t" blanchet@48750: by (simp (no_asm_use) add: init_seg_of_def) (metis (no_types) in_mono order_trans) nipkow@26191: nipkow@26191: lemma antisym_init_seg_of: nipkow@26191: "r initial_segment_of s \ s initial_segment_of r \ r=s" huffman@35175: unfolding init_seg_of_def by safe nipkow@26191: nipkow@26191: lemma Chain_init_seg_of_Union: nipkow@26191: "R \ Chain init_seg_of \ r\R \ r initial_segment_of \R" blanchet@48750: by(simp add: init_seg_of_def Chain_def Ball_def) blast nipkow@26191: nipkow@26272: lemma chain_subset_trans_Union: nipkow@26272: "chain\<^bsub>\\<^esub> R \ \r\R. trans r \ trans(\R)" blanchet@48750: apply(simp add:chain_subset_def) nipkow@26191: apply(simp (no_asm_use) add:trans_def) blanchet@48750: by (metis subsetD) nipkow@26191: nipkow@26272: lemma chain_subset_antisym_Union: nipkow@26272: "chain\<^bsub>\\<^esub> R \ \r\R. antisym r \ antisym(\R)" blanchet@48750: by (simp add:chain_subset_def antisym_def) (metis subsetD) nipkow@26191: nipkow@26272: lemma chain_subset_Total_Union: nipkow@26272: assumes "chain\<^bsub>\\<^esub> R" "\r\R. Total r" nipkow@26191: shows "Total (\R)" nipkow@26295: proof (simp add: total_on_def Ball_def, auto del:disjCI) nipkow@26191: fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a\b" nipkow@26272: from `chain\<^bsub>\\<^esub> R` `r:R` `s:R` have "r\s \ s\r" nipkow@26272: by(simp add:chain_subset_def) nipkow@26191: thus "(\r\R. (a,b) \ r) \ (\r\R. (b,a) \ r)" nipkow@26191: proof nipkow@26191: assume "r\s" hence "(a,b):s \ (b,a):s" using assms(2) A nipkow@26295: by(simp add:total_on_def)(metis mono_Field subsetD) nipkow@26191: thus ?thesis using `s:R` by blast nipkow@26191: next nipkow@26191: assume "s\r" hence "(a,b):r \ (b,a):r" using assms(2) A nipkow@26295: by(simp add:total_on_def)(metis mono_Field subsetD) nipkow@26191: thus ?thesis using `r:R` by blast nipkow@26191: qed nipkow@26191: qed nipkow@26191: nipkow@26191: lemma wf_Union_wf_init_segs: nipkow@26191: assumes "R \ Chain init_seg_of" and "\r\R. wf r" shows "wf(\R)" nipkow@26191: proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto) nipkow@26191: fix f assume 1: "\i. \r\R. (f(Suc i), f i) \ r" nipkow@26191: then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto nipkow@26191: { fix i have "(f(Suc i), f i) \ r" nipkow@26191: proof(induct i) nipkow@26191: case 0 show ?case by fact nipkow@26191: next nipkow@26191: case (Suc i) nipkow@26191: moreover obtain s where "s\R" and "(f(Suc(Suc i)), f(Suc i)) \ s" wenzelm@32960: using 1 by auto nipkow@26191: moreover hence "s initial_segment_of r \ r initial_segment_of s" wenzelm@32960: using assms(1) `r:R` by(simp add: Chain_def) nipkow@26191: ultimately show ?case by(simp add:init_seg_of_def) blast nipkow@26191: qed nipkow@26191: } nipkow@26191: thus False using assms(2) `r:R` nipkow@26191: by(simp add:wf_iff_no_infinite_down_chain) blast nipkow@26191: qed nipkow@26191: huffman@27476: lemma initial_segment_of_Diff: huffman@27476: "p initial_segment_of q \ p - s initial_segment_of q - s" huffman@27476: unfolding init_seg_of_def by blast huffman@27476: nipkow@26191: lemma Chain_inits_DiffI: nipkow@26191: "R \ Chain init_seg_of \ {r - s |r. r \ R} \ Chain init_seg_of" huffman@27476: unfolding Chain_def by (blast intro: initial_segment_of_Diff) nipkow@26191: nipkow@26272: theorem well_ordering: "\r::('a*'a)set. Well_order r \ Field r = UNIV" nipkow@26191: proof- nipkow@26191: -- {*The initial segment relation on well-orders: *} nipkow@26191: let ?WO = "{r::('a*'a)set. Well_order r}" nipkow@26191: def I \ "init_seg_of \ ?WO \ ?WO" blanchet@48750: have I_init: "I \ init_seg_of" by(simp add: I_def) nipkow@26272: hence subch: "!!R. R : Chain I \ chain\<^bsub>\\<^esub> R" nipkow@26272: by(auto simp:init_seg_of_def chain_subset_def Chain_def) nipkow@26191: have Chain_wo: "!!R r. R \ Chain I \ r \ R \ Well_order r" nipkow@26191: by(simp add:Chain_def I_def) blast nipkow@26191: have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def) nipkow@26191: hence 0: "Partial_order I" nipkow@30198: 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) nipkow@26191: -- {*I-chains have upper bounds in ?WO wrt I: their Union*} nipkow@26191: { fix R assume "R \ Chain I" nipkow@26191: hence Ris: "R \ Chain init_seg_of" using mono_Chain[OF I_init] by blast nipkow@26272: have subch: "chain\<^bsub>\\<^esub> R" using `R : Chain I` I_init nipkow@26272: by(auto simp:init_seg_of_def chain_subset_def Chain_def) nipkow@26191: have "\r\R. Refl r" "\r\R. trans r" "\r\R. antisym r" "\r\R. Total r" nipkow@26191: "\r\R. wf(r-Id)" nipkow@26295: using Chain_wo[OF `R \ Chain I`] by(simp_all add:order_on_defs) nipkow@30198: have "Refl (\R)" using `\r\R. Refl r` by(auto simp:refl_on_def) nipkow@26191: moreover have "trans (\R)" nipkow@26272: by(rule chain_subset_trans_Union[OF subch `\r\R. trans r`]) nipkow@26191: moreover have "antisym(\R)" nipkow@26272: by(rule chain_subset_antisym_Union[OF subch `\r\R. antisym r`]) nipkow@26191: moreover have "Total (\R)" nipkow@26272: by(rule chain_subset_Total_Union[OF subch `\r\R. Total r`]) nipkow@26191: moreover have "wf((\R)-Id)" nipkow@26191: proof- nipkow@26191: have "(\R)-Id = \{r-Id|r. r \ R}" by blast nipkow@26191: with `\r\R. wf(r-Id)` wf_Union_wf_init_segs[OF Chain_inits_DiffI[OF Ris]] nipkow@26191: show ?thesis by (simp (no_asm_simp)) blast nipkow@26191: qed nipkow@26295: ultimately have "Well_order (\R)" by(simp add:order_on_defs) nipkow@26191: moreover have "\r \ R. r initial_segment_of \R" using Ris nipkow@26191: by(simp add: Chain_init_seg_of_Union) nipkow@26191: ultimately have "\R : ?WO \ (\r\R. (r,\R) : I)" nipkow@26191: using mono_Chain[OF I_init] `R \ Chain I` blanchet@48750: by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo) nipkow@26191: } nipkow@26191: hence 1: "\R \ Chain I. \u\Field I. \r\R. (r,u) : I" by (subst FI) blast nipkow@26191: --{*Zorn's Lemma yields a maximal well-order m:*} nipkow@26191: then obtain m::"('a*'a)set" where "Well_order m" and nipkow@26191: max: "\r. Well_order r \ (m,r):I \ r=m" nipkow@26191: using Zorns_po_lemma[OF 0 1] by (auto simp:FI) nipkow@26191: --{*Now show by contradiction that m covers the whole type:*} nipkow@26191: { fix x::'a assume "x \ Field m" nipkow@26191: --{*We assume that x is not covered and extend m at the top with x*} nipkow@26191: have "m \ {}" nipkow@26191: proof nipkow@26191: assume "m={}" nipkow@26191: moreover have "Well_order {(x,x)}" haftmann@46752: by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_unfold Domain_converse [symmetric]) nipkow@26191: ultimately show False using max wenzelm@32960: by (auto simp:I_def init_seg_of_def simp del:Field_insert) nipkow@26191: qed nipkow@26191: hence "Field m \ {}" by(auto simp:Field_def) nipkow@26295: moreover have "wf(m-Id)" using `Well_order m` nipkow@26295: by(simp add:well_order_on_def) nipkow@26191: --{*The extension of m by x:*} nipkow@26191: let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s" nipkow@26191: have Fm: "Field ?m = insert x (Field m)" nipkow@26191: apply(simp add:Field_insert Field_Un) nipkow@26191: unfolding Field_def by auto nipkow@26191: have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)" nipkow@26295: using `Well_order m` by(simp_all add:order_on_defs) nipkow@26191: --{*We show that the extension is a well-order*} nipkow@30198: have "Refl ?m" using `Refl m` Fm by(auto simp:refl_on_def) nipkow@26191: moreover have "trans ?m" using `trans m` `x \ Field m` haftmann@46752: unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast nipkow@26191: moreover have "antisym ?m" using `antisym m` `x \ Field m` haftmann@46752: unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast nipkow@26295: moreover have "Total ?m" using `Total m` Fm by(auto simp: total_on_def) nipkow@26191: moreover have "wf(?m-Id)" nipkow@26191: proof- nipkow@26191: have "wf ?s" using `x \ Field m` blanchet@48750: by(simp add:wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis nipkow@26191: thus ?thesis using `wf(m-Id)` `x \ Field m` wenzelm@32960: wf_subset[OF `wf ?s` Diff_subset] nipkow@44890: by (fastforce intro!: wf_Un simp add: Un_Diff Field_def) nipkow@26191: qed nipkow@26295: ultimately have "Well_order ?m" by(simp add:order_on_defs) nipkow@26191: --{*We show that the extension is above m*} nipkow@26191: moreover hence "(m,?m) : I" using `Well_order m` `x \ Field m` haftmann@46752: by(fastforce simp:I_def init_seg_of_def Field_def Domain_unfold Domain_converse [symmetric]) nipkow@26191: ultimately nipkow@26191: --{*This contradicts maximality of m:*} nipkow@26191: have False using max `x \ Field m` unfolding Field_def by blast nipkow@26191: } nipkow@26191: hence "Field m = UNIV" by auto nipkow@26272: moreover with `Well_order m` have "Well_order m" by simp nipkow@26272: ultimately show ?thesis by blast nipkow@26272: qed nipkow@26272: nipkow@26295: corollary well_order_on: "\r::('a*'a)set. well_order_on A r" nipkow@26272: proof - nipkow@26272: obtain r::"('a*'a)set" where wo: "Well_order r" and univ: "Field r = UNIV" nipkow@26272: using well_ordering[where 'a = "'a"] by blast nipkow@26272: let ?r = "{(x,y). x:A & y:A & (x,y):r}" nipkow@26272: have 1: "Field ?r = A" using wo univ haftmann@46752: by(fastforce simp: Field_def Domain_unfold Domain_converse [symmetric] order_on_defs refl_on_def) nipkow@26272: have "Refl r" "trans r" "antisym r" "Total r" "wf(r-Id)" nipkow@26295: using `Well_order r` by(simp_all add:order_on_defs) nipkow@30198: have "Refl ?r" using `Refl r` by(auto simp:refl_on_def 1 univ) nipkow@26272: moreover have "trans ?r" using `trans r` nipkow@26272: unfolding trans_def by blast nipkow@26272: moreover have "antisym ?r" using `antisym r` nipkow@26272: unfolding antisym_def by blast nipkow@26295: moreover have "Total ?r" using `Total r` by(simp add:total_on_def 1 univ) nipkow@26272: moreover have "wf(?r - Id)" by(rule wf_subset[OF `wf(r-Id)`]) blast nipkow@26295: ultimately have "Well_order ?r" by(simp add:order_on_defs) nipkow@26295: with 1 show ?thesis by metis nipkow@26191: qed nipkow@26191: paulson@13551: end