710 moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf (r - Id)`]) blast |
711 moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf (r - Id)`]) blast |
711 ultimately have "Well_order ?r" by (simp add: order_on_defs) |
712 ultimately have "Well_order ?r" by (simp add: order_on_defs) |
712 with 1 show ?thesis by metis |
713 with 1 show ?thesis by metis |
713 qed |
714 qed |
714 |
715 |
|
716 subsection {* Extending Well-founded Relations to Well-Orders *} |
|
717 |
|
718 text {*A \emph{downset} (also lower set, decreasing set, initial segment, or |
|
719 downward closed set) is closed w.r.t.\ smaller elements.*} |
|
720 definition downset_on where |
|
721 "downset_on A r = (\<forall>x y. (x, y) \<in> r \<and> y \<in> A \<longrightarrow> x \<in> A)" |
|
722 |
|
723 (* |
|
724 text {*Connection to order filters of the @{theory Cardinals} theory.*} |
|
725 lemma (in wo_rel) ofilter_downset_on_conv: |
|
726 "ofilter A \<longleftrightarrow> downset_on A r \<and> A \<subseteq> Field r" |
|
727 by (auto simp: downset_on_def ofilter_def under_def) |
|
728 *) |
|
729 |
|
730 lemma downset_onI: |
|
731 "(\<And>x y. (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A) \<Longrightarrow> downset_on A r" |
|
732 by (auto simp: downset_on_def) |
|
733 |
|
734 lemma downset_onD: |
|
735 "downset_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A \<Longrightarrow> x \<in> A" |
|
736 by (auto simp: downset_on_def) |
|
737 |
|
738 text {*Extensions of relations w.r.t.\ a given set.*} |
|
739 definition extension_on where |
|
740 "extension_on A r s = (\<forall>x\<in>A. \<forall>y\<in>A. (x, y) \<in> s \<longrightarrow> (x, y) \<in> r)" |
|
741 |
|
742 lemma extension_onI: |
|
743 "(\<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" |
|
744 by (auto simp: extension_on_def) |
|
745 |
|
746 lemma extension_onD: |
|
747 "extension_on A r s \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> (x, y) \<in> s \<Longrightarrow> (x, y) \<in> r" |
|
748 by (auto simp: extension_on_def) |
|
749 |
|
750 lemma downset_on_Union: |
|
751 assumes "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" |
|
752 shows "downset_on (Field (\<Union>R)) p" |
|
753 using assms by (auto intro: downset_onI dest: downset_onD) |
|
754 |
|
755 lemma chain_subset_extension_on_Union: |
|
756 assumes "chain\<^sub>\<subseteq> R" and "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p" |
|
757 shows "extension_on (Field (\<Union>R)) (\<Union>R) p" |
|
758 using assms |
|
759 by (simp add: chain_subset_def extension_on_def) |
|
760 (metis Field_def mono_Field set_mp) |
|
761 |
|
762 lemma downset_on_empty [simp]: "downset_on {} p" |
|
763 by (auto simp: downset_on_def) |
|
764 |
|
765 lemma extension_on_empty [simp]: "extension_on {} p q" |
|
766 by (auto simp: extension_on_def) |
|
767 |
|
768 text {*Every well-founded relation can be extended to a well-order.*} |
|
769 theorem well_order_extension: |
|
770 assumes "wf p" |
|
771 shows "\<exists>w. p \<subseteq> w \<and> Well_order w" |
|
772 proof - |
|
773 let ?K = "{r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p}" |
|
774 def I \<equiv> "init_seg_of \<inter> ?K \<times> ?K" |
|
775 have I_init: "I \<subseteq> init_seg_of" by (simp add: I_def) |
|
776 then have subch: "\<And>R. R \<in> Chains I \<Longrightarrow> chain\<^sub>\<subseteq> R" |
|
777 by (auto simp: init_seg_of_def chain_subset_def Chains_def) |
|
778 have Chains_wo: "\<And>R r. R \<in> Chains I \<Longrightarrow> r \<in> R \<Longrightarrow> |
|
779 Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p" |
|
780 by (simp add: Chains_def I_def) blast |
|
781 have FI: "Field I = ?K" by (auto simp: I_def init_seg_of_def Field_def) |
|
782 then have 0: "Partial_order I" |
|
783 by (auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def |
|
784 trans_def I_def elim: trans_init_seg_of) |
|
785 { fix R assume "R \<in> Chains I" |
|
786 then have Ris: "R \<in> Chains init_seg_of" using mono_Chains [OF I_init] by blast |
|
787 have subch: "chain\<^sub>\<subseteq> R" using `R \<in> Chains I` I_init |
|
788 by (auto simp: init_seg_of_def chain_subset_def Chains_def) |
|
789 have "\<forall>r\<in>R. Refl r" and "\<forall>r\<in>R. trans r" and "\<forall>r\<in>R. antisym r" and |
|
790 "\<forall>r\<in>R. Total r" and "\<forall>r\<in>R. wf (r - Id)" and |
|
791 "\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p" and |
|
792 "\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p" |
|
793 using Chains_wo [OF `R \<in> Chains I`] by (simp_all add: order_on_defs) |
|
794 have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by (auto simp: refl_on_def) |
|
795 moreover have "trans (\<Union>R)" |
|
796 by (rule chain_subset_trans_Union [OF subch `\<forall>r\<in>R. trans r`]) |
|
797 moreover have "antisym (\<Union>R)" |
|
798 by (rule chain_subset_antisym_Union [OF subch `\<forall>r\<in>R. antisym r`]) |
|
799 moreover have "Total (\<Union>R)" |
|
800 by (rule chain_subset_Total_Union [OF subch `\<forall>r\<in>R. Total r`]) |
|
801 moreover have "wf ((\<Union>R) - Id)" |
|
802 proof - |
|
803 have "(\<Union>R) - Id = \<Union>{r - Id | r. r \<in> R}" by blast |
|
804 with `\<forall>r\<in>R. wf (r - Id)` wf_Union_wf_init_segs [OF Chains_inits_DiffI [OF Ris]] |
|
805 show ?thesis by (simp (no_asm_simp)) blast |
|
806 qed |
|
807 ultimately have "Well_order (\<Union>R)" by (simp add: order_on_defs) |
|
808 moreover have "\<forall>r\<in>R. r initial_segment_of \<Union>R" using Ris |
|
809 by (simp add: Chains_init_seg_of_Union) |
|
810 moreover have "downset_on (Field (\<Union>R)) p" |
|
811 by (rule downset_on_Union [OF `\<And>r. r \<in> R \<Longrightarrow> downset_on (Field r) p`]) |
|
812 moreover have "extension_on (Field (\<Union>R)) (\<Union>R) p" |
|
813 by (rule chain_subset_extension_on_Union [OF subch `\<And>r. r \<in> R \<Longrightarrow> extension_on (Field r) r p`]) |
|
814 ultimately have "\<Union>R \<in> ?K \<and> (\<forall>r\<in>R. (r,\<Union>R) \<in> I)" |
|
815 using mono_Chains [OF I_init] and `R \<in> Chains I` |
|
816 by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo) |
|
817 } |
|
818 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 |
|
819 txt {*Zorn's Lemma yields a maximal well-order m.*} |
|
820 from Zorns_po_lemma [OF 0 1] obtain m :: "('a \<times> 'a) set" |
|
821 where "Well_order m" and "downset_on (Field m) p" and "extension_on (Field m) m p" and |
|
822 max: "\<forall>r. Well_order r \<and> downset_on (Field r) p \<and> extension_on (Field r) r p \<and> |
|
823 (m, r) \<in> I \<longrightarrow> r = m" |
|
824 by (auto simp: FI) |
|
825 have "Field p \<subseteq> Field m" |
|
826 proof (rule ccontr) |
|
827 let ?Q = "Field p - Field m" |
|
828 assume "\<not> (Field p \<subseteq> Field m)" |
|
829 with assms [unfolded wf_eq_minimal, THEN spec, of ?Q] |
|
830 obtain x where "x \<in> Field p" and "x \<notin> Field m" and |
|
831 min: "\<forall>y. (y, x) \<in> p \<longrightarrow> y \<notin> ?Q" by blast |
|
832 txt {*Add @{term x} as topmost element to @{term m}.*} |
|
833 let ?s = "{(y, x) | y. y \<in> Field m}" |
|
834 let ?m = "insert (x, x) m \<union> ?s" |
|
835 have Fm: "Field ?m = insert x (Field m)" by (auto simp: Field_def) |
|
836 have "Refl m" and "trans m" and "antisym m" and "Total m" and "wf (m - Id)" |
|
837 using `Well_order m` by (simp_all add: order_on_defs) |
|
838 txt {*We show that the extension is a well-order.*} |
|
839 have "Refl ?m" using `Refl m` Fm by (auto simp: refl_on_def) |
|
840 moreover have "trans ?m" using `trans m` `x \<notin> Field m` |
|
841 unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast |
|
842 moreover have "antisym ?m" using `antisym m` `x \<notin> Field m` |
|
843 unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast |
|
844 moreover have "Total ?m" using `Total m` Fm by (auto simp: Relation.total_on_def) |
|
845 moreover have "wf (?m - Id)" |
|
846 proof - |
|
847 have "wf ?s" using `x \<notin> Field m` |
|
848 by (simp add: wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis |
|
849 thus ?thesis using `wf (m - Id)` `x \<notin> Field m` |
|
850 wf_subset [OF `wf ?s` Diff_subset] |
|
851 by (fastforce intro!: wf_Un simp add: Un_Diff Field_def) |
|
852 qed |
|
853 ultimately have "Well_order ?m" by (simp add: order_on_defs) |
|
854 moreover have "extension_on (Field ?m) ?m p" |
|
855 using `extension_on (Field m) m p` `downset_on (Field m) p` |
|
856 by (subst Fm) (auto simp: extension_on_def dest: downset_onD) |
|
857 moreover have "downset_on (Field ?m) p" |
|
858 using `downset_on (Field m) p` and min |
|
859 by (subst Fm, simp add: downset_on_def Field_def) (metis Domain_iff) |
|
860 moreover have "(m, ?m) \<in> I" |
|
861 using `Well_order m` and `Well_order ?m` and |
|
862 `downset_on (Field m) p` and `downset_on (Field ?m) p` and |
|
863 `extension_on (Field m) m p` and `extension_on (Field ?m) ?m p` and |
|
864 `Refl m` and `x \<notin> Field m` |
|
865 by (auto simp: I_def init_seg_of_def refl_on_def) |
|
866 ultimately |
|
867 --{*This contradicts maximality of m:*} |
|
868 show False using max and `x \<notin> Field m` unfolding Field_def by blast |
|
869 qed |
|
870 have "p \<subseteq> m" |
|
871 using `Field p \<subseteq> Field m` and `extension_on (Field m) m p` |
|
872 by (force simp: Field_def extension_on_def) |
|
873 with `Well_order m` show ?thesis by blast |
|
874 qed |
|
875 |
|
876 text {*Every well-founded relation can be extended to a total well-order.*} |
|
877 corollary total_well_order_extension: |
|
878 assumes "wf p" |
|
879 shows "\<exists>w. p \<subseteq> w \<and> Well_order w \<and> Field w = UNIV" |
|
880 proof - |
|
881 from well_order_extension [OF assms] obtain w |
|
882 where "p \<subseteq> w" and wo: "Well_order w" by blast |
|
883 let ?A = "UNIV - Field w" |
|
884 from well_order_on [of ?A] obtain w' where wo': "well_order_on ?A w'" .. |
|
885 have [simp]: "Field w' = ?A" using rel.well_order_on_Well_order [OF wo'] by simp |
|
886 have *: "Field w \<inter> Field w' = {}" by simp |
|
887 let ?w = "w \<union>o w'" |
|
888 have "p \<subseteq> ?w" using `p \<subseteq> w` by (auto simp: Osum_def) |
|
889 moreover have "Well_order ?w" using Osum_Well_order [OF * wo] and wo' by simp |
|
890 moreover have "Field ?w = UNIV" by (simp add: Field_Osum) |
|
891 ultimately show ?thesis by blast |
|
892 qed |
|
893 |
|
894 corollary well_order_on_extension: |
|
895 assumes "wf p" and "Field p \<subseteq> A" |
|
896 shows "\<exists>w. p \<subseteq> w \<and> well_order_on A w" |
|
897 proof - |
|
898 from total_well_order_extension [OF `wf p`] obtain r |
|
899 where "p \<subseteq> r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast |
|
900 let ?r = "{(x, y). x \<in> A \<and> y \<in> A \<and> (x, y) \<in> r}" |
|
901 from `p \<subseteq> r` have "p \<subseteq> ?r" using `Field p \<subseteq> A` by (auto simp: Field_def) |
|
902 have 1: "Field ?r = A" using wo univ |
|
903 by (fastforce simp: Field_def order_on_defs refl_on_def) |
|
904 have "Refl r" "trans r" "antisym r" "Total r" "wf (r - Id)" |
|
905 using `Well_order r` by (simp_all add: order_on_defs) |
|
906 have "refl_on A ?r" using `Refl r` by (auto simp: refl_on_def univ) |
|
907 moreover have "trans ?r" using `trans r` |
|
908 unfolding trans_def by blast |
|
909 moreover have "antisym ?r" using `antisym r` |
|
910 unfolding antisym_def by blast |
|
911 moreover have "total_on A ?r" using `Total r` by (simp add: total_on_def univ) |
|
912 moreover have "wf (?r - Id)" by (rule wf_subset [OF `wf(r - Id)`]) blast |
|
913 ultimately have "well_order_on A ?r" by (simp add: order_on_defs) |
|
914 with `p \<subseteq> ?r` show ?thesis by blast |
|
915 qed |
|
916 |
715 end |
917 end |
716 |
918 |