# HG changeset patch # User wenzelm # Date 1610998539 -3600 # Node ID 480521bdaf3abc416108cd1f9d45e79a15ca040d # Parent 51c53a7c6473d2aae9d3074380d7331c81989fb9 tuned proofs; diff -r 51c53a7c6473 -r 480521bdaf3a src/HOL/Cardinals/Wellorder_Extension.thy --- a/src/HOL/Cardinals/Wellorder_Extension.thy Mon Jan 18 20:17:58 2021 +0100 +++ b/src/HOL/Cardinals/Wellorder_Extension.thy Mon Jan 18 20:35:39 2021 +0100 @@ -77,8 +77,9 @@ 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 \ Chains I" - then have Ris: "R \ Chains init_seg_of" using mono_Chains [OF I_init] by blast + have "\R \ ?K \ (\r\R. (r,\R) \ I)" if "R \ Chains I" for R + proof - + from that have Ris: "R \ Chains init_seg_of" using mono_Chains [OF I_init] by blast have subch: "chain\<^sub>\ R" using \R \ Chains I\ I_init by (auto simp: init_seg_of_def chain_subset_def Chains_def) have "\r\R. Refl r" and "\r\R. trans r" and "\r\R. antisym r" and @@ -106,10 +107,10 @@ by (rule downset_on_Union [OF \\r. r \ R \ downset_on (Field r) p\]) moreover have "extension_on (Field (\R)) (\R) p" by (rule chain_subset_extension_on_Union [OF subch \\r. r \ R \ extension_on (Field r) r p\]) - ultimately have "\R \ ?K \ (\r\R. (r,\R) \ I)" + ultimately show ?thesis using mono_Chains [OF I_init] and \R \ Chains I\ by (simp (no_asm) add: I_def del: Field_Union) (metis Chains_wo) - } + qed then have 1: "\u\Field I. \r\R. (r, u) \ I" if "R\Chains I" for R using that by (subst FI) blast txt \Zorn's Lemma yields a maximal wellorder m.\ @@ -196,8 +197,6 @@ where "p \ r" and wo: "Well_order r" and univ: "Field r = UNIV" by blast let ?r = "{(x, y). x \ A \ y \ A \ (x, y) \ r}" from \p \ r\ have "p \ ?r" using \Field p \ 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)