blanchet@54545: (* Title: HOL/Cardinals/Wellorder_Extension.thy blanchet@54545: Author: Christian Sternagel, JAIST blanchet@54545: *) blanchet@54545: blanchet@54545: header {* Extending Well-founded Relations to Wellorders *} blanchet@54545: blanchet@54545: theory Wellorder_Extension blanchet@55018: imports Zorn Order_Union blanchet@54545: begin blanchet@54545: blanchet@54545: subsection {* Extending Well-founded Relations to Wellorders *} blanchet@54545: blanchet@54545: text {*A \emph{downset} (also lower set, decreasing set, initial segment, or blanchet@54545: downward closed set) is closed w.r.t.\ smaller elements.*} blanchet@54545: definition downset_on where blanchet@54545: "downset_on A r = (\x y. (x, y) \ r \ y \ A \ x \ A)" blanchet@54545: blanchet@54545: (* blanchet@54545: text {*Connection to order filters of the @{theory Cardinals} theory.*} blanchet@54545: lemma (in wo_rel) ofilter_downset_on_conv: blanchet@54545: "ofilter A \ downset_on A r \ A \ Field r" blanchet@54545: by (auto simp: downset_on_def ofilter_def under_def) blanchet@54545: *) blanchet@54545: blanchet@54545: lemma downset_onI: blanchet@54545: "(\x y. (x, y) \ r \ y \ A \ x \ A) \ downset_on A r" blanchet@54545: by (auto simp: downset_on_def) blanchet@54545: blanchet@54545: lemma downset_onD: blanchet@54545: "downset_on A r \ (x, y) \ r \ y \ A \ x \ A" blanchet@54545: unfolding downset_on_def by blast blanchet@54545: blanchet@54545: text {*Extensions of relations w.r.t.\ a given set.*} blanchet@54545: definition extension_on where blanchet@54545: "extension_on A r s = (\x\A. \y\A. (x, y) \ s \ (x, y) \ r)" blanchet@54545: blanchet@54545: lemma extension_onI: blanchet@54545: "(\x y. \x \ A; y \ A; (x, y) \ s\ \ (x, y) \ r) \ extension_on A r s" blanchet@54545: by (auto simp: extension_on_def) blanchet@54545: blanchet@54545: lemma extension_onD: blanchet@54545: "extension_on A r s \ x \ A \ y \ A \ (x, y) \ s \ (x, y) \ r" blanchet@54545: by (auto simp: extension_on_def) blanchet@54545: blanchet@54545: lemma downset_on_Union: blanchet@54545: assumes "\r. r \ R \ downset_on (Field r) p" blanchet@54545: shows "downset_on (Field (\R)) p" blanchet@54545: using assms by (auto intro: downset_onI dest: downset_onD) blanchet@54545: blanchet@54545: lemma chain_subset_extension_on_Union: blanchet@54545: assumes "chain\<^sub>\ R" and "\r. r \ R \ extension_on (Field r) r p" blanchet@54545: shows "extension_on (Field (\R)) (\R) p" blanchet@54545: using assms blanchet@54545: by (simp add: chain_subset_def extension_on_def) blanchet@54545: (metis (no_types) mono_Field set_mp) blanchet@54545: blanchet@54545: lemma downset_on_empty [simp]: "downset_on {} p" blanchet@54545: by (auto simp: downset_on_def) blanchet@54545: blanchet@54545: lemma extension_on_empty [simp]: "extension_on {} p q" blanchet@54545: by (auto simp: extension_on_def) blanchet@54545: blanchet@54545: text {*Every well-founded relation can be extended to a wellorder.*} blanchet@54545: theorem well_order_extension: blanchet@54545: assumes "wf p" blanchet@54545: shows "\w. p \ w \ Well_order w" blanchet@54545: proof - blanchet@54545: let ?K = "{r. Well_order r \ downset_on (Field r) p \ extension_on (Field r) r p}" blanchet@54545: def I \ "init_seg_of \ ?K \ ?K" blanchet@54545: have I_init: "I \ init_seg_of" by (simp add: I_def) blanchet@54545: then have subch: "\R. R \ Chains I \ chain\<^sub>\ R" blanchet@54545: by (auto simp: init_seg_of_def chain_subset_def Chains_def) blanchet@54545: have Chains_wo: "\R r. R \ Chains I \ r \ R \ blanchet@54545: Well_order r \ downset_on (Field r) p \ extension_on (Field r) r p" blanchet@54545: by (simp add: Chains_def I_def) blast blanchet@54545: have FI: "Field I = ?K" by (auto simp: I_def init_seg_of_def Field_def) blanchet@54545: then have 0: "Partial_order I" blanchet@54545: by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def blanchet@54545: trans_def I_def elim: trans_init_seg_of) blanchet@54545: { fix R assume "R \ Chains I" blanchet@54545: then have Ris: "R \ Chains init_seg_of" using mono_Chains [OF I_init] by blast blanchet@54545: have subch: "chain\<^sub>\ R" using `R \ Chains I` I_init blanchet@54545: by (auto simp: init_seg_of_def chain_subset_def Chains_def) blanchet@54545: have "\r\R. Refl r" and "\r\R. trans r" and "\r\R. antisym r" and blanchet@54545: "\r\R. Total r" and "\r\R. wf (r - Id)" and blanchet@54545: "\r. r \ R \ downset_on (Field r) p" and blanchet@54545: "\r. r \ R \ extension_on (Field r) r p" blanchet@54545: using Chains_wo [OF `R \ Chains I`] by (simp_all add: order_on_defs) blanchet@54545: have "Refl (\R)" using `\r\R. Refl r` unfolding refl_on_def by fastforce blanchet@54545: moreover have "trans (\R)" blanchet@54545: by (rule chain_subset_trans_Union [OF subch `\r\R. trans r`]) blanchet@54545: moreover have "antisym (\R)" blanchet@54545: by (rule chain_subset_antisym_Union [OF subch `\r\R. antisym r`]) blanchet@54545: moreover have "Total (\R)" blanchet@54545: by (rule chain_subset_Total_Union [OF subch `\r\R. Total r`]) blanchet@54545: moreover have "wf ((\R) - Id)" blanchet@54545: proof - blanchet@54545: have "(\R) - Id = \{r - Id | r. r \ R}" by blast blanchet@54545: with `\r\R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]] blanchet@54545: show ?thesis by fastforce blanchet@54545: qed blanchet@54545: ultimately have "Well_order (\R)" by (simp add: order_on_defs) blanchet@54545: moreover have "\r\R. r initial_segment_of \R" using Ris blanchet@54545: by (simp add: Chains_init_seg_of_Union) blanchet@54545: moreover have "downset_on (Field (\R)) p" blanchet@54545: by (rule downset_on_Union [OF `\r. r \ R \ downset_on (Field r) p`]) blanchet@54545: moreover have "extension_on (Field (\R)) (\R) p" blanchet@54545: by (rule chain_subset_extension_on_Union [OF subch `\r. r \ R \ extension_on (Field r) r p`]) blanchet@54545: ultimately have "\R \ ?K \ (\r\R. (r,\R) \ I)" blanchet@54545: using mono_Chains [OF I_init] and `R \ Chains I` blanchet@54545: by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo) blanchet@54545: } blanchet@54545: then have 1: "\R\Chains I. \u\Field I. \r\R. (r, u) \ I" by (subst FI) blast blanchet@54545: txt {*Zorn's Lemma yields a maximal wellorder m.*} blanchet@54545: from Zorns_po_lemma [OF 0 1] obtain m :: "('a \ 'a) set" blanchet@54545: where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and blanchet@54545: max: "\r. Well_order r \ downset_on (Field r) p \ extension_on (Field r) r p \ blanchet@54545: (m, r) \ I \ r = m" blanchet@54545: by (auto simp: FI) blanchet@54545: have "Field p \ Field m" blanchet@54545: proof (rule ccontr) blanchet@54545: let ?Q = "Field p - Field m" blanchet@54545: assume "\ (Field p \ Field m)" blanchet@54545: with assms [unfolded wf_eq_minimal, THEN spec, of ?Q] blanchet@54545: obtain x where "x \ Field p" and "x \ Field m" and blanchet@54545: min: "\y. (y, x) \ p \ y \ ?Q" by blast blanchet@54545: txt {*Add @{term x} as topmost element to @{term m}.*} blanchet@54545: let ?s = "{(y, x) | y. y \ Field m}" blanchet@54545: let ?m = "insert (x, x) m \ ?s" blanchet@54545: have Fm: "Field ?m = insert x (Field m)" by (auto simp: Field_def) blanchet@54545: have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)" blanchet@54545: using `Well_order m` by (simp_all add: order_on_defs) blanchet@54545: txt {*We show that the extension is a wellorder.*} blanchet@54545: have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def) blanchet@54545: moreover have "trans ?m" using `trans m` `x \ Field m` blanchet@54545: unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast blanchet@54545: moreover have "antisym ?m" using `antisym m` `x \ Field m` blanchet@54545: unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast blanchet@54545: moreover have "Total ?m" using `Total m` Fm by (auto simp: Relation.total_on_def) blanchet@54545: moreover have "wf (?m - Id)" blanchet@54545: proof - blanchet@54545: have "wf ?s" using `x \ Field m` blanchet@54545: by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis blanchet@54545: thus ?thesis using `wf (m - Id)` `x \ Field m` blanchet@54545: wf_subset [OF `wf ?s` Diff_subset] blanchet@54545: by (fastforce intro!: wf_Un simp add: Un_Diff Field_def) blanchet@54545: qed blanchet@54545: ultimately have "Well_order ?m" by (simp add: order_on_defs) blanchet@54545: moreover have "extension_on (Field ?m) ?m p" blanchet@54545: using `extension_on (Field m) m p` `downset_on (Field m) p` blanchet@54545: by (subst Fm) (auto simp: extension_on_def dest: downset_onD) blanchet@54545: moreover have "downset_on (Field ?m) p" blanchet@54545: apply (subst Fm) blanchet@54545: using `downset_on (Field m) p` and min blanchet@54545: unfolding downset_on_def Field_def by blast blanchet@54545: moreover have "(m, ?m) \ I" blanchet@54545: using `Well_order m` and `Well_order ?m` and blanchet@54545: `downset_on (Field m) p` and `downset_on (Field ?m) p` and blanchet@54545: `extension_on (Field m) m p` and `extension_on (Field ?m) ?m p` and blanchet@54545: `Refl m` and `x \ Field m` blanchet@54545: by (auto simp: I_def init_seg_of_def refl_on_def) blanchet@54545: ultimately blanchet@54545: --{*This contradicts maximality of m:*} blanchet@54545: show False using max and `x \ Field m` unfolding Field_def by blast blanchet@54545: qed blanchet@54545: have "p \ m" blanchet@54545: using `Field p \ Field m` and `extension_on (Field m) m p` blanchet@54545: unfolding Field_def extension_on_def by auto fast blanchet@54545: with `Well_order m` show ?thesis by blast blanchet@54545: qed blanchet@54545: blanchet@54545: text {*Every well-founded relation can be extended to a total wellorder.*} blanchet@54545: corollary total_well_order_extension: blanchet@54545: assumes "wf p" blanchet@54545: shows "\w. p \ w \ Well_order w \ Field w = UNIV" blanchet@54545: proof - blanchet@54545: from well_order_extension [OF assms] obtain w blanchet@54545: where "p \ w" and wo: "Well_order w" by blast blanchet@54545: let ?A = "UNIV - Field w" blanchet@54545: from well_order_on [of ?A] obtain w' where wo': "well_order_on ?A w'" .. blanchet@54545: have [simp]: "Field w' = ?A" using rel.well_order_on_Well_order [OF wo'] by simp blanchet@54545: have *: "Field w \ Field w' = {}" by simp blanchet@54545: let ?w = "w \o w'" blanchet@54545: have "p \ ?w" using `p \ w` by (auto simp: Osum_def) blanchet@54545: moreover have "Well_order ?w" using Osum_Well_order [OF * wo] and wo' by simp blanchet@54545: moreover have "Field ?w = UNIV" by (simp add: Field_Osum) blanchet@54545: ultimately show ?thesis by blast blanchet@54545: qed blanchet@54545: blanchet@54545: corollary well_order_on_extension: blanchet@54545: assumes "wf p" and "Field p \ A" blanchet@54545: shows "\w. p \ w \ well_order_on A w" blanchet@54545: proof - blanchet@54545: from total_well_order_extension [OF `wf p`] obtain r blanchet@54545: where "p \ r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast blanchet@54545: let ?r = "{(x, y). x \ A \ y \ A \ (x, y) \ r}" blanchet@54545: from `p \ r` have "p \ ?r" using `Field p \ A` by (auto simp: Field_def) blanchet@54545: have 1: "Field ?r = A" using wo univ blanchet@54545: by (fastforce simp: Field_def order_on_defs refl_on_def) blanchet@54545: have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)" blanchet@54545: using `Well_order r` by (simp_all add: order_on_defs) blanchet@54545: have "refl_on A ?r" using `Refl r` by (auto simp: refl_on_def univ) blanchet@54545: moreover have "trans ?r" using `trans r` blanchet@54545: unfolding trans_def by blast blanchet@54545: moreover have "antisym ?r" using `antisym r` blanchet@54545: unfolding antisym_def by blast blanchet@54545: moreover have "total_on A ?r" using `Total r` by (simp add: total_on_def univ) blanchet@54545: moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf(r - Id)`]) blast blanchet@54545: ultimately have "well_order_on A ?r" by (simp add: order_on_defs) blanchet@54545: with `p \ ?r` show ?thesis by blast blanchet@54545: qed blanchet@54545: blanchet@54545: end