diff --git a/thys/Complete_Non_Orders/Binary_Relations.thy b/thys/Complete_Non_Orders/Binary_Relations.thy --- a/thys/Complete_Non_Orders/Binary_Relations.thy +++ b/thys/Complete_Non_Orders/Binary_Relations.thy @@ -1,2246 +1,2366 @@ (* Author: Akihisa Yamada (2018-2021) License: LGPL (see file COPYING.LESSER) *) section \Binary Relations\ text \We start with basic properties of binary relations.\ theory Binary_Relations imports (* To verify that we don't use the axiom of choice, import HOL.Complete_Partial_Order HOL.Wellfounded - instead of *) Main + instead of *) + Main begin +unbundle lattice_syntax + +lemma conj_iff_conj_iff_imp_iff: "Trueprop (x \ y \ x \ z) \ (x \ (y \ z))" + by (auto intro!: equal_intr_rule) + lemma conj_imp_eq_imp_imp: "(P \ Q \ PROP R) \ (P \ Q \ PROP R)" by standard simp_all lemma tranclp_trancl: "r\<^sup>+\<^sup>+ = (\x y. (x,y) \ {(a,b). r a b}\<^sup>+)" by (auto simp: tranclp_trancl_eq[symmetric]) lemma tranclp_id[simp]: "transp r \ tranclp r = r" using trancl_id[of "{(x,y). r x y}", folded transp_trans] by (auto simp:tranclp_trancl) lemma transp_tranclp[simp]: "transp (tranclp r)" by (auto simp: tranclp_trancl transp_trans) lemma funpow_dom: "f ` A \ A \ (f^^n) ` A \ A" by (induct n, auto) +lemma image_subsetD: "f ` A \ B \ a \ A \ f a \ B" by auto + text \Below we introduce an Isabelle-notation for $\{ \ldots x\ldots \mid x \in X \}$.\ syntax "_range" :: "'a \ idts \ 'a set" ("(1{_ /|./ _})") "_image" :: "'a \ pttrn \ 'a set \ 'a set" ("(1{_ /|./ (_/ \ _)})") translations "{e |. p}" \ "{e | p. CONST True}" "{e |. p \ A}" \ "CONST image (\p. e) A" lemma image_constant: assumes "\i. i \ I \ f i = y" shows "f ` I = (if I = {} then {} else {y})" using assms by auto subsection \Various Definitions\ text \Here we introduce various definitions for binary relations. The first one is our abbreviation for the dual of a relation.\ abbreviation(input) dual ("(_\<^sup>-)" [1000] 1000) where "r\<^sup>- x y \ r y x" lemma conversep_is_dual[simp]: "conversep = dual" by auto +lemma dual_inf: "(r \ s)\<^sup>- = r\<^sup>- \ s\<^sup>-" by (auto intro!: ext) + text \Monotonicity is already defined in the library, but we want one restricted to a domain.\ -definition monotone_on where - "monotone_on X r s f \ \x y. x \ X \ y \ X \ r x y \ s (f x) (f y)" - -lemmas monotone_onI = monotone_on_def[unfolded atomize_eq, THEN iffD2, rule_format] -lemma monotone_onD: "monotone_on X r s f \ r x y \ x \ X \ y \ X \ s (f x) (f y)" - by (auto simp: monotone_on_def) - lemmas monotone_onE = monotone_on_def[unfolded atomize_eq, THEN iffD1, elim_format, rule_format] -lemma monotone_on_UNIV[simp]: "monotone_on UNIV = monotone" - by (intro ext, auto simp: monotone_on_def monotone_def) - lemma monotone_on_dual: "monotone_on X r s f \ monotone_on X r\<^sup>- s\<^sup>- f" by (auto simp: monotone_on_def) lemma monotone_on_id: "monotone_on X r r id" by (auto simp: monotone_on_def) lemma monotone_on_cmono: "A \ B \ monotone_on B \ monotone_on A" by (intro le_funI, auto simp: monotone_on_def) text \Here we define the following notions in a standard manner\ text \The symmetric part of a relation:\ definition sympartp where "sympartp r x y \ r x y \ r y x" lemma sympartpI[intro]: fixes r (infix "\" 50) assumes "x \ y" and "y \ x" shows "sympartp (\) x y" using assms by (auto simp: sympartp_def) lemma sympartpE[elim]: fixes r (infix "\" 50) assumes "sympartp (\) x y" and "x \ y \ y \ x \ thesis" shows thesis using assms by (auto simp: sympartp_def) lemma sympartp_dual: "sympartp r\<^sup>- = sympartp r" by (auto intro!:ext simp: sympartp_def) lemma sympartp_eq[simp]: "sympartp (=) = (=)" by auto +lemma sympartp_sympartp[simp]: "sympartp (sympartp r) = sympartp r" by (auto intro!:ext) + lemma reflclp_sympartp[simp]: "(sympartp r)\<^sup>=\<^sup>= = sympartp r\<^sup>=\<^sup>=" by auto definition "equivpartp r x y \ x = y \ r x y \ r y x" lemma sympartp_reflclp_equivp[simp]: "sympartp r\<^sup>=\<^sup>= = equivpartp r" by (auto intro!:ext simp: equivpartp_def) lemma equivpartI[simp]: "equivpartp r x x" and sympartp_equivpartpI: "sympartp r x y \ equivpartp r x y" and equivpartpCI[intro]: "(x \ y \ sympartp r x y) \ equivpartp r x y" by (auto simp:equivpartp_def) lemma equivpartpE[elim]: assumes "equivpartp r x y" and "x = y \ thesis" and "r x y \ r y x \ thesis" shows "thesis" using assms by (auto simp: equivpartp_def) lemma equivpartp_eq[simp]: "equivpartp (=) = (=)" by auto lemma sympartp_equivpartp[simp]: "sympartp (equivpartp r) = (equivpartp r)" and equivpartp_equivpartp[simp]: "equivpartp (equivpartp r) = (equivpartp r)" and equivpartp_sympartp[simp]: "equivpartp (sympartp r) = (equivpartp r)" by (auto 0 5 intro!:ext) lemma equivpartp_dual: "equivpartp r\<^sup>- = equivpartp r" by (auto intro!:ext simp: equivpartp_def) text \The asymmetric part:\ definition "asympartp r x y \ r x y \ \ r y x" lemma asympartpE[elim]: fixes r (infix "\" 50) shows "asympartp (\) x y \ (x \ y \ \y \ x \ thesis) \ thesis" by (auto simp: asympartp_def) lemmas asympartpI[intro] = asympartp_def[unfolded atomize_eq, THEN iffD2, unfolded conj_imp_eq_imp_imp, rule_format] lemma asympartp_eq[simp]: "asympartp (=) = bot" by auto lemma asympartp_sympartp [simp]: "asympartp (sympartp r) = bot" and sympartp_asympartp [simp]: "sympartp (asympartp r) = bot" by (auto intro!: ext) +lemma asympartp_dual: "asympartp r\<^sup>- = (asympartp r)\<^sup>-" by auto + text \Restriction to a set:\ definition Restrp (infixl "\" 60) where "(r \ A) a b \ a \ A \ b \ A \ r a b" lemmas RestrpI[intro!] = Restrp_def[unfolded atomize_eq, THEN iffD2, unfolded conj_imp_eq_imp_imp] lemmas RestrpE[elim!] = Restrp_def[unfolded atomize_eq, THEN iffD1, elim_format, unfolded conj_imp_eq_imp_imp] +lemma Restrp_simp[simp]: "a \ A \ b \ A \ (r \ A) a b \ r a b" by auto + lemma Restrp_UNIV[simp]: "r \ UNIV \ r" by (auto simp: atomize_eq) lemma Restrp_Restrp[simp]: "r \ A \ B \ r \ A \ B" by (auto simp: atomize_eq Restrp_def) lemma sympartp_Restrp[simp]: "sympartp (r \ A) \ sympartp r \ A" by (auto simp: atomize_eq) text \Relational images:\ definition Imagep (infixr "```" 59) where "r ``` A \ {b. \a \ A. r a b}" lemma Imagep_Image: "r ``` A = {(a,b). r a b} `` A" by (auto simp: Imagep_def) lemma in_Imagep: "b \ r ``` A \ (\a \ A. r a b)" by (auto simp: Imagep_def) lemma ImagepI: "a \ A \ r a b \ b \ r ``` A" by (auto simp: in_Imagep) lemma subset_Imagep: "B \ r ``` A \ (\b\B. \a\A. r a b)" by (auto simp: Imagep_def) text \Bounds of a set:\ -definition "bound X r b \ \x \ X. r x b" +definition "bound X (\) b \ \x \ X. x \ b" for r (infix "\" 50) lemma fixes r (infix "\" 50) shows boundI[intro!]: "(\x. x \ X \ x \ b) \ bound X (\) b" and boundE[elim]: "bound X (\) b \ ((\x. x \ X \ x \ b) \ thesis) \ thesis" + and boundD: "bound X (\) b \ a \ X \ a \ b" by (auto simp: bound_def) lemma bound_empty: "bound {} = (\r x. True)" by auto +lemma bound_cmono: assumes "X \ Y" shows "bound Y \ bound X" + using assms by auto + +lemmas bound_subset = bound_cmono[THEN le_funD, THEN le_funD, THEN le_boolD, folded atomize_imp] + +lemma bound_un: "bound (A \ B) = bound A \ bound B" + by auto + lemma bound_insert[simp]: fixes r (infix "\" 50) shows "bound (insert x X) (\) b \ x \ b \ bound X (\) b" by auto +lemma bound_cong: + assumes "A = A'" + and "b = b'" + and "\a. a \ A' \ le a b' = le' a b'" + shows "bound A le b = bound A' le' b'" + by (auto simp: assms) + +lemma bound_subsel: "le \ le' \ bound A le \ bound A le'" + by (auto simp add: bound_def) + text \Extreme (greatest) elements in a set:\ -definition "extreme X r e \ e \ X \ (\x \ X. r x e)" +definition "extreme X (\) e \ e \ X \ (\x \ X. x \ e)" for r (infix "\" 50) lemma fixes r (infix "\" 50) shows extremeI[intro]: "e \ X \ (\x. x \ X \ x \ e) \ extreme X (\) e" and extremeD: "extreme X (\) e \ e \ X" "extreme X (\) e \ (\x. x \ X \ x \ e)" and extremeE[elim]: "extreme X (\) e \ (e \ X \ (\x. x \ X \ x \ e) \ thesis) \ thesis" by (auto simp: extreme_def) lemma fixes r (infix "\" 50) shows extreme_UNIV[simp]: "extreme UNIV (\) t \ (\x. x \ t)" by auto +lemma extreme_iff_bound: "extreme X r e \ bound X r e \ e \ X" by auto + +lemma extreme_imp_bound: "extreme X r x \ bound X r x" by auto + +lemma extreme_inf: "extreme X (r \ s) x \ extreme X r x \ extreme X s x" by auto + lemma extremes_equiv: "extreme X r b \ extreme X r c \ sympartp r b c" by blast -lemma bound_cmono: assumes "X \ Y" shows "bound Y \ bound X" - using assms by auto - -lemma sympartp_sympartp[simp]: "sympartp (sympartp r) = sympartp r" by (auto intro!:ext) +lemma extreme_cong: + assumes "A = A'" + and "b = b'" + and "\a. a \ A' \ b' \ A' \ le a b' = le' a b'" + shows "extreme A le b = extreme A' le' b'" + by (auto simp: assms extreme_def) + +lemma extreme_subset: "X \ Y \ extreme X r x \ extreme Y r y \ r x y" by blast + +lemma extreme_subrel: + "le \ le' \ extreme A le \ extreme A le'" by (auto simp: extreme_def) text \Now suprema and infima are given uniformly as follows. The definition is restricted to a given set. \ + +definition + "extreme_bound A (\) X \ extreme {b \ A. bound X (\) b} (\)\<^sup>-" for r (infix "\" 50) + +lemmas extreme_boundI_extreme = extreme_bound_def[unfolded atomize_eq, THEN fun_cong, THEN iffD2] + +lemmas extreme_boundD_extreme = extreme_bound_def[unfolded atomize_eq, THEN fun_cong, THEN iffD1] + context fixes A :: "'a set" and less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin -abbreviation "extreme_bound X \ extreme {b \ A. bound X (\) b} (\x y. y \ x)" - lemma extreme_boundI[intro]: assumes "\b. bound X (\) b \ b \ A \ s \ b" and "\x. x \ X \ x \ s" and "s \ A" - shows "extreme_bound X s" - using assms by auto - -lemma extreme_bound_bound: "extreme_bound X y \ x \ X \ x \ y" by auto - -lemma extreme_bound_mono: + shows "extreme_bound A (\) X s" + using assms by (auto simp: extreme_bound_def) + +lemma extreme_boundD: + assumes "extreme_bound A (\) X s" + shows "x \ X \ x \ s" + and "bound X (\) b \ b \ A \ s \ b" + and extreme_bound_in: "s \ A" + using assms by (auto simp: extreme_bound_def) + +lemma extreme_boundE[elim]: + assumes "extreme_bound A (\) X s" + and "s \ A \ bound X (\) s \ (\b. bound X (\) b \ b \ A \ s \ b) \ thesis" + shows thesis + using assms by (auto simp: extreme_bound_def) + +lemma extreme_bound_imp_bound: "extreme_bound A (\) X s \ bound X (\) s" by auto + +lemma extreme_imp_extreme_bound: + assumes Xs: "extreme X (\) s" and XA: "X \ A" shows "extreme_bound A (\) X s" + using assms by force + +lemma extreme_bound_subset_bound: assumes XY: "X \ Y" - and sX: "extreme_bound X sX" - and sY: "extreme_bound Y sY" + and sX: "extreme_bound A (\) X s" + and b: "bound Y (\) b" and bA: "b \ A" + shows "s \ b" + using bound_subset[OF XY b] sX bA by auto + +lemma extreme_bound_subset: + assumes XY: "X \ Y" + and sX: "extreme_bound A (\) X sX" + and sY: "extreme_bound A (\) Y sY" shows "sX \ sY" -proof- - have "bound X (\) sY" using XY sY by force - with sX sY show ?thesis by (auto 0 4) -qed + using extreme_bound_subset_bound[OF XY sX] sY by auto lemma extreme_bound_iff: - shows "extreme_bound X s \ s \ A \ (\c \ A. (\x \ X. x \ c) \ s \ c) \ (\x \ X. x \ s)" - by (auto simp: extreme_def) + "extreme_bound A (\) X s \ s \ A \ (\c \ A. (\x \ X. x \ c) \ s \ c) \ (\x \ X. x \ s)" + by (auto simp: extreme_bound_def extreme_def) + +lemma extreme_bound_empty: "extreme_bound A (\) {} x \ extreme A (\)\<^sup>- x" + by auto lemma extreme_bound_singleton_refl[simp]: - "extreme_bound {x} x \ x \ A \ x \ x" by auto + "extreme_bound A (\) {x} x \ x \ A \ x \ x" by auto lemma extreme_bound_image_const: - "x \ x \ I \ {} \ (\i. i \ I \ f i = x) \ x \ A \ extreme_bound (f ` I) x" + "x \ x \ I \ {} \ (\i. i \ I \ f i = x) \ x \ A \ extreme_bound A (\) (f ` I) x" by (auto simp: image_constant) lemma extreme_bound_UN_const: "x \ x \ I \ {} \ (\i y. i \ I \ P i y \ x = y) \ x \ A \ - extreme_bound (\i\I. {y. P i y}) x" + extreme_bound A (\) (\i\I. {y. P i y}) x" by auto +lemma extreme_bounds_equiv: + assumes s: "extreme_bound A (\) X s" and s': "extreme_bound A (\) X s'" + shows "sympartp (\) s s'" + using s s' + apply (unfold extreme_bound_def) + apply (subst sympartp_dual) + by (rule extremes_equiv) + +lemma extreme_bound_sqweeze: + assumes XY: "X \ Y" and YZ: "Y \ Z" + and Xs: "extreme_bound A (\) X s" and Zs: "extreme_bound A (\) Z s" + shows "extreme_bound A (\) Y s" +proof + from Xs show "s \ A" by auto + fix b assume Yb: "bound Y (\) b" and bA: "b \ A" + from bound_subset[OF XY Yb] have "bound X (\) b". + with Xs bA + show "s \ b" by auto +next + fix y assume yY: "y \ Y" + with YZ have "y \ Z" by auto + with Zs show "y \ s" by auto +qed + +lemma bound_closed_imp_extreme_bound_eq_extreme: + assumes closed: "\b \ A. bound X (\) b \ b \ X" and XA: "X \ A" + shows "extreme_bound A (\) X = extreme X (\)" +proof (intro ext iffI extreme_boundI extremeI) + fix e + assume "extreme_bound A (\) X e" + then have Xe: "bound X (\) e" and "e \ A" by auto + with closed show "e \ X" by auto + fix x assume "x \ X" + with Xe show "x \ e" by auto +next + fix e + assume Xe: "extreme X (\) e" + then have eX: "e \ X" by auto + with XA show "e \ A" by auto + { fix b assume Xb: "bound X (\) b" and "b \ A" + from eX Xb show "e \ b" by auto + } + fix x assume xX: "x \ X" with Xe show "x \ e" by auto +qed + +end + +lemma extreme_bound_cong: + assumes "A = A'" + and "X = X'" + and "\a b. a \ A' \ b \ A' \ le a b \ le' a b" + and "\a b. a \ X' \ b \ A' \ le a b \ le' a b" + shows "extreme_bound A le X s = extreme_bound A le' X s" + apply (unfold extreme_bound_def) + apply (rule extreme_cong) + by (auto simp: assms) + + +text \Maximal or Minimal\ + +definition "extremal X (\) x \ x \ X \ (\y \ X. x \ y \ y \ x)" for r (infix "\" 50) + +context + fixes r :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma extremalI: + assumes "x \ X" "\y. y \ X \ x \ y \ y \ x" + shows "extremal X (\) x" + using assms by (auto simp: extremal_def) + +lemma extremalE: + assumes "extremal X (\) x" + and "x \ X \ (\y. y \ X \ x \ y \ y \ x) \ thesis" + shows thesis + using assms by (auto simp: extremal_def) + +lemma extremalD: + assumes "extremal X (\) x" shows "x \ X" "y \ X \ x \ y \ y \ x" + using assms by (auto elim!: extremalE) + +end + +context + fixes ir (infix "\" 50) and r (infix "\" 50) and I f + assumes mono: "monotone_on I (\) (\) f" +begin + +lemma monotone_image_bound: + assumes "X \ I" and "b \ I" and "bound X (\) b" + shows "bound (f ` X) (\) (f b)" + using assms monotone_onD[OF mono] + by (auto simp: bound_def) + +lemma monotone_image_extreme: + assumes e: "extreme I (\) e" + shows "extreme (f ` I) (\) (f e)" + using e[unfolded extreme_iff_bound] monotone_image_bound[of I e] by auto + end context fixes ir :: "'i \ 'i \ bool" (infix "\" 50) and r :: "'a \ 'a \ bool" (infix "\" 50) and f and A and e and I assumes fIA: "f ` I \ A" and mono: "monotone_on I (\) (\) f" and e: "extreme I (\) e" begin lemma monotone_extreme_imp_extreme_bound: "extreme_bound A (\) (f ` I) (f e)" using monotone_onD[OF mono] e fIA by (intro extreme_boundI, auto simp: image_def elim!: extremeE) lemma monotone_extreme_extreme_boundI: "x = f e \ extreme_bound A (\) (f ` I) x" using monotone_extreme_imp_extreme_bound by auto end subsection \Locales for Binary Relations\ text \We now define basic properties of binary relations, -in form of \emph{locales}~\<^cite>\"Kammuller00" and "locale"\.\ +in form of \emph{locales}~\cite{Kammuller00,locale}.\ subsubsection \Syntactic Locales\ text \The following locales do not assume anything, but provide infix notations for relations.\ locale less_eq_syntax = fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) locale less_syntax = fixes less :: "'a \ 'a \ bool" (infix "\" 50) locale equivalence_syntax = fixes equiv :: "'a \ 'a \ bool" (infix "\" 50) begin abbreviation equiv_class ("[_]\<^sub>\") where "[x]\<^sub>\ \ { y. x \ y }" end text \Next ones introduce abbreviations for dual etc. To avoid needless constants, one should be careful when declaring them as sublocales.\ -locale less_eq_notations = less_eq_syntax +locale less_eq_dualize = less_eq_syntax begin abbreviation (input) greater_eq (infix "\" 50) where "x \ y \ y \ x" + +end + +locale less_eq_symmetrize = less_eq_dualize +begin + abbreviation sym (infix "\" 50) where "(\) \ sympartp (\)" +abbreviation equiv (infix "(\)" 50) where "(\) \ equivpartp (\)" + +end + +locale less_eq_asymmetrize = less_eq_symmetrize +begin + abbreviation less (infix "\" 50) where "(\) \ asympartp (\)" abbreviation greater (infix "\" 50) where "(\) \ (\)\<^sup>-" -abbreviation equiv (infix "(\)" 50) where "(\) \ equivpartp (\)" lemma asym_cases[consumes 1, case_names asym sym]: assumes "x \ y" and "x \ y \ thesis" and "x \ y \ thesis" shows thesis using assms by auto end -locale less_notations = less_syntax +locale less_dualize = less_syntax begin abbreviation (input) greater (infix "\" 50) where "x \ y \ y \ x" end locale related_set = fixes A :: "'a set" and less_eq :: "'a \ 'a \ bool" (infix "\" 50) subsubsection \Basic Properties of Relations\ text \In the following we define basic properties in form of locales.\ text \Reflexivity restricted on a set:\ locale reflexive = related_set + assumes refl[intro]: "x \ A \ x \ x" begin lemma eq_implies: "x = y \ x \ A \ x \ y" by auto +lemma reflexive_subset: "B \ A \ reflexive B (\)" apply unfold_locales by auto + lemma extreme_singleton[simp]: "x \ A \ extreme {x} (\) y \ x = y" by auto lemma extreme_bound_singleton: "x \ A \ extreme_bound A (\) {x} x" by auto -lemma reflexive_subset: "B \ A \ reflexive B (\)" apply unfold_locales by auto +lemma extreme_bound_cone: "x \ A \ extreme_bound A (\) {a \ A. a \ x} x" by auto end -declare reflexive.intro[intro!] +lemmas reflexiveI[intro!] = reflexive.intro lemma reflexiveE[elim]: assumes "reflexive A r" and "(\x. x \ A \ r x x) \ thesis" shows thesis using assms by (auto simp: reflexive.refl) lemma reflexive_cong: "(\a b. a \ A \ b \ A \ r a b \ r' a b) \ reflexive A r \ reflexive A r'" by (simp add: reflexive_def) locale irreflexive = related_set A "(\)" for A and less (infix "\" 50) + assumes irrefl: "x \ A \ \ x \ x" begin lemma irreflD[simp]: "x \ x \ \x \ A" by (auto simp: irrefl) lemma implies_not_eq: "x \ y \ x \ A \ x \ y" by auto lemma Restrp_irreflexive: "irreflexive UNIV ((\)\A)" apply unfold_locales by auto lemma irreflexive_subset: "B \ A \ irreflexive B (\)" apply unfold_locales by auto end -declare irreflexive.intro[intro!] +lemmas irreflexiveI[intro!] = irreflexive.intro lemma irreflexive_cong: "(\a b. a \ A \ b \ A \ r a b \ r' a b) \ irreflexive A r \ irreflexive A r'" by (simp add: irreflexive_def) +context reflexive begin + +interpretation less_eq_asymmetrize. + +lemma asympartp_irreflexive: "irreflexive A (\)" by auto + +end + locale transitive = related_set + assumes trans[trans]: "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" begin -interpretation less_eq_notations. - lemma Restrp_transitive: "transitive UNIV ((\)\A)" apply unfold_locales by (auto intro: trans) lemma bound_trans[trans]: "bound X (\) b \ b \ c \ X \ A \ b \ A \ c \ A \ bound X (\) c" by (auto 0 4 dest: trans) +lemma extreme_bound_mono: + assumes XY: "\x\X. \y\Y. x \ y" and XA: "X \ A" and YA: "Y \ A" + and sX: "extreme_bound A (\) X sX" + and sY: "extreme_bound A (\) Y sY" + shows "sX \ sY" +proof (intro extreme_boundD(2)[OF sX] CollectI conjI boundI) + from sY show sYA: "sY \ A" by auto + from sY have "bound Y (\) sY" by auto + fix x assume xX: "x \ X" with XY obtain y where yY: "y \ Y" and xy: "x \ y" by auto + from yY sY have "y \ sY" by auto + from trans[OF xy this] xX XA yY YA sYA show "x \ sY" by auto +qed + lemma transitive_subset: assumes BA: "B \ A" shows "transitive B (\)" apply unfold_locales using trans BA by blast -lemma asympartp_transitive: "transitive A (\)" +lemma asympartp_transitive: "transitive A (asympartp (\))" apply unfold_locales by (auto dest:trans) lemma reflclp_transitive: "transitive A (\)\<^sup>=\<^sup>=" apply unfold_locales by (auto dest: trans) text \The symmetric part is also transitive, but this is done in the later semiattractive locale\ end -declare transitive.intro[intro?] +lemmas transitiveI = transitive.intro + +lemma transitive_ball[code]: + "transitive A (\) \ (\x \ A. \y \ A. \z \ A. x \ y \ y \ z \ x \ z)" + for less_eq (infix "\" 50) + by (auto simp: transitive_def) lemma transitive_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "transitive A r \ transitive A r'" proof (intro iffI) show "transitive A r \ transitive A r'" apply (intro transitive.intro) apply (unfold r[symmetric]) using transitive.trans. show "transitive A r' \ transitive A r" apply (intro transitive.intro) apply (unfold r) using transitive.trans. qed +lemma transitive_empty[intro!]: "transitive {} r" by (auto intro!: transitive.intro) + lemma tranclp_transitive: "transitive A (tranclp r)" using tranclp_trans by unfold_locales locale symmetric = related_set A "(\)" for A and equiv (infix "\" 50) + assumes sym[sym]: "x \ y \ x \ A \ y \ A \ y \ x" begin lemma sym_iff: "x \ A \ y \ A \ x \ y \ y \ x" by (auto dest: sym) lemma Restrp_symmetric: "symmetric UNIV ((\)\A)" apply unfold_locales by (auto simp: sym_iff) lemma symmetric_subset: "B \ A \ symmetric B (\)" apply unfold_locales by (auto dest: sym) end -declare symmetric.intro[intro] +lemmas symmetricI[intro] = symmetric.intro lemma symmetric_cong: "(\a b. a \ A \ b \ A \ r a b \ r' a b) \ symmetric A r \ symmetric A r'" by (auto simp: symmetric_def) +lemma symmetric_empty[intro!]: "symmetric {} r" by auto global_interpretation sympartp: symmetric UNIV "sympartp r" rewrites "\r. r \ UNIV \ r" and "\x. x \ UNIV \ True" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" by auto lemma sympartp_symmetric: "symmetric A (sympartp r)" by auto locale antisymmetric = related_set + assumes antisym: "x \ y \ y \ x \ x \ A \ y \ A \ x = y" begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. lemma sym_iff_eq_refl: "x \ A \ y \ A \ x \ y \ x = y \ y \ y" by (auto dest: antisym) lemma equiv_iff_eq[simp]: "x \ A \ y \ A \ x \ y \ x = y" by (auto dest: antisym elim: equivpartpE) lemma extreme_unique: "X \ A \ extreme X (\) x \ extreme X (\) y \ x = y" by (elim extremeE, auto dest!: antisym[OF _ _ subsetD]) lemma ex_extreme_iff_ex1: "X \ A \ Ex (extreme X (\)) \ Ex1 (extreme X (\))" by (auto simp: extreme_unique) lemma ex_extreme_iff_the: "X \ A \ Ex (extreme X (\)) \ extreme X (\) (The (extreme X (\)))" apply (rule iffI) apply (rule theI') using extreme_unique by auto +lemma eq_The_extreme: "X \ A \ extreme X (\) x \ x = The (extreme X (\))" + by (rule the1_equality[symmetric], auto simp: ex_extreme_iff_ex1[symmetric]) + lemma Restrp_antisymmetric: "antisymmetric UNIV ((\)\A)" apply unfold_locales by (auto dest: antisym) lemma antisymmetric_subset: "B \ A \ antisymmetric B (\)" apply unfold_locales using antisym by auto end -declare antisymmetric.intro[intro] +lemmas antisymmetricI[intro] = antisymmetric.intro lemma antisymmetric_cong: "(\a b. a \ A \ b \ A \ r a b \ r' a b) \ antisymmetric A r \ antisymmetric A r'" by (auto simp: antisymmetric_def) +lemma antisymmetric_empty[intro!]: "antisymmetric {} r" by auto + lemma antisymmetric_union: fixes less_eq (infix "\" 50) assumes A: "antisymmetric A (\)" and B: "antisymmetric B (\)" and AB: "\a \ A. \b \ B. a \ b \ b \ a \ a = b" shows "antisymmetric (A \ B) (\)" proof- interpret A: antisymmetric A "(\)" using A. interpret B: antisymmetric B "(\)" using B. show ?thesis by (auto dest: AB[rule_format] A.antisym B.antisym) qed text \The following notion is new, generalizing antisymmetry and transitivity.\ locale semiattractive = related_set + assumes attract: "x \ y \ y \ x \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. lemma equiv_order_trans[trans]: assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" shows "x \ z" using attract[OF _ _ _ x y z] xy yz by (auto elim: equivpartpE) lemma equiv_transitive: "transitive A (\)" proof unfold_locales fix x y z assume x: "x \ A" and y: "y \ A" and z: "z \ A" and xy: "x \ y" and yz: "y \ z" show "x \ z" using equiv_order_trans[OF xy _ x y z] attract[OF _ _ _ z y x] xy yz by (auto simp:equivpartp_def) qed lemma sym_order_trans[trans]: assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" shows "x \ z" using attract[OF _ _ _ x y z] xy yz by auto interpretation sym: transitive A "(\)" proof unfold_locales fix x y z assume x: "x \ A" and y: "y \ A" and z: "z \ A" and xy: "x \ y" and yz: "y \ z" show "x \ z" using sym_order_trans[OF xy _ x y z] attract[OF _ _ _ z y x] xy yz by auto qed lemmas sym_transitive = sym.transitive_axioms lemma extreme_bound_quasi_const: assumes C: "C \ A" and x: "x \ A" and C0: "C \ {}" and const: "\y \ C. y \ x" shows "extreme_bound A (\) C x" proof (intro extreme_boundI x) from C0 obtain c where cC: "c \ C" by auto with C have c: "c \ A" by auto from cC const have cx: "c \ x" by auto fix b assume b: "b \ A" and "bound C (\) b" with cC have cb: "c \ b" by auto from attract[OF _ _ cb x c b] cx show "x \ b" by auto next fix c assume "c \ C" with const show "c \ x" by auto qed lemma extreme_bound_quasi_const_iff: assumes C: "C \ A" and x: "x \ A" and y: "y \ A" and C0: "C \ {}" and const: "\z \ C. z \ x" shows "extreme_bound A (\) C y \ x \ y" proof (intro iffI) assume y: "extreme_bound A (\) C y" note x = extreme_bound_quasi_const[OF C x C0 const] - from extremes_equiv[OF y x] + from extreme_bounds_equiv[OF y x] show "x \ y" by auto next assume xy: "x \ y" with const C sym.trans[OF _ xy _ x y] have Cy: "\z \ C. z \ y" by auto show "extreme_bound A (\) C y" using extreme_bound_quasi_const[OF C y C0 Cy]. qed lemma Restrp_semiattractive: "semiattractive UNIV ((\)\A)" apply unfold_locales by (auto dest: attract) lemma semiattractive_subset: "B \ A \ semiattractive B (\)" apply unfold_locales using attract by blast end +lemmas semiattractiveI = semiattractive.intro + lemma semiattractive_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "semiattractive A r \ semiattractive A r'" (is "?l \ ?r") proof show "?l \ ?r" apply (intro semiattractive.intro) apply (unfold r[symmetric]) using semiattractive.attract. show "?r \ ?l" apply (intro semiattractive.intro) apply (unfold r) using semiattractive.attract. qed +lemma semiattractive_empty[intro!]: "semiattractive {} r" + by (auto intro!: semiattractiveI) + locale attractive = semiattractive + assumes "semiattractive A (\)\<^sup>-" begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. sublocale dual: semiattractive A "(\)\<^sup>-" rewrites "\r. sympartp (r \ A) \ sympartp r \ A" and "\r. sympartp (sympartp r) \ sympartp r" and "sympartp ((\) \ A)\<^sup>- \ (\) \ A" and "sympartp (\)\<^sup>- \ (\)" and "equivpartp (\)\<^sup>- \ (\)" using attractive_axioms[unfolded attractive_def] by (auto intro!: ext simp: attractive_axioms_def atomize_eq equivpartp_def) lemma order_equiv_trans[trans]: assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" shows "x \ z" using dual.attract[OF _ _ _ z y x] xy yz by auto lemma order_sym_trans[trans]: assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" shows "x \ z" using dual.attract[OF _ _ _ z y x] xy yz by auto +lemma extreme_bound_sym_trans: + assumes XA: "X \ A" and Xx: "extreme_bound A (\) X x" + and xy: "x \ y" and yA: "y \ A" + shows "extreme_bound A (\) X y" +proof (intro extreme_boundI yA) + from Xx have xA: "x \ A" by auto + { + fix b assume Xb: "bound X (\) b" and bA: "b \ A" + with Xx have xb: "x \ b" by auto + from sym_order_trans[OF _ xb yA xA bA] xy show "y \ b" by auto + } + fix a assume aX: "a \ X" + with Xx have ax: "a \ x" by auto + from aX XA have aA: "a \ A" by auto + from order_sym_trans[OF ax xy aA xA yA] show "a \ y" by auto +qed + interpretation Restrp: semiattractive UNIV "(\)\A" using Restrp_semiattractive. interpretation dual.Restrp: semiattractive UNIV "(\)\<^sup>-\A" using dual.Restrp_semiattractive. lemma Restrp_attractive: "attractive UNIV ((\)\A)" apply unfold_locales using dual.Restrp.attract by auto lemma attractive_subset: "B \ A \ attractive B (\)" apply (intro attractive.intro attractive_axioms.intro) using semiattractive_subset dual.semiattractive_subset by auto end +lemmas attractiveI = attractive.intro[OF _ attractive_axioms.intro] + lemma attractive_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "attractive A r \ attractive A r'" by (simp add: attractive_def attractive_axioms_def r cong: semiattractive_cong) +lemma attractive_empty[intro!]: "attractive {} r" + by (auto intro!: attractiveI) + context antisymmetric begin sublocale attractive apply unfold_locales by (auto dest: antisym) end context transitive begin sublocale attractive rewrites "\r. sympartp (r \ A) \ sympartp r \ A" and "\r. sympartp (sympartp r) \ sympartp r" and "sympartp (\)\<^sup>- \ sympartp (\)" and "(sympartp (\))\<^sup>- \ sympartp (\)" and "(sympartp (\) \ A)\<^sup>- \ sympartp (\) \ A" and "asympartp (asympartp (\)) = asympartp (\)" and "asympartp (sympartp (\)) = bot" and "asympartp (\) \ A = asympartp ((\) \ A)" apply unfold_locales by (auto intro!:ext dest: trans simp: atomize_eq) end subsection \Combined Properties\ text \Some combinations of the above basic properties are given names.\ locale asymmetric = related_set A "(\)" for A and less (infix "\" 50) + assumes asym: "x \ y \ y \ x \ x \ A \ y \ A \ False" begin sublocale irreflexive apply unfold_locales by (auto dest: asym) lemma antisymmetric_axioms: "antisymmetric A (\)" apply unfold_locales by (auto dest: asym) lemma Restrp_asymmetric: "asymmetric UNIV ((\)\A)" apply unfold_locales by (auto dest:asym) lemma asymmetric_subset: "B \ A \ asymmetric B (\)" apply unfold_locales using asym by auto end +lemmas asymmetricI = asymmetric.intro + lemma asymmetric_iff_irreflexive_antisymmetric: fixes less (infix "\" 50) shows "asymmetric A (\) \ irreflexive A (\) \ antisymmetric A (\)" (is "?l \ ?r") proof assume ?l then interpret asymmetric. show ?r by (auto dest: asym) next assume ?r then interpret irreflexive + antisymmetric A "(\)" by auto - show ?l by (auto intro!:asymmetric.intro dest: antisym irrefl) + show ?l by (auto intro!:asymmetricI dest: antisym irrefl) qed lemma asymmetric_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "asymmetric A r \ asymmetric A r'" by (simp add: asymmetric_iff_irreflexive_antisymmetric r cong: irreflexive_cong antisymmetric_cong) +lemma asymmetric_empty: "asymmetric {} r" + by (auto simp: asymmetric_iff_irreflexive_antisymmetric) + locale quasi_ordered_set = reflexive + transitive begin lemma quasi_ordered_subset: "B \ A \ quasi_ordered_set B (\)" apply intro_locales using reflexive_subset transitive_subset by auto end +lemmas quasi_ordered_setI = quasi_ordered_set.intro + lemma quasi_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "quasi_ordered_set A r \ quasi_ordered_set A r'" by (simp add: quasi_ordered_set_def r cong: reflexive_cong transitive_cong) +lemma quasi_ordered_set_empty[intro!]: "quasi_ordered_set {} r" + by (auto intro!: quasi_ordered_set.intro) + +lemma rtranclp_quasi_ordered: "quasi_ordered_set A (rtranclp r)" + by (unfold_locales, auto) + locale near_ordered_set = antisymmetric + transitive begin interpretation Restrp: antisymmetric UNIV "(\)\A" using Restrp_antisymmetric. interpretation Restrp: transitive UNIV "(\)\A" using Restrp_transitive. lemma Restrp_near_order: "near_ordered_set UNIV ((\)\A)".. lemma near_ordered_subset: "B \ A \ near_ordered_set B (\)" apply intro_locales using antisymmetric_subset transitive_subset by auto end +lemmas near_ordered_setI = near_ordered_set.intro + lemma near_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "near_ordered_set A r \ near_ordered_set A r'" by (simp add: near_ordered_set_def r cong: antisymmetric_cong transitive_cong) +lemma near_ordered_set_empty[intro!]: "near_ordered_set {} r" + by (auto intro!: near_ordered_set.intro) + locale pseudo_ordered_set = reflexive + antisymmetric begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. lemma sym_eq[simp]: "x \ A \ y \ A \ x \ y \ x = y" by (auto simp: refl dest: antisym) lemma extreme_bound_singleton_eq[simp]: "x \ A \ extreme_bound A (\) {x} y \ x = y" by (auto intro!: antisym) lemma eq_iff: "x \ A \ y \ A \ x = y \ x \ y \ y \ x" by (auto dest: antisym simp:refl) lemma extreme_order_iff_eq: "e \ A \ extreme {x \ A. x \ e} (\) s \ e = s" by (auto intro!: antisym) lemma pseudo_ordered_subset: "B \ A \ pseudo_ordered_set B (\)" apply intro_locales using reflexive_subset antisymmetric_subset by auto end +lemmas pseudo_ordered_setI = pseudo_ordered_set.intro + lemma pseudo_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "pseudo_ordered_set A r \ pseudo_ordered_set A r'" by (simp add: pseudo_ordered_set_def r cong: reflexive_cong antisymmetric_cong) +lemma pseudo_ordered_set_empty[intro!]: "pseudo_ordered_set {} r" + by (auto intro!: pseudo_ordered_setI) + locale partially_ordered_set = reflexive + antisymmetric + transitive begin sublocale pseudo_ordered_set + quasi_ordered_set + near_ordered_set .. lemma partially_ordered_subset: "B \ A \ partially_ordered_set B (\)" apply intro_locales using reflexive_subset transitive_subset antisymmetric_subset by auto end +lemmas partially_ordered_setI = partially_ordered_set.intro + lemma partially_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "partially_ordered_set A r \ partially_ordered_set A r'" by (simp add: partially_ordered_set_def r cong: reflexive_cong antisymmetric_cong transitive_cong) +lemma partially_ordered_set_empty[intro!]: "partially_ordered_set {} r" + by (auto intro!: partially_ordered_setI) + locale strict_ordered_set = irreflexive + transitive A "(\)" begin sublocale asymmetric proof fix x y assume x: "x \ A" and y: "y \ A" assume xy: "x \ y" also assume yx: "y \ x" finally have "x \ x" using x y by auto with x show False by auto qed lemma near_ordered_set_axioms: "near_ordered_set A (\)" using antisymmetric_axioms by intro_locales interpretation Restrp: asymmetric UNIV "(\)\A" using Restrp_asymmetric. interpretation Restrp: transitive UNIV "(\)\A" using Restrp_transitive. lemma Restrp_strict_order: "strict_ordered_set UNIV ((\)\A)".. lemma strict_ordered_subset: "B \ A \ strict_ordered_set B (\)" apply intro_locales using irreflexive_subset transitive_subset by auto end +lemmas strict_ordered_setI = strict_ordered_set.intro + lemma strict_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "strict_ordered_set A r \ strict_ordered_set A r'" by (simp add: strict_ordered_set_def r cong: irreflexive_cong transitive_cong) +lemma strict_ordered_set_empty[intro!]: "strict_ordered_set {} r" + by (auto intro!: strict_ordered_set.intro) + locale tolerance = symmetric + reflexive A "(\)" begin lemma tolerance_subset: "B \ A \ tolerance B (\)" apply intro_locales using symmetric_subset reflexive_subset by auto end +lemmas toleranceI = tolerance.intro + lemma tolerance_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "tolerance A r \ tolerance A r'" by (simp add: tolerance_def r cong: reflexive_cong symmetric_cong) +lemma tolerance_empty[intro!]: "tolerance {} r" by (auto intro!: toleranceI) + global_interpretation equiv: tolerance UNIV "equivpartp r" rewrites "\r. r \ UNIV \ r" and "\x. x \ UNIV \ True" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" by unfold_locales (auto simp:equivpartp_def) locale partial_equivalence = symmetric + assumes "transitive A (\)" begin sublocale transitive A "(\)" rewrites "sympartp (\)\A \ (\)\A" and "sympartp ((\)\A) \ (\)\A" using partial_equivalence_axioms unfolding partial_equivalence_axioms_def partial_equivalence_def by (auto simp: atomize_eq sym intro!:ext) lemma partial_equivalence_subset: "B \ A \ partial_equivalence B (\)" apply (intro partial_equivalence.intro partial_equivalence_axioms.intro) using symmetric_subset transitive_subset by auto end +lemmas partial_equivalenceI = partial_equivalence.intro[OF _ partial_equivalence_axioms.intro] + lemma partial_equivalence_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "partial_equivalence A r \ partial_equivalence A r'" by (simp add: partial_equivalence_def partial_equivalence_axioms_def r cong: transitive_cong symmetric_cong) +lemma partial_equivalence_empty[intro!]: "partial_equivalence {} r" + by (auto intro!: partial_equivalenceI) + locale equivalence = symmetric + reflexive A "(\)" + transitive A "(\)" begin sublocale tolerance + partial_equivalence + quasi_ordered_set A "(\)".. lemma equivalence_subset: "B \ A \ equivalence B (\)" apply (intro equivalence.intro) using symmetric_subset transitive_subset by auto end +lemmas equivalenceI = equivalence.intro + lemma equivalence_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "equivalence A r \ equivalence A r'" by (simp add: equivalence_def r cong: reflexive_cong transitive_cong symmetric_cong) text \Some combinations lead to uninteresting relations.\ context fixes r :: "'a \ 'a \ bool" (infix "\" 50) begin proposition reflexive_irreflexive_is_empty: assumes r: "reflexive A (\)" and ir: "irreflexive A (\)" shows "A = {}" proof (rule ccontr) interpret irreflexive A "(\)" using ir. interpret reflexive A "(\)" using r. assume "A \ {}" then obtain a where a: "a \ A" by auto from a refl have "a \ a" by auto with irrefl a show False by auto qed proposition symmetric_antisymmetric_imp_eq: assumes s: "symmetric A (\)" and as: "antisymmetric A (\)" shows "(\)\A \ (=)" proof- interpret symmetric A "(\)" + antisymmetric A "(\)" using assms by auto show "?thesis" using antisym by (auto dest: sym) qed proposition nontolerance: shows "irreflexive A (\) \ symmetric A (\) \ tolerance A (\x y. \ x \ y)" proof (intro iffI conjI, elim conjE) assume "irreflexive A (\)" and "symmetric A (\)" then interpret irreflexive A "(\)" + symmetric A "(\)". show "tolerance A (\x y. \ x \ y)" by (unfold_locales, auto dest: sym irrefl) next assume "tolerance A (\x y. \ x \ y)" then interpret tolerance A "\x y. \ x \ y". show "irreflexive A (\)" by (auto simp: eq_implies) show "symmetric A (\)" using sym by auto qed proposition irreflexive_transitive_symmetric_is_empty: assumes irr: "irreflexive A (\)" and tr: "transitive A (\)" and sym: "symmetric A (\)" shows "(\)\A = bot" proof (intro ext, unfold bot_fun_def bot_bool_def eq_False, rule notI, erule RestrpE) interpret strict_ordered_set A "(\)" using assms by (unfold strict_ordered_set_def, auto) interpret symmetric A "(\)" using assms by auto fix x y assume x: "x \ A" and y: "y \ A" assume xy: "x \ y" also note sym[OF xy x y] finally have "x \ x" using x y by auto with x show False by auto qed end subsection \Totality\ -locale semiconnex = related_set A "(\)" for A and less (infix "\" 50) + +locale semiconnex = related_set _ "(\)" + less_syntax + assumes semiconnex: "x \ A \ y \ A \ x \ y \ x = y \ y \ x" begin lemma cases[consumes 2, case_names less eq greater]: assumes "x \ A" and "y \ A" and "x \ y \ P" and "x = y \ P" and "y \ x \ P" shows "P" using semiconnex assms by auto lemma neqE: assumes "x \ A" and "y \ A" shows "x \ y \ (x \ y \ P) \ (y \ x \ P) \ P" by (cases rule: cases[OF assms], auto) lemma semiconnex_subset: "B \ A \ semiconnex B (\)" apply (intro semiconnex.intro) using semiconnex by auto end -declare semiconnex.intro[intro] - -text \Totality is negated antisymmetry \<^cite>\\Proposition 2.2.4\ in "Schmidt1993"\.\ +lemmas semiconnexI[intro] = semiconnex.intro + +text \Totality is negated antisymmetry \cite[Proposition 2.2.4]{Schmidt1993}.\ proposition semiconnex_iff_neg_antisymmetric: fixes less (infix "\" 50) shows "semiconnex A (\) \ antisymmetric A (\x y. \ x \ y)" (is "?l \ ?r") -proof (intro iffI semiconnex.intro antisymmetric.intro) +proof (intro iffI semiconnexI antisymmetricI) assume ?l then interpret semiconnex. fix x y assume "x \ A" "y \ A" "\ x \ y" and "\ y \ x" then show "x = y" by (cases rule: cases, auto) next assume ?r then interpret neg: antisymmetric A "(\x y. \ x \ y)". fix x y show "x \ A \ y \ A \ x \ y \ x = y \ y \ x" using neg.antisym by auto qed lemma semiconnex_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "semiconnex A r \ semiconnex A r'" by (simp add: semiconnex_iff_neg_antisymmetric r cong: antisymmetric_cong) locale semiconnex_irreflexive = semiconnex + irreflexive begin lemma neq_iff: "x \ A \ y \ A \ x \ y \ x \ y \ y \ x" by (auto elim:neqE dest: irrefl) lemma semiconnex_irreflexive_subset: "B \ A \ semiconnex_irreflexive B (\)" apply (intro semiconnex_irreflexive.intro) using semiconnex_subset irreflexive_subset by auto end +lemmas semiconnex_irreflexiveI = semiconnex_irreflexive.intro + lemma semiconnex_irreflexive_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "semiconnex_irreflexive A r \ semiconnex_irreflexive A r'" by (simp add: semiconnex_irreflexive_def r cong: semiconnex_cong irreflexive_cong) locale connex = related_set + assumes comparable: "x \ A \ y \ A \ x \ y \ y \ x" begin -interpretation less_eq_notations. +interpretation less_eq_asymmetrize. sublocale reflexive apply unfold_locales using comparable by auto lemma comparable_cases[consumes 2, case_names le ge]: assumes "x \ A" and "y \ A" and "x \ y \ P" and "y \ x \ P" shows "P" using assms comparable by auto lemma comparable_three_cases[consumes 2, case_names less eq greater]: assumes "x \ A" and "y \ A" and "x \ y \ P" and "x \ y \ P" and "y \ x \ P" shows "P" using assms comparable by auto lemma assumes x: "x \ A" and y: "y \ A" shows not_iff_asym: "\x \ y \ y \ x" - and not_asym_iff[simp]: "\x \ y \ y \ x" + and not_asym_iff: "\x \ y \ y \ x" using comparable[OF x y] by auto lemma connex_subset: "B \ A \ connex B (\)" by (intro connex.intro comparable, auto) +interpretation less_eq_asymmetrize. + end +lemmas connexI[intro] = connex.intro + lemmas connexE = connex.comparable_cases -lemmas connexI[intro] = connex.intro +lemma connex_empty: "connex {} A" by auto context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin lemma connex_iff_semiconnex_reflexive: "connex A (\) \ semiconnex A (\) \ reflexive A (\)" (is "?c \ ?t \ ?r") proof (intro iffI conjI; (elim conjE)?) assume ?c then interpret connex. show ?t apply unfold_locales using comparable by auto show ?r by unfold_locales next assume ?t then interpret semiconnex A "(\)". assume ?r then interpret reflexive. from semiconnex show ?c by auto qed lemma chain_connect: "Complete_Partial_Order.chain r A \ connex A r" by (auto intro!: ext simp: atomize_eq connex_def Complete_Partial_Order.chain_def) lemma connex_union: assumes "connex X (\)" and "connex Y (\)" and "\x \ X. \y \ Y. x \ y \ y \ x" shows "connex (X\Y) (\)" using assms by (auto simp: connex_def) end lemma connex_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "connex A r \ connex A r'" by (simp add: connex_iff_semiconnex_reflexive r cong: semiconnex_cong reflexive_cong) locale total_pseudo_ordered_set = connex + antisymmetric begin sublocale pseudo_ordered_set .. lemma not_weak_iff: assumes x: "x \ A" and y: "y \ A" shows "\ y \ x \ x \ y \ x \ y" using x y by (cases rule: comparable_cases, auto intro:antisym) lemma total_pseudo_ordered_subset: "B \ A \ total_pseudo_ordered_set B (\)" apply (intro_locales) using antisymmetric_subset connex_subset by auto +interpretation less_eq_asymmetrize. + +interpretation asympartp: semiconnex_irreflexive A "(\)" +proof (intro semiconnex_irreflexive.intro asympartp_irreflexive semiconnexI) + fix x y assume xA: "x \ A" and yA: "y \ A" + with comparable antisym + show "x \ y \ x = y \ y \ x" by (auto simp: asympartp_def) +qed + +lemmas asympartp_semiconnex = asympartp.semiconnex_axioms +lemmas asympartp_semiconnex_irreflexive = asympartp.semiconnex_irreflexive_axioms + end +lemmas total_pseudo_ordered_setI = total_pseudo_ordered_set.intro + lemma total_pseudo_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "total_pseudo_ordered_set A r \ total_pseudo_ordered_set A r'" by (simp add: total_pseudo_ordered_set_def r cong: connex_cong antisymmetric_cong) locale total_quasi_ordered_set = connex + transitive begin sublocale quasi_ordered_set .. lemma total_quasi_ordered_subset: "B \ A \ total_quasi_ordered_set B (\)" using transitive_subset connex_subset by intro_locales end +lemmas total_quasi_ordered_setI = total_quasi_ordered_set.intro + lemma total_quasi_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "total_quasi_ordered_set A r \ total_quasi_ordered_set A r'" by (simp add: total_quasi_ordered_set_def r cong: connex_cong transitive_cong) locale total_ordered_set = total_quasi_ordered_set + antisymmetric begin sublocale partially_ordered_set + total_pseudo_ordered_set .. lemma total_ordered_subset: "B \ A \ total_ordered_set B (\)" using total_quasi_ordered_subset antisymmetric_subset by (intro total_ordered_set.intro) +lemma weak_semiconnex: "semiconnex A (\)" + using connex_axioms by (simp add: connex_iff_semiconnex_reflexive) + +interpretation less_eq_asymmetrize. + end +lemmas total_ordered_setI = total_ordered_set.intro[OF total_quasi_ordered_setI] + lemma total_ordered_set_cong: assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" shows "total_ordered_set A r \ total_ordered_set A r'" by (simp add: total_ordered_set_def r cong: total_quasi_ordered_set_cong antisymmetric_cong) -subsection \Well-Foundedness\ - -locale well_founded = related_set A "(\)" for A and less (infix "\" 50) + - assumes induct[consumes 1, case_names less, induct set]: - "a \ A \ (\x. x \ A \ (\y. y \ A \ y \ x \ P y) \ P x) \ P a" -begin - -sublocale asymmetric -proof (intro asymmetric.intro notI) - fix x y - assume xA: "x \ A" - then show "y \ A \ x \ y \ y \ x \ False" - by (induct arbitrary: y rule: induct, auto) -qed - -lemma prefixed_Imagep_imp_empty: - assumes a: "X \ ((\) ``` X) \ A" shows "X = {}" -proof - - from a have XA: "X \ A" by auto - have "x \ A \ x \ X" for x - proof (induct x rule: induct) - case (less x) - with a show ?case by (auto simp: Imagep_def) - qed - with XA show ?thesis by auto -qed - -lemma nonempty_imp_ex_extremal: - assumes QA: "Q \ A" and Q: "Q \ {}" - shows "\z \ Q. \y \ Q. \ y \ z" - using Q prefixed_Imagep_imp_empty[of Q] QA by (auto simp: Imagep_def) - -interpretation Restrp: well_founded UNIV "(\)\A" - rewrites "\x. x \ UNIV \ True" - and "(\)\A\UNIV = (\)\A" - and "\P1. (True \ PROP P1) \ PROP P1" - and "\P1. (True \ P1) \ Trueprop P1" - and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" -proof - - have "(\x. (\y. ((\) \ A) y x \ P y) \ P x) \ P a" for a P - using induct[of a P] by (auto simp: Restrp_def) - then show "well_founded UNIV ((\)\A)" apply unfold_locales by auto -qed auto - -lemmas Restrp_well_founded = Restrp.well_founded_axioms -lemmas Restrp_induct[consumes 0, case_names less] = Restrp.induct - -interpretation Restrp.tranclp: well_founded UNIV "((\)\A)\<^sup>+\<^sup>+" - rewrites "\x. x \ UNIV \ True" - and "((\)\A)\<^sup>+\<^sup>+ \ UNIV = ((\)\A)\<^sup>+\<^sup>+" - and "(((\)\A)\<^sup>+\<^sup>+)\<^sup>+\<^sup>+ = ((\)\A)\<^sup>+\<^sup>+" - and "\P1. (True \ PROP P1) \ PROP P1" - and "\P1. (True \ P1) \ Trueprop P1" - and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" -proof- - { fix P x - assume induct_step: "\x. (\y. ((\)\A)\<^sup>+\<^sup>+ y x \ P y) \ P x" - have "P x" - proof (rule induct_step) - show "\y. ((\)\A)\<^sup>+\<^sup>+ y x \ P y" - proof (induct x rule: Restrp_induct) - case (less x) - from \((\)\A)\<^sup>+\<^sup>+ y x\ - show ?case - proof (cases rule: tranclp.cases) - case r_into_trancl - with induct_step less show ?thesis by auto - next - case (trancl_into_trancl b) - with less show ?thesis by auto - qed - qed - qed - } - then show "well_founded UNIV ((\)\A)\<^sup>+\<^sup>+" by unfold_locales auto -qed auto - -lemmas Restrp_tranclp_well_founded = Restrp.tranclp.well_founded_axioms -lemmas Restrp_tranclp_induct[consumes 0, case_names less] = Restrp.tranclp.induct - -end - -context - fixes A :: "'a set" and less :: "'a \ 'a \ bool" (infix "\" 50) -begin - -lemma well_foundedI_pf: - assumes pre: "\X. X \ A \ X \ ((\) ``` X) \ A \ X = {}" - shows "well_founded A (\)" -proof - fix P a assume aA: "a \ A" and Ind: "\x. x \ A \ (\y. y \ A \ y \ x \ P y) \ P x" - from Ind have "{a\A. \P a} \ ((\) ``` {a\A. \P a}) \ A" by (auto simp: Imagep_def) - from pre[OF _ this] aA - show "P a" by auto -qed - -lemma well_foundedI_extremal: - assumes a: "\X. X \ A \ X \ {} \ \x \ X. \y \ X. \ y \ x" - shows "well_founded A (\)" -proof (rule well_foundedI_pf) - fix X assume XA: "X \ A" and pf: "X \ ((\) ``` X) \ A" - from a[OF XA] pf show "X = {}" by (auto simp: Imagep_def) -qed - -lemma well_founded_iff_ex_extremal: - "well_founded A (\) \ (\X \ A. X \ {} \ (\x \ X. \z \ X. \ z \ x))" - using well_founded.nonempty_imp_ex_extremal well_foundedI_extremal by blast - -end - -lemma well_founded_cong: - assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" - and A: "\a b. r' a b \ a \ A \ a \ A'" - and B: "\a b. r' a b \ b \ A \ b \ A'" - shows "well_founded A r \ well_founded A' r'" -proof (intro iffI) - assume wf: "well_founded A r" - show "well_founded A' r'" - proof (intro well_foundedI_extremal) - fix X - assume X: "X \ A'" and X0: "X \ {}" - show "\x\X. \y\X. \ r' y x" - proof (cases "X \ A = {}") - case True - from X0 obtain x where xX: "x \ X" by auto - with True have "x \ A" by auto - with xX X have "\y\X. \ r' y x" by (auto simp: B) - with xX show ?thesis by auto - next - case False - from well_founded.nonempty_imp_ex_extremal[OF wf _ this] - obtain x where x: "x \ X \ A" and Ar: "\y. y \ X \ y \ A \ \ r y x" by auto - have "\y \ X. \ r' y x" - proof (intro ballI notI) - fix y assume yX: "y \ X" and yx: "r' y x" - from yX X have yA': "y \ A'" by auto - show False - proof (cases "y \ A") - case True with x Ar[OF yX] yx r show ?thesis by auto - next - case False with yA' x A[OF yx] r X show ?thesis by (auto simp:) - qed - qed - with x show "\x \ X. \y \ X. \ r' y x" by auto - qed - qed -next - assume wf: "well_founded A' r'" - show "well_founded A r" - proof (intro well_foundedI_extremal) - fix X - assume X: "X \ A" and X0: "X \ {}" - show "\x\X. \y\X. \ r y x" - proof (cases "X \ A' = {}") - case True - from X0 obtain x where xX: "x \ X" by auto - with True have "x \ A'" by auto - with xX X B have "\y\X. \ r y x" by (auto simp: r in_mono) - with xX show ?thesis by auto - next - case False - from well_founded.nonempty_imp_ex_extremal[OF wf _ this] - obtain x where x: "x \ X \ A'" and Ar: "\y. y \ X \ y \ A' \ \ r' y x" by auto - have "\y \ X. \ r y x" - proof (intro ballI notI) - fix y assume yX: "y \ X" and yx: "r y x" - from yX X have y: "y \ A" by auto - show False - proof (cases "y \ A'") - case True with x Ar[OF yX] yx r X y show ?thesis by auto - next - case False with y x A yx r X show ?thesis by auto - qed - qed - with x show "\x \ X. \y \ X. \ r y x" by auto - qed - qed -qed - -lemma wfP_iff_well_founded_UNIV: "wfP r \ well_founded UNIV r" - by (auto simp: wfP_def wf_def well_founded_def) - -lemma well_founded_singleton: - assumes "\r x x" shows "well_founded {x} r" - using assms by (auto simp: well_founded_iff_ex_extremal) - -lemma well_founded_Restrp[simp]: "well_founded A (r\B) \ well_founded (A\B) r" (is "?l \ ?r") -proof (intro iffI well_foundedI_extremal) - assume l: ?l - fix X assume XAB: "X \ A \ B" and X0: "X \ {}" - with l[THEN well_founded.nonempty_imp_ex_extremal] - have "\x\X. \z\X. \ (r \ B) z x" by auto - with XAB show "\x\X. \y\X. \ r y x" by (auto simp: Restrp_def) -next - assume r: ?r - fix X assume XA: "X \ A" and X0: "X \ {}" - show "\x\X. \y\X. \ (r \ B) y x" - proof (cases "X \ B") - case True - with r[THEN well_founded.nonempty_imp_ex_extremal, of X] XA X0 - have "\z\X. \y\X. \ r y z" by auto - then show ?thesis by auto - next - case False - then obtain x where x: "x \ X - B" by auto - then have "\y\X. \ (r \ B) y x" by auto - with x show ?thesis by auto - qed -qed - -lemma (in well_founded) well_founded_subset: - assumes "B \ A" shows "well_founded B (\)" - using assms well_founded_axioms by (auto simp: well_founded_iff_ex_extremal) - -lemma well_founded_extend: - fixes less (infix "\" 50) - assumes A: "well_founded A (\)" - assumes B: "well_founded B (\)" - assumes AB: "\a \ A. \b \ B. \b \ a" - shows "well_founded (A \ B) (\)" -proof (intro well_foundedI_extremal) - interpret A: well_founded A "(\)" using A. - interpret B: well_founded B "(\)" using B. - fix X assume XAB: "X \ A \ B" and X0: "X \ {}" - show "\x\X. \y\X. \ y \ x" - proof (cases "X \ A = {}") - case True - with XAB have XB: "X \ B" by auto - from B.nonempty_imp_ex_extremal[OF XB X0] show ?thesis. - next - case False - with A.nonempty_imp_ex_extremal[OF _ this] - obtain e where XAe: "e \ X \ A" "\y\X \ A. \ y \ e" by auto - then have eX: "e \ X" and eA: "e \ A" by auto - { fix x assume xX: "x \ X" - have "\x \ e" - proof (cases "x \ A") - case True with XAe xX show ?thesis by auto - next - case False - with xX XAB have "x \ B" by auto - with AB eA show ?thesis by auto - qed - } - with eX show ?thesis by auto - qed -qed - -lemma closed_UN_well_founded: - fixes r (infix "\" 50) - assumes XX: "\X\XX. well_founded X (\) \ (\x\X. \y\\XX. y \ x \ y \ X)" - shows "well_founded (\XX) (\)" -proof (intro well_foundedI_extremal) - have *: "X \ XX \ x\X \ y \ \XX \ y \ x \ y \ X" for X x y using XX by blast - fix S - assume S: "S \ \XX" and S0: "S \ {}" - from S0 obtain x where xS: "x \ S" by auto - with S obtain X where X: "X \ XX" and xX: "x \ X" by auto - from xS xX have Sx0: "S \ X \ {}" by auto - from X XX interpret well_founded X "(\)" by auto - from nonempty_imp_ex_extremal[OF _ Sx0] - obtain z where zS: "z \ S" and zX: "z \ X" and min: "\y \ S \ X. \ y \ z" by auto - show "\x\S. \y\S. \ y \ x" - proof (intro bexI[OF _ zS] ballI notI) - fix y - assume yS: "y \ S" and yz: "y \ z" - have yXX: "y \ \ XX" using S yS by auto - from *[OF X zX yXX yz] yS have "y \ X \ S" by auto - with min yz show False by auto - qed -qed - -lemma well_founded_cmono: - assumes r': "r' \ r" and wf: "well_founded A r" - shows "well_founded A r'" -proof (intro well_foundedI_extremal) - fix X assume "X \ A" and "X \ {}" - from well_founded.nonempty_imp_ex_extremal[OF wf this] - show "\x\X. \y\X. \ r' y x" using r' by auto -qed - -locale well_founded_ordered_set = well_founded + transitive _ "(\)" -begin - -sublocale strict_ordered_set.. - -interpretation Restrp: strict_ordered_set UNIV "(\)\A" + Restrp: well_founded UNIV "(\)\A" - using Restrp_strict_order Restrp_well_founded . - -lemma Restrp_well_founded_order: "well_founded_ordered_set UNIV ((\)\A)".. - -lemma well_founded_ordered_subset: "B \ A \ well_founded_ordered_set B (\)" - apply intro_locales - using well_founded_subset transitive_subset by auto - -end - -lemma (in well_founded) Restrp_tranclp_well_founded_ordered: "well_founded_ordered_set UNIV ((\)\A)\<^sup>+\<^sup>+" - using Restrp_tranclp_well_founded tranclp_transitive by intro_locales - -locale well_related_set = related_set + - assumes nonempty_imp_ex_extreme: "X \ A \ X \ {} \ \e. extreme X (\)\<^sup>- e" -begin - -sublocale connex -proof - fix x y assume "x \ A" and "y \ A" - with nonempty_imp_ex_extreme[of "{x,y}"] show "x \ y \ y \ x" by auto -qed - -lemmas connex_axioms = connex_axioms - -interpretation less_eq_notations. - -sublocale asym: well_founded A "(\)" -proof (unfold well_founded_iff_ex_extremal, intro allI impI) - fix X - assume XA: "X \ A" and X0: "X \ {}" - from nonempty_imp_ex_extreme[OF XA X0] obtain e where "extreme X (\)\<^sup>- e" by auto - then show "\x\X. \z\X. \z \ x" by (auto intro!: bexI[of _ e]) -qed - -lemma well_related_subset: "B \ A \ well_related_set B (\)" - by (auto intro!: well_related_set.intro nonempty_imp_ex_extreme) - -end - -context - fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) -begin - -lemma well_related_iff_neg_well_founded: - "well_related_set A (\) \ well_founded A (\x y. \ y \ x)" - by (simp add: well_related_set_def well_founded_iff_ex_extremal extreme_def Bex_def) - -lemma well_related_singleton_refl: - assumes "x \ x" shows "well_related_set {x} (\)" - by (intro well_related_set.intro exI[of _ x], auto simp: subset_singleton_iff assms) - -lemma closed_UN_well_related: - assumes XX: "\X\XX. well_related_set X (\) \ (\x\X. \y\\XX. \x \ y \ y \ X)" - shows "well_related_set (\XX) (\)" - using XX - apply (unfold well_related_iff_neg_well_founded) - using closed_UN_well_founded[of _ "\x y. \ y \ x"]. - -end - -lemma well_related_extend: - fixes r (infix "\" 50) - assumes "well_related_set A (\)" and "well_related_set B (\)" and "\a \ A. \b \ B. a \ b" - shows "well_related_set (A \ B) (\)" - using well_founded_extend[of _ "\x y. \ y \ x", folded well_related_iff_neg_well_founded] - using assms by auto - -locale pre_well_ordered_set = semiattractive + well_related_set -begin - -interpretation less_eq_notations. - -sublocale transitive -proof - fix x y z assume xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" - from x y z have "\e. extreme {x,y,z} (\) e" (is "\e. ?P e") by (auto intro!: nonempty_imp_ex_extreme) - then have "?P x \ ?P y \ ?P z" by auto - then show "x \ z" - proof (elim disjE) - assume "?P x" - then show ?thesis by auto - next - assume "?P y" - then have "y \ x" by auto - from attract[OF xy this yz] x y z show ?thesis by auto - next - assume "?P z" - then have zx: "z \ x" and zy: "z \ y" by auto - from attract[OF yz zy zx] x y z have yx: "y \ x" by auto - from attract[OF xy yx yz] x y z show ?thesis by auto - qed -qed - -sublocale total_quasi_ordered_set.. - -lemmas connex_axioms = connex_axioms - -lemma strict_weak_trans[trans]: - assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" - shows "x \ z" -proof (intro asympartpI notI) - from trans xy yz x y z show "x \ z" by auto - assume "z \ x" - from trans[OF yz this] y z x have "y \ x" by auto - with xy show False by auto -qed - -lemma weak_strict_trans[trans]: - assumes xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" - shows "x \ z" -proof (intro asympartpI notI) - from trans xy yz x y z show "x \ z" by auto - assume "z \ x" - from trans[OF this xy] z x y have "z \ y" by auto - with yz show False by auto -qed - -end - -lemma pre_well_ordered_iff: - "pre_well_ordered_set A r \ total_quasi_ordered_set A r \ well_founded A (asympartp r)" - (is "?p \ ?t \ ?w") -proof safe - assume ?p - then interpret pre_well_ordered_set A r. - show ?t ?w by unfold_locales -next - assume ?t - then interpret total_quasi_ordered_set A r. - assume ?w - then have "well_founded UNIV (asympartp r \ A)" by simp - also have "asympartp r \ A = (\x y. \ r y x) \ A" by (intro ext, auto simp: not_iff_asym) - finally have "well_related_set A r" by (simp add: well_related_iff_neg_well_founded) - then show ?p by intro_locales -qed - -lemma (in semiattractive) pre_well_ordered_iff_well_related: - assumes XA: "X \ A" - shows "pre_well_ordered_set X (\) \ well_related_set X (\)" (is "?l \ ?r") -proof - interpret X: semiattractive X using semiattractive_subset[OF XA]. - { assume ?l - then interpret X: pre_well_ordered_set X. - show ?r by unfold_locales - } - assume ?r - then interpret X: well_related_set X. - show ?l by unfold_locales -qed - -lemma semiattractive_extend: - fixes r (infix "\" 50) - assumes A: "semiattractive A (\)" and B: "semiattractive B (\)" - and AB: "\a \ A. \b \ B. a \ b \ \ b \ a" - shows "semiattractive (A \ B) (\)" -proof- - interpret A: semiattractive A "(\)" using A. - interpret B: semiattractive B "(\)" using B. - { - fix x y z - assume yB: "y \ B" and zA: "z \ A" and yz: "y \ z" - have False using AB[rule_format, OF zA yB] yz by auto - } - note * = this - show ?thesis - by (auto intro!: semiattractive.intro dest:* AB[rule_format] A.attract B.attract) -qed - -lemma pre_well_order_extend: - fixes r (infix "\" 50) - assumes A: "pre_well_ordered_set A (\)" and B: "pre_well_ordered_set B (\)" - and AB: "\a \ A. \b \ B. a \ b \ \ b \ a" - shows "pre_well_ordered_set (A\B) (\)" -proof- - interpret A: pre_well_ordered_set A "(\)" using A. - interpret B: pre_well_ordered_set B "(\)" using B. - show ?thesis - apply (intro pre_well_ordered_set.intro well_related_extend semiattractive_extend) - apply unfold_locales - by (auto dest: AB[rule_format]) -qed - -locale well_ordered_set = antisymmetric + well_related_set -begin - -sublocale pre_well_ordered_set.. - -sublocale total_ordered_set.. - -lemma well_ordered_subset: "B \ A \ well_ordered_set B (\)" - using well_related_subset antisymmetric_subset by (intro well_ordered_set.intro) - -lemmas connex_axioms = connex_axioms - -end - -lemma (in antisymmetric) well_ordered_iff_well_related: - assumes XA: "X \ A" - shows "well_ordered_set X (\) \ well_related_set X (\)" (is "?l \ ?r") -proof - interpret X: antisymmetric X using antisymmetric_subset[OF XA]. - { assume ?l - then interpret X: well_ordered_set X. - show ?r by unfold_locales - } - assume ?r - then interpret X: well_related_set X. - show ?l by unfold_locales -qed - -context - fixes A and less_eq (infix "\" 50) - assumes A: "\a \ A. \b \ A. a \ b" -begin - -interpretation well_related_set A "(\)" - apply unfold_locales - using A by blast - -lemmas trivial_well_related = well_related_set_axioms - -lemma trivial_pre_well_order: "pre_well_ordered_set A (\)" - apply unfold_locales - using A by blast - -end - -context - fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) -begin - -interpretation less_eq_notations. - -lemma well_order_extend: - assumes A: "well_ordered_set A (\)" and B: "well_ordered_set B (\)" - and ABa: "\a \ A. \b \ B. a \ b \ b \ a \ a = b" - and AB: "\a \ A. \b \ B. a \ b" - shows "well_ordered_set (A\B) (\)" -proof- - interpret A: well_ordered_set A "(\)" using A. - interpret B: well_ordered_set B "(\)" using B. - show ?thesis - apply (intro well_ordered_set.intro antisymmetric_union well_related_extend ABa AB) - by unfold_locales -qed - -interpretation singleton: antisymmetric "{a}" "(\)" for a apply unfold_locales by auto - -lemmas singleton_antisymmetric[intro!] = singleton.antisymmetric_axioms - -lemma singleton_well_ordered[intro!]: "a \ a \ well_ordered_set {a} (\)" - apply unfold_locales by auto - -lemma closed_UN_well_ordered: - assumes anti: "antisymmetric (\ XX) (\)" - and XX: "\X\XX. well_ordered_set X (\) \ (\x\X. \y\\XX. \ x \ y \ y \ X)" - shows "well_ordered_set (\XX) (\)" - apply (intro well_ordered_set.intro closed_UN_well_related anti) - using XX well_ordered_set.axioms by fast - -end - -text \Directed sets:\ - -definition "directed A r \ \x \ A. \y \ A. \z \ A. r x z \ r y z" - -lemmas directedI[intro] = directed_def[unfolded atomize_eq, THEN iffD2, rule_format] - -lemmas directedD = directed_def[unfolded atomize_eq, THEN iffD1, rule_format] - -context - fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) -begin - -lemma directedE: - assumes "directed A (\)" and "x \ A" and "y \ A" - and "\z. z \ A \ x \ z \ y \ z \ thesis" - shows "thesis" - using assms by (auto dest: directedD) - -lemma directed_empty[simp]: "directed {} (\)" by auto - -lemma directed_union: - assumes dX: "directed X (\)" and dY: "directed Y (\)" - and XY: "\x\X. \y\Y. \z \ X \ Y. x \ z \ y \ z" - shows "directed (X \ Y) (\)" - using directedD[OF dX] directedD[OF dY] XY - apply (intro directedI) by blast - -lemma directed_extend: - assumes X: "directed X (\)" and Y: "directed Y (\)" and XY: "\x\X. \y\Y. x \ y" - shows "directed (X \ Y) (\)" -proof - - { fix x y - assume xX: "x \ X" and yY: "y \ Y" - let ?g = "\z\X \ Y. x \ z \ y \ z" - from directedD[OF Y yY yY] obtain z where zY: "z \ Y" and yz: "y \ z" by auto - from xX XY zY yz have ?g by auto - } - then show ?thesis by (auto intro!: directed_union[OF X Y]) -qed - -end - -context connex begin - -lemma directed: "directed A (\)" -proof - fix x y - assume x: "x \ A" and y: "y \ A" - then show "\z\A. x \ z \ y \ z" - proof (cases rule: comparable_cases) - case le - with refl[OF y] y show ?thesis by (intro bexI[of _ y], auto) - next - case ge - with refl[OF x] x show ?thesis by (intro bexI[of _ x], auto) - qed -qed - -end - -context - fixes ir :: "'i \ 'i \ bool" (infix "\" 50) - fixes r :: "'a \ 'a \ bool" (infix "\" 50) -begin lemma monotone_connex_image: - assumes mono: "monotone_on I (\) (\) f" - assumes connex: "connex I (\)" + fixes ir (infix "\" 50) and r (infix "\" 50) + assumes mono: "monotone_on I (\) (\) f" and connex: "connex I (\)" shows "connex (f ` I) (\)" proof (rule connexI) fix x y assume "x \ f ` I" and "y \ f ` I" then obtain i j where ij: "i \ I" "j \ I" and [simp]: "x = f i" "y = f j" by auto from connex ij have "i \ j \ j \ i" by (auto elim: connexE) with ij mono show "x \ y \ y \ x" by (elim disjE, auto dest: monotone_onD) qed -lemma monotone_directed_image: - assumes mono: "monotone_on I (\) (\) f" - assumes dir: "directed I (\)" shows "directed (f ` I) (\)" -proof (rule directedI, safe) - fix x y assume x: "x \ I" and y: "y \ I" - with dir obtain z where z: "z \ I" and "x \ z" and "y \ z" by (auto elim: directedE) - with mono x y have "f x \ f z" and "f y \ f z" by (auto dest: monotone_onD) - with z show "\fz \ f ` I. f x \ fz \ f y \ fz" by auto -qed - -end - subsection \Order Pairs\ -locale compatible = related_set + related_set A "(\)" for less (infix "\" 50) + - assumes compat_right[trans]: "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" - assumes compat_left[trans]: "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" - -locale compatible_ordering = reflexive + irreflexive + compatible + +text \We pair a relation (weak part) with a well-behaving ``strict'' part. Here no assumption is + put on the ``weak'' part.\ + +locale compatible_ordering = + related_set + irreflexive + assumes strict_implies_weak: "x \ y \ x \ A \ y \ A \ x \ y" + assumes weak_strict_trans[trans]: "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" + assumes strict_weak_trans[trans]: "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" begin -text \The strict part is necessarily transitive.\ - text \The following sequence of declarations are in order to obtain fact names in a manner similar to the Isabelle/HOL facts of orders.\ +text \The strict part is necessarily transitive.\ + sublocale strict: transitive A "(\)" - using compat_right[OF strict_implies_weak] by unfold_locales + using weak_strict_trans[OF strict_implies_weak] by unfold_locales sublocale strict_ordered_set A "(\)" .. thm strict.trans asym irrefl +lemma Restrp_compatible_ordering: "compatible_ordering UNIV ((\)\A) ((\)\A)" + apply (unfold_locales) + by (auto dest: weak_strict_trans strict_weak_trans strict_implies_weak) + lemma strict_implies_not_weak: "x \ y \ x \ A \ y \ A \ \ y \ x" - using irrefl compat_left by blast + using irrefl weak_strict_trans by blast + +lemma weak_implies_not_strict: + assumes xy: "x \ y" and [simp]: "x \ A" "y \ A" + shows "\y \ x" +proof + assume "y \ x" + also note xy + finally show False using irrefl by auto +qed + +lemma compatible_ordering_subset: assumes "X \ A" shows "compatible_ordering X (\) (\)" + apply unfold_locales + using assms strict_implies_weak by (auto intro: strict_weak_trans weak_strict_trans) end context transitive begin -interpretation less_eq_notations. +interpretation less_eq_asymmetrize. lemma asym_trans[trans]: shows "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" and "x \ y \ y \ z \ x \ A \ y \ A \ z \ A \ x \ z" by (auto 0 3 dest: trans) -end - -locale attractive_ordering = compatible_ordering + attractive - -locale pseudo_ordering = compatible_ordering + pseudo_ordered_set -begin - -sublocale attractive_ordering .. - -end - -locale quasi_ordering = compatible_ordering + quasi_ordered_set -begin - -sublocale attractive_ordering .. +lemma asympartp_compatible_ordering: "compatible_ordering A (\) (\)" + apply unfold_locales + by (auto dest: asym_trans) end -locale partial_ordering = compatible_ordering + partially_ordered_set +locale reflexive_ordering = reflexive + compatible_ordering + +locale reflexive_attractive_ordering = reflexive_ordering + attractive + +locale pseudo_ordering = pseudo_ordered_set + compatible_ordering begin -sublocale pseudo_ordering + quasi_ordering .. +sublocale reflexive_attractive_ordering.. end -locale well_founded_ordering = quasi_ordering + well_founded - -locale total_ordering = compatible_ordering + total_ordered_set +locale quasi_ordering = quasi_ordered_set + compatible_ordering begin -sublocale partial_ordering .. +sublocale reflexive_attractive_ordering.. + +lemma quasi_ordering_subset: assumes "X \ A" shows "quasi_ordering X (\) (\)" + by (intro quasi_ordering.intro quasi_ordered_subset compatible_ordering_subset assms) end -locale strict_total_ordering = partial_ordering + semiconnex A "(\)" +context quasi_ordered_set begin + +interpretation less_eq_asymmetrize. + +lemma asympartp_quasi_ordering: "quasi_ordering A (\) (\)" + by (intro quasi_ordering.intro quasi_ordered_set_axioms asympartp_compatible_ordering) + +end + +locale partial_ordering = partially_ordered_set + compatible_ordering +begin + +sublocale quasi_ordering + pseudo_ordering.. + +lemma partial_ordering_subset: assumes "X \ A" shows "partial_ordering X (\) (\)" + by (intro partial_ordering.intro partially_ordered_subset compatible_ordering_subset assms) + +end + +context partially_ordered_set begin + +interpretation less_eq_asymmetrize. + +lemma asympartp_partial_ordering: "partial_ordering A (\) (\)" + by (intro partial_ordering.intro partially_ordered_set_axioms asympartp_compatible_ordering) + +end + +locale total_quasi_ordering = total_quasi_ordered_set + compatible_ordering +begin + +sublocale quasi_ordering.. + +lemma total_quasi_ordering_subset: assumes "X \ A" shows "total_quasi_ordering X (\) (\)" + by (intro total_quasi_ordering.intro total_quasi_ordered_subset compatible_ordering_subset assms) + +end + +context total_quasi_ordered_set begin + +interpretation less_eq_asymmetrize. + +lemma asympartp_total_quasi_ordering: "total_quasi_ordering A (\) (\)" + by (intro total_quasi_ordering.intro total_quasi_ordered_set_axioms asympartp_compatible_ordering) + +end + + +text \Fixing the definition of the strict part is very common, though it looks restrictive +to the author.\ +locale strict_quasi_ordering = quasi_ordered_set + less_syntax + + assumes strict_iff: "x \ A \ y \ A \ x \ y \ x \ y \ \y \ x" +begin + +sublocale compatible_ordering +proof unfold_locales + fix x y z + show "x \ A \ \ x \ x" by (auto simp: strict_iff) + { assume xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" + from yz y z have ywz: "y \ z" and zy: "\z \ y" by (auto simp: strict_iff) + from trans[OF xy ywz]x y z have xz: "x \ z" by auto + from trans[OF _ xy] x y z zy have zx: "\z \ x" by auto + from xz zx x z show "x \ z" by (auto simp: strict_iff) + } + { assume xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" + from xy x y have xwy: "x \ y" and yx: "\y \ x" by (auto simp: strict_iff) + from trans[OF xwy yz]x y z have xz: "x \ z" by auto + from trans[OF yz] x y z yx have zx: "\z \ x" by auto + from xz zx x z show "x \ z" by (auto simp: strict_iff) + } + { show "x \ y \ x \ A \ y \ A \ x \ y" by (auto simp: strict_iff) } +qed + +end + +locale strict_partial_ordering = strict_quasi_ordering + antisymmetric +begin + +sublocale partial_ordering.. + +lemma strict_iff_neq: "x \ A \ y \ A \ x \ y \ x \ y \ x \ y" + by (auto simp: strict_iff antisym) + +end + +locale total_ordering = reflexive + compatible_ordering + semiconnex A "(\)" begin sublocale semiconnex_irreflexive .. sublocale connex proof fix x y assume x: "x \ A" and y: "y \ A" then show "x \ y \ y \ x" - apply (cases rule: cases[OF x y]) - by (auto dest: strict_implies_weak) + by (cases rule: cases, auto dest: strict_implies_weak) qed -sublocale total_ordering .. -(* -sublocale old: ordering "(\)" "(\)" -proof- - have "a \ b \ a \ b \ a \ b" for a b - by (cases a b rule: cases, auto dest: strict_implies_weak) - then show "ordering (\) (\)" - by (unfold_locales, auto dest:strict_implies_weak trans) -qed -*) - -lemma not_weak[simp]: +lemma not_weak: assumes "x \ A" and "y \ A" shows "\ x \ y \ y \ x" using assms by (cases rule:cases, auto simp: strict_implies_not_weak dest: strict_implies_weak) -lemma not_strict[simp]: "x \ A \ y \ A \ \ x \ y \ y \ x" +lemma not_strict: "x \ A \ y \ A \ \ x \ y \ y \ x" using not_weak by blast +sublocale strict_partial_ordering +proof + fix a b + assume a: "a \ A" and b: "b \ A" + then show "a \ b \ a \ b \ \ b \ a" by (auto simp: not_strict[symmetric] dest: asym) +next + fix x y z assume xy: "x \ y" and yz: "y \ z" and xA: "x \ A" and yA: "y \ A" and zA: "z \ A" + with weak_strict_trans[OF yz] show "x \ z" by (auto simp: not_strict[symmetric]) +next + fix x y assume xy: "x \ y" and yx: "y \ x" and xA: "x \ A" and yA: "y \ A" + with semiconnex show "x = y" by (auto dest: weak_implies_not_strict) +qed + +sublocale total_ordered_set.. + +context + fixes s + assumes s: "\x \ A. x \ s \ (\z \ A. x \ z \ z \ s)" and sA: "s \ A" +begin + +lemma dense_weakI: + assumes bound: "\x. x \ s \ x \ A \ x \ y" and yA: "y \ A" + shows "s \ y" +proof (rule ccontr) + assume "\ ?thesis" + with yA sA have "y \ s" by (simp add: not_weak) + from s[rule_format, OF yA this] + obtain x where xA: "x \ A" and xs: "x \ s" and yx: "y \ x" by safe + have xy: "x \ y" using bound[OF xs xA] . + from yx xy xA yA + show False by (simp add: weak_implies_not_strict) +qed + +lemma dense_bound_iff: + assumes bA: "b \ A" shows "bound {x\A. x \ s} (\) b \ s \ b" + using assms sA + by (auto simp: bound_def intro: strict_implies_weak strict_weak_trans dense_weakI) + +lemma dense_extreme_bound: + "extreme_bound A (\) {x \ A. x \ s} s" + by (auto intro!: extreme_boundI intro: strict_implies_weak simp: dense_bound_iff sA) + end +lemma ordinal_cases[consumes 1, case_names suc lim]: + assumes aA: "a \ A" + and suc: "\p. extreme {x \ A. x \ a} (\) p \ thesis" + and lim: "extreme_bound A (\) {x \ A. x \ a} a \ thesis" + shows "thesis" +proof (cases "\p. extreme {x \ A. x \ a} (\) p") + case True + with suc show ?thesis by auto +next + case False + show ?thesis + proof (rule lim, rule dense_extreme_bound, safe intro!: aA) + fix x assume xA: "x \ A" and xa: "x \ a" + show "\z\A. x \ z \ z \ a" + proof (rule ccontr) + assume "\?thesis" + with xA xa have "extreme {x \ A. x \ a} (\) x" by (auto simp: not_strict) + with False show False by auto + qed + qed +qed + +end + +context total_ordered_set begin + +interpretation less_eq_asymmetrize. + +lemma asympartp_total_ordering: "total_ordering A (\) (\)" + by (intro total_ordering.intro reflexive_axioms asympartp_compatible_ordering asympartp_semiconnex) + +end + +subsection \Functions\ + +definition "pointwise I r f g \ \i \ I. r (f i) (g i)" + +lemmas pointwiseI = pointwise_def[unfolded atomize_eq, THEN iffD2, rule_format] + +lemmas pointwiseD[simp] = pointwise_def[unfolded atomize_eq, THEN iffD1, rule_format] + +lemma pointwise_cong: + assumes "r = r'" "\i. i \ I \ f i = f' i" "\i. i \ I \ g i = g' i" + shows "pointwise I r f g = pointwise I r' f' g'" + using assms by (auto simp: pointwise_def) + +lemma pointwise_empty[simp]: "pointwise {} = \" by (auto intro!: ext pointwiseI) + +lemma dual_pointwise[simp]: "(pointwise I r)\<^sup>- = pointwise I r\<^sup>-" + by (auto intro!: ext pointwiseI dest: pointwiseD) + +lemma pointwise_dual: "pointwise I r\<^sup>- f g \ pointwise I r g f" by (auto simp: pointwise_def) + +lemma pointwise_un: "pointwise (I\J) r = pointwise I r \ pointwise J r" + by (auto intro!: ext pointwiseI) + +lemma pointwise_unI[intro!]: "pointwise I r f g \ pointwise J r f g \ pointwise (I \ J) r f g" + by (auto simp: pointwise_un) + +lemma pointwise_bound: "bound F (pointwise I r) f \ (\i \ I. bound {f i |. f \ F} r (f i))" + by (auto intro!:pointwiseI elim!: boundE) + +lemma pointwise_extreme: + shows "extreme F (pointwise X r) e \ e \ F \ (\x \ X. extreme {f x |. f \ F} r (e x))" + by (auto intro!: pointwiseI extremeI elim!: extremeE) + +lemma pointwise_extreme_bound: + fixes r (infix "\" 50) + assumes F: "F \ {f. f ` X \ A}" + shows "extreme_bound {f. f ` X \ A} (pointwise X (\)) F s \ + (\x \ X. extreme_bound A (\) {f x |. f \ F} (s x))" (is "?p \ ?a") +proof (safe intro!: extreme_boundI pointwiseI) + fix x + assume s: ?p and xX: "x \ X" + { fix b + assume b: "bound {f x |. f \ F} (\) b" and bA: "b \ A" + have "pointwise X (\) s (s(x:=b))" + proof (rule extreme_boundD(2)[OF s], safe intro!: pointwiseI) + fix f y + assume fF: "f \ F" and yX: "y \ X" + show "f y \ (s(x:=b)) y" + proof (cases "x = y") + case True + with b fF show "?thesis" by auto + next + case False + with s[THEN extreme_bound_imp_bound] fF yX show ?thesis by (auto dest: boundD) + qed + next + fix y assume "y \ X" with bA s show "(s(x := b)) y \ A" by auto + qed + with xX show "s x \ b" by (auto dest: pointwiseD) + next + fix f assume "f \ F" + from extreme_boundD(1)[OF s this] F xX + show "f x \ s x" by auto + next + show "s x \ A" using s xX by auto + } +next + fix x + assume s: ?a and xX: "x \ X" + { from s xX show "s x \ A" by auto + next + fix b assume b: "bound F (pointwise X (\)) b" and bA: "b ` X \ A" + with xX have "bound {f x |. f \ F} (\) (b x)" by (auto simp: pointwise_bound) + with s[rule_format, OF xX] bA xX show "s x \ b x" by auto + next + fix f assume "f \ F" + with s[rule_format, OF xX] show "f x \ s x" by auto + } +qed + +lemma dual_pointwise_extreme_bound: + "extreme_bound FA (pointwise X r)\<^sup>- F = extreme_bound FA (pointwise X r\<^sup>-) F" + by (simp) + +lemma pointwise_monotone_on: + fixes less_eq (infix "\" 50) and prec_eq (infix "\" 50) + shows "monotone_on I (\) (pointwise A (\)) f \ + (\a \ A. monotone_on I (\) (\) (\i. f i a))" (is "?l \ ?r") +proof (safe intro!: monotone_onI pointwiseI) + fix a i j assume aA: "a \ A" and *: ?l "i \ j" "i \ I" "j \ I" + then + show "f i a \ f j a" by (auto dest: monotone_onD) +next + fix a i j assume ?r and "a \ A" and ij: "i \ j" "i \ I" "j \ I" + then have "monotone_on I (\) (\) (\i. f i a)" by auto + from monotone_onD[OF this]ij + show "f i a \ f j a" by auto +qed + +lemmas pointwise_monotone = pointwise_monotone_on[of UNIV] + +lemma (in reflexive) pointwise_reflexive: "reflexive {f. f ` I \ A} (pointwise I (\))" + apply unfold_locales by (auto intro!: pointwiseI simp: subsetD[OF _ imageI]) + +lemma (in irreflexive) pointwise_irreflexive: + assumes I0: "I \ {}" shows "irreflexive {f. f ` I \ A} (pointwise I (\))" +proof (safe intro!: irreflexive.intro) + fix f + assume f: "f ` I \ A" and ff: "pointwise I (\) f f" + from I0 obtain i where i: "i \ I" by auto + with ff have "f i \ f i" by auto + with f i show False by auto +qed + +lemma (in semiattractive) pointwise_semiattractive: "semiattractive {f. f ` I \ A} (pointwise I (\))" +proof (unfold_locales, safe intro!: pointwiseI) + fix f g h i + assume fg: "pointwise I (\) f g" and gf: "pointwise I (\) g f" and gh: "pointwise I (\) g h" + and [simp]: "i \ I" and f: "f ` I \ A" and g: "g ` I \ A" and h: "h ` I \ A" + show "f i \ h i" + proof (rule attract) + from fg show "f i \ g i" by auto + from gf show "g i \ f i" by auto + from gh show "g i \ h i" by auto + qed (insert f g h, auto simp: subsetD[OF _ imageI]) +qed + +lemma (in attractive) pointwise_attractive: "attractive {f. f ` I \ A} (pointwise I (\))" + apply (intro attractive.intro attractive_axioms.intro) + using pointwise_semiattractive dual.pointwise_semiattractive by auto + +text \Antisymmetry will not be preserved by pointwise extension over restricted domain.\ +lemma (in antisymmetric) pointwise_antisymmetric: + "antisymmetric {f. f ` I \ A} (pointwise I (\))" + oops + +lemma (in transitive) pointwise_transitive: "transitive {f. f ` I \ A} (pointwise I (\))" +proof (unfold_locales, safe intro!: pointwiseI) + fix f g h i + assume fg: "pointwise I (\) f g" and gh: "pointwise I (\) g h" + and [simp]: "i \ I" and f: "f ` I \ A" and g: "g ` I \ A" and h: "h ` I \ A" + from fg have "f i \ g i" by auto + also from gh have "g i \ h i" by auto + finally show "f i \ h i" using f g h by (auto simp: subsetD[OF _ imageI]) +qed + +lemma (in quasi_ordered_set) pointwise_quasi_order: + "quasi_ordered_set {f. f ` I \ A} (pointwise I (\))" + by (intro quasi_ordered_setI pointwise_transitive pointwise_reflexive) + +lemma (in compatible_ordering) pointwise_compatible_ordering: + assumes I0: "I \ {}" + shows "compatible_ordering {f. f ` I \ A} (pointwise I (\)) (pointwise I (\))" +proof (intro compatible_ordering.intro compatible_ordering_axioms.intro pointwise_irreflexive[OF I0], safe intro!: pointwiseI) + fix f g h i + assume fg: "pointwise I (\) f g" and gh: "pointwise I (\) g h" + and [simp]: "i \ I" and f: "f ` I \ A" and g: "g ` I \ A" and h: "h ` I \ A" + from fg have "f i \ g i" by auto + also from gh have "g i \ h i" by auto + finally show "f i \ h i" using f g h by (auto simp: subsetD[OF _ imageI]) +next + fix f g h i + assume fg: "pointwise I (\) f g" and gh: "pointwise I (\) g h" + and [simp]: "i \ I" and f: "f ` I \ A" and g: "g ` I \ A" and h: "h ` I \ A" + from fg have "f i \ g i" by auto + also from gh have "g i \ h i" by auto + finally show "f i \ h i" using f g h by (auto simp: subsetD[OF _ imageI]) +next + fix f g i + assume fg: "pointwise I (\) f g" + and [simp]: "i \ I" + and f: "f ` I \ A" and g: "g ` I \ A" + from fg have "f i \ g i" by auto + with f g show "f i \ g i" by (auto simp: subsetD[OF _ imageI] strict_implies_weak) +qed + subsection \Relating to Classes\ text \In Isabelle 2020, we should declare sublocales in class before declaring dual sublocales, since otherwise facts would be prefixed by ``dual.dual.''\ context ord begin abbreviation least where "least X \ extreme X (\x y. y \ x)" abbreviation greatest where "greatest X \ extreme X (\)" abbreviation supremum where "supremum X \ least (Collect (bound X (\)))" abbreviation infimum where "infimum X \ greatest (Collect (bound X (\x y. y \ x)))" +lemma supremumI: "bound X (\) s \ (\b. bound X (\) b \ s \ b) \ supremum X s" + and infimumI: "bound X (\) i \ (\b. bound X (\) b \ b \ i) \ infimum X i" + by (auto intro!: extremeI) + +lemma supremumE: "supremum X s \ + (bound X (\) s \ (\b. bound X (\) b \ s \ b) \ thesis) \ thesis" + and infimumE: "infimum X i \ + (bound X (\) i \ (\b. bound X (\) b \ b \ i) \ thesis) \ thesis" + by (auto) + +lemma extreme_bound_supremum[simp]: "extreme_bound UNIV (\) = supremum" by (auto intro!: ext) +lemma extreme_bound_infimum[simp]: "extreme_bound UNIV (\) = infimum" by (auto intro!: ext) + lemma Least_eq_The_least: "Least P = The (least {x. P x})" by (auto simp: Least_def extreme_def[unfolded atomize_eq, THEN ext]) lemma Greatest_eq_The_greatest: "Greatest P = The (greatest {x. P x})" by (auto simp: Greatest_def extreme_def[unfolded atomize_eq, THEN ext]) end lemma Ball_UNIV[simp]: "Ball UNIV = All" by auto lemma Bex_UNIV[simp]: "Bex UNIV = Ex" by auto -class compat = ord + assumes "compatible_ordering UNIV (\) (<)" +lemma pointwise_UNIV_le[simp]: "pointwise UNIV (\) = (\)" by (intro ext, simp add: pointwise_def le_fun_def) +lemma pointwise_UNIV_ge[simp]: "pointwise UNIV (\) = (\)" by (intro ext, simp add: pointwise_def le_fun_def) + +lemma fun_supremum_iff: "supremum F e \ (\x. supremum {f x |. f \ F} (e x))" + using pointwise_extreme_bound[of F UNIV UNIV "(\)"] by simp + +lemma fun_infimum_iff: "infimum F e \ (\x. infimum {f x |. f \ F} (e x))" + using pointwise_extreme_bound[of F UNIV UNIV "(\)"] by simp + + +class reflorder = ord + assumes "reflexive_ordering UNIV (\) (<)" begin -sublocale order: compatible_ordering UNIV +sublocale order: reflexive_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - using compat_axioms unfolding class.compat_def by (auto 0 4 simp:atomize_eq) + using reflorder_axioms unfolding class.reflorder_def by (auto 0 4 simp:atomize_eq) end text \We should have imported locale-based facts in classes, e.g.:\ thm order.trans order.strict.trans order.refl order.irrefl order.asym order.extreme_bound_singleton -class attractive_order = ord + assumes "attractive_ordering UNIV (\) (<)" +class attrorder = ord + + assumes "reflexive_attractive_ordering UNIV (\) (<)" begin text \We need to declare subclasses before sublocales in order to preserve facts for superclasses.\ -interpretation attractive_ordering UNIV - using attractive_order_axioms unfolding class.attractive_order_def. - -subclass compat .. - -sublocale order: attractive_ordering UNIV +subclass reflorder +proof- + interpret reflexive_attractive_ordering UNIV + using attrorder_axioms unfolding class.attrorder_def by auto + show "class.reflorder (\) (<)".. +qed + +sublocale order: reflexive_attractive_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - apply unfold_locales by (auto simp:atomize_eq) + using attrorder_axioms unfolding class.attrorder_def + by (auto simp:atomize_eq) end thm order.extreme_bound_quasi_const class psorder = ord + assumes "pseudo_ordering UNIV (\) (<)" begin -interpretation pseudo_ordering UNIV using psorder_axioms unfolding class.psorder_def. - -subclass attractive_order .. +subclass attrorder +proof- + interpret pseudo_ordering UNIV + using psorder_axioms unfolding class.psorder_def by auto + show "class.attrorder (\) (<)".. +qed sublocale order: pseudo_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - apply unfold_locales by (auto simp:atomize_eq) + using psorder_axioms unfolding class.psorder_def by (auto simp:atomize_eq) end class qorder = ord + assumes "quasi_ordering UNIV (\) (<)" begin -interpretation quasi_ordering UNIV using qorder_axioms unfolding class.qorder_def. - -subclass attractive_order .. +subclass attrorder +proof- + interpret quasi_ordering UNIV + using qorder_axioms unfolding class.qorder_def by auto + show "class.attrorder (\) (<)".. +qed sublocale order: quasi_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - apply unfold_locales by (auto simp:atomize_eq) + using qorder_axioms unfolding class.qorder_def by (auto simp:atomize_eq) + +lemmas [intro!] = order.quasi_ordered_subset end class porder = ord + assumes "partial_ordering UNIV (\) (<)" begin interpretation partial_ordering UNIV - using porder_axioms unfolding class.porder_def. - -subclass psorder .. -subclass qorder .. + using porder_axioms unfolding class.porder_def by auto + +subclass psorder.. + +subclass qorder.. sublocale order: partial_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" apply unfold_locales by (auto simp:atomize_eq) end -class wf_qorder = ord + assumes "well_founded_ordering UNIV (\) (<)" +class linqorder = ord + assumes "total_quasi_ordering UNIV (\) (<)" begin -interpretation well_founded_ordering UNIV - using wf_qorder_axioms unfolding class.wf_qorder_def. - -subclass qorder .. - -sublocale order: well_founded_ordering UNIV +interpretation total_quasi_ordering UNIV + using linqorder_axioms unfolding class.linqorder_def by auto + +subclass qorder.. + +sublocale order: total_quasi_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - apply unfold_locales by (auto simp:atomize_eq) + using linqorder_axioms unfolding class.linqorder_def + by (auto simp:atomize_eq) + +lemmas asympartp_le = order.not_iff_asym[symmetric, abs_def] end -class totalorder = ord + assumes "total_ordering UNIV (\) (<)" -begin - -interpretation total_ordering UNIV - using totalorder_axioms unfolding class.totalorder_def. - -subclass porder .. + +text \Isabelle/HOL's @{class preorder} belongs to @{class qorder}, but not vice versa.\ + +context preorder begin + +text \The relation @{term "(<)"} is defined as the antisymmetric part of @{term "(\)"}.\ +lemma [simp]: + shows asympartp_le: "asympartp (\) = (<)" + and asympartp_ge: "asympartp (\) = (>)" + by (intro ext, auto simp: asympartp_def less_le_not_le) + +interpretation strict_quasi_ordering UNIV "(\)" "(<)" + apply unfold_locales + using order_refl apply assumption + using order_trans apply assumption + using less_le_not_le apply assumption + done + +subclass qorder.. + +sublocale order: strict_quasi_ordering UNIV + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + apply unfold_locales + by (auto simp:atomize_eq) + +end + +context order begin + +interpretation strict_partial_ordering UNIV "(\)" "(<)" + apply unfold_locales + using order_antisym by assumption + +subclass porder.. + +sublocale order: strict_partial_ordering UNIV + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + apply unfold_locales + by (auto simp:atomize_eq) + +end + +text \Isabelle/HOL's @{class linorder} is equivalent to our locale @{locale total_ordering}.\ + +context linorder begin + +subclass linqorder apply unfold_locales by auto sublocale order: total_ordering UNIV rewrites "\x. x \ UNIV \ True" and "\X. X \ UNIV \ True" and "\r. r \ UNIV \ r" and "\P. True \ P \ P" and "Ball UNIV \ All" and "Bex UNIV \ Ex" and "sympartp (\)\<^sup>- \ sympartp (\)" and "\P1. (True \ PROP P1) \ PROP P1" and "\P1. (True \ P1) \ Trueprop P1" and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" apply unfold_locales by (auto simp:atomize_eq) end -text \Isabelle/HOL's @{class preorder} belongs to @{class qorder}, but not vice versa.\ - -subclass (in preorder) qorder - apply unfold_locales - using order_refl apply assumption - apply simp - using le_less_trans apply assumption - using less_le_trans apply assumption - using less_imp_le apply assumption - using order_trans apply assumption - done - -subclass (in order) porder by (unfold_locales, auto) - -subclass (in wellorder) wf_qorder - apply (unfold_locales) - using less_induct by auto - -text \Isabelle/HOL's @{class linorder} is equivalent to our locale @{locale strict_total_ordering}.\ - -context linorder begin - -interpretation strict_total_ordering UNIV - apply unfold_locales by auto - -subclass totalorder .. - -sublocale order: strict_total_ordering UNIV - rewrites "\x. x \ UNIV \ True" - and "\X. X \ UNIV \ True" - and "\r. r \ UNIV \ r" - and "\P. True \ P \ P" - and "Ball UNIV \ All" - and "Bex UNIV \ Ex" - and "sympartp (\)\<^sup>- \ sympartp (\)" - and "\P1. (True \ PROP P1) \ PROP P1" - and "\P1. (True \ P1) \ Trueprop P1" - and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" - apply unfold_locales by (auto simp:atomize_eq) - -end - text \Tests: facts should be available in the most general classes.\ -thm order.strict.trans[where 'a="'a::compat"] -thm order.extreme_bound_quasi_const[where 'a="'a::attractive_order"] +thm order.strict.trans[where 'a="'a::reflorder"] +thm order.extreme_bound_quasi_const[where 'a="'a::attrorder"] thm order.extreme_bound_singleton_eq[where 'a="'a::psorder"] thm order.trans[where 'a="'a::qorder"] -thm order.comparable_cases[where 'a="'a::totalorder"] +thm order.comparable_cases[where 'a="'a::linqorder"] thm order.cases[where 'a="'a::linorder"] subsection \Declaring Duals\ sublocale reflexive \ sym: reflexive A "sympartp (\)" rewrites "sympartp (\)\<^sup>- \ sympartp (\)" and "\r. sympartp (sympartp r) \ sympartp r" and "\r. sympartp r \ A \ sympartp (r \ A)" by (auto 0 4 simp:atomize_eq) sublocale quasi_ordered_set \ sym: quasi_ordered_set A "sympartp (\)" rewrites "sympartp (\)\<^sup>- = sympartp (\)" and "sympartp (sympartp (\)) = sympartp (\)" apply unfold_locales by (auto 0 4 dest: trans) text \At this point, we declare dual as sublocales. In the following, ``rewrites'' eventually cleans up redundant facts.\ sublocale reflexive \ dual: reflexive A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- \ sympartp (\)" and "\r. sympartp (r \ A) \ sympartp r \ A" and "(\)\<^sup>- \ A \ ((\) \ A)\<^sup>-" by (auto simp: atomize_eq) context attractive begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. sublocale dual: attractive A "(\)" rewrites "sympartp (\) = (\)" and "equivpartp (\) \ (\)" and "\r. sympartp (r \ A) \ sympartp r \ A" and "\r. sympartp (sympartp r) \ sympartp r" and "(\)\<^sup>- \ A \ ((\) \ A)\<^sup>-" apply unfold_locales by (auto intro!: ext dest: attract dual.attract simp: atomize_eq) end context irreflexive begin sublocale dual: irreflexive A "(\)\<^sup>-" rewrites "(\)\<^sup>- \ A \ ((\) \ A)\<^sup>-" apply unfold_locales by (auto dest: irrefl simp: atomize_eq) end sublocale transitive \ dual: transitive A "(\)\<^sup>-" rewrites "(\)\<^sup>- \ A \ ((\) \ A)\<^sup>-" and "sympartp (\)\<^sup>- = sympartp (\)" and "asympartp (\)\<^sup>- = (asympartp (\))\<^sup>-" apply unfold_locales by (auto dest: trans simp: atomize_eq intro!:ext) sublocale antisymmetric \ dual: antisymmetric A "(\)\<^sup>-" rewrites "(\)\<^sup>- \ A \ ((\) \ A)\<^sup>-" and "sympartp (\)\<^sup>- = sympartp (\)" by (auto dest: antisym simp: atomize_eq) +context antisymmetric begin + +lemma extreme_bound_unique: + "extreme_bound A (\) X x \ extreme_bound A (\) X y \ x = y" + apply (unfold extreme_bound_def) + apply (rule dual.extreme_unique) by auto + +lemma ex_extreme_bound_iff_ex1: + "Ex (extreme_bound A (\) X) \ Ex1 (extreme_bound A (\) X)" + apply (unfold extreme_bound_def) + apply (rule dual.ex_extreme_iff_ex1) by auto + +lemma ex_extreme_bound_iff_the: + "Ex (extreme_bound A (\) X) \ extreme_bound A (\) X (The (extreme_bound A (\) X))" + apply (rule iffI) + apply (rule theI') + using extreme_bound_unique by auto + +end + sublocale semiconnex \ dual: semiconnex A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" using semiconnex by auto sublocale connex \ dual: connex A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by (auto intro!: chainI dest:comparable) sublocale semiconnex_irreflexive \ dual: semiconnex_irreflexive A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales auto sublocale pseudo_ordered_set \ dual: pseudo_ordered_set A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales (auto 0 4) sublocale quasi_ordered_set \ dual: quasi_ordered_set A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales auto sublocale partially_ordered_set \ dual: partially_ordered_set A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales (auto 0 4) sublocale total_pseudo_ordered_set \ dual: total_pseudo_ordered_set A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales (auto 0 4) sublocale total_quasi_ordered_set \ dual: total_quasi_ordered_set A "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales auto sublocale compatible_ordering \ dual: compatible_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" apply unfold_locales - by (auto dest: compat_left compat_right strict_implies_weak) - -sublocale attractive_ordering \ dual: attractive_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + by (auto dest: strict_implies_weak strict_weak_trans weak_strict_trans) + +lemmas(in qorder) [intro!] = order.dual.quasi_ordered_subset + +sublocale reflexive_ordering \ dual: reflexive_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + rewrites "sympartp (\)\<^sup>- = sympartp (\)" + by unfold_locales auto + +sublocale reflexive_attractive_ordering \ dual: reflexive_attractive_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales auto sublocale pseudo_ordering \ dual: pseudo_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" - by unfold_locales (auto 0 4) + by unfold_locales auto sublocale quasi_ordering \ dual: quasi_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" by unfold_locales auto sublocale partial_ordering \ dual: partial_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" - by unfold_locales (auto 0 4) + by unfold_locales auto + +sublocale total_quasi_ordering \ dual: total_quasi_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + rewrites "sympartp (\)\<^sup>- = sympartp (\)" + by unfold_locales auto sublocale total_ordering \ dual: total_ordering A "(\)\<^sup>-" "(\)\<^sup>-" rewrites "sympartp (\)\<^sup>- = sympartp (\)" - by unfold_locales (auto 0 4) + by unfold_locales auto + +sublocale strict_quasi_ordering \ dual: strict_quasi_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + rewrites "sympartp (\)\<^sup>- = sympartp (\)" + by unfold_locales (auto simp: strict_iff) + +sublocale strict_partial_ordering \ dual: strict_partial_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + rewrites "sympartp (\)\<^sup>- = sympartp (\)" + by unfold_locales auto + +sublocale total_ordering \ dual: total_ordering A "(\)\<^sup>-" "(\)\<^sup>-" + rewrites "sympartp (\)\<^sup>- = sympartp (\)" + by unfold_locales auto lemma(in antisymmetric) monotone_extreme_imp_extreme_bound_iff: fixes ir (infix "\" 50) assumes "f ` C \ A" and "monotone_on C (\) (\) f" and i: "extreme C (\) i" shows "extreme_bound A (\) (f ` C) x \ f i = x" - using dual.extreme_unique monotone_extreme_extreme_boundI[OF assms] by auto - + using dual.extreme_unique monotone_extreme_extreme_boundI[OF assms] + by (auto simp: extreme_bound_def) subsection \Instantiations\ text \Finally, we instantiate our classes for sanity check.\ instance nat :: linorder .. text \Pointwise ordering of functions are compatible only if the weak part is transitive.\ -instance "fun" :: (type,qorder) compat +instance "fun" :: (type,qorder) reflorder proof (intro_classes, unfold_locales) note [simp] = le_fun_def less_fun_def fix f g h :: "'a \ 'b" { assume fg: "f \ g" and gh: "g < h" show "f < h" proof (unfold less_fun_def, intro conjI le_funI notI) from fg have "f x \ g x" for x by auto also from gh have "g x \ h x" for x by auto finally (order.trans) show "f x \ h x" for x. assume hf: "h \ f" then have "h x \ f x" for x by auto also from fg have "f x \ g x" for x by auto finally have "h \ g" by auto with gh show False by auto qed } { assume fg: "f < g" and gh: "g \ h" show "f < h" proof (unfold less_fun_def, intro conjI le_funI notI) from fg have "f x \ g x" for x by auto also from gh have "g x \ h x" for x by auto finally show "f x \ h x" for x. from gh have "g x \ h x" for x by auto also assume hf: "h \ f" then have "h x \ f x" for x by auto finally have "g \ f" by auto with fg show False by auto qed } show "f < g \ f \ g" by auto show "\f < f" by auto show "f \ f" by auto qed instance "fun" :: (type,qorder) qorder apply intro_classes apply unfold_locales by (auto simp: le_fun_def dest: order.trans) instance "fun" :: (type,porder) porder apply intro_classes apply unfold_locales proof (intro ext) fix f g :: "'a \ 'b" and x :: 'a assume fg: "f \ g" and gf: "g \ f" then have "f x \ g x" and "g x \ f x" by (auto elim: le_funE) from order.antisym[OF this] show "f x = g x" by auto qed end diff --git a/thys/Complete_Non_Orders/Complete_Relations.thy b/thys/Complete_Non_Orders/Complete_Relations.thy --- a/thys/Complete_Non_Orders/Complete_Relations.thy +++ b/thys/Complete_Non_Orders/Complete_Relations.thy @@ -1,590 +1,822 @@ (* Author: Akihisa Yamada (2018-2021) License: LGPL (see file COPYING.LESSER) *) section \ Completeness of Relations \ text \Here we formalize various order-theoretic completeness conditions.\ theory Complete_Relations - imports Binary_Relations + imports Well_Relations Directedness begin subsection \Completeness Conditions\ text \Order-theoretic completeness demands certain subsets of elements to admit suprema or infima.\ definition complete ("_-complete"[999]1000) where - "CC-complete A r \ \X \ A. X \ CC \ (\s. extreme_bound A r X s)" + "\-complete A (\) \ \X \ A. \ X (\) \ (\s. extreme_bound A (\) X s)" for less_eq (infix "\" 50) lemmas completeI = complete_def[unfolded atomize_eq, THEN iffD2, rule_format] lemmas completeD = complete_def[unfolded atomize_eq, THEN iffD1, rule_format] +lemmas completeE = complete_def[unfolded atomize_eq, THEN iffD1, rule_format, THEN exE] -lemma complete_cmono: "CC \ DD \ DD-complete \ CC-complete" +lemma complete_cmono: "CC \ DD \ DD-complete \ CC-complete" by (force simp: complete_def) -lemma complete_empty[simp]: "CC-complete {} r \ {} \ CC" by (auto simp: complete_def) +lemma complete_subclass: + fixes less_eq (infix "\" 50) + assumes "\-complete A (\)" and "\X \ A. \ X (\) \ \ X (\)" + shows "\-complete A (\)" + using assms by (auto simp: complete_def) -text \ -A related set $\tp{A,\SLE}$ is called \emph{topped} if there is a ``top'' element $\top \in A$, -a greatest element in $A$. -Note that there might be multiple tops if $(\SLE)$ is not antisymmetric.\ -definition "extremed A r \ \e. extreme A r e" - -lemma extremed_imp_ex_bound: "extremed A r \ X \ A \ \b \ A. bound X r b" - by (auto simp: extremed_def) +lemma complete_empty[simp]: "\-complete {} r \ \ \ {} r" by (auto simp: complete_def) context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin text\Toppedness can be also seen as a completeness condition, since it is equivalent to saying that the universe has a supremum.\ -lemma extremed_iff_UNIV_complete: "extremed A (\) \ {A}-complete A (\)" (is "?l \ ?r") +lemma extremed_iff_UNIV_complete: "extremed A (\) \ (\X r. X = A)-complete A (\)" (is "?l \ ?r") proof assume ?l - then obtain e where "extreme A (\) e" by (auto simp: extremed_def) + then obtain e where "extreme A (\) e" by (erule extremedE) then have "extreme_bound A (\) A e" by auto then show ?r by (auto intro!: completeI) next assume ?r from completeD[OF this] obtain s where "extreme_bound A (\) A s" by auto then have "extreme A (\) s" by auto then show ?l by (auto simp: extremed_def) qed text \The dual notion of topped is called ``pointed'', equivalently ensuring a supremum of the empty set.\ -lemma pointed_iff_empty_complete: "extremed A (\) \ {{}}-complete A (\)\<^sup>-" +lemma pointed_iff_empty_complete: "extremed A (\) \ (\X r. X = {})-complete A (\)\<^sup>-" by (auto simp: complete_def extremed_def) +text \Downward closure is topped.\ +lemma dual_closure_is_extremed: + assumes bA: "b \ A" and "b \ b" + shows "extremed {a \ A. a \ b} (\)" + using assms by (auto intro!: extremedI[of _ _ b]) + +text \Downward closure preserves completeness.\ +lemma dual_closure_is_complete: + assumes A: "\-complete A (\)" and bA: "b \ A" + shows "\-complete {x\A. x \ b} (\)" +proof (intro completeI) + fix X assume XAb:"X \ {x \ A. x \ b}" and X: "\ X (\)" + with completeD[OF A] obtain x where x: "extreme_bound A (\) X x" by auto + then have xA: "x \ A" by auto + from XAb have "x \ b" by (auto intro!: extreme_boundD[OF x] bA) + with xA x show "\x. extreme_bound {x \ A. x \ b} (\) X x" by (auto intro!: exI[of _ x]) +qed + +interpretation less_eq_dualize. + +text \Upward closure preserves completeness, under a condition.\ +lemma closure_is_complete: + assumes A: "\-complete A (\)" and bA: "b \ A" + and Cb: "\X \ A. \ X (\) \ bound X (\) b \ \ (X \ {b}) (\)" + shows "\-complete {x\A. b \ x} (\)" +proof (intro completeI) + fix X assume XAb:"X \ {x \ A. b \ x}" and X: "\ X (\)" + with Cb have XbC: "\ (X \ {b}) (\)" by auto + from XAb bA have XbA: "X \ {b} \ A" by auto + with completeD[OF A XbA] XbC + obtain x where x: "extreme_bound A (\) (X\{b}) x" by auto + then show "\x. extreme_bound {x \ A. b \ x} (\) X x" + by (auto intro!: exI[of _ x] extreme_boundI) +qed + +lemma biclosure_is_complete: + assumes A: "\-complete A (\)" and aA: "a \ A" and bA: "b \ A" and ab: "a \ b" + and Ca: "\X \ A. \ X (\) \ bound X (\) a \ \ (X \ {a}) (\)" + shows "\-complete {x\A. a \ x \ x \ b} (\)" +proof- + note closure_is_complete[OF A aA Ca] + from dual_closure_is_complete[OF this, of b] bA ab show ?thesis by auto +qed + end + text \One of the most well-studied notion of completeness would be the semilattice condition: every pair of elements $x$ and $y$ has a supremum $x \sqcup y$ (not necessarily unique if the underlying relation is not antisymmetric).\ -definition "pair_complete \ {{x,y} |. x y :: 'a}-complete" +definition "pair_complete \ (\X r. \x y. X = {x,y})-complete" lemma pair_completeI: assumes "\x y. x \ A \ y \ A \ \s. extreme_bound A r {x,y} s" shows "pair_complete A r" using assms by (auto simp: pair_complete_def complete_def) lemma pair_completeD: assumes "pair_complete A r" shows "x \ A \ y \ A \ \s. extreme_bound A r {x,y} s" by (intro completeD[OF assms[unfolded pair_complete_def]], auto) context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin lemma pair_complete_imp_directed: assumes comp: "pair_complete A (\)" shows "directed A (\)" proof fix x y :: 'a assume "x \ A" "y \ A" with pair_completeD[OF comp this] show "\z \ A. x \ z \ y \ z" by auto qed end lemma (in connex) pair_complete: "pair_complete A (\)" proof (safe intro!: pair_completeI) fix x y assume x: "x \ A" and y: "y \ A" then show "\s. extreme_bound A (\) {x, y} s" proof (cases rule:comparable_cases) case le with x y show ?thesis by (auto intro!: exI[of _ y]) next case ge with x y show ?thesis by (auto intro!: exI[of _ x]) qed qed text \The next one assumes that every nonempty finite set has a supremum.\ -abbreviation "finite_complete \ {X. finite X \ X \ {}}-complete" +abbreviation "finite_complete \ (\X r. finite X \ X \ {})-complete" lemma finite_complete_le_pair_complete: "finite_complete \ pair_complete" by (unfold pair_complete_def, rule complete_cmono, auto) text \The next one assumes that every nonempty bounded set has a supremum. It is also called the Dedekind completeness.\ -abbreviation "conditionally_complete A r \ {X. \b \ A. bound X r b \ X \ {}}-complete A r" +abbreviation "conditionally_complete A \ (\X r. \b \ A. bound X r b \ X \ {})-complete A" lemma conditionally_complete_imp_nonempty_imp_ex_extreme_bound_iff_ex_bound: assumes comp: "conditionally_complete A r" assumes "X \ A" and "X \ {}" shows "(\s. extreme_bound A r X s) \ (\b\A. bound X r b)" using assms by (auto 0 4 intro!:completeD[OF comp]) -text \The $\omega$-completeness condition demands a supremum for an $\omega$-chain, - $a_1 \sqsubseteq a_2 \sqsubseteq \dots$. - We model $\omega$-chain as the range of a monotone map $f : i \mapsto a_i$.\ - -definition "omega_complete A r \ {range f | f :: nat \ 'a. monotone (\) r f}-complete A r" +text \The $\omega$-completeness condition demands a supremum for an $\omega$-chain.\ -lemma (in transitive) local_chain: - assumes CA: "range C \ A" - shows "(\i::nat. C i \ C (Suc i)) \ monotone (<) (\) C" -proof (intro iffI allI monotoneI) - fix i j :: nat - assume loc: "\i. C i \ C (Suc i)" and ij: "i < j" - have "C i \ C (i+k+1)" for k - proof (induct k) - case 0 - from loc show ?case by auto - next - case (Suc k) - also have "C (i+k+1) \ C (i+k+Suc 1)" using loc by auto - finally show ?case using CA by auto - qed - from this[of "j-i-1"] ij show "C i \ C j" by auto -next - fix i - assume "monotone (<) (\) C" - then show "C i \ C (Suc i)" by (auto dest: monotoneD) -qed - -text\\emph{Directed completeness} is an important notion in domain theory~\<^cite>\"abramski94"\, +text\\emph{Directed completeness} is an important notion in domain theory~\cite{abramski94}, asserting that every nonempty directed set has a supremum. Here, a set $X$ is \emph{directed} if any pair of two elements in $X$ has a bound in $X$.\ -definition "directed_complete A r \ {X. directed X r \ X \ {}}-complete A r" +definition "directed_complete \ (\X r. directed X r \ X \ {})-complete" lemma monotone_directed_complete: assumes comp: "directed_complete A r" assumes fI: "f ` I \ A" and dir: "directed I ri" and I0: "I \ {}" and mono: "monotone_on I ri r f" shows "\s. extreme_bound A r (f ` I) s" apply (rule completeD[OF comp[unfolded directed_complete_def], OF fI]) using monotone_directed_image[OF mono dir] I0 by auto +lemma (in reflexive) dual_closure_is_directed_complete: + assumes comp: "directed-complete A (\)" and bA: "b \ A" + shows "directed-complete {X \ A. b \ X} (\)" + apply (rule closure_is_complete[OF comp bA]) +proof (intro allI impI directedI CollectI) + interpret less_eq_dualize. + fix X x y assume Xdir: "directed X (\)" and X: "X \ A" + and bX: "bound X (\) b" and x: "x \ X \ {b}" and y: "y \ X \ {b}" + from x y X bA have xA: "x \ A" and yA: "y \ A" by auto + show "\z\X \ {b}. x \ z \ y \ z" + proof (cases "x = b") + case [simp]: True + with y bX bA have "b \ y" by auto + with y yA bA show ?thesis by (auto intro!: bexI[of _ y]) + next + case False + with x have x: "x \ X" by auto + show ?thesis + proof (cases "y = b") + case [simp]: True + with x bX have "b \ x" by auto + with x xA bA show ?thesis by (auto intro!: bexI[of _ x]) + next + case False + with y have y: "y \ X" by auto + from directedD[OF Xdir x y] show ?thesis by simp + qed + qed +qed + + text \The next one is quite complete, only the empty set may fail to have a supremum. -The terminology follows \<^cite>\"Bergman2015"\, +The terminology follows \cite{Bergman2015}, although there it is defined more generally depending on a cardinal $\alpha$ such that a nonempty set $X$ of cardinality below $\alpha$ has a supremum.\ -abbreviation "semicomplete \ {X. X \ {}}-complete" +abbreviation "semicomplete \ (\X r. X \ {})-complete" lemma semicomplete_nonempty_imp_extremed: "semicomplete A r \ A \ {} \ extremed A r" unfolding extremed_iff_UNIV_complete - using complete_cmono[of "{A}" "{X. X \ {}}"] + using complete_cmono[of "\X r. X = A" "\X r. X \ {}"] by auto lemma connex_dual_semicomplete: "semicomplete {C. connex C r} (\)" -proof (intro completeI, elim CollectE) +proof (intro completeI) fix X assume "X \ {C. connex C r}" and "X \ {}" then have "connex (\X) r" by (auto simp: connex_def) then have "extreme_bound {C. connex C r} (\) X (\ X)" by auto then show "\S. extreme_bound {C. connex C r} (\) X S" by auto qed subsection \Pointed Ones\ text \The term `pointed' refers to the dual notion of toppedness, i.e., there is a global least element. This serves as the supremum of the empty set.\ -lemma complete_union: "(CC \ CC')-complete A r \ CC-complete A r \ CC'-complete A r" +lemma complete_sup: "(CC \ CC')-complete A r \ CC-complete A r \ CC'-complete A r" by (auto simp: complete_def) lemma pointed_directed_complete: - "{X. directed X r}-complete A r \ directed_complete A r \ extremed A r\<^sup>-" + "directed-complete A r \ directed_complete A r \ extremed A r\<^sup>-" proof- - have [simp]: "{X. directed X r \ X \ {} \ X = {}} = {X. directed X r}" by auto + have [simp]: "(\X r. directed X r \ X \ {} \ X = {}) = directed" by auto show ?thesis - by (auto simp: directed_complete_def pointed_iff_empty_complete complete_union[symmetric] Un_def) + by (auto simp: directed_complete_def pointed_iff_empty_complete complete_sup[symmetric] + sup_fun_def) qed text \``Bounded complete'' refers to pointed conditional complete, but this notion is just the dual of semicompleteness. We prove this later.\ -abbreviation "bounded_complete A r \ {X. \b\A. bound X r b}-complete A r" +abbreviation "bounded_complete A \ (\X r. \b\A. bound X r b)-complete A" subsection \Relations between Completeness Conditions\ context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin -interpretation less_eq_notations. +interpretation less_eq_dualize. text \Pair-completeness implies that the universe is directed. Thus, with directed completeness implies toppedness.\ proposition directed_complete_pair_complete_imp_extremed: assumes dc: "directed_complete A (\)" and pc: "pair_complete A (\)" and A: "A \ {}" shows "extremed A (\)" proof- have "\s. extreme_bound A (\) A s" apply (rule completeD[OF dc[unfolded directed_complete_def]]) using pair_complete_imp_directed[OF pc] A by auto then obtain t where "extreme_bound A (\) A t" by auto then have "\x \ A. x \ t" and "t \ A" by auto then show ?thesis by (auto simp: extremed_def) qed text \Semicomplete is conditional complete and topped.\ proposition semicomplete_iff_conditionally_complete_extremed: assumes A: "A \ {}" shows "semicomplete A (\) \ conditionally_complete A (\) \ extremed A (\)" (is "?l \ ?r") proof (intro iffI) assume r: "?r" then have cc: "conditionally_complete A (\)" and e: "extremed A (\)" by auto show ?l - proof (intro completeI, unfold mem_Collect_eq) + proof (intro completeI) fix X assume X: "X \ A" and "X \ {}" with extremed_imp_ex_bound[OF e X] show "\s. extreme_bound A (\) X s" by (auto intro!: completeD[OF cc X]) qed next assume l: ?l with semicomplete_nonempty_imp_extremed[OF l] A show ?r by (auto intro!: completeI dest: completeD) qed proposition complete_iff_pointed_semicomplete: - "UNIV-complete A (\) \ semicomplete A (\) \ extremed A (\)" (is "?l \ ?r") - by (unfold pointed_iff_empty_complete complete_union[symmetric] Un_def, auto) + "\-complete A (\) \ semicomplete A (\) \ extremed A (\)" (is "?l \ ?r") + by (unfold pointed_iff_empty_complete complete_sup[symmetric], auto simp: sup_fun_def top_fun_def) text \Conditional completeness only lacks top and bottom to be complete.\ proposition complete_iff_conditionally_complete_extremed_pointed: - "UNIV-complete A (\) \ conditionally_complete A (\) \ extremed A (\) \ extremed A (\)" + "\-complete A (\) \ conditionally_complete A (\) \ extremed A (\) \ extremed A (\)" unfolding complete_iff_pointed_semicomplete apply (cases "A = {}") apply (auto intro!: completeI dest: extremed_imp_ex_bound)[1] unfolding semicomplete_iff_conditionally_complete_extremed apply (auto intro:completeI) done text \If the universe is directed, then every pair is bounded, and thus has a supremum. On the other hand, supremum gives an upper bound, witnessing directedness.\ proposition conditionally_complete_imp_pair_complete_iff_directed: assumes comp: "conditionally_complete A (\)" shows "pair_complete A (\) \ directed A (\)" (is "?l \ ?r") proof(intro iffI) assume ?r then show ?l by (auto intro!: pair_completeI completeD[OF comp] elim: directedE) next assume pc: ?l show ?r proof (intro directedI) fix x y assume "x \ A" and "y \ A" with pc obtain z where "extreme_bound A (\) {x,y} z" by (auto dest: pair_completeD) then show "\z\A. x \ z \ y \ z" by auto qed qed end -text \Following is a new generalization of (weak) chain-completeness:\ -abbreviation "well_complete A r \ {X. well_related_set X r}-complete A r" - subsection \Duality of Completeness Conditions\ text \Conditional completeness is symmetric.\ context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) begin -interpretation less_eq_notations. +interpretation less_eq_dualize. lemma conditionally_complete_dual: assumes comp: "conditionally_complete A (\)" shows "conditionally_complete A (\)" -proof (intro completeI, elim CollectE) +proof (intro completeI) fix X assume XA: "X \ A" define B where [simp]: "B \ {b\A. bound X (\) b}" assume bound: "\b\A. bound X (\) b \ X \ {}" with in_mono[OF XA] have B: "B \ A" and "\b \ A. bound B (\) b" and "B \ {}" by auto from comp[THEN completeD, OF B] this obtain s where "s \ A" and "extreme_bound A (\) B s" by auto - with in_mono[OF XA] show "\s. extreme_bound A (\) X s" by (intro exI[of _ s] extremeI, auto) + with in_mono[OF XA] show "\s. extreme_bound A (\) X s" + by (intro exI[of _ s] extreme_boundI, auto) qed text \Full completeness is symmetric.\ lemma complete_dual: - "UNIV-complete A (\) \ UNIV-complete A (\)" + "\-complete A (\) \ \-complete A (\)" apply (unfold complete_iff_conditionally_complete_extremed_pointed) using conditionally_complete_dual by auto text \Now we show that bounded completeness is the dual of semicompleteness.\ -lemma pointed_conditionally_complete_iff_bounded_complete: +lemma bounded_complete_iff_pointed_conditionally_complete: assumes A: "A \ {}" - shows "conditionally_complete A (\) \ extremed A (\) \ bounded_complete A (\)" + shows "bounded_complete A (\) \ conditionally_complete A (\) \ extremed A (\)" apply (unfold pointed_iff_empty_complete) - apply (fold complete_union) - apply (unfold Un_def) + apply (fold complete_sup) + apply (unfold sup_fun_def) apply (rule arg_cong[of _ _ "\CC. CC-complete A (\)"]) using A by auto proposition bounded_complete_iff_dual_semicomplete: "bounded_complete A (\) \ semicomplete A (\)" proof (cases "A = {}") case True then show ?thesis by auto next case False then show ?thesis - apply (fold pointed_conditionally_complete_iff_bounded_complete[OF False]) + apply (unfold bounded_complete_iff_pointed_conditionally_complete[OF False]) apply (unfold semicomplete_iff_conditionally_complete_extremed) using Complete_Relations.conditionally_complete_dual by auto qed +lemma bounded_complete_imp_conditionally_complete: + assumes "bounded_complete A (\)" shows "conditionally_complete A (\)" + using assms by (cases "A = {}", auto simp:bounded_complete_iff_pointed_conditionally_complete) + +text \Completeness in downward-closure:\ +lemma conditionally_complete_imp_semicomplete_in_dual_closure: + assumes A: "conditionally_complete A (\)" and bA: "b \ A" + shows "semicomplete {a \ A. a \ b} (\)" +proof (intro completeI) + fix X assume X: "X \ {a \ A. a \ b}" and X0: "X \ {}" + then have "X \ A" and Xb: "bound X (\) b" by auto + with bA completeD[OF A] X0 obtain s where Xs: "extreme_bound A (\) X s" by auto + with Xb bA have sb: "s \ b" by auto + with Xs have "extreme_bound {a \ A. a \ b} (\) X s" + by (intro extreme_boundI, auto) + then show "\s. extreme_bound {a \ A. a \ b} (\) X s" by auto +qed + end +text \Completeness in intervals:\ +lemma conditionally_complete_imp_complete_in_interval: + fixes less_eq (infix "\" 50) + assumes comp: "conditionally_complete A (\)" and aA: "a \ A" and bA: "b \ A" + and aa: "a \ a" and ab: "a \ b" + shows "\-complete {x \ A. a \ x \ x \ b} (\)" +proof (intro completeI) + fix X assume X: "X \ {x \ A. a \ x \ x \ b}" + note conditionally_complete_imp_semicomplete_in_dual_closure[OF comp bA] + from closure_is_complete[OF this, of a,simplified] aA ab + have semi: "semicomplete {x \ A. a \ x \ x \ b} (\)" by (simp add: conj_commute cong: Collect_cong) + show "Ex (extreme_bound {x \ A. a \ x \ x \ b} (\) X)" + proof (cases "X = {}") + case True + with aA aa ab have "extreme_bound {x \ A. a \ x \ x \ b} (\) X a" by (auto simp: bound_empty) + then show ?thesis by auto + next + case False + with completeD[OF semi X] show ?thesis by simp + qed +qed + lemmas connex_bounded_complete = connex_dual_semicomplete[folded bounded_complete_iff_dual_semicomplete] subsection \Completeness Results Requiring Order-Like Properties\ text \Above results hold without any assumption on the relation. This part demands some order-like properties.\ text \It is well known that in a semilattice, i.e., a pair-complete partial order, every finite nonempty subset of elements has a supremum. We prove the result assuming transitivity, but only that.\ lemma (in transitive) pair_complete_iff_finite_complete: "pair_complete A (\) \ finite_complete A (\)" (is "?l \ ?r") proof (intro iffI completeI, elim CollectE conjE) fix X assume pc: ?l show "finite X \ X \ A \ X \ {} \ Ex (extreme_bound A (\) X)" proof (induct X rule: finite_induct) case empty then show ?case by auto next case (insert x X) then have x: "x \ A" and X: "X \ A" by auto show ?case proof (cases "X = {}") case True obtain x' where "extreme_bound A (\) {x,x} x'" using pc[THEN pair_completeD, OF x x] by auto with True show ?thesis by (auto intro!: exI[of _ x']) next case False with insert obtain b where b: "extreme_bound A (\) X b" by auto with pc[THEN pair_completeD] x obtain c where c: "extreme_bound A (\) {x,b} c" by auto from c have cA: "c \ A" and xc: "x \ c" and bc: "b \ c" by auto from b have bA: "b \ A" and bX: "bound X (\) b" by auto show ?thesis proof (intro exI extreme_boundI) fix xb assume xb: "xb \ insert x X" from bound_trans[OF bX bc X bA cA] have "bound X (\) c". with xb xc show "xb \ c" by auto next fix d assume "bound (insert x X) (\) d" and dA: "d \ A" with b have "bound {x,b} (\) d" by auto with c show "c \ d" using dA by auto qed (fact cA) qed qed qed (insert finite_complete_le_pair_complete, auto) -text \Gierz et al.~\<^cite>\"gierz03"\ showed that a directed complete partial order is semicomplete +text \Gierz et al.~\cite{gierz03} showed that a directed complete partial order is semicomplete if and only if it is also a semilattice. We generalize the claim so that the underlying relation is only transitive.\ proposition(in transitive) semicomplete_iff_directed_complete_pair_complete: "semicomplete A (\) \ directed_complete A (\) \ pair_complete A (\)" (is "?l \ ?r") proof (intro iffI) assume l: ?l then show ?r by (auto simp: directed_complete_def intro!: completeI pair_completeI completeD[OF l]) next assume ?r then have dc: "directed_complete A (\)" and pc: "pair_complete A (\)" by auto with pair_complete_iff_finite_complete have fc: "finite_complete A (\)" by auto show ?l - proof (intro completeI, elim CollectE) + proof (intro completeI) fix X assume XA: "X \ A" have 1: "directed {x. \Y \ X. finite Y \ Y \ {} \ extreme_bound A (\) Y x} (\)" (is "directed ?B _") proof (intro directedI) fix a b assume a: "a \ ?B" and b: "b \ ?B" from a obtain Y where Y: "extreme_bound A (\) Y a" "finite Y" "Y \ {}" "Y \ X" by auto from b obtain B where B: "extreme_bound A (\) B b" "finite B" "B \ {}" "B \ X" by auto from XA Y B have AB: "Y \ A" "B \ A" "finite (Y \ B)" "Y \ B \ {}" "Y \ B \ X" by auto with fc[THEN completeD] have "Ex (extreme_bound A (\) (Y \ B))" by auto then obtain c where c: "extreme_bound A (\) (Y \ B) c" by auto show "\c \ ?B. a \ c \ b \ c" proof (intro bexI conjI) from Y B c show "a \ c" and "b \ c" by (auto simp: extreme_bound_iff) from AB c show "c \ ?B" by (auto intro!: exI[of _ "Y \ B"]) qed qed have B: "?B \ A" by auto assume "X \ {}" then obtain x where xX: "x \ X" by auto with fc[THEN completeD, of "{x}"] XA obtain x' where "extreme_bound A (\) {x} x'" by auto with xX have x'B: "x' \ ?B" by (auto intro!: exI[of _ "{x}"] extreme_boundI) then have 2: "?B \ {}" by auto from dc[unfolded directed_complete_def, THEN completeD, of ?B] 1 2 obtain b where b: "extreme_bound A (\) ?B b" by auto then have bA: "b \ A" by auto show "Ex (extreme_bound A (\) X)" proof (intro exI extreme_boundI UNIV_I) fix x assume xX: "x \ X" with XA fc[THEN completeD, of "{x}"] obtain c where c: "extreme_bound A (\) {x} c" by auto then have cA: "c \ A" and xc: "x \ c" by auto from c xX have cB: "c \ ?B" by (auto intro!: exI[of _ "{x}"] extreme_boundI) with b have bA: "b \ A" and cb: "c \ b" by auto from xX XA cA bA trans[OF xc cb] show "x \ b" by auto text\ Here transitivity is needed. \ next fix x assume xA: "x \ A" and Xx: "bound X (\) x" have "bound ?B (\) x" proof (intro boundI UNIV_I, clarify) fix c Y assume "finite Y" and YX: "Y \ X" and "Y \ {}" and c: "extreme_bound A (\) Y c" from YX Xx have "bound Y (\) x" by auto with c xA show "c \ x" by auto qed with b xA show "b \ x" by auto qed (fact bA) qed qed text\The last argument in the above proof requires transitivity, but if we had reflexivity then $x$ itself is a supremum of $\set{x}$ (see @{thm reflexive.extreme_bound_singleton}) and so $x \SLE s$ would be immediate. Thus we can replace transitivity by reflexivity, but then pair-completeness does not imply finite completeness. We obtain the following result.\ proposition (in reflexive) semicomplete_iff_directed_complete_finite_complete: "semicomplete A (\) \ directed_complete A (\) \ finite_complete A (\)" (is "?l \ ?r") proof (intro iffI) assume l: ?l then show ?r by (auto simp: directed_complete_def intro!: completeI pair_completeI completeD[OF l]) next assume ?r then have dc: "directed_complete A (\)" and fc: "finite_complete A (\)" by auto show ?l - proof (intro completeI, elim CollectE) + proof (intro completeI) fix X assume XA: "X \ A" have 1: "directed {x. \Y \ X. finite Y \ Y \ {} \ extreme_bound A (\) Y x} (\)" (is "directed ?B _") proof (intro directedI) fix a b assume a: "a \ ?B" and b: "b \ ?B" from a obtain Y where Y: "extreme_bound A (\) Y a" "finite Y" "Y \ {}" "Y \ X" by auto from b obtain B where B: "extreme_bound A (\) B b" "finite B" "B \ {}" "B \ X" by auto from XA Y B have AB: "Y \ A" "B \ A" "finite (Y \ B)" "Y \ B \ {}" "Y \ B \ X" by auto with fc[THEN completeD] have "Ex (extreme_bound A (\) (Y \ B))" by auto then obtain c where c: "extreme_bound A (\) (Y \ B) c" by auto show "\c \ ?B. a \ c \ b \ c" proof (intro bexI conjI) from Y B c show "a \ c" and "b \ c" by (auto simp: extreme_bound_iff) from AB c show "c \ ?B" by (auto intro!: exI[of _ "Y \ B"]) qed qed have B: "?B \ A" by auto assume "X \ {}" then obtain x where xX: "x \ X" by auto with XA have "extreme_bound A (\) {x} x" by (intro extreme_bound_singleton, auto) with xX have xB: "x \ ?B" by (auto intro!: exI[of _ "{x}"]) then have 2: "?B \ {}" by auto from dc[unfolded directed_complete_def, THEN completeD, of ?B] B 1 2 obtain b where b: "extreme_bound A (\) ?B b" by auto then have bA: "b \ A" by auto show "Ex (extreme_bound A (\) X)" proof (intro exI extreme_boundI UNIV_I) fix x assume xX: "x \ X" with XA have x: "extreme_bound A (\) {x} x" by (intro extreme_bound_singleton, auto) from x xX have cB: "x \ ?B" by (auto intro!: exI[of _ "{x}"]) with b show "x \ b" by auto next fix x assume Xx: "bound X (\) x" and xA: "x \ A" have "bound ?B (\) x" proof (intro boundI UNIV_I, clarify) fix c Y assume "finite Y" and YX: "Y \ X" and "Y \ {}" and c: "extreme_bound A (\) Y c" from YX Xx have "bound Y (\) x" by auto with YX XA xA c show "c \ x" by auto qed with xA b show "b \ x" by auto qed (fact bA) qed qed + subsection \Relating to Classes\ (* HOL.Conditionally_Complete_Lattices imports choice. text \Isabelle's class @{class conditionally_complete_lattice} satisfies @{const conditionally_complete}. The other direction does not hold, since for the former, @{term "Sup {}"} and @{term "Inf {}"} are arbitrary even if there are top or bottom elements.\ lemma (in conditionally_complete_lattice) "conditionally_complete UNIV (\)" proof (intro completeI, elim CollectE) fix X :: "'a set" assume "\b\UNIV. bound X (\) b \ X \ {}" then have bdd:"bdd_above X" and X0: "X \ {}" by (auto 0 3) from cSup_upper[OF _ bdd] cSup_least[OF X0] have "supremum X (Sup X)" by (intro extremeI boundI, auto) then show "\s. least {b \ UNIV. bound X (\) b} s" by auto qed *) -text \Isabelle's class @{class complete_lattice} is @{term "UNIV-complete"}.\ +text \Isabelle's class @{class complete_lattice} is @{term "\-complete"}.\ -lemma (in complete_lattice) "UNIV-complete UNIV (\)" - by (auto intro!: completeI Sup_upper Sup_least Inf_lower Inf_greatest) +lemma (in complete_lattice) "\-complete UNIV (\)" + by (auto intro!: completeI extreme_boundI Sup_upper Sup_least Inf_lower Inf_greatest) subsection \Set-wise Completeness\ +lemma Pow_extreme_bound: "X \ Pow A \ extreme_bound (Pow A) (\) X (\X)" + by (intro extreme_boundI, auto 2 3) + +lemma Pow_complete: "\-complete (Pow A) (\)" + by (auto intro!: completeI dest: Pow_extreme_bound) + +lemma directed_directed_Un: +assumes ch: "XX \ {X. directed X r}" and dir: "directed XX (\)" + shows "directed (\XX) r" +proof (intro directedI, elim UnionE) + fix x y X Y assume xX: "x \ X" and X: "X \ XX" and yY: "y \ Y" and Y: "Y \ XX" + from directedD[OF dir X Y] + obtain Z where "X \ Z" "Y \ Z" and Z: "Z \ XX" by auto + with ch xX yY have "directed Z r" "x \ Z" "y \ Z" by auto + then obtain z where "z \ Z" "r x z \ r y z" by (auto elim:directedE) + with Z show "\z\\ XX. r x z \ r y z" by auto +qed + +lemmas directed_connex_Un = directed_directed_Un[OF _ connex.directed] + lemma directed_sets_directed_complete: assumes cl: "\DC. DC \ AA \ (\X\DC. directed X r) \ (\DC) \ AA" - shows "{DC. directed DC (\)}-complete {X \ AA. directed X r} (\)" -proof (intro completeI, safe) + shows "directed-complete {X \ AA. directed X r} (\)" +proof (intro completeI) fix XX assume ch: "XX \ {X \ AA. directed X r}" and dir: "directed XX (\)" with cl have "(\XX) \ AA" by auto moreover have "directed (\XX) r" - proof (intro directedI, elim UnionE) - fix x y X Y assume xX: "x \ X" and X: "X \ XX" and yY: "y \ Y" and Y: "Y \ XX" - from directedD[OF dir X Y] - obtain Z where "X \ Z" "Y \ Z" and Z: "Z \ XX" by auto - with ch xX yY have "directed Z r" "x \ Z" "y \ Z" by auto - then obtain z where "z \ Z" "r x z \ r y z" by (auto elim:directedE) - with Z show "\z\\ XX. r x z \ r y z" by auto - qed - ultimately have "extreme_bound {X \ AA. directed X r} (\) XX (\XX)" by auto - then show "Ex (extreme_bound {X \ AA. directed X r} (\) XX)" by auto + apply (rule directed_directed_Un) using ch by (auto simp: dir) + ultimately show "Ex (extreme_bound {X \ AA. directed X r} (\) XX)" + by (auto intro!: exI[of _ "\XX"]) qed lemma connex_directed_Un: assumes ch: "CC \ {C. connex C r}" and dir: "directed CC (\)" shows "connex (\CC) r" proof (intro connexI, elim UnionE) fix x y X Y assume xX: "x \ X" and X: "X \ CC" and yY: "y \ Y" and Y: "Y \ CC" from directedD[OF dir X Y] obtain Z where "X \ Z" "Y \ Z" "Z \ CC" by auto with xX yY ch have "x \ Z" "y \ Z" "connex Z r" by auto then show "r x y \ r y x" by (auto elim:connexE) qed -lemma connex_directed_complete: "{DC. directed DC (\)}-complete {C. connex C r} (\)" - by (intro completeI, force dest!: connex_directed_Un) +lemma connex_is_directed_complete: "directed-complete {C. C \ A \ connex C r} (\)" +proof (intro completeI) + fix CC assume CC: "CC \ {C. C \ A \ connex C r}" and "directed CC (\)" + with connex_directed_Un have Scon: "connex (\CC) r" by auto + from CC have SA: "\CC \ A" by auto + from Scon SA show "\S. extreme_bound {C. C \ A \ connex C r} (\) CC S" + by (auto intro!: exI[of _ "\CC"] extreme_boundI) +qed + +lemma (in well_ordered_set) well_ordered_set_insert: + assumes aA: "total_ordered_set (insert a A) (\)" + shows "well_ordered_set (insert a A) (\)" +proof- + interpret less_eq_asymmetrize. + interpret aA: total_ordered_set "insert a A" "(\)" using aA. + show ?thesis + proof (intro well_ordered_set.intro aA.antisymmetric_axioms well_related_setI) + fix X assume XaA: "X \ insert a A" and X0: "X \ {}" + show "\e. extreme X (\) e" + proof (cases "a \ X") + case False + with XaA have "X \ A" by auto + from nonempty_imp_ex_extreme[OF this X0] show ?thesis. + next + case aX: True + show ?thesis + proof (cases "X - {a} = {}") + case True + with aX XaA have Xa: "X = {a}" by auto + from aA.refl[of a] + have "a \ a" by auto + then show ?thesis by (auto simp: Xa) + next + case False + from nonempty_imp_ex_extreme[OF _ False] XaA + obtain e where Xae: "extreme (X-{a}) (\) e" by auto + with Xae XaA have eaA: "e \ insert a A" by auto + then have "e \ a \ a \ e" by (intro aA.comparable, auto) + then show ?thesis + proof (elim disjE) + assume ea: "e \ a" + with Xae show ?thesis by (auto intro!:exI[of _ e]) + next + assume ae: "a \ e" + show ?thesis + proof (intro exI[of _ a] extremeI aX) + fix x assume xX: "x \ X" + show "a \ x" + proof (cases "a = x") + case True with aA.refl[of a] show ?thesis by auto + next + case False + with xX have "x \ X - {a}" by auto + with Xae have "e \ x" by auto + from aA.trans[OF ae this _ eaA] xX XaA + show ?thesis by auto + qed + qed + qed + qed + qed + qed +qed + +text \The following should be true in general, but here we use antisymmetry to avoid + the axiom of choice.\ + +lemma (in antisymmetric) pointwise_connex_complete: + assumes comp: "connex-complete A (\)" + shows "connex-complete {f. f ` X \ A} (pointwise X (\))" +proof (safe intro!: completeI exI) + fix F + assume FXA: "F \ {f. f ` X \ A}" and F: "connex F (pointwise X (\))" + show "extreme_bound {f. f ` X \ A} (pointwise X (\)) F (\x. The (extreme_bound A (\) {f x |. f \ F}))" + proof (unfold pointwise_extreme_bound[OF FXA], safe) + fix x assume xX: "x \ X" + from FXA xX have FxA: "{f x |. f \ F} \ A" by auto + have "Ex (extreme_bound A (\) {f x |. f \ F})" + proof (intro completeD[OF comp] FxA CollectI connexI, elim imageE, fold atomize_eq) + fix f g assume fF: "f \ F" and gF: "g \ F" + from connex.comparable[OF F this] xX show "f x \ g x \ g x \ f x" by auto + qed + also note ex_extreme_bound_iff_the + finally + show "extreme_bound A (\) {f x |. f \ F} (The (extreme_bound A (\) {f x |. f \ F}))". + qed +qed + +text \Our supremum/infimum coincides with those of Isabelle's @{class complete_lattice}.\ + +lemma complete_UNIV: "\-complete (UNIV::'a::complete_lattice set) (\)" +proof- + have "Ex (supremum X)" for X :: "'a set" + by (auto intro!: exI[of _ "\X"] supremumI simp:Sup_upper Sup_least) + then show ?thesis by (auto intro!: completeI ) +qed + +context + fixes X :: "'a :: complete_lattice set" +begin + +lemma supremum_Sup: "supremum X (\X)" +proof- + define it where "it \ The (supremum X)" + note completeD[OF complete_UNIV,simplified, of X] + from this[unfolded order.dual.ex_extreme_iff_the] + have 1: "supremum X it" by (simp add: it_def) + then have "\X = it" by (intro Sup_eqI, auto) + with 1 show ?thesis by auto +qed + +lemmas Sup_eq_The_supremum = order.dual.eq_The_extreme[OF supremum_Sup] + +lemma supremum_eq_Sup: "supremum X x \ \X = x" + using order.dual.eq_The_extreme supremum_Sup by auto + +lemma infimum_Inf: + shows "infimum X (\X)" +proof- + define it where "it \ The (infimum X)" + note completeD[OF complete_dual[OF complete_UNIV],simplified, of X] + from this[unfolded order.ex_extreme_iff_the] + have 1: "infimum X it" by (simp add: it_def) + then have "\X = it" by (intro Inf_eqI, auto) + with 1 show ?thesis by auto +qed + +lemmas Inf_eq_The_infimum = order.eq_The_extreme[OF infimum_Inf] + +lemma infimum_eq_Inf: "infimum X x \ \X = x" + using order.eq_The_extreme infimum_Inf by auto + +end end \ No newline at end of file diff --git a/thys/Complete_Non_Orders/Continuity.thy b/thys/Complete_Non_Orders/Continuity.thy new file mode 100755 --- /dev/null +++ b/thys/Complete_Non_Orders/Continuity.thy @@ -0,0 +1,166 @@ +theory Continuity + imports Complete_Relations +begin + +subsection \Scott Continuity, $\omega$-Continuity\ + +text \In this Section, we formalize Scott continuity and $\omega$-continuity. +We then prove that a Scott continuous map is $\omega$-continuous and that an $\omega$-continuous +map is ``nearly'' monotone.\ + +definition continuous ("_-continuous" [1000]1000) where + "\-continuous A (\) B (\) f \ + f ` A \ B \ + (\X s. \ X (\) \ X \ {} \ X \ A \ extreme_bound A (\) X s \ extreme_bound B (\) (f`X) (f s))" + for leA (infix "\" 50) and leB (infix "\" 50) + +lemmas continuousI[intro?] = + continuous_def[unfolded atomize_eq, THEN iffD2, unfolded conj_imp_eq_imp_imp, rule_format] + +lemmas continuousE = + continuous_def[unfolded atomize_eq, THEN iffD1, elim_format, unfolded conj_imp_eq_imp_imp, rule_format] + +lemma + fixes prec_eq (infix "\" 50) and less_eq (infix "\" 50) + assumes "\-continuous I (\) A (\) f" + shows continuous_carrierD[dest]: "f ` I \ A" + and continuousD: "\ X (\) \ X \ {} \ X \ I \ extreme_bound I (\) X b \ extreme_bound A (\) (f ` X) (f b)" + using assms by (auto elim!: continuousE) + +lemma continuous_comp: + fixes leA (infix "\\<^sub>A" 50) and leB (infix "\\<^sub>B" 50) and leC (infix "\\<^sub>C" 50) + assumes KfL: "\X \ A. \ X (\\<^sub>A) \ \ (f ` X) (\\<^sub>B)" + assumes f: "\-continuous A (\\<^sub>A) B (\\<^sub>B) f" and g: "\-continuous B (\\<^sub>B) C (\\<^sub>C) g" + shows "\-continuous A (\\<^sub>A) C (\\<^sub>C) (g \ f)" + apply (intro continuousI) +proof - + from f g have fAB: "f ` A \ B" and gBC: "g ` B \ C" by auto + then show "(g \ f) ` A \ C" by auto + fix X s + assume XA: "X \ A" and X0: "X \ {}" and XK: "\ X (\\<^sub>A)" and Xs: "extreme_bound A (\\<^sub>A) X s" + from fAB XA have fXB: "f ` X \ B" by auto + from X0 have fX0: "f ` X \ {}" by auto + from KfL XA XK have fXL: "\ (f ` X) (\\<^sub>B)" by auto + from continuousD[OF f XK X0 XA Xs] + have "extreme_bound B (\\<^sub>B) (f ` X) (f s)". + from continuousD[OF g fXL fX0 fXB this] + show "extreme_bound C (\\<^sub>C) ((g\f)`X) ((g\f) s)" by (auto simp: image_comp) +qed + +lemma continuous_comp_top: + fixes leA (infix "\\<^sub>A" 50) and leB (infix "\\<^sub>B" 50) and leC (infix "\\<^sub>C" 50) + assumes f: "\-continuous A (\\<^sub>A) B (\\<^sub>B) f" and g: "\-continuous B (\\<^sub>B) C (\\<^sub>C) g" + shows "\-continuous A (\\<^sub>A) C (\\<^sub>C) (g \ f)" + by (rule continuous_comp[OF _ f g], auto) + +lemma id_continuous: + fixes leA (infix "\\<^sub>A" 50) + shows "\-continuous A (\\<^sub>A) A (\\<^sub>A) (\x. x)" + by (auto intro: continuousI) + +lemma cst_continuous: + fixes leA (infix "\\<^sub>A" 50) and leB (infix "\\<^sub>B" 50) + assumes "b \ B" and bb: "b \\<^sub>B b" + shows "\-continuous A (\\<^sub>A) B (\\<^sub>B) (\x. b)" + apply (intro continuousI) + using assms(1) apply auto + using assms extreme_bound_singleton_refl[of B "(\\<^sub>B)" b] by blast + + +lemma continuous_cmono: + assumes CD: "\ \ \" shows "\-continuous \ \-continuous" +proof (safe intro!: le_funI le_boolI) + fix I A f and prec_eq (infix "\" 50) and less_eq (infix "\" 50) + assume cont: "\-continuous I (\) A (\) f" + show "\-continuous I (\) A (\) f" + proof (rule continuousI) + from cont show "f ` I \ A" by auto + fix X s assume X: "\ X (\)" and X0: "X \ {}" and XI: "X \ I" and Xs: "extreme_bound I (\) X s" + from CD X have "\ X (\)" by auto + from continuousD[OF cont, OF this X0 XI Xs] + show "extreme_bound A (\) (f ` X) (f s)". + qed +qed + +context + fixes prec_eq :: "'i \ 'i \ bool" (infix "\" 50) and less_eq :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma continuous_subclass: + assumes CD: "\X\I. X \ {} \ \ X (\) \ \ X (\)" and cont: "\-continuous I (\) A (\) f" + shows "\-continuous I (\) A (\) f" + using cont by (auto simp: continuous_def CD[rule_format]) + +lemma chain_continuous_imp_well_continuous: + assumes cont: "connex-continuous I (\) A (\) f" + shows "well_related_set-continuous I (\) A (\) f" + by (rule continuous_subclass[OF _ cont], auto simp: well_related_set.connex) + +lemma well_continuous_imp_omega_continous: + assumes cont: "well_related_set-continuous I (\) A (\) f" + shows "omega_chain-continuous I (\) A (\) f" + by (rule continuous_subclass[OF _ cont], auto simp: omega_chain_imp_well_related) + +end + +abbreviation "scott_continuous I (\) \ directed_set-continuous I (\)" + for prec_eq (infix "\" 50) + +lemma scott_continuous_imp_well_continuous: + fixes prec_eq :: "'i \ 'i \ bool" (infix "\" 50) and less_eq :: "'a \ 'a \ bool" (infix "\" 50) + assumes cont: "scott_continuous I (\) A (\) f" + shows "well_related_set-continuous I (\) A (\) f" + by (rule continuous_subclass[OF _ cont], auto simp: well_related_set.directed_set) + +lemmas scott_continuous_imp_omega_continuous = + scott_continuous_imp_well_continuous[THEN well_continuous_imp_omega_continous] + + +subsubsection \Continuity implies monotonicity\ + +lemma continuous_imp_mono_refl: + fixes prec_eq (infix "\" 50) and less_eq (infix "\" 50) + assumes cont: "\-continuous I (\) A (\) f" and xyC: "\ {x,y} (\)" + and xy: "x \ y" and yy: "y \ y" + and x: "x \ I" and y: "y \ I" + shows "f x \ f y" +proof- + have fboy: "extreme_bound A (\) (f ` {x,y}) (f y)" + proof (intro continuousD[OF cont] xyC) + from x y show CI: "{x,y} \ I" by auto + show Cy: "extreme_bound I (\) {x,y} y" using xy yy x y by auto + qed auto + then show ?thesis by auto +qed + +lemma omega_continuous_imp_mono_refl: + fixes prec_eq (infix "\" 50) and less_eq (infix "\" 50) + assumes cont: "omega_chain-continuous I (\) A (\) f" + and xx: "x \ x" and xy: "x \ y" and yy: "y \ y" + and x: "x \ I" and y: "y \ I" + shows "f x \ f y" + apply (rule continuous_imp_mono_refl[OF cont, OF pair_omega_chain]) + using assms by auto + +context reflexive begin + +lemma continuous_imp_monotone_on: + fixes leB (infix "\" 50) + assumes cont: "\-continuous A (\) B (\) f" + and II: "\i \ A. \ j \ A. i \ j \ \ {i,j} (\)" + shows "monotone_on A (\) (\) f" +proof- + show ?thesis + apply (intro monotone_onI continuous_imp_mono_refl[OF cont]) + using II by auto +qed + +lemma well_complete_imp_monotone_on: + fixes leB (infix "\" 50) + assumes cont: "well_related_set-continuous A (\) B (\) f" + shows "monotone_on A (\) (\) f" + by (auto intro!: continuous_imp_monotone_on cont pair_well_related) + +end + +end \ No newline at end of file diff --git a/thys/Complete_Non_Orders/Directedness.thy b/thys/Complete_Non_Orders/Directedness.thy new file mode 100755 --- /dev/null +++ b/thys/Complete_Non_Orders/Directedness.thy @@ -0,0 +1,261 @@ +theory Directedness + imports Binary_Relations Well_Relations +begin + +text \Directed sets:\ + +locale directed = + fixes A and less_eq (infix "\" 50) + assumes pair_bounded: "x \ A \ y \ A \ \z \ A. x \ z \ y \ z" + +lemmas directedI[intro] = directed.intro + +lemmas directedD = directed_def[unfolded atomize_eq, THEN iffD1, rule_format] + +context + fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma directedE: + assumes "directed A (\)" and "x \ A" and "y \ A" + and "\z. z \ A \ x \ z \ y \ z \ thesis" + shows "thesis" + using assms by (auto dest: directedD) + +lemma directed_empty[simp]: "directed {} (\)" by auto + +lemma directed_union: + assumes dX: "directed X (\)" and dY: "directed Y (\)" + and XY: "\x\X. \y\Y. \z \ X \ Y. x \ z \ y \ z" + shows "directed (X \ Y) (\)" + using directedD[OF dX] directedD[OF dY] XY + apply (intro directedI) by blast + +lemma directed_extend: + assumes X: "directed X (\)" and Y: "directed Y (\)" and XY: "\x\X. \y\Y. x \ y" + shows "directed (X \ Y) (\)" +proof - + { fix x y + assume xX: "x \ X" and yY: "y \ Y" + let ?g = "\z\X \ Y. x \ z \ y \ z" + from directedD[OF Y yY yY] obtain z where zY: "z \ Y" and yz: "y \ z" by auto + from xX XY zY yz have ?g by auto + } + then show ?thesis by (auto intro!: directed_union[OF X Y]) +qed + +end + +sublocale connex \ directed +proof + fix x y + assume x: "x \ A" and y: "y \ A" + then show "\z\A. x \ z \ y \ z" + proof (cases rule: comparable_cases) + case le + with refl[OF y] y show ?thesis by (intro bexI[of _ y], auto) + next + case ge + with refl[OF x] x show ?thesis by (intro bexI[of _ x], auto) + qed +qed + +lemmas(in connex) directed = directed_axioms + +lemma monotone_directed_image: + fixes ir (infix "\" 50) and r (infix "\" 50) + assumes mono: "monotone_on I (\) (\) f" and dir: "directed I (\)" + shows "directed (f ` I) (\)" +proof (rule directedI, safe) + fix x y assume x: "x \ I" and y: "y \ I" + with dir obtain z where z: "z \ I" and "x \ z" and "y \ z" by (auto elim: directedE) + with mono x y have "f x \ f z" and "f y \ f z" by (auto dest: monotone_onD) + with z show "\fz \ f ` I. f x \ fz \ f y \ fz" by auto +qed + + +definition "directed_set A (\) \ \X \ A. finite X \ (\b \ A. bound X (\) b)" + for less_eq (infix "\" 50) + +lemmas directed_setI = directed_set_def[unfolded atomize_eq, THEN iffD2, rule_format] +lemmas directed_setD = directed_set_def[unfolded atomize_eq, THEN iffD1, rule_format] + +lemma directed_imp_nonempty: + fixes less_eq (infix "\" 50) + shows "directed_set A (\) \ A \ {}" + by (auto simp: directed_set_def) + +lemma directedD2: + fixes less_eq (infix "\" 50) + assumes dir: "directed_set A (\)" and xA: "x \ A" and yA: "y \ A" + shows "\z \ A. x \ z \ y \ z" + using directed_setD[OF dir, of "{x,y}"] xA yA by auto + +lemma monotone_directed_set_image: + fixes ir (infix "\" 50) and r (infix "\" 50) + assumes mono: "monotone_on I (\) (\) f" and dir: "directed_set I (\)" + shows "directed_set (f ` I) (\)" +proof (rule directed_setI) + fix X assume "finite X" and "X \ f ` I" + from finite_subset_image[OF this] + obtain J where JI: "J \ I" and Jfin: "finite J" and X: "X = f ` J" by auto + from directed_setD[OF dir JI Jfin] + obtain b where bI: "b \ I" and Jb: "bound J (\) b" by auto + from monotone_image_bound[OF mono JI bI Jb] bI + show "Bex (f ` I) (bound X (\))" by (auto simp: X) +qed + + +lemma directed_set_iff_extremed: + fixes less_eq (infix "\" 50) + assumes Dfin: "finite D" + shows "directed_set D (\) \ extremed D (\)" +proof (intro iffI directed_setI conjI) + assume "directed_set D (\)" + from directed_setD[OF this order.refl Dfin] + show "extremed D (\)" by (auto intro: extremedI) +next + fix X assume XD: "X \ D" and Xfin: "finite X" + assume "extremed D (\)" + then obtain b where "b \ D" and "extreme D (\) b" by (auto elim!: extremedE) + with XD show "\b \ D. bound X (\) b" by auto +qed + +lemma (in transitive) directed_iff_nonempty_pair_bounded: + "directed_set A (\) \ A \ {} \ (\x\A. \y\A. \z\A. x \ z \ y \ z)" + (is "?l \ _ \ ?r") +proof (safe intro!: directed_setI del: notI subset_antisym) + assume dir: ?l + from directed_imp_nonempty[OF dir] show "A \ {}". + fix x y assume "x \ A" "y \ A" + with directed_setD[OF dir, of "{x,y}"] + show "\z\A. x \ z \ y \ z" by auto +next + fix X + assume A0: "A \ {}" and r: ?r + assume "finite X" and "X \ A" + then show "Bex A (bound X (\))" + proof (induct) + case empty + with A0 show ?case by (auto simp: bound_empty ex_in_conv) + next + case (insert x X) + then obtain y where xA: "x \ A" and XA: "X \ A" and yA: "y \ A" and Xy: "bound X (\) y" by auto + from r yA xA obtain z where zA: "z \ A" and xz: "x \ z" and yz: "y \ z" by auto + from bound_trans[OF Xy yz XA yA zA] xz zA + show ?case by auto + qed +qed + +lemma (in transitive) directed_set_iff_nonempty_directed: + "directed_set A (\) \ A \ {} \ directed A (\)" + apply (unfold directed_iff_nonempty_pair_bounded) + by (auto simp: directed_def) + +lemma (in well_related_set) finite_sets_extremed: + assumes fin: "finite X" and X0: "X \ {}" and XA: "X \ A" + shows "extremed X (\)" +proof- + interpret less_eq_asymmetrize. + from fin X0 XA show ?thesis + proof (induct "card X" arbitrary: X) + case 0 + then show ?case by auto + next + case (Suc n) + note XA = \X \ A\ and X0 = \X \ {}\ and Sn = \Suc n = card X\ and finX = \finite X\ + note IH = Suc(1) + from nonempty_imp_ex_extreme[OF XA X0] + obtain l where l: "extreme X (\) l" by auto + from l have lX: "l \ X" by auto + with XA have lA: "l \ A" by auto + from Sn lX have n: "n = card (X-{l})" by auto + show ?case + proof (cases "X - {l} = {}") + case True + with lA lX show ?thesis by (auto intro!: extremedI) + next + case False + from IH[OF n _ this] finX XA + obtain e where e: "extreme (X - {l}) (\) e" by (auto elim!: extremedE) + with l show ?thesis by (auto intro!: extremedI[of _ _ e] extremeI) + qed + qed +qed + +lemma (in well_related_set) directed_set: + assumes A0: "A \ {}" shows "directed_set A (\)" +proof (intro directed_setI) + fix X assume XA: "X \ A" and Xfin: "finite X" + show "Bex A (bound X (\))" + proof (cases "X = {}") + case True + with A0 show ?thesis by (auto simp: bound_empty ex_in_conv) + next + case False + from finite_sets_extremed[OF Xfin this] XA + show ?thesis by (auto elim!: extremedE) + qed +qed + +lemma prod_directed: + fixes leA (infix "\\<^sub>A" 50) and leB (infix "\\<^sub>B" 50) + assumes dir: "directed X (rel_prod (\\<^sub>A) (\\<^sub>B))" + shows "directed (fst ` X) (\\<^sub>A)" and "directed (snd ` X) (\\<^sub>B)" +proof (safe intro!: directedI) + fix a b x y assume "(a,x) \ X" and "(b,y) \ X" + from directedD[OF dir this] + obtain c z where cz: "(c,z) \ X" and ac: "a \\<^sub>A c" and bc: "b \\<^sub>A c" and "x \\<^sub>B z" and "y \\<^sub>B z" by auto + then show "\z\fst ` X. fst (a,x) \\<^sub>A z \ fst (b,y) \\<^sub>A z" + and "\z\snd ` X. snd (a,x) \\<^sub>B z \ snd (b,y) \\<^sub>B z" + by (auto intro!: bexI[OF _ cz]) +qed + +class dir = ord + assumes "directed UNIV (\)" +begin + +sublocale order: directed UNIV "(\)" + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + using dir_axioms[unfolded class.dir_def] + by (auto simp: atomize_eq) + +end + +class filt = ord + + assumes "directed UNIV (\)" +begin + +sublocale order.dual: directed UNIV "(\)" + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + using filt_axioms[unfolded class.filt_def] + by (auto simp: atomize_eq) + +end + +subclass (in linqorder) dir.. + +subclass (in linqorder) filt.. + +thm order.directed_axioms[where 'a = "'a ::dir"] + +thm order.dual.directed_axioms[where 'a = "'a ::filt"] + +end \ No newline at end of file diff --git a/thys/Complete_Non_Orders/Fixed_Points.thy b/thys/Complete_Non_Orders/Fixed_Points.thy --- a/thys/Complete_Non_Orders/Fixed_Points.thy +++ b/thys/Complete_Non_Orders/Fixed_Points.thy @@ -1,1236 +1,1258 @@ (* Author: Jérémy Dubut (2019-2021) Author: Akihisa Yamada (2019-2021) License: LGPL (see file COPYING.LESSER) *) theory Fixed_Points - imports Complete_Relations + imports Complete_Relations Directedness begin section \Existence of Fixed Points in Complete Related Sets\ text \\label{sec:qfp-exists}\ text \The following proof is simplified and generalized from - Stouti--Maaden \<^cite>\"SM13"\. We construct some set whose extreme bounds + Stouti--Maaden \cite{SM13}. We construct some set whose extreme bounds -- if they exist, typically when the underlying related set is complete -- are fixed points of a monotone or inflationary function on any related set. When the related set is attractive, those are actually the least fixed points. - This generalizes \<^cite>\"SM13"\, relaxing reflexivity and antisymmetry.\ + This generalizes \cite{SM13}, relaxing reflexivity and antisymmetry.\ locale fixed_point_proof = related_set + fixes f assumes f: "f ` A \ A" begin -sublocale less_eq_notations. +sublocale less_eq_asymmetrize. definition AA where "AA \ {X. X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X)}" lemma AA_I: "X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X) \ X \ AA" by (unfold AA_def, safe) lemma AA_E: "X \ AA \ (X \ A \ f ` X \ X \ (\Y s. Y \ X \ extreme_bound A (\) Y s \ s \ X) \ thesis) \ thesis" by (auto simp: AA_def) definition C where "C \ \ AA" lemma A_AA: "A \ AA" by (auto intro!:AA_I f) lemma C_AA: "C \ AA" proof (intro AA_I) show "C \ A" using C_def A_AA f by auto show "f ` C \ C" unfolding C_def AA_def by auto fix B b assume B: "B \ C" "extreme_bound A (\) B b" { fix X assume X: "X \ AA" with B have "B \ X" by (auto simp: C_def) with X B have "b\X" by (auto elim!: AA_E) } then show "b \ C" by (auto simp: C_def AA_def) qed lemma CA: "C \ A" using A_AA by (auto simp: C_def) lemma fC: "f ` C \ C" using C_AA by (auto elim!: AA_E) context fixes c assumes Cc: "extreme_bound A (\) C c" begin private lemma cA: "c \ A" using Cc by auto private lemma cC: "c \ C" using Cc C_AA by (blast elim!:AA_E) private lemma fcC: "f c \ C" using cC AA_def C_AA by auto private lemma fcA: "f c \ A" using fcC CA by auto lemma qfp_as_extreme_bound: assumes infl_mono: "\x \ A. x \ f x \ (\y \ A. y \ x \ f y \ f x)" shows "f c \ c" proof (intro conjI bexI sympartpI) show "f c \ c" using fcC Cc by auto from infl_mono[rule_format, OF cA] show "c \ f c" proof (safe) text \Monotone case:\ assume mono: "\b\A. b \ c \ f b \ f c" define D where "D \ {x \ C. x \ f c}" have "D \ AA" proof (intro AA_I) show "D \ A" unfolding D_def C_def using A_AA f by auto have fxC: "x \ C \ x \ f c \ f x \ C" for x using C_AA by (auto simp: AA_def) show "f ` D \ D" proof (unfold D_def, safe intro!: fxC) fix x assume xC: "x \ C" have "x \ c" "x \ A" using Cc xC CA by auto then show "f x \ f c" using mono by (auto dest:monotoneD) qed have DC: "D \ C" unfolding D_def by auto fix B b assume BD: "B \ D" and Bb: "extreme_bound A (\) B b" have "B \ C" using DC BD by auto then have bC: "b \ C" using C_AA Bb BD by (auto elim!: AA_E) have bfc: "\a\B. a \ f c" using BD unfolding D_def by auto with f cA Bb have "b \ f c" by (auto simp: extreme_def image_subset_iff) with bC show "b \ D" unfolding D_def by auto qed then have "C \ D" unfolding C_def by auto then show "c \ f c" using cC unfolding D_def by auto qed qed lemma extreme_qfp: assumes attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "extreme {q \ A. f q \ q \ f q = q} (\) c" proof- have fcc: "f c \ c" apply (rule qfp_as_extreme_bound) using mono by (auto elim!: monotone_onE) define L where [simp]: "L \ {a \ A. \s \ A. (f s \ s \ f s = s) \ a \ s}" have "L \ AA" proof (unfold AA_def, intro CollectI conjI allI impI) show XA: "L \ A" by auto show "f ` L \ L" proof safe fix x assume xL: "x \ L" show "f x \ L" proof (unfold L_def, safe) have xA: "x \ A" using xL by auto then show fxA: "f x \ A" using f by auto { fix s assume sA: "s \ A" and sf: "f s \ s \ f s = s" then have "x \ s" using xL sA sf by auto then have "f x \ f s" using mono fxA sA xA by (auto elim!:monotone_onE)} note fxfs = this { fix s assume sA: "s \ A" and sf: "f s \ s" then show "f x \ s" using fxfs attract mono sf fxA sA xA by (auto elim!:monotone_onE) } { fix s assume sA: "s \ A" and sf: "f s = s" with fxfs[OF sA] show "f x \ s" by simp} qed qed fix B b assume BL: "B \ L" and b: "extreme_bound A (\) B b" then have BA: "B \ A" by auto with BL b have bA: "b \ A" by auto show "b \ L" proof (unfold L_def, safe intro!: bA) { fix s assume sA: "s \ A" and sf: "f s \ s \ f s = s" have "bound B (\) s" using sA BL b sf by auto } note Bs = this { fix s assume sA: "s \ A" and sf: "f s \ s" with b sA Bs show "b \ s" by auto } { fix s assume sA: "s \ A" and sf: "f s = s" with b sA Bs show "b \ s" by auto } qed qed then have "C \ L" by (simp add: C_def Inf_lower) with cC have "c \ L" by auto with L_def fcc show ?thesis by auto qed end lemma ex_qfp: - assumes comp: "CC-complete A (\)" and C: "C \ CC" + assumes comp: "CC-complete A (\)" and C: "CC C (\)" and infl_mono: "\a \ A. a \ f a \ (\b \ A. b \ a \ f b \ f a)" shows "\s \ A. f s \ s" - using qfp_as_extreme_bound[OF _ infl_mono] completeD[OF comp CA C] by auto + using qfp_as_extreme_bound[OF _ infl_mono] completeD[OF comp CA, OF C] by auto lemma ex_extreme_qfp_fp: - assumes comp: "CC-complete A (\)" and C: "C \ CC" + assumes comp: "CC-complete A (\)" and C: "CC C (\)" and attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {q \ A. f q \ q \ f q = q} (\) c" - using extreme_qfp[OF _ attract mono] completeD[OF comp CA C] by auto + using extreme_qfp[OF _ attract mono] completeD[OF comp CA, OF C] by auto lemma ex_extreme_qfp: - assumes comp: "CC-complete A (\)" and C: "C \ CC" + assumes comp: "CC-complete A (\)" and C: "CC C (\)" and attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {q \ A. f q \ q} (\) c" proof- - from completeD[OF comp CA C] + from completeD[OF comp CA, OF C] obtain c where Cc: "extreme_bound A (\) C c" by auto from extreme_qfp[OF Cc attract mono] have Qc: "bound {q \ A. f q \ q} (\) c" by auto have fcc: "f c \ c" apply (rule qfp_as_extreme_bound[OF Cc]) using mono by (auto simp: monotone_onD) from Cc CA have cA: "c \ A" by auto from Qc fcc cA show ?thesis by (auto intro!: exI[of _ c]) qed end context fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) and A :: "'a set" and f assumes f: "f ` A \ A" begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. interpretation fixed_point_proof A "(\)" f using f by unfold_locales theorem complete_infl_mono_imp_ex_qfp: - assumes comp: "UNIV-complete A (\)" and infl_mono: "\a\A. a \ f a \ (\b\A. b \ a \ f b \ f a)" + assumes comp: "\-complete A (\)" and infl_mono: "\a\A. a \ f a \ (\b\A. b \ a \ f b \ f a)" shows "\s\A. f s \ s" apply (rule ex_qfp[OF comp _ infl_mono]) by auto end corollary (in antisymmetric) complete_infl_mono_imp_ex_fp: - assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" + assumes comp: "\-complete A (\)" and f: "f ` A \ A" and infl_mono: "\a\A. a \ f a \ (\b\A. b \ a \ f b \ f a)" shows "\s \ A. f s = s" proof- - interpret less_eq_notations. + interpret less_eq_symmetrize. from complete_infl_mono_imp_ex_qfp[OF f comp infl_mono] obtain s where sA: "s \ A" and fss: "f s \ s" by auto from f sA have fsA: "f s \ A" by auto have "f s = s" using antisym fsA sA fss by auto with sA show ?thesis by auto qed context semiattractive begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. theorem complete_mono_imp_ex_extreme_qfp: - assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" + assumes comp: "\-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "\s. extreme {p \ A. f p \ p} (\) s" proof- interpret dual: fixed_point_proof A "(\)" rewrites "dual.sym = (\)" using f by unfold_locales (auto intro!:ext) show ?thesis apply (rule dual.ex_extreme_qfp[OF complete_dual[OF comp] _ _ monotone_on_dual[OF mono]]) apply simp using f sym_order_trans by blast qed end corollary (in antisymmetric) complete_mono_imp_ex_extreme_fp: - assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" + assumes comp: "\-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" shows "\s. extreme {s \ A. f s = s} (\)\<^sup>- s" proof- - interpret less_eq_notations. + interpret less_eq_symmetrize. interpret fixed_point_proof A "(\)" f using f by unfold_locales have "\c. extreme {q \ A. f q \ q \ f q = q} (\) c" apply (rule ex_extreme_qfp_fp[OF comp _ _ mono]) using antisym f by (auto dest: order_sym_trans) then obtain c where c: "extreme {q \ A. f q \ q \ f q = q} (\) c" by auto then have "f c = c" using antisym f by blast with c have "extreme {q \ A. f q = q} (\) c" by auto then show ?thesis by auto qed section \Fixed Points in Well-Complete Antisymmetric Sets\ text \\label{sec:well-complete}\ text \In this section, we prove that an inflationary or monotone map over a well-complete antisymmetric set has a fixed point. In order to formalize such a theorem in Isabelle, -we followed Grall's~\<^cite>\"grall10"\ elementary proof for Bourbaki--Witt and Markowsky's theorems. +we followed Grall's~\cite{grall10} elementary proof for Bourbaki--Witt and Markowsky's theorems. His idea is to consider well-founded derivation trees over $A$, where from a set $C \subseteq A$ of premises one can derive $f\:(\bigsqcup C)$ if $C$ is a chain. The main observation is as follows: Let $D$ be the set of all the derivable elements; that is, for each $d \in D$ there exists a well-founded derivation whose root is $d$. It is shown that $D$ is a chain, and hence one can build a derivation yielding $f\:(\bigsqcup D)$, and $f\:(\bigsqcup D)$ is shown to be a fixed point.\ lemma bound_monotone_on: assumes mono: "monotone_on A r s f" and XA: "X \ A" and aA: "a \ A" and rXa: "bound X r a" shows "bound (f`X) s (f a)" proof (safe) fix x assume xX: "x \ X" from rXa xX have "r x a" by auto with xX XA mono aA show "s (f x) (f a)" by (auto elim!:monotone_onE) qed context fixed_point_proof begin text \To avoid the usage of the axiom of choice, we carefully define derivations so that any derivable element determines its lower set. This led to the following definition:\ definition "derivation X \ X \ A \ well_ordered_set X (\) \ (\x \ X. let Y = {y \ X. y \ x} in (\y. extreme Y (\) y \ x = f y) \ f ` Y \ Y \ extreme_bound A (\) Y x)" +lemma empty_derivation: "derivation {}" by (auto simp: derivation_def) + lemma assumes "derivation P" shows derivation_A: "P \ A" and derivation_well_ordered: "well_ordered_set P (\)" using assms by (auto simp: derivation_def) lemma derivation_cases[consumes 2, case_names suc lim]: assumes "derivation X" and "x \ X" and "\Y y. Y = {y \ X. y \ x} \ extreme Y (\) y \ x = f y \ thesis" and "\Y. Y = {y \ X. y \ x} \ f ` Y \ Y \ extreme_bound A (\) Y x \ thesis" shows thesis using assms unfolding derivation_def Let_def by auto definition "derivable x \ \X. derivation X \ x \ X" lemma derivableI[intro?]: "derivation X \ x \ X \ derivable x" by (auto simp: derivable_def) lemma derivableE: "derivable x \ (\P. derivation P \ x \ P \ thesis) \ thesis" by (auto simp: derivable_def) lemma derivable_A: "derivable x \ x \ A" by (auto elim: derivableE dest:derivation_A) lemma UN_derivations_eq_derivable: "(\{P. derivation P}) = {x. derivable x}" by (auto simp: derivable_def) -context - assumes ord: "antisymmetric A (\)" +end + +locale fixed_point_proof2 = fixed_point_proof + antisymmetric + + assumes derivation_infl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" + and derivation_f_refl: "\X x. derivation X \ x \ X \ f x \ f x" begin -interpretation antisymmetric using ord. - lemma derivation_lim: assumes P: "derivation P" and fP: "f ` P \ P" and Pp: "extreme_bound A (\) P p" shows "derivation (P \ {p})" proof (cases "p \ P") case True with P show ?thesis by (auto simp: insert_absorb) next case pP: False interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. have PA: "P \ A" using derivation_A[OF P]. from Pp have pA: "p \ A" by auto have bp: "bound P (\) p" using Pp by auto then have pp: "p \ p" using Pp by auto have 1: "y \ P \ {x. (x = p \ x \ P) \ x \ y} = {x \ P. x \ y}" for y - using Pp by (auto dest:extreme_bound_bound) + using Pp by (auto dest!: extreme_bound_imp_bound) { fix x assume xP: "x \ P" and px: "p \ x" from xP Pp have "x \ p" by auto with px have "p = x" using xP PA pA by (auto intro!: antisym) with xP pP have "False" by auto } note 2 = this then have 3: "{x. (x = p \ x \ P) \ x \ p} = P" using Pp by (auto intro!: asympartpI) have wr: "well_ordered_set (P \ {p}) (\)" apply (rule well_order_extend[OF P.well_ordered_set_axioms]) using pp bp pP 2 by auto from P fP Pp show "derivation (P \ {p})" by (auto simp: derivation_def pA wr[simplified] 1 3) qed -context - assumes derivation_infl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" - and derivation_f_refl: "\X x. derivation X \ x \ X \ f x \ f x" -begin - lemma derivation_suc: assumes P: "derivation P" and Pp: "extreme P (\) p" shows "derivation (P \ {f p})" proof (cases "f p \ P") case True with P show ?thesis by (auto simp: insert_absorb) next case fpP: False interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. have PA: "P \ A" using derivation_A[OF P]. with Pp have pP: "p \ P" and pA: "p \ A" by auto with f have fpA: "f p \ A" by auto from pP have pp: "p \ p" by auto from derivation_infl[rule_format, OF P pP pP pp] have "p \ f p". { fix x assume xP: "x \ P" then have xA: "x \ A" using PA by auto have xp: "x \ p" using xP Pp by auto from derivation_infl[rule_format, OF P xP pP this] have "x \ f p". } note Pfp = this then have bfp: "bound P (\) (f p)" by auto { fix y assume yP: "y \ P" note yfp = Pfp[OF yP] { assume fpy: "f p \ y" with yfp have "f p = y" using yP PA pA fpA antisym by auto with yP fpP have "False" by auto } with Pfp yP have "y \ f p" by auto } note Pfp = this have 1: "\y. y \ P \ {x. (x = f p \ x \ P) \ x \ y} = {x \ P. x \ y}" and 2: "{x. (x = f p \ x \ P) \ x \ f p} = P" using Pfp by auto have wr: "well_ordered_set (P \ {f p}) (\)" apply (rule well_order_extend[OF P.well_ordered_set_axioms singleton_well_ordered]) using Pfp derivation_f_refl[rule_format, OF P pP] by auto from P Pp show "derivation (P \ {f p})" by (auto simp: derivation_def wr[simplified] 1 2 fpA) qed lemma derivable_closed: assumes x: "derivable x" shows "derivable (f x)" proof (insert x, elim derivableE) fix P assume P: "derivation P" and xP: "x \ P" note PA = derivation_A[OF P] then have xA: "x \ A" using xP by auto interpret P: well_ordered_set P "(\)" using derivation_well_ordered[OF P]. interpret P.asympartp: transitive P "(\)" using P.asympartp_transitive. define Px where "Px \ {y. y \ P \ y \ x} \ {x}" then have PxP: "Px \ P" using xP by auto have "x \ x" using xP by auto then have Pxx: "extreme Px (\) x" using xP PA by (auto simp: Px_def) have wr: "well_ordered_set Px (\)" using P.well_ordered_subset[OF PxP]. { fix z y assume zPx: "z \ Px" and yP: "y \ P" and yz: "y \ z" then have zP: "z \ P" using PxP by auto have "y \ x" proof (cases "z = x") case True then show ?thesis using yz by auto next case False then have zx: "z \ x" using zPx by (auto simp: Px_def) - from P.asympartp.trans[OF yz zx yP zP xP] show ?thesis. + from P.asym.trans[OF yz zx yP zP xP] show ?thesis. qed } then have 1: "\z. z \ Px \ {y \ Px. y \ z} = {y \ P. y \ z}" using Px_def by blast have Px: "derivation Px" using PxP PA P by (auto simp: wr derivation_def 1) from derivation_suc[OF Px Pxx] show ?thesis by (auto intro!: derivableI) qed -text \The following lemma is derived from Grall's proof. We simplify the claim so that we +text \The following lemma is derived from Grall’s proof. We simplify the claim so that we consider two elements from one derivation, instead of two derivations.\ lemma derivation_useful: assumes X: "derivation X" and xX: "x \ X" and yX: "y \ X" and xy: "x \ y" shows "f x \ y" proof- interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. note XA = derivation_A[OF X] { fix x y assume xX: "x \ X" and yX: "y \ X" from xX yX have "(x \ y \ f x \ y \ f x \ X) \ (y \ x \ f y \ x \ f y \ X)" proof (induct x arbitrary: y) case (less x) note xX = \x \ X\ and IHx = this(2) with XA have xA: "x \ A" by auto from \y \ X\ show ?case proof (induct y) case (less y) note yX = \y \ X\ and IHy = this(2) with XA have yA: "y \ A" by auto show ?case proof (rule conjI; intro impI) assume xy: "x \ y" from X yX show "f x \ y \ f x \ X" proof (cases rule:derivation_cases) case (suc Z z) with XA have zX: "z \ X" and zA: "z \ A" and zy: "z \ y" and yfz: "y = f z" by auto from xX zX show ?thesis proof (cases rule: X.comparable_three_cases) case xz: less with IHy[OF zX zy] have fxz: "f x \ z" and fxX: "f x \ X" by auto from derivation_infl[rule_format, OF X fxX zX fxz] have "f x \ y" by (auto simp: yfz) with fxX show ?thesis by auto next case eq with xX zX have "x = z" by auto with yX yfz show ?thesis by auto next case zx: greater with IHy[OF zX zy] yfz xy have False by auto then show ?thesis by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ y}\ and fZ = \f ` Z \ Z\ from xX xy have "x \ Z" by (auto simp: Z) with fZ have "f x \ Z" by auto then have "f x \ y" and "f x \ X" by (auto simp: Z) then show ?thesis by auto qed next assume yx: "y \ x" from X xX show "f y \ x \ f y \ X" proof (cases rule:derivation_cases) case (suc Z z) with XA have zX: "z \ X" and zA: "z \ A" and zx: "z \ x" and xfz: "x = f z" by auto from yX zX show ?thesis proof (cases rule: X.comparable_three_cases) case yz: less with IHx[OF zX zx yX] have fyz: "f y \ z" and fyX: "f y \ X" by auto from derivation_infl[rule_format, OF X fyX zX fyz] have "f y \ x" by (auto simp: xfz) with fyX show ?thesis by auto next case eq with yX zX have "y = z" by auto with xX xfz show ?thesis by auto next case greater with IHx[OF zX zx yX] xfz yx have False by auto then show ?thesis by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ x}\ and fZ = \f ` Z \ Z\ from yX yx have "y \ Z" by (auto simp: Z) with fZ have "f y \ Z" by auto then have "f y \ x" and "f y \ X" by (auto simp: Z) then show ?thesis by auto qed qed qed qed } with assms show "f x \ y" by auto qed text \Next one is the main lemma of this section, stating that elements from two possibly different derivations are comparable, and moreover the lower one is in the derivation of the upper one. -The latter claim, not found in Grall's proof, is crucial in proving that the union of all +The latter claim, not found in Grall’s proof, is crucial in proving that the union of all derivations is well-related.\ lemma derivations_cross_compare: assumes X: "derivation X" and Y: "derivation Y" and xX: "x \ X" and yY: "y \ Y" shows "(x \ y \ x \ Y) \ x = y \ (y \ x \ y \ X)" proof- { fix X Y x y assume X: "derivation X" and Y: "derivation Y" and xX: "x \ X" and yY: "y \ Y" interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. interpret X.asympartp: transitive X "(\)" using X.asympartp_transitive. interpret Y: well_ordered_set Y "(\)" using derivation_well_ordered[OF Y]. have XA: "X \ A" using derivation_A[OF X]. then have xA: "x \ A" using xX by auto with f have fxA: "f x \ A" by auto have YA: "Y \ A" using derivation_A[OF Y]. then have yA: "y \ A" using yY by auto with f have fyA: "f y \ A" by auto { fix Z assume Z: "Z = {z \ X. z \ x}" and fZ: "f ` Z \ Z" and Zx: "extreme_bound A (\) Z x" and IHx: "\z \ X. z \ x \ (z \ y \ z \ Y) \ z = y \ (y \ z \ y \ X)" have "(y \ x \ y \ X) \ x \ y" proof (cases "\z \ Z. y \ z") case True then obtain z where zZ: "z \ Z" and yz: "y \ z" by auto from zZ Z have zX: "z \ X" and zx: "z \ x" by auto from IHx[rule_format, OF zX zx] yz have yX: "y \ X" by auto - from X.asympartp.trans[OF yz zx yX zX xX] have "y \ x". + from X.asym.trans[OF yz zx yX zX xX] have "y \ x". with yX show ?thesis by auto next case False have "bound Z (\) y" proof fix z assume "z \ Z" then have zX: "z \ X" and zx: "z \ x" and nyz: "\ y \ z" using Z False by auto with IHx[rule_format, OF zX zx] X show "z \ y" by auto qed with yA Zx have xy: "x \ y" by auto then show ?thesis by auto qed } note lim_any = this { fix z Z assume Z: "Z = {z \ X. z \ x}" and Zz: "extreme Z (\) z" and xfz: "x = f z" and IHx: "(z \ y \ z \ Y) \ z = y \ (y \ z \ y \ X)" have zX: "z \ X" and zx: "z \ x" using Zz Z by (auto simp: extreme_def) then have zA: "z \ A" using XA by auto from IHx have "(y \ x \ y \ X) \ x \ y" proof (elim disjE conjE) assume zy: "z \ y" and zY: "z \ Y" from derivation_useful[OF Y zY yY zy] xfz have xy: "x \ y" by auto then show ?thesis by auto next assume zy: "z = y" then have "y \ x" using zx by auto with zy zX show ?thesis by auto next assume yz: "y \ z" and yX: "y \ X" - from X.asympartp.trans[OF yz zx yX zX xX] have "y \ x". + from X.asym.trans[OF yz zx yX zX xX] have "y \ x". with yX show ?thesis by auto qed } note lim_any this } note lim_any = this(1) and suc_any = this(2) interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. interpret Y: well_ordered_set Y "(\)" using derivation_well_ordered[OF Y]. have XA: "X \ A" using derivation_A[OF X]. have YA: "Y \ A" using derivation_A[OF Y]. from xX yY show ?thesis proof (induct x arbitrary: y) case (less x) note xX = \x \ X\ and IHx = this(2) from xX XA f have xA: "x \ A" and fxA: "f x \ A" by auto from \y \ Y\ show ?case proof (induct y) case (less y) note yY = \y \ Y\ and IHy = less(2) from yY YA f have yA: "y \ A" and fyA: "f y \ A" by auto from X xX show ?case proof (cases rule: derivation_cases) case (suc Z z) note Z = \Z = {z \ X. z \ x}\ and Zz = \extreme Z (\) z\ and xfz = \x = f z\ then have zx: "z \ x" and zX: "z \ X" by auto note IHz = IHx[OF zX zx yY] have 1: "y \ x \ y \ X \ x \ y" using suc_any[OF X Y xX yY Z Zz xfz IHz] IHy by auto from Y yY show ?thesis proof (cases rule: derivation_cases) case (suc W w) note W = \W = {w \ Y. w \ y}\ and Ww = \extreme W (\) w\ and yfw = \y = f w\ then have wY: "w \ Y" and wy: "w \ y" by auto have IHw: "w \ x \ w \ X \ w = x \ x \ w \ x \ Y" using IHy[OF wY wy] by auto have "x \ y \ x \ Y \ y \ x" using suc_any[OF Y X yY xX W Ww yfw IHw] by auto with 1 show ?thesis using antisym xA yA by auto next case (lim W) note W = \W = {w \ Y. w \ y}\ and fW = \f ` W \ W\ and Wy = \extreme_bound A (\) W y\ have "x \ y \ x \ Y \ y \ x" using lim_any[OF Y X yY xX W fW Wy] IHy by auto with 1 show ?thesis using antisym xA yA by auto qed next case (lim Z) note Z = \Z = {z \ X. z \ x}\ and fZ = \f ` Z \ Z\ and Zx = \extreme_bound A (\) Z x\ have 1: "y \ x \ y \ X \ x \ y" using lim_any[OF X Y xX yY Z fZ Zx] IHx[OF _ _ yY] by auto from Y yY show ?thesis proof (cases rule: derivation_cases) case (suc W w) note W = \W = {w \ Y. w \ y}\ and Ww = \extreme W (\) w\ and yfw = \y = f w\ then have wY: "w \ Y" and wy: "w \ y" by auto have IHw: "w \ x \ w \ X \ w = x \ x \ w \ x \ Y" using IHy[OF wY wy] by auto have "x \ y \ x \ Y \ y \ x" using suc_any[OF Y X yY xX W Ww yfw IHw] by auto with 1 show ?thesis using antisym xA yA by auto next case (lim W) note W = \W = {w \ Y. w \ y}\ and fW = \f ` W \ W\ and Wy = \extreme_bound A (\) W y\ have "x \ y \ x \ Y \ y \ x" using lim_any[OF Y X yY xX W fW Wy] IHy by auto with 1 show ?thesis using antisym xA yA by auto qed qed qed qed qed -interpretation derivable: well_ordered_set "{x. derivable x}" "(\)" +sublocale derivable: well_ordered_set "{x. derivable x}" "(\)" proof (rule well_ordered_set.intro) show "antisymmetric {x. derivable x} (\)" apply unfold_locales by (auto dest: derivable_A antisym) show "well_related_set {x. derivable x} (\)" apply (fold UN_derivations_eq_derivable) apply (rule closed_UN_well_related) by (auto dest: derivation_well_ordered derivations_cross_compare well_ordered_set.axioms) qed -lemmas derivable_well_ordered = derivable.well_ordered_set_axioms -lemmas derivable_total_ordered = derivable.total_ordered_set_axioms -lemmas derivable_well_related = derivable.well_related_set_axioms - lemma pred_unique: assumes X: "derivation X" and xX: "x \ X" shows "{z \ X. z \ x} = {z. derivable z \ z \ x}" proof { fix z assume "z \ X" and "z \ x" then have "derivable z \ z \ x" using X by (auto simp: derivable_def) } then show "{z \ X. z \ x} \ {z. derivable z \ z \ x}" by auto { fix z assume "derivable z" and zx: "z \ x" then obtain Y where Y: "derivation Y" and zY: "z \ Y" by (auto simp: derivable_def) then have "z \ X" using derivations_cross_compare[OF X Y xX zY] zx by auto } then show "{z \ X. z \ x} \ {z. derivable z \ z \ x}" by auto qed text \The set of all derivable elements is itself a derivation.\ lemma derivation_derivable: "derivation {x. derivable x}" apply (unfold derivation_def) apply (safe intro!: derivable_A derivable.well_ordered_set_axioms elim!: derivableE) apply (unfold mem_Collect_eq pred_unique[symmetric]) by (auto simp: derivation_def) text \Finally, if the set of all derivable elements admits a supremum, then it is a fixed point.\ -lemma +context + fixes p assumes p: "extreme_bound A (\) {x. derivable x} p" - shows sup_derivable_qfp: "f p \ p" and sup_derivable_fp: "f p = p" -proof (intro antisym sympartpI) - define P where "P \ {x. derivable x}" - have pA: "p \ A" using p by auto - have P: "derivation P" using derivation_derivable by (simp add: P_def) - from derivable_closed have fP: "f ` P \ P" by (auto simp: P_def) - from derivation_lim[OF P fP] p - have pP: "p \ P" by (auto simp: P_def intro:derivableI) - with fP have "f p \ P" by auto - with p show fpp: "f p \ p" by (auto simp: P_def) - show pfp: "p \ f p" apply (rule derivation_infl[rule_format, OF P]) using pP by (auto simp: P_def) - from fpp pfp p f show "f p = p" by (auto intro!: antisym) -qed +begin + +lemma sup_derivable_derivable: "derivable p" + using derivation_lim[OF derivation_derivable _ p] derivable_closed + by (auto intro: derivableI) + +private lemmas sucp = sup_derivable_derivable[THEN derivable_closed] + +lemma sup_derivable_prefixed: "f p \ p" using sucp p by auto + +lemma sup_derivable_postfixed: "p \ f p" + apply (rule derivation_infl[rule_format, OF derivation_derivable]) + using sup_derivable_derivable by auto + +lemma sup_derivable_qfp: "f p \ p" + using sup_derivable_prefixed sup_derivable_postfixed by auto + +lemma sup_derivable_fp: "f p = p" + using sup_derivable_derivable sucp + by (auto intro!: antisym sup_derivable_prefixed sup_derivable_postfixed simp: derivable_A) + +end end text "The assumptions are satisfied by monotone functions." +context fixed_point_proof begin + +context + assumes ord: "antisymmetric A (\)" +begin + +interpretation antisymmetric using ord. + context assumes mono: "monotone_on A (\) (\) f" begin -lemma mono_imp_derivation_infl: - "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" -proof (intro allI impI) - fix X x y - assume X: "derivation X" and xX: "x \ X" and yX: "y \ X" and xy: "x \ y" - interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. - note XA = derivation_A[OF X] - from xX yX xy show "x \ f y" - proof (induct x) - case (less x) - note IH = this(2) and xX = \x \ X\ and yX = \y \ X\ and xy = \x \ y\ - from xX yX XA have xA: "x \ A" and yA: "y \ A" by auto - from X xX show ?case - proof (cases rule: derivation_cases) - case (suc Z z) - then have zX: "z \ X" and zsx: "z \ x" and xfz: "x = f z" by auto - then have zx: "z \ x" by auto - from X.trans[OF zx xy zX xX yX] have zy: "z \ y". - from zX XA have zA: "z \ A" by auto - from zy monotone_onD[OF mono] zA yA xfz show "x \ f y" by auto - next - case (lim Z) - note Z = \Z = {z \ X. z \ x}\ and Zx = \extreme_bound A (\) Z x\ - from f yA have fyA: "f y \ A" by auto - have "bound Z (\) (f y)" - proof - fix z assume zZ: "z \ Z" - with Z xX have zsx: "z \ x" and zX: "z \ X" by auto +interpretation fixed_point_proof2 +proof + show mono_imp_derivation_infl: + "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" + proof (intro allI impI) + fix X x y + assume X: "derivation X" and xX: "x \ X" and yX: "y \ X" and xy: "x \ y" + interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. + note XA = derivation_A[OF X] + from xX yX xy show "x \ f y" + proof (induct x) + case (less x) + note IH = this(2) and xX = \x \ X\ and yX = \y \ X\ and xy = \x \ y\ + from xX yX XA have xA: "x \ A" and yA: "y \ A" by auto + from X xX show ?case + proof (cases rule: derivation_cases) + case (suc Z z) + then have zX: "z \ X" and zsx: "z \ x" and xfz: "x = f z" by auto then have zx: "z \ x" by auto from X.trans[OF zx xy zX xX yX] have zy: "z \ y". - from IH[OF zX zsx yX] zy show "z \ f y" by auto + from zX XA have zA: "z \ A" by auto + from zy monotone_onD[OF mono] zA yA xfz show "x \ f y" by auto + next + case (lim Z) + note Z = \Z = {z \ X. z \ x}\ and Zx = \extreme_bound A (\) Z x\ + from f yA have fyA: "f y \ A" by auto + have "bound Z (\) (f y)" + proof + fix z assume zZ: "z \ Z" + with Z xX have zsx: "z \ x" and zX: "z \ X" by auto + then have zx: "z \ x" by auto + from X.trans[OF zx xy zX xX yX] have zy: "z \ y". + from IH[OF zX zsx yX] zy show "z \ f y" by auto + qed + with Zx fyA show ?thesis by auto qed - with Zx fyA show ?thesis by auto qed qed + show mono_imp_derivation_f_refl: + "\X x. derivation X \ x \ X \ f x \ f x" + proof (intro allI impI) + fix X x + assume X: "derivation X" and xX: "x \ X" + interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. + note XA = derivation_A[OF X] + from monotone_onD[OF mono] xX XA show "f x \ f x" by auto + qed qed -lemma mono_imp_derivation_f_refl: - "\X x. derivation X \ x \ X \ f x \ f x" -proof (intro allI impI) - fix X x - assume X: "derivation X" and xX: "x \ X" - interpret X: well_ordered_set X "(\)" using derivation_well_ordered[OF X]. - note XA = derivation_A[OF X] - from xX have "x \ x" by auto - from monotone_onD[OF mono this] xX XA show "f x \ f x" by auto -qed +lemmas mono_imp_fixed_point_proof2 = fixed_point_proof2_axioms corollary mono_imp_sup_derivable_fp: assumes p: "extreme_bound A (\) {x. derivable x} p" shows "f p = p" - by (simp add: sup_derivable_fp[OF mono_imp_derivation_infl mono_imp_derivation_f_refl p]) + by (simp add: sup_derivable_fp[OF p]) lemma mono_imp_sup_derivable_lfp: assumes p: "extreme_bound A (\) {x. derivable x} p" shows "extreme {q \ A. f q = q} (\) p" proof (safe intro!: extremeI) from p show "p \ A" by auto - from sup_derivable_fp[OF mono_imp_derivation_infl mono_imp_derivation_f_refl p] + from sup_derivable_fp[OF p] show "f p = p". fix q assume qA: "q \ A" and fqq: "f q = q" have "bound {x. derivable x} (\) q" proof (safe intro!: boundI elim!:derivableE) fix x X assume X: "derivation X" and xX: "x \ X" from X interpret well_ordered_set X "(\)" by (rule derivation_well_ordered) from xX show "x \ q" proof (induct x) case (less x) note xP = this(1) and IH = this(2) with X show ?case proof (cases rule: derivation_cases) case (suc Z z) with IH[of z] have zq: "z \ q" and zX: "z \ X" by auto - from monotone_onD[OF mono zq] zX qA derivation_A[OF X] + from monotone_onD[OF mono _ qA zq] zX derivation_A[OF X] show ?thesis by (auto simp: fqq suc) next case lim with IH have "bound {z \ X. z \ x} (\) q" by auto with lim qA show ?thesis by auto qed qed qed with p qA show "p \ q" by auto qed lemma mono_imp_ex_least_fp: - assumes comp: "well_complete A (\)" + assumes comp: "well_related_set-complete A (\)" shows "\p. extreme {q \ A. f q = q} (\) p" proof- interpret fixed_point_proof using f by unfold_locales - note as = antisymmetric_axioms - note infl = mono_imp_derivation_infl - note refl = mono_imp_derivation_f_refl - have wr: "well_related_set {x. derivable x} (\)" - using derivable_well_related[OF infl refl]. have "\p. extreme_bound A (\) {x. derivable x} p" apply (rule completeD[OF comp]) - using derivable_A wr by auto + using derivable_A derivable.well_related_set_axioms by auto then obtain p where p: "extreme_bound A (\) {x. derivable x} p" by auto - from p mono_imp_sup_derivable_lfp[OF p] sup_derivable_qfp[OF infl refl p] + from p mono_imp_sup_derivable_lfp[OF p] sup_derivable_qfp[OF p] show ?thesis by auto qed end end end text \Bourbaki-Witt Theorem on well-complete pseudo-ordered set:\ -theorem (in pseudo_ordered_set) well_complete_infl_imp_ex_fp: - assumes comp: "well_complete A (\)" +theorem (in pseudo_ordered_set) well_complete_infl'_imp_ex_fp: + assumes comp: "well_related_set-complete A (\)" and f: "f ` A \ A" and infl: "\x \ A. \y \ A. x \ y \ x \ f y" shows "\p \ A. f p = p" proof- - note as = antisymmetric_axioms interpret fixed_point_proof using f by unfold_locales - have dinfl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" - using infl by (auto dest!:derivation_A) - have drefl: "\X x. derivation X \ x \ X \ f x \ f x" using f by (auto dest!: derivation_A) + interpret fixed_point_proof2 + proof + show dinfl: "\X x y. derivation X \ x \ X \ y \ X \ x \ y \ x \ f y" + using infl by (auto dest!:derivation_A) + show drefl: "\X x. derivation X \ x \ X \ f x \ f x" + using f by (auto dest!: derivation_A) + qed have "\p. extreme_bound A (\) {x. derivable x} p" apply (rule completeD[OF comp]) - using derivable_well_related[OF as dinfl drefl] derivable_A by auto - with sup_derivable_fp[OF as dinfl drefl] + using derivable.well_related_set_axioms derivable_A by auto + with sup_derivable_fp show ?thesis by auto qed +text \Bourbaki-Witt Theorem on posets:\ +corollary (in partially_ordered_set) well_complete_infl_imp_ex_fp: + assumes comp: "well_related_set-complete A (\)" + and f: "f ` A \ A" and infl: "\x \ A. x \ f x" + shows "\p \ A. f p = p" +proof (intro well_complete_infl'_imp_ex_fp[OF comp f] ballI impI) + fix x y assume x: "x \ A" and y: "y \ A" and xy: "x \ y" + from y infl have "y \ f y" by auto + from trans[OF xy this x y] f y show "x \ f y" by auto +qed section \Completeness of (Quasi-)Fixed Points\ text \We now prove that, under attractivity, the set of quasi-fixed points is complete.\ definition setwise where "setwise r X Y \ \x\X. \y\Y. r x y" lemmas setwiseI[intro] = setwise_def[unfolded atomize_eq, THEN iffD2, rule_format] lemmas setwiseE[elim] = setwise_def[unfolded atomize_eq, THEN iffD1, elim_format, rule_format] context fixed_point_proof begin abbreviation setwise_less_eq (infix "\\<^sup>s" 50) where "(\\<^sup>s) \ setwise (\)" subsection \Least Quasi-Fixed Points for Attractive Relations.\ lemma attract_mono_imp_least_qfp: assumes attract: "attractive A (\)" - and comp: "well_complete A (\)" + and comp: "well_related_set-complete A (\)" and mono: "monotone_on A (\) (\) f" shows "\c. extreme {p \ A. f p \ p \ f p = p} (\) c \ f c \ c" proof- interpret attractive using attract by auto interpret sym: transitive A "(\)" using sym_transitive. define ecl ("[_]\<^sub>\") where "[x]\<^sub>\ \ {y \ A. x \ y} \ {x}" for x define Q where "Q \ {[x]\<^sub>\ |. x \ A}" { fix X x assume XQ: "X \ Q" and xX: "x \ X" then have XA: "X \ A" by (auto simp: Q_def ecl_def) then have xA: "x \ A" using xX by auto obtain q where qA: "q \ A" and X: "X = [q]\<^sub>\" using XQ by (auto simp: Q_def) have xqqx: "x \ q \ x = q" using X xX by (auto simp: ecl_def) {fix y assume yX: "y \ X" then have yA: "y \ A" using XA by auto have "y \ q \ y = q" using yX X by (auto simp: ecl_def) then have "x \ y \ y = x" using sym_order_trans xqqx xA qA yA by blast } then have 1: "X \ [x]\<^sub>\" using X qA by (auto simp: ecl_def) { fix y assume "y \ A" and "x \ y \ y = x" then have "q \ y \ y = q" using sym_order_trans xqqx xA qA by blast } then have 2: "X \ [x]\<^sub>\" using X xX by (auto simp: ecl_def) from 1 2 have "X = [x]\<^sub>\" by auto } then have XQx: "\X \ Q. \x \ X. X = [x]\<^sub>\" by auto have RSLE_eq: "X \ Q \ Y \ Q \ x \ X \ y \ Y \ x \ y \ X \\<^sup>s Y" for X Y x y proof- assume XQ: "X \ Q" and YQ: "Y \ Q" and xX: "x \ X" and yY: "y \ Y" and xy: "x \ y" then have XA: "X \ A" and YA: "Y \ A" by (auto simp: Q_def ecl_def) then have xA: "x \ A" and yA: "y \ A" using xX yY by auto { fix xp yp assume xpX: "xp \ X" and ypY: "yp \ Y" then have xpA: "xp \ A" and ypA: "yp \ A" using XA YA by auto then have "xp \ x \ xp = x" using xpX XQx xX XQ by (auto simp: ecl_def) then have xpy: "xp \ y" using attract[OF _ _ xy xpA xA yA] xy by blast have "yp \ y \ yp = y" using ypY XQx yY YQ by (auto simp: ecl_def) then have "xp \ yp" using dual.attract[OF _ _ xpy ypA yA xpA] xpy by blast } then show "X \\<^sup>s Y" using XQ YQ XA YA by auto qed - have compQ: "well_complete Q (\\<^sup>s)" - proof (intro completeI, safe) + have compQ: "well_related_set-complete Q (\\<^sup>s)" + proof (intro completeI) fix XX assume XXQ: "XX \ Q" and XX: "well_related_set XX (\\<^sup>s)" have BA: "\XX \ A" using XXQ by (auto simp: Q_def ecl_def) from XX interpret XX: well_related_set XX "(\\<^sup>s)". interpret UXX: semiattractive "\XX" "(\)" by (rule semiattractive_subset[OF BA]) have "well_related_set (\XX) (\)" proof(unfold_locales) fix Y assume YXX: "Y \ \XX" and Y0: "Y \ {}" have "{X \ XX. X \ Y \ {}} \ {}" using YXX Y0 by auto from XX.nonempty_imp_ex_extreme[OF _ this] obtain E where E: "extreme {X \ XX. X \ Y \ {}} (\\<^sup>s)\<^sup>- E" by auto then have "E \ Y \ {}" by auto then obtain e where eE: "e \ E" and eX: "e \ Y" by auto have "extreme Y (\) e" proof (intro extremeI eX) fix x assume xY: "x \ Y" with YXX obtain X where XXX: "X \ XX" and xX: "x \ X" by auto with xY E XXX have "E \\<^sup>s X" by auto with eE xX show "e \ x" by auto qed then show "\e. extreme Y (\) e" by auto qed with completeD[OF comp BA] obtain b where extb: "extreme_bound A (\) (\XX) b" by auto then have bb: "b \ b" using extreme_def bound_def by auto have bA: "b \ A" using extb extreme_def by auto then have XQ: "[b]\<^sub>\ \ Q" using Q_def bA by auto have bX: "b \ [b]\<^sub>\" by (auto simp: ecl_def) have "extreme_bound Q (\\<^sup>s) XX [b]\<^sub>\" proof(intro extreme_boundI) show "[b]\<^sub>\ \ Q" using XQ. next fix Y assume YXX: "Y \ XX" then have YQ: "Y \ Q" using XXQ by auto then obtain y where yA: "y \ A" and Yy: "Y = [y]\<^sub>\" by (auto simp: Q_def) then have yY: "y \ Y" by (auto simp: ecl_def) then have "y \ \XX" using yY YXX by auto then have "y \ b" using extb by auto then show "Y \\<^sup>s [b]\<^sub>\" using RSLE_eq[OF YQ XQ yY bX] by auto next fix Z assume boundZ: "bound XX (\\<^sup>s) Z" and ZQ: "Z \ Q" then obtain z where zA: "z \ A" and Zz: "Z = [z]\<^sub>\" by (auto simp: Q_def) then have zZ: "z \ Z" by (auto simp: ecl_def) { fix y assume "y \ \XX" then obtain Y where yY: "y \ Y" and YXX: "Y \ XX" by auto then have YA: "Y \ A" using XXQ Q_def by (auto simp: ecl_def) then have "Y \\<^sup>s Z" using YXX boundZ bound_def by auto then have "y \ z" using yY zZ by auto } then have "bound (\XX) (\) z" by auto then have "b \ z" using extb zA by auto then show "[b]\<^sub>\ \\<^sup>s Z" using RSLE_eq[OF XQ ZQ bX zZ] by auto qed then show "Ex (extreme_bound Q (\\<^sup>s) XX)" by auto qed interpret Q: antisymmetric Q "(\\<^sup>s)" proof fix X Y assume XY: "X \\<^sup>s Y" and YX: "Y \\<^sup>s X" and XQ: "X \ Q" and YQ: "Y \ Q" then obtain q where qA: "q \ A" and X: "X = [q]\<^sub>\" using Q_def by auto then have qX: "q \ X" using X by (auto simp: ecl_def) then obtain p where pA: "p \ A" and Y: "Y = [p]\<^sub>\" using YQ Q_def by auto then have pY: "p \ Y" using X by (auto simp: ecl_def) have pq: "p \ q" using XQ YQ YX qX pY by auto have "q \ p" using XQ YQ XY qX pY by auto then have "p \ X" using pq X pA by (auto simp: ecl_def) then have "X = [p]\<^sub>\" using XQ XQx by auto then show "X = Y" using Y by (auto simp: ecl_def) qed define F where "F X \ {y \ A. \x \ X. y \ f x} \ f ` X" for X have XQFXQ: "\X. X \ Q \ F X \ Q" proof- fix X assume XQ: "X \ Q" then obtain x where xA: "x \ A" and X: "X = [x]\<^sub>\" using Q_def by auto then have xX: "x \ X" by (auto simp: ecl_def) have fxA: "f x \ A" using xA f by auto have FXA: "F X \ A" using f fxA X by (auto simp: F_def ecl_def) have "F X = [f x]\<^sub>\" proof (unfold X, intro equalityI subsetI) fix z assume zFX: "z \ F [x]\<^sub>\" then obtain y where yX: "y \ [x]\<^sub>\" and zfy: "z \ f y \ z = f y" by (auto simp: F_def) have yA: "y \ A" using yX xA by (auto simp: ecl_def) with f have fyA: "f y \ A" by auto have zA: "z \ A" using zFX FXA by (auto simp: X) have "y \ x \ y = x" using X yX by (auto simp: ecl_def) then have "f y \ f x \ f y = f x" using mono xA yA by (auto simp: monotone_on_def) then have "z \ f x \ z = f x" using zfy sym.trans[OF _ _ zA fyA fxA] by (auto simp:) with zA show "z \ [f x]\<^sub>\" by (auto simp: ecl_def) qed (auto simp: xX F_def ecl_def) with FXA show "F X \ Q" by (auto simp: Q_def ecl_def) qed then have F: "F ` Q \ Q" by auto then interpret Q: fixed_point_proof Q "(\\<^sup>s)" F by unfold_locales have monoQ: "monotone_on Q (\\<^sup>s) (\\<^sup>s) F" proof (intro monotone_onI) fix X Y assume XQ: "X \ Q" and YQ: "Y \ Q" and XY: "X \\<^sup>s Y" then obtain x y where xX: "x \ X" and yY: "y \ Y" using Q_def by (auto simp: ecl_def) then have xA: "x \ A" and yA: "y \ A" using XQ YQ by (auto simp: Q_def ecl_def) have "x \ y" using XY xX yY by auto then have fxfy: "f x \ f y" using monotone_on_def[of A "(\)" "(\)" f] xA yA mono by auto have fxgX: "f x \ F X" using xX F_def by blast have fygY: "f y \ F Y" using yY F_def by blast show "F X \\<^sup>s F Y" using RSLE_eq[OF XQFXQ[OF XQ] XQFXQ[OF YQ] fxgX fygY fxfy]. qed have QdA: "{x. Q.derivable x} \ Q" using Q.derivable_A by auto - note asQ = Q.antisymmetric_axioms - note infl = Q.mono_imp_derivation_infl[OF asQ monoQ] - note f_refl = Q.mono_imp_derivation_f_refl[OF asQ monoQ] - from Q.mono_imp_ex_least_fp[OF asQ monoQ compQ] + interpret Q: fixed_point_proof2 Q "(\\<^sup>s)" F + using Q.mono_imp_fixed_point_proof2[OF Q.antisymmetric_axioms monoQ]. + from Q.mono_imp_ex_least_fp[OF Q.antisymmetric_axioms monoQ compQ] obtain P where P: "extreme {q \ Q. F q = q} (\\<^sup>s)\<^sup>- P" by auto then have PQ: "P \ Q" by (auto simp: extreme_def) from P have FPP: "F P = P" using PQ by auto with P have PP: "P \\<^sup>s P" by auto from P obtain p where pA: "p \ A" and Pp: "P = [p]\<^sub>\" using Q_def by auto then have pP: "p \ P" by (auto simp: ecl_def) then have fpA: "f p \ A" using pA f by auto have "f p \ F P" using pP F_def fpA by auto then have "F P = [f p]\<^sub>\" using XQx XQFXQ[OF PQ] by auto then have fp: "f p \ p \ f p = p" using pP FPP by (auto simp: ecl_def) have "p \ p" using PP pP by auto with fp have fpp: "f p \ p" by auto have e: "extreme {p \ A. f p \ p \ f p = p} (\) p" proof (intro extremeI CollectI conjI pA fp, elim CollectE conjE) fix q assume qA: "q \ A" and fq: "f q \ q \ f q = q" define Z where "Z \ {z \ A. q \ z}\{q}" then have qZ: "q \ Z" using qA by auto then have ZQ: "Z \ Q" using qA by (auto simp: Z_def Q_def ecl_def) have fqA: "f q \ A" using qA f by auto then have "f q \ Z" using fq by (auto simp: Z_def) then have 1: "Z = [f q]\<^sub>\" using XQx ZQ by auto then have "f q \ F Z" using qZ fqA by (auto simp: F_def) then have "F Z = [f q]\<^sub>\" using XQx XQFXQ[OF ZQ] by auto with 1 have "Z = F Z" by auto then have "P \\<^sup>s Z" using P ZQ by auto then show "p \ q" using pP qZ by auto qed with fpp show ?thesis using e by auto qed subsection \General Completeness\ lemma attract_mono_imp_fp_qfp_complete: assumes attract: "attractive A (\)" and comp: "CC-complete A (\)" - and wr_CC: "\C \ A. well_related_set C (\) \ C \ CC" - and extend: "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" + and wr_CC: "\C \ A. well_related_set C (\) \ CC C (\)" + and extend: "\X Y. CC X (\) \ CC Y (\) \ X \\<^sup>s Y \ CC (X \ Y) (\)" and mono: "monotone_on A (\) (\) f" and P: "P \ {x \ A. f x = x}" shows "CC-complete ({q \ A. f q \ q} \ P) (\)" proof (intro completeI) interpret attractive using attract. - fix X assume Xfix: "X \ {q \ A. f q \ q} \ P" and XCC: "X \ CC" + fix X assume Xfix: "X \ {q \ A. f q \ q} \ P" and XCC: "CC X (\)" with P have XA: "X \ A" by auto define B where "B \ {b \ A. \a \ X. a \ b}" { fix s a assume sA: "s \ A" and as: "\a \ X. a \ s" and aX: "a \ X" then have aA: "a \ A" using XA by auto then have fafs: "f a \ f s" using mono f aX sA as by (auto elim!: monotone_onE) have "a \ f s" proof (cases "f a = a") case True then show ?thesis using fafs by auto next case False then have "a \ f a" using P aX Xfix by auto also from fafs have "f a \ f s" by auto finally show ?thesis using f aA sA by auto qed } with f have fBB: "f ` B \ B" unfolding B_def by auto have BA: "B \ A" by (auto simp: B_def) have compB: "CC-complete B (\)" proof (unfold complete_def, intro allI impI) - fix Y assume YS: "Y \ B" and YCC: "Y \ CC" + fix Y assume YS: "Y \ B" and YCC: "CC Y (\)" with BA have YA: "Y \ A" by auto define C where "C \ X\Y" then have CA: "C \ A" using XA YA C_def by auto have XY: "X \\<^sup>s Y" using B_def YS by auto - then have CCC: "C \ CC" using extend XA YA XCC YCC C_def by auto + then have CCC: "CC C (\)" using extend XA YA XCC YCC C_def by auto then obtain s where s: "extreme_bound A (\) C s" - using completeD[OF comp CA CCC] by auto + using completeD[OF comp CA, OF CCC] by auto then have sA: "s \ A" by auto show "Ex (extreme_bound B (\) Y)" proof (intro exI extreme_boundI) { fix x assume "x \ X" then have "x \ s" using s C_def by auto } then show "s \ B" using sA B_def by auto next fix y assume "y \ Y" then show "y \ s" using s C_def using extremeD by auto next fix c assume cS: "c \ B" and "bound Y (\) c" then have "bound C (\) c" using C_def B_def by auto then show "s \ c" using s BA cS by auto qed qed from fBB interpret B: fixed_point_proof B "(\)" f by unfold_locales from BA have *: "{x \ A. f x \ x} \ B = {x \ B. f x \ x}" by auto have asB: "attractive B (\)" using attractive_subset[OF BA] by auto have monoB: "monotone_on B (\) (\) f" using monotone_on_cmono[OF BA] mono by (auto dest!: le_funD) - have compB: "well_complete B (\)" + have compB: "well_related_set-complete B (\)" using wr_CC compB BA by (simp add: complete_def) from B.attract_mono_imp_least_qfp[OF asB compB monoB] obtain l where "extreme {p \ B. f p \ p \ f p = p} (\) l" and fll: "f l \ l" by auto with P have l: "extreme ({p \ B. f p \ p} \ P \ B) (\) l" by auto show "Ex (extreme_bound ({q \ A. f q \ q} \ P) (\) X)" proof (intro exI extreme_boundI) show "l \ {q \ A. f q \ q} \ P" using l BA by auto fix a assume "a \ X" with l show "a \ l" by (auto simp: B_def) next fix c assume c: "bound X (\) c" and cfix: "c \ {q \ A. f q \ q} \ P" with P have cA: "c \ A" by auto with c have "c \ B" by (auto simp: B_def) with cfix l show "l \ c" by auto qed qed lemma attract_mono_imp_qfp_complete: assumes "attractive A (\)" and "CC-complete A (\)" - and "\C \ A. well_related_set C (\) \ C \ CC" - and "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" + and "\C \ A. well_related_set C (\) \ CC C (\)" + and "\X Y. CC X (\) \ CC Y (\) \ X \\<^sup>s Y \ CC (X \ Y) (\)" and "monotone_on A (\) (\) f" shows "CC-complete {p \ A. f p \ p} (\)" using attract_mono_imp_fp_qfp_complete[OF assms, of "{}"] by simp lemma antisym_mono_imp_fp_complete: assumes anti: "antisymmetric A (\)" and comp: "CC-complete A (\)" - and wr_CC: "\C \ A. well_related_set C (\) \ C \ CC" - and extend: "\X \ CC. \Y \ CC. X \\<^sup>s Y \ X \ Y \ CC" + and wr_CC: "\C \ A. well_related_set C (\) \ CC C (\)" + and extend: "\X Y. CC X (\) \ CC Y (\) \ X \\<^sup>s Y \ CC (X \ Y) (\)" and mono: "monotone_on A (\) (\) f" shows "CC-complete {p \ A. f p = p} (\)" proof- interpret antisymmetric using anti. have *: "{q \ A. f q \ q} \ {p \ A. f p = p}" using f by (auto intro!: antisym) from * attract_mono_imp_fp_qfp_complete[OF attractive_axioms comp wr_CC extend mono, of "{p\A. f p = p}"] show ?thesis by (auto simp: subset_Un_eq) qed end subsection \Instances\ subsubsection \Instances under attractivity\ context attractive begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. text \Full completeness\ -theorem mono_imp_qfp_UNIV_complete: - assumes comp: "UNIV-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "UNIV-complete {p \ A. f p \ p} (\)" +theorem mono_imp_qfp_complete: + assumes comp: "\-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" + shows "\-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete comp mono) apply unfold_locales by (auto simp: f) text \Connex completeness\ theorem mono_imp_qfp_connex_complete: - assumes comp: "{X. connex X (\)}-complete A (\)" + assumes comp: "connex-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "{X. connex X (\)}-complete {p \ A. f p \ p} (\)" + shows "connex-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales - by (auto simp: f intro: connex_union well_related_set.connex_axioms) + by (auto simp: f intro: connex_union well_related_set.connex) text \Directed completeness\ theorem mono_imp_qfp_directed_complete: - assumes comp: "{X. directed X (\)}-complete A (\)" + assumes comp: "directed-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "{X. directed X (\)}-complete {p \ A. f p \ p} (\)" + shows "directed-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales - by (auto simp: f intro!: directed_extend intro: well_related_set.connex_axioms connex.directed) + by (auto simp: f intro!: directed_extend intro: well_related_set.connex connex.directed) text \Well Completeness\ theorem mono_imp_qfp_well_complete: - assumes comp: "well_complete A (\)" + assumes comp: "well_related_set-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "well_complete {p \ A. f p \ p} (\)" + shows "well_related_set-complete {p \ A. f p \ p} (\)" apply (intro fixed_point_proof.attract_mono_imp_qfp_complete mono comp) apply unfold_locales by (auto simp: f well_related_extend) end subsubsection \Usual instances under antisymmetry \ context antisymmetric begin text \Knaster--Tarski\ theorem mono_imp_fp_complete: - assumes comp: "UNIV-complete A (\)" + assumes comp: "\-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "UNIV-complete {p \ A. f p = p} (\)" + shows "\-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) by auto qed text \Markowsky 1976\ theorem mono_imp_fp_connex_complete: - assumes comp: "{X. connex X (\)}-complete A (\)" + assumes comp: "connex-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "{X. connex X (\)}-complete {p \ A. f p = p} (\)" + shows "connex-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete antisymmetric_axioms mono comp) - by (auto intro: connex_union well_related_set.connex_axioms) + by (auto intro: connex_union well_related_set.connex) qed text \Pataraia\ theorem mono_imp_fp_directed_complete: - assumes comp: "{X. directed X (\)}-complete A (\)" + assumes comp: "directed-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "{X. directed X (\)}-complete {p \ A. f p = p} (\)" + shows "directed-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) - by (auto intro: directed_extend connex.directed well_related_set.connex_axioms) + by (auto intro: directed_extend connex.directed well_related_set.connex) qed text \Bhatta \& George 2011\ theorem mono_imp_fp_well_complete: - assumes comp: "well_complete A (\)" + assumes comp: "well_related_set-complete A (\)" and f: "f ` A \ A" and mono: "monotone_on A (\) (\) f" - shows "well_complete {p \ A. f p = p} (\)" + shows "well_related_set-complete {p \ A. f p = p} (\)" proof- interpret fixed_point_proof using f by unfold_locales show ?thesis apply (intro antisym_mono_imp_fp_complete mono antisymmetric_axioms comp) by (auto intro!: antisym well_related_extend) qed end end diff --git a/thys/Complete_Non_Orders/Kleene_Fixed_Point.thy b/thys/Complete_Non_Orders/Kleene_Fixed_Point.thy --- a/thys/Complete_Non_Orders/Kleene_Fixed_Point.thy +++ b/thys/Complete_Non_Orders/Kleene_Fixed_Point.thy @@ -1,291 +1,210 @@ theory Kleene_Fixed_Point - imports Complete_Relations + imports Complete_Relations Continuity begin section \Iterative Fixed Point Theorem\ text \Kleene's fixed-point theorem states that, for a pointed directed complete partial order $\tp{A,\SLE}$ and a Scott-continuous map $f: A \to A$, the supremum of $\set{f^n(\bot) \mid n\in\Nat}$ exists in $A$ and is a least fixed point. -Mashburn \<^cite>\"mashburn83"\ generalized the result so that +Mashburn \cite{mashburn83} generalized the result so that $\tp{A,\SLE}$ is a $\omega$-complete partial order and $f$ is $\omega$-continuous. In this section we further generalize the result and show that for $\omega$-complete relation $\tp{A,\SLE}$ and for every bottom element $\bot \in A$, the set $\set{f^n(\bot) \mid n\in\Nat}$ has suprema (not necessarily unique, of course) and, they are quasi-fixed points. Moreover, if $(\SLE)$ is attractive, then the suprema are precisely the least quasi-fixed points.\ -subsection \Scott Continuity, $\omega$-Completeness, $\omega$-Continuity\ - -text \In this Section, we formalize $\omega$-completeness, Scott continuity and $\omega$-continuity. -We then prove that a Scott continuous map is $\omega$-continuous and that an $\omega$-continuous -map is ``nearly'' monotone.\ - -context - fixes A :: "'a set" and less_eq :: "'a \ 'a \ bool" (infix "\" 50) -begin - -definition "omega_continuous f \ - f ` A \ A \ - (\c :: nat \ 'a. \ b \ A. - range c \ A \ - monotone (\) (\) c \ - extreme_bound A (\) (range c) b \ extreme_bound A (\) (f ` range c) (f b))" - -lemmas omega_continuousI[intro?] = - omega_continuous_def[unfolded atomize_eq, THEN iffD2, unfolded conj_imp_eq_imp_imp, rule_format] - -lemmas omega_continuousDdom = - omega_continuous_def[unfolded atomize_eq, THEN iffD1, unfolded conj_imp_eq_imp_imp, THEN conjunct1] - -lemmas omega_continuousD = - omega_continuous_def[unfolded atomize_eq, THEN iffD1, unfolded conj_imp_eq_imp_imp, THEN conjunct2, rule_format] - -lemmas omega_continuousE[elim] = - omega_continuous_def[unfolded atomize_eq, THEN iffD1, elim_format, unfolded conj_imp_eq_imp_imp, rule_format] - -lemma omega_continuous_imp_mono_refl: - assumes cont: "omega_continuous f" - and x: "x \ A" and y: "y \ A" - and xy: "x \ y" and xx: "x \ x" and yy: "y \ y" - shows "f x \ f y" -proof- - define c :: "nat \ 'a" where "c \ \i. if i = 0 then x else y" - from x y xx xy yy have c: "range c \ A" "monotone (\) (\) c" - by (auto simp: c_def intro!: monotoneI) - have "extreme_bound A (\) (range c) y" using xy yy x y by (auto simp: c_def) - then have fboy: "extreme_bound A (\) (f ` range c) (f y)" using c cont y by auto - then show "f x \ f y" by (auto simp: c_def) -qed - -definition "scott_continuous f \ - f ` A \ A \ - (\X s. X \ A \ directed X (\) \ X \ {} \ extreme_bound A (\) X s \ extreme_bound A (\) (f ` X) (f s))" - -lemmas scott_continuousI[intro?] = - scott_continuous_def[unfolded atomize_eq, THEN iffD2, unfolded conj_imp_eq_imp_imp, rule_format] - -lemmas scott_continuousE = - scott_continuous_def[unfolded atomize_eq, THEN iffD1, elim_format, unfolded conj_imp_eq_imp_imp, rule_format] - -lemma scott_continuous_imp_mono_refl: - assumes scott: "scott_continuous f" - and x: "x \ A" and y: "y \ A" and xy: "x \ y" and yy: "y \ y" - shows "f x \ f y" -proof- - define D where "D \ {x,y}" - from x y xy yy have dir_D: "D \ A" "directed D (\)" "D \ {}" - by (auto simp: D_def intro!: bexI[of _ y] directedI) - have "extreme_bound A (\) D y" using xy yy x y by (auto simp: D_def) - then have fboy: "extreme_bound A (\) (f ` D) (f y)" using dir_D scott by (auto elim!: scott_continuousE) - then show "f x \ f y" by (auto simp: D_def) -qed - -lemma scott_continuous_imp_omega_continuous: - assumes scott: "scott_continuous f" shows "omega_continuous f" -proof - from scott show "f ` A \ A" by (auto elim!: scott_continuousE) - fix c :: "nat \ 'a" - assume mono: "monotone (\) (\) c" and c: "range c \ A" - from monotone_directed_image[OF mono[folded monotone_on_UNIV] order.directed] scott c - show "extreme_bound A (\) (range c) b \ extreme_bound A (\) (f ` range c) (f b)" for b - by (auto elim!: scott_continuousE) -qed - -end - subsection \Existence of Iterative Fixed Points\ text \The first part of Kleene's theorem demands to prove that the set $\set{f^n(\bot) \mid n \in \Nat}$ has a supremum and that all such are quasi-fixed points. We prove this claim without assuming anything on the relation $\SLE$ besides $\omega$-completeness and one bottom element.\ -(* -no_notation power (infixr "^" 80) -*) notation compower ("_^_"[1000,1000]1000) -lemma mono_funpow: assumes f: "f ` A \ A" and mono: "monotone_on A r r f" +lemma monotone_on_funpow: assumes f: "f ` A \ A" and mono: "monotone_on A r r f" shows "monotone_on A r r (f^n)" proof (induct n) case 0 show ?case using monotone_on_id by (auto simp: id_def) next case (Suc n) with funpow_dom[OF f] show ?case by (auto intro!: monotone_onI monotone_onD[OF mono] elim!:monotone_onE) qed no_notation bot ("\") context fixes A and less_eq (infix "\" 50) and bot ("\") and f assumes bot: "\ \ A" "\q \ A. \ \ q" - assumes cont: "omega_continuous A (\) f" + assumes cont: "omega_chain-continuous A (\) A (\) f" begin -interpretation less_eq_notations. +interpretation less_eq_symmetrize. private lemma f: "f ` A \ A" using cont by auto private abbreviation(input) "Fn \ {f^n \ |. n :: nat}" private lemma fn_ref: "f^n \ \ f^n \" and fnA: "f^n \ \ A" proof (atomize(full), induct n) case 0 from bot show ?case by simp next case (Suc n) then have fn: "f^n \ \ A" and fnfn: "f^n \ \ f^n \" by auto - from f fn omega_continuous_imp_mono_refl[OF cont fn fn fnfn fnfn fnfn] + from f fn omega_continuous_imp_mono_refl[OF cont fnfn fnfn fnfn] show ?case by auto qed private lemma FnA: "Fn \ A" using fnA by auto -private lemma fn_monotone: "monotone (\) (\) (\n. f^n \)" -proof - fix n m :: nat - assume "n \ m" - from le_Suc_ex[OF this] obtain k where m: "m = n + k" by auto - from bot fn_ref fnA omega_continuous_imp_mono_refl[OF cont] - show "f^n \ \ f^m \" by (unfold m, induct n, auto) -qed +private lemma Fn_chain: "omega_chain Fn (\)" +proof (intro omega_chainI) + show fn_monotone: "monotone (\) (\) (\n. f^n \)" + proof + fix n m :: nat + assume "n \ m" + from le_Suc_ex[OF this] obtain k where m: "m = n + k" by auto + from bot fn_ref fnA omega_continuous_imp_mono_refl[OF cont] + show "f^n \ \ f^m \" by (unfold m, induct n, auto) + qed +qed auto private lemma Fn: "Fn = range (\n. f^n \)" by auto theorem kleene_qfp: assumes q: "extreme_bound A (\) Fn q" shows "f q \ q" proof have fq: "extreme_bound A (\) (f ` Fn) (f q)" - apply (unfold Fn) - apply (rule omega_continuousD[OF cont]) - using FnA fn_monotone q by (unfold Fn, auto) - with bot have nq: "f^n \ \ f q" for n - by(induct n, auto simp: extreme_bound_iff) + apply (rule continuousD[OF cont _ _ FnA q]) + using Fn_chain by auto + with bot have nq: "f^n \ \ f q" for n by (induct n, auto simp: extreme_bound_iff) then show "q \ f q" using f q by blast have "f (f^n \) \ Fn" for n by (auto intro!: exI[of _ "Suc n"]) then have "f ` Fn \ Fn" by auto - from extreme_bound_mono[OF this fq q] + from extreme_bound_subset[OF this fq q] show "f q \ q". qed lemma ex_kleene_qfp: - assumes comp: "omega_complete A (\)" - shows "\p. extreme_bound A (\) Fn p" - using fn_monotone - apply (intro comp[unfolded omega_complete_def, THEN completeD, OF FnA]) - by fast + assumes comp: "omega_chain-complete A (\)" + shows "\p. extreme_bound A (\) Fn p" + apply (intro comp[THEN completeD, OF FnA]) + using Fn_chain + by auto subsection \Iterative Fixed Points are Least.\ text \Kleene's theorem also states that the quasi-fixed point found this way is a least one. Again, attractivity is needed to prove this statement.\ lemma kleene_qfp_is_least: assumes attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" assumes q: "extreme_bound A (\) Fn q" shows "extreme {s \ A. f s \ s} (\) q" proof(safe intro!: extremeI kleene_qfp[OF q]) from q show "q \ A" by auto fix c assume c: "c \ A" and cqfp: "f c \ c" { fix n::nat have "f^n \ \ c" proof(induct n) case 0 show ?case using bot c by auto next case IH: (Suc n) have "c \ c" using attract cqfp c by auto with IH have "f^(Suc n) \ \ f c" using omega_continuous_imp_mono_refl[OF cont] fn_ref fnA c by auto then show ?case using attract cqfp c fnA by blast qed } then show "q \ c" using q c by auto qed lemma kleene_qfp_iff_least: - assumes comp: "omega_complete A (\)" + assumes comp: "omega_chain-complete A (\)" assumes attract: "\q \ A. \x \ A. f q \ q \ x \ f q \ x \ q" assumes dual_attract: "\p \ A. \q \ A. \x \ A. p \ q \ q \ x \ p \ x" shows "extreme_bound A (\) Fn = extreme {s \ A. f s \ s} (\)" proof (intro ext iffI kleene_qfp_is_least[OF attract]) fix q assume q: "extreme {s \ A. f s \ s} (\) q" from q have qA: "q \ A" by auto from q have qq: "q \ q" by auto from q have fqq: "f q \ q" by auto from ex_kleene_qfp[OF comp] obtain k where k: "extreme_bound A (\) Fn k" by auto have qk: "q \ k" proof from kleene_qfp[OF k] q k show "q \ k" by auto from kleene_qfp_is_least[OF _ k] q attract show "k \ q" by blast qed show "extreme_bound A (\) Fn q" proof (intro extreme_boundI,safe) fix n show "f^n \ \ q" proof (induct n) case 0 from bot q show ?case by auto next case S:(Suc n) from fnA f have fsnbA: "f (f^n \) \ A" by auto have fnfn: "f^n \ \ f^n \" using fn_ref by auto have "f (f^n \) \ f q" - using omega_continuous_imp_mono_refl[OF cont fnA qA S fnfn qq] by auto + using omega_continuous_imp_mono_refl[OF cont] fnA qA S fnfn qq by auto then show ?case using fsnbA qA attract fqq by auto qed next fix x assume "bound Fn (\) x" and x: "x \ A" with k have kx: "k \ x" by auto with dual_attract[rule_format, OF _ _ x qk] q k show "q \ x" by auto next from q show "q \ A" by auto qed qed end context attractive begin -interpretation less_eq_notations. +interpretation less_eq_dualize + less_eq_symmetrize. theorem kleene_qfp_is_dual_extreme: - assumes comp: "omega_complete A (\)" - and cont: "omega_continuous A (\) f" and bA: "b \ A" and bot: "\x \ A. b \ x" + assumes comp: "omega_chain-complete A (\)" + and cont: "omega_chain-continuous A (\) A (\) f" and bA: "b \ A" and bot: "\x \ A. b \ x" shows "extreme_bound A (\) {f^n b |. n :: nat} = extreme {s \ A. f s \ s} (\)" apply (rule kleene_qfp_iff_least[OF bA bot cont comp]) - using cont[THEN omega_continuousDdom] + using continuous_carrierD[OF cont] by (auto dest: sym_order_trans order_sym_trans) end corollary(in antisymmetric) kleene_fp: - assumes cont: "omega_continuous A (\) f" + assumes cont: "omega_chain-continuous A (\) A (\) f" and b: "b \ A" "\x \ A. b \ x" and p: "extreme_bound A (\) {f^n b |. n :: nat} p" shows "f p = p" - using kleene_qfp[OF b cont] p cont[THEN omega_continuousDdom] + using kleene_qfp[OF b cont] p cont[THEN continuous_carrierD] by (auto 2 3 intro!:antisym) no_notation compower ("_^_"[1000,1000]1000) end \ No newline at end of file diff --git a/thys/Complete_Non_Orders/ROOT b/thys/Complete_Non_Orders/ROOT --- a/thys/Complete_Non_Orders/ROOT +++ b/thys/Complete_Non_Orders/ROOT @@ -1,10 +1,12 @@ chapter AFP session Complete_Non_Orders (AFP) = HOL + description "Fixed-Point Theorems for Non-Transitive Relations" options [timeout = 300] + sessions + "HOL-Library" theories Fixed_Points Kleene_Fixed_Point document_files "root.tex" "root.bib" diff --git a/thys/Complete_Non_Orders/Well_Relations.thy b/thys/Complete_Non_Orders/Well_Relations.thy new file mode 100755 --- /dev/null +++ b/thys/Complete_Non_Orders/Well_Relations.thy @@ -0,0 +1,845 @@ +theory Well_Relations + imports Binary_Relations +begin + +section \Well-Relations\ + +text \ +A related set $\tp{A,\SLE}$ is called \emph{topped} if there is a ``top'' element $\top \in A$, +a greatest element in $A$. +Note that there might be multiple tops if $(\SLE)$ is not antisymmetric.\ + +definition "extremed A r \ \e. extreme A r e" + +lemma extremedI: "extreme A r e \ extremed A r" + by (auto simp: extremed_def) + +lemma extremedE: "extremed A r \ (\e. extreme A r e \ thesis) \ thesis" + by (auto simp: extremed_def) + +lemma extremed_imp_ex_bound: "extremed A r \ X \ A \ \b \ A. bound X r b" + by (auto simp: extremed_def) + +locale well_founded = related_set _ "(\)" + less_syntax + + assumes induct[consumes 1, case_names less, induct set]: + "a \ A \ (\x. x \ A \ (\y. y \ A \ y \ x \ P y) \ P x) \ P a" +begin + +sublocale asymmetric +proof (intro asymmetric.intro notI) + fix x y + assume xA: "x \ A" + then show "y \ A \ x \ y \ y \ x \ False" + by (induct arbitrary: y rule: induct, auto) +qed + +lemma prefixed_Imagep_imp_empty: + assumes a: "X \ ((\) ``` X) \ A" shows "X = {}" +proof - + from a have XA: "X \ A" by auto + have "x \ A \ x \ X" for x + proof (induct x rule: induct) + case (less x) + with a show ?case by (auto simp: Imagep_def) + qed + with XA show ?thesis by auto +qed + +lemma nonempty_imp_ex_extremal: + assumes QA: "Q \ A" and Q: "Q \ {}" + shows "\z \ Q. \y \ Q. \ y \ z" + using Q prefixed_Imagep_imp_empty[of Q] QA by (auto simp: Imagep_def) + +interpretation Restrp: well_founded UNIV "(\)\A" + rewrites "\x. x \ UNIV \ True" + and "(\)\A\UNIV = (\)\A" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" +proof - + have "(\x. (\y. ((\) \ A) y x \ P y) \ P x) \ P a" for a P + using induct[of a P] by (auto simp: Restrp_def) + then show "well_founded UNIV ((\)\A)" apply unfold_locales by auto +qed auto + +lemmas Restrp_well_founded = Restrp.well_founded_axioms +lemmas Restrp_induct[consumes 0, case_names less] = Restrp.induct + +interpretation Restrp.tranclp: well_founded UNIV "((\)\A)\<^sup>+\<^sup>+" + rewrites "\x. x \ UNIV \ True" + and "((\)\A)\<^sup>+\<^sup>+ \ UNIV = ((\)\A)\<^sup>+\<^sup>+" + and "(((\)\A)\<^sup>+\<^sup>+)\<^sup>+\<^sup>+ = ((\)\A)\<^sup>+\<^sup>+" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" +proof- + { fix P x + assume induct_step: "\x. (\y. ((\)\A)\<^sup>+\<^sup>+ y x \ P y) \ P x" + have "P x" + proof (rule induct_step) + show "\y. ((\)\A)\<^sup>+\<^sup>+ y x \ P y" + proof (induct x rule: Restrp_induct) + case (less x) + from \((\)\A)\<^sup>+\<^sup>+ y x\ + show ?case + proof (cases rule: tranclp.cases) + case r_into_trancl + with induct_step less show ?thesis by auto + next + case (trancl_into_trancl b) + with less show ?thesis by auto + qed + qed + qed + } + then show "well_founded UNIV ((\)\A)\<^sup>+\<^sup>+" by unfold_locales auto +qed auto + +lemmas Restrp_tranclp_well_founded = Restrp.tranclp.well_founded_axioms +lemmas Restrp_tranclp_induct[consumes 0, case_names less] = Restrp.tranclp.induct + +end + +context + fixes A :: "'a set" and less :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma well_foundedI_pf: + assumes pre: "\X. X \ A \ X \ ((\) ``` X) \ A \ X = {}" + shows "well_founded A (\)" +proof + fix P a assume aA: "a \ A" and Ind: "\x. x \ A \ (\y. y \ A \ y \ x \ P y) \ P x" + from Ind have "{a\A. \P a} \ ((\) ``` {a\A. \P a}) \ A" by (auto simp: Imagep_def) + from pre[OF _ this] aA + show "P a" by auto +qed + +lemma well_foundedI_extremal: + assumes a: "\X. X \ A \ X \ {} \ \x \ X. \y \ X. \ y \ x" + shows "well_founded A (\)" +proof (rule well_foundedI_pf) + fix X assume XA: "X \ A" and pf: "X \ ((\) ``` X) \ A" + from a[OF XA] pf show "X = {}" by (auto simp: Imagep_def) +qed + +lemma well_founded_iff_ex_extremal: + "well_founded A (\) \ (\X \ A. X \ {} \ (\x \ X. \z \ X. \ z \ x))" + using well_founded.nonempty_imp_ex_extremal well_foundedI_extremal by blast + +end + +lemma well_founded_cong: + assumes r: "\a b. a \ A \ b \ A \ r a b \ r' a b" + and A: "\a b. r' a b \ a \ A \ a \ A'" + and B: "\a b. r' a b \ b \ A \ b \ A'" + shows "well_founded A r \ well_founded A' r'" +proof (intro iffI) + assume wf: "well_founded A r" + show "well_founded A' r'" + proof (intro well_foundedI_extremal) + fix X + assume X: "X \ A'" and X0: "X \ {}" + show "\x\X. \y\X. \ r' y x" + proof (cases "X \ A = {}") + case True + from X0 obtain x where xX: "x \ X" by auto + with True have "x \ A" by auto + with xX X have "\y\X. \ r' y x" by (auto simp: B) + with xX show ?thesis by auto + next + case False + from well_founded.nonempty_imp_ex_extremal[OF wf _ this] + obtain x where x: "x \ X \ A" and Ar: "\y. y \ X \ y \ A \ \ r y x" by auto + have "\y \ X. \ r' y x" + proof (intro ballI notI) + fix y assume yX: "y \ X" and yx: "r' y x" + from yX X have yA': "y \ A'" by auto + show False + proof (cases "y \ A") + case True with x Ar[OF yX] yx r show ?thesis by auto + next + case False with yA' x A[OF yx] r X show ?thesis by (auto simp:) + qed + qed + with x show "\x \ X. \y \ X. \ r' y x" by auto + qed + qed +next + assume wf: "well_founded A' r'" + show "well_founded A r" + proof (intro well_foundedI_extremal) + fix X + assume X: "X \ A" and X0: "X \ {}" + show "\x\X. \y\X. \ r y x" + proof (cases "X \ A' = {}") + case True + from X0 obtain x where xX: "x \ X" by auto + with True have "x \ A'" by auto + with xX X B have "\y\X. \ r y x" by (auto simp: r in_mono) + with xX show ?thesis by auto + next + case False + from well_founded.nonempty_imp_ex_extremal[OF wf _ this] + obtain x where x: "x \ X \ A'" and Ar: "\y. y \ X \ y \ A' \ \ r' y x" by auto + have "\y \ X. \ r y x" + proof (intro ballI notI) + fix y assume yX: "y \ X" and yx: "r y x" + from yX X have y: "y \ A" by auto + show False + proof (cases "y \ A'") + case True with x Ar[OF yX] yx r X y show ?thesis by auto + next + case False with y x A yx r X show ?thesis by auto + qed + qed + with x show "\x \ X. \y \ X. \ r y x" by auto + qed + qed +qed + +lemma wfP_iff_well_founded_UNIV: "wfP r \ well_founded UNIV r" + by (auto simp: wfP_def wf_def well_founded_def) + +lemma well_founded_empty[intro!]: "well_founded {} r" + by (auto simp: well_founded_iff_ex_extremal) + +lemma well_founded_singleton: + assumes "\r x x" shows "well_founded {x} r" + using assms by (auto simp: well_founded_iff_ex_extremal) + +lemma well_founded_Restrp[simp]: "well_founded A (r\B) \ well_founded (A\B) r" (is "?l \ ?r") +proof (intro iffI well_foundedI_extremal) + assume l: ?l + fix X assume XAB: "X \ A \ B" and X0: "X \ {}" + with l[THEN well_founded.nonempty_imp_ex_extremal] + have "\x\X. \z\X. \ (r \ B) z x" by auto + with XAB show "\x\X. \y\X. \ r y x" by (auto simp: Restrp_def) +next + assume r: ?r + fix X assume XA: "X \ A" and X0: "X \ {}" + show "\x\X. \y\X. \ (r \ B) y x" + proof (cases "X \ B") + case True + with r[THEN well_founded.nonempty_imp_ex_extremal, of X] XA X0 + have "\z\X. \y\X. \ r y z" by auto + then show ?thesis by auto + next + case False + then obtain x where x: "x \ X - B" by auto + then have "\y\X. \ (r \ B) y x" by auto + with x show ?thesis by auto + qed +qed + +lemma Restrp_tranclp_well_founded_iff: + fixes less (infix "\" 50) + shows "well_founded UNIV ((\) \ A)\<^sup>+\<^sup>+ \ well_founded A (\)" (is "?l \ ?r") +proof (rule iffI) + show "?r \ ?l" by (fact well_founded.Restrp_tranclp_well_founded) + assume ?l + then interpret well_founded UNIV "((\) \ A)\<^sup>+\<^sup>+". + show ?r + proof (unfold well_founded_iff_ex_extremal, intro allI impI) + fix X assume XA: "X \ A" and X0: "X \ {}" + from nonempty_imp_ex_extremal[OF _ X0] + obtain z where zX: "z \ X" and Xz: "\y\X. \ ((\) \ A)\<^sup>+\<^sup>+ y z" by auto + show "\z \ X. \y \ X. \ y \ z" + proof (intro bexI[OF _ zX] ballI notI) + fix y assume yX: "y \ X" and yz: "y \ z" + from yX yz zX XA have "((\) \ A) y z" by auto + with yX Xz show False by auto + qed + qed +qed + +lemma (in well_founded) well_founded_subset: + assumes "B \ A" shows "well_founded B (\)" + using assms well_founded_axioms by (auto simp: well_founded_iff_ex_extremal) + +lemma well_founded_extend: + fixes less (infix "\" 50) + assumes A: "well_founded A (\)" + assumes B: "well_founded B (\)" + assumes AB: "\a \ A. \b \ B. \b \ a" + shows "well_founded (A \ B) (\)" +proof (intro well_foundedI_extremal) + interpret A: well_founded A "(\)" using A. + interpret B: well_founded B "(\)" using B. + fix X assume XAB: "X \ A \ B" and X0: "X \ {}" + show "\x\X. \y\X. \ y \ x" + proof (cases "X \ A = {}") + case True + with XAB have XB: "X \ B" by auto + from B.nonempty_imp_ex_extremal[OF XB X0] show ?thesis. + next + case False + with A.nonempty_imp_ex_extremal[OF _ this] + obtain e where XAe: "e \ X \ A" "\y\X \ A. \ y \ e" by auto + then have eX: "e \ X" and eA: "e \ A" by auto + { fix x assume xX: "x \ X" + have "\x \ e" + proof (cases "x \ A") + case True with XAe xX show ?thesis by auto + next + case False + with xX XAB have "x \ B" by auto + with AB eA show ?thesis by auto + qed + } + with eX show ?thesis by auto + qed +qed + +lemma closed_UN_well_founded: + fixes r (infix "\" 50) + assumes XX: "\X\XX. well_founded X (\) \ (\x\X. \y\\XX. y \ x \ y \ X)" + shows "well_founded (\XX) (\)" +proof (intro well_foundedI_extremal) + have *: "X \ XX \ x\X \ y \ \XX \ y \ x \ y \ X" for X x y using XX by blast + fix S + assume S: "S \ \XX" and S0: "S \ {}" + from S0 obtain x where xS: "x \ S" by auto + with S obtain X where X: "X \ XX" and xX: "x \ X" by auto + from xS xX have Sx0: "S \ X \ {}" by auto + from X XX interpret well_founded X "(\)" by auto + from nonempty_imp_ex_extremal[OF _ Sx0] + obtain z where zS: "z \ S" and zX: "z \ X" and min: "\y \ S \ X. \ y \ z" by auto + show "\x\S. \y\S. \ y \ x" + proof (intro bexI[OF _ zS] ballI notI) + fix y + assume yS: "y \ S" and yz: "y \ z" + have yXX: "y \ \ XX" using S yS by auto + from *[OF X zX yXX yz] yS have "y \ X \ S" by auto + with min yz show False by auto + qed +qed + +lemma well_founded_cmono: + assumes r': "r' \ r" and wf: "well_founded A r" + shows "well_founded A r'" +proof (intro well_foundedI_extremal) + fix X assume "X \ A" and "X \ {}" + from well_founded.nonempty_imp_ex_extremal[OF wf this] + show "\x\X. \y\X. \ r' y x" using r' by auto +qed + +locale well_founded_ordered_set = well_founded + transitive _ "(\)" +begin + +sublocale strict_ordered_set.. + +interpretation Restrp: strict_ordered_set UNIV "(\)\A" + Restrp: well_founded UNIV "(\)\A" + using Restrp_strict_order Restrp_well_founded . + +lemma Restrp_well_founded_order: "well_founded_ordered_set UNIV ((\)\A)".. + +lemma well_founded_ordered_subset: "B \ A \ well_founded_ordered_set B (\)" + apply intro_locales + using well_founded_subset transitive_subset by auto + +end + +lemmas well_founded_ordered_setI = well_founded_ordered_set.intro + +lemma well_founded_ordered_set_empty[intro!]: "well_founded_ordered_set {} r" + by (auto intro!: well_founded_ordered_setI) + + +locale well_related_set = related_set + + assumes nonempty_imp_ex_extreme: "X \ A \ X \ {} \ \e. extreme X (\)\<^sup>- e" +begin + +sublocale connex +proof + fix x y assume "x \ A" and "y \ A" + with nonempty_imp_ex_extreme[of "{x,y}"] show "x \ y \ y \ x" by auto +qed + +lemmas connex = connex_axioms(* I'd like to access well_related_set.connex_axioms, + but it's not visible without interpretation. *) + +interpretation less_eq_asymmetrize. + +sublocale asym: well_founded A "(\)" +proof (unfold well_founded_iff_ex_extremal, intro allI impI) + fix X + assume XA: "X \ A" and X0: "X \ {}" + from nonempty_imp_ex_extreme[OF XA X0] obtain e where "extreme X (\)\<^sup>- e" by auto + then show "\x\X. \z\X. \z \ x" by (auto intro!: bexI[of _ e]) +qed + +lemma well_related_subset: "B \ A \ well_related_set B (\)" + by (auto intro!: well_related_set.intro nonempty_imp_ex_extreme) + +lemma monotone_image_well_related: + fixes leB (infix "\" 50) + assumes mono: "monotone_on A (\) (\) f" shows "well_related_set (f ` A) (\)" +proof (intro well_related_set.intro) + interpret less_eq_dualize. + fix X' assume X'fA: "X' \ f ` A" and X'0: "X' \ {}" + then obtain X where XA: "X \ A" and X': "X' = f ` X" and X0: "X \ {}" + by (auto elim!: subset_imageE) + from nonempty_imp_ex_extreme[OF XA X0] + obtain e where Xe: "extreme X (\) e" by auto + note monotone_on_subset[OF mono XA] + note monotone_on_dual[OF this] + from monotone_image_extreme[OF this Xe] + show "\e'. extreme X' (\)\<^sup>- e'" by (auto simp: X') +qed + +end + +sublocale well_related_set \ reflexive using local.reflexive_axioms. + +lemmas well_related_setI = well_related_set.intro + +lemmas well_related_iff_ex_extreme = well_related_set_def + +lemma well_related_set_empty[intro!]: "well_related_set {} r" + by (auto intro!: well_related_setI) + +context + fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma well_related_iff_neg_well_founded: + "well_related_set A (\) \ well_founded A (\x y. \ y \ x)" + by (simp add: well_related_set_def well_founded_iff_ex_extremal extreme_def Bex_def) + +lemma well_related_singleton_refl: + assumes "x \ x" shows "well_related_set {x} (\)" + by (intro well_related_set.intro exI[of _ x], auto simp: subset_singleton_iff assms) + +lemma closed_UN_well_related: + assumes XX: "\X\XX. well_related_set X (\) \ (\x\X. \y\\XX. \x \ y \ y \ X)" + shows "well_related_set (\XX) (\)" + using XX + apply (unfold well_related_iff_neg_well_founded) + using closed_UN_well_founded[of _ "\x y. \ y \ x"]. + +end + +lemma well_related_extend: + fixes r (infix "\" 50) + assumes "well_related_set A (\)" and "well_related_set B (\)" and "\a \ A. \b \ B. a \ b" + shows "well_related_set (A \ B) (\)" + using well_founded_extend[of _ "\x y. \ y \ x", folded well_related_iff_neg_well_founded] + using assms by auto + +lemma pair_well_related: + fixes less_eq (infix "\" 50) + assumes "i \ i" and "i \ j" and "j \ j" + shows "well_related_set {i, j} (\)" +proof (intro well_related_setI) + fix X assume "X \ {i,j}" and "X \ {}" + then have "X = {i,j} \ X = {i} \ X = {j}" by auto + with assms show "\e. extreme X (\)\<^sup>- e" by auto +qed + +locale pre_well_ordered_set = semiattractive + well_related_set +begin + +interpretation less_eq_asymmetrize. + +sublocale transitive +proof + fix x y z assume xy: "x \ y" and yz: "y \ z" and x: "x \ A" and y: "y \ A" and z: "z \ A" + from x y z have "\e. extreme {x,y,z} (\) e" (is "\e. ?P e") by (auto intro!: nonempty_imp_ex_extreme) + then have "?P x \ ?P y \ ?P z" by auto + then show "x \ z" + proof (elim disjE) + assume "?P x" + then show ?thesis by auto + next + assume "?P y" + then have "y \ x" by auto + from attract[OF xy this yz] x y z show ?thesis by auto + next + assume "?P z" + then have zx: "z \ x" and zy: "z \ y" by auto + from attract[OF yz zy zx] x y z have yx: "y \ x" by auto + from attract[OF xy yx yz] x y z show ?thesis by auto + qed +qed + +sublocale total_quasi_ordered_set.. + +end + +lemmas pre_well_ordered_iff_semiattractive_well_related = + pre_well_ordered_set_def[unfolded atomize_eq] + +lemma pre_well_ordered_set_empty[intro!]: "pre_well_ordered_set {} r" + by (auto simp: pre_well_ordered_iff_semiattractive_well_related) + +lemma pre_well_ordered_iff: + "pre_well_ordered_set A r \ total_quasi_ordered_set A r \ well_founded A (asympartp r)" + (is "?p \ ?t \ ?w") +proof safe + assume ?p + then interpret pre_well_ordered_set A r. + show ?t ?w by unfold_locales +next + assume ?t + then interpret total_quasi_ordered_set A r. + assume ?w + then have "well_founded UNIV (asympartp r \ A)" by simp + also have "asympartp r \ A = (\x y. \ r y x) \ A" by (intro ext, auto simp: not_iff_asym) + finally have "well_related_set A r" by (simp add: well_related_iff_neg_well_founded) + then show ?p by intro_locales +qed + +lemma (in semiattractive) pre_well_ordered_iff_well_related: + assumes XA: "X \ A" + shows "pre_well_ordered_set X (\) \ well_related_set X (\)" (is "?l \ ?r") +proof + interpret X: semiattractive X using semiattractive_subset[OF XA]. + { assume ?l + then interpret X: pre_well_ordered_set X. + show ?r by unfold_locales + } + assume ?r + then interpret X: well_related_set X. + show ?l by unfold_locales +qed + +lemma semiattractive_extend: + fixes r (infix "\" 50) + assumes A: "semiattractive A (\)" and B: "semiattractive B (\)" + and AB: "\a \ A. \b \ B. a \ b \ \ b \ a" + shows "semiattractive (A \ B) (\)" +proof- + interpret A: semiattractive A "(\)" using A. + interpret B: semiattractive B "(\)" using B. + { + fix x y z + assume yB: "y \ B" and zA: "z \ A" and yz: "y \ z" + have False using AB[rule_format, OF zA yB] yz by auto + } + note * = this + show ?thesis + by (auto intro!: semiattractive.intro dest:* AB[rule_format] A.attract B.attract) +qed + +lemma pre_well_order_extend: + fixes r (infix "\" 50) + assumes A: "pre_well_ordered_set A (\)" and B: "pre_well_ordered_set B (\)" + and AB: "\a \ A. \b \ B. a \ b \ \ b \ a" + shows "pre_well_ordered_set (A\B) (\)" +proof- + interpret A: pre_well_ordered_set A "(\)" using A. + interpret B: pre_well_ordered_set B "(\)" using B. + show ?thesis + apply (intro pre_well_ordered_set.intro well_related_extend semiattractive_extend) + apply unfold_locales + by (auto dest: AB[rule_format]) +qed + +lemma (in well_related_set) monotone_image_pre_well_ordered: + fixes leB (infix "\''" 50) + assumes mono: "monotone_on A (\) (\') f" + and image: "semiattractive (f ` A) (\')" + shows "pre_well_ordered_set (f ` A) (\')" + by (intro pre_well_ordered_set.intro monotone_image_well_related[OF mono] image) + +locale well_ordered_set = antisymmetric + well_related_set +begin + +sublocale pre_well_ordered_set.. + +sublocale total_ordered_set.. + +lemma well_ordered_subset: "B \ A \ well_ordered_set B (\)" + using well_related_subset antisymmetric_subset by (intro well_ordered_set.intro) + +sublocale asym: well_founded_ordered_set A "asympartp (\)" + by (intro well_founded_ordered_set.intro asym.well_founded_axioms asympartp_transitive) + +end + +lemmas well_ordered_iff_antisymmetric_well_related = well_ordered_set_def[unfolded atomize_eq] + +lemma well_ordered_set_empty[intro!]: "well_ordered_set {} r" + by (auto simp: well_ordered_iff_antisymmetric_well_related) + +lemma (in antisymmetric) well_ordered_iff_well_related: + assumes XA: "X \ A" + shows "well_ordered_set X (\) \ well_related_set X (\)" (is "?l \ ?r") +proof + interpret X: antisymmetric X using antisymmetric_subset[OF XA]. + { assume ?l + then interpret X: well_ordered_set X. + show ?r by unfold_locales + } + assume ?r + then interpret X: well_related_set X. + show ?l by unfold_locales +qed + +context + fixes A :: "'a set" and less_eq :: "'a \ 'a \ bool" (infix "\" 50) +begin + +context + assumes A: "\a \ A. \b \ A. a \ b" +begin + +interpretation well_related_set A "(\)" + apply unfold_locales + using A by blast + +lemmas trivial_well_related = well_related_set_axioms + +lemma trivial_pre_well_order: "pre_well_ordered_set A (\)" + apply unfold_locales + using A by blast + +end + +interpretation less_eq_asymmetrize. + +lemma well_ordered_iff_well_founded_total_ordered: + "well_ordered_set A (\) \ total_ordered_set A (\) \ well_founded A (\)" +proof (safe) + assume "well_ordered_set A (\)" + then interpret well_ordered_set A "(\)". + show "total_ordered_set A (\)" "well_founded A (\)" by unfold_locales +next + assume "total_ordered_set A (\)" and "well_founded A (\)" + then interpret total_ordered_set A "(\)" + asympartp: well_founded A "(\)". + show "well_ordered_set A (\)" + proof + fix X assume XA: "X \ A" and "X \ {}" + from XA asympartp.nonempty_imp_ex_extremal[OF this] + show "\e. extreme X (\) e" by (auto simp: not_asym_iff subsetD) + qed +qed + +end + +context + fixes less_eq :: "'a \ 'a \ bool" (infix "\" 50) +begin + +lemma well_order_extend: + assumes A: "well_ordered_set A (\)" and B: "well_ordered_set B (\)" + and ABa: "\a \ A. \b \ B. a \ b \ b \ a \ a = b" + and AB: "\a \ A. \b \ B. a \ b" + shows "well_ordered_set (A\B) (\)" +proof- + interpret A: well_ordered_set A "(\)" using A. + interpret B: well_ordered_set B "(\)" using B. + show ?thesis + apply (intro well_ordered_set.intro antisymmetric_union well_related_extend ABa AB) + by unfold_locales +qed + +interpretation singleton: antisymmetric "{a}" "(\)" for a apply unfold_locales by auto + +lemmas singleton_antisymmetric[intro!] = singleton.antisymmetric_axioms + +lemma singleton_well_ordered[intro!]: "a \ a \ well_ordered_set {a} (\)" + apply unfold_locales by auto + +lemma closed_UN_well_ordered: + assumes anti: "antisymmetric (\ XX) (\)" + and XX: "\X\XX. well_ordered_set X (\) \ (\x\X. \y\\XX. \ x \ y \ y \ X)" + shows "well_ordered_set (\XX) (\)" + apply (intro well_ordered_set.intro closed_UN_well_related anti) + using XX well_ordered_set.axioms by fast + +end + +lemma (in well_related_set) monotone_image_well_ordered: + fixes leB (infix "\''" 50) + assumes mono: "monotone_on A (\) (\') f" + and image: "antisymmetric (f ` A) (\')" + shows "well_ordered_set (f ` A) (\')" + by (intro well_ordered_set.intro monotone_image_well_related[OF mono] image) + + +subsection \Relating to Classes\ + +locale well_founded_quasi_ordering = quasi_ordering + well_founded +begin + +lemma well_founded_quasi_ordering_subset: + assumes "X \ A" shows "well_founded_quasi_ordering X (\) (\)" + by (intro well_founded_quasi_ordering.intro quasi_ordering_subset well_founded_subset assms) + +end + +class wf_qorder = ord + + assumes "well_founded_quasi_ordering UNIV (\) (<)" +begin + +interpretation well_founded_quasi_ordering UNIV + using wf_qorder_axioms unfolding class.wf_qorder_def by auto + +subclass qorder .. + +sublocale order: well_founded_quasi_ordering UNIV + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + apply unfold_locales by (auto simp:atomize_eq) + +end + +context wellorder begin + +subclass wf_qorder + apply (unfold_locales) + using less_induct by auto + +sublocale order: well_ordered_set UNIV + rewrites "\x. x \ UNIV \ True" + and "\X. X \ UNIV \ True" + and "\r. r \ UNIV \ r" + and "\P. True \ P \ P" + and "Ball UNIV \ All" + and "Bex UNIV \ Ex" + and "sympartp (\)\<^sup>- \ sympartp (\)" + and "\P1. (True \ PROP P1) \ PROP P1" + and "\P1. (True \ P1) \ Trueprop P1" + and "\P1 P2. (True \ PROP P1 \ PROP P2) \ (PROP P1 \ PROP P2)" + apply (unfold well_ordered_iff_well_founded_total_ordered) + apply (intro conjI order.total_ordered_set_axioms) + by (auto simp: order.well_founded_axioms atomize_eq) + +end + +thm order.nonempty_imp_ex_extreme + +subsection \omega-Chains\ + +definition "omega_chain A r \ \f :: nat \ 'a. monotone (\) r f \ range f = A" + +lemma omega_chainI: + fixes f :: "nat \ 'a" + assumes "monotone (\) r f" "range f = A" shows "omega_chain A r" + using assms by (auto simp: omega_chain_def) + +lemma omega_chainE: + assumes "omega_chain A r" + and "\f :: nat \ 'a. monotone (\) r f \ range f = A \ thesis" + shows thesis + using assms by (auto simp: omega_chain_def) + +lemma (in transitive) local_chain: + assumes CA: "range C \ A" + shows "(\i::nat. C i \ C (Suc i)) \ monotone (<) (\) C" +proof (intro iffI allI monotoneI) + fix i j :: nat + assume loc: "\i. C i \ C (Suc i)" and ij: "i < j" + have "C i \ C (i+k+1)" for k + proof (induct k) + case 0 + from loc show ?case by auto + next + case (Suc k) + also have "C (i+k+1) \ C (i+k+Suc 1)" using loc by auto + finally show ?case using CA by auto + qed + from this[of "j-i-1"] ij show "C i \ C j" by auto +next + fix i + assume "monotone (<) (\) C" + then show "C i \ C (Suc i)" by (auto dest: monotoneD) +qed + +lemma pair_omega_chain: + assumes "r a a" "r b b" "r a b" shows "omega_chain {a,b} r" + using assms by (auto intro!: omega_chainI[of r "\i. if i = 0 then a else b"] monotoneI) + +text \Every omega-chain is a well-order.\ + +lemma omega_chain_imp_well_related: + fixes less_eq (infix "\" 50) + assumes A: "omega_chain A (\)" shows "well_related_set A (\)" +proof + interpret less_eq_dualize. + from A obtain f :: "nat \ 'a" where mono: "monotone_on UNIV (\) (\) f" and A: "A = range f" + by (auto elim!: omega_chainE) + fix X assume XA: "X \ A" and "X \ {}" + then obtain I where X: "X = f ` I" and I0: "I \ {}" by (auto simp: A subset_image_iff) + from order.nonempty_imp_ex_extreme[OF I0] + obtain i where "least I i" by auto + with mono + show "\e. extreme X (\) e" by (auto intro!: exI[of _ "f i"] extremeI simp: X monotoneD) +qed + +lemma (in semiattractive) omega_chain_imp_pre_well_ordered: + assumes "omega_chain A (\)" shows "pre_well_ordered_set A (\)" + apply (intro pre_well_ordered_set.intro omega_chain_imp_well_related assms).. + +lemma (in antisymmetric) omega_chain_imp_well_ordered: + assumes "omega_chain A (\)" shows "well_ordered_set A (\)" + by (intro well_ordered_set.intro omega_chain_imp_well_related assms antisymmetric_axioms) + +subsubsection \Relation image that preserves well-orderedness.\ + +definition "well_image f A (\) fa fb \ + \a b. extreme {x\A. fa = f x} (\)\<^sup>- a \ extreme {y\A. fb = f y} (\)\<^sup>- b \ a \ b" + for less_eq (infix "\" 50) + +lemmas well_imageI = well_image_def[unfolded atomize_eq, THEN iffD2, rule_format] +lemmas well_imageD = well_image_def[unfolded atomize_eq, THEN iffD1, rule_format] + +lemma (in pre_well_ordered_set) + well_image_well_related: "pre_well_ordered_set (f`A) (well_image f A (\))" +proof- + interpret less_eq_dualize. + interpret im: transitive "f`A" "well_image f A (\)" + proof (safe intro!: transitiveI well_imageI) + interpret less_eq_dualize. + fix x y z a c + assume fxfy: "well_image f A (\) (f x) (f y)" + and fyfz: "well_image f A (\) (f y) (f z)" + and xA: "x \ A" and yA: "y \ A" and zA: "z \ A" + and a: "extreme {a \ A. f x = f a} (\) a" + and c: "extreme {c \ A. f z = f c} (\) c" + have "\b. extreme {b \ A. f y = f b} (\) b" + apply (rule nonempty_imp_ex_extreme) using yA by auto + then obtain b where b: "extreme {b \ A. f y = f b} (\) b" by auto + from trans[OF well_imageD[OF fxfy a b] well_imageD[OF fyfz b c]] a b c + show "a \ c" by auto + qed + interpret im: well_related_set "f`A" "well_image f A (\)" + proof + fix fX + assume fXfA: "fX \ f ` A" and fX0: "fX \ {}" + define X where "X \ {x\A. f x \ fX}" + with fXfA fX0 have XA: "X \ A" and "X \ {}" by (auto simp: ex_in_conv[symmetric]) + from nonempty_imp_ex_extreme[OF this] obtain e where Xe: "extreme X (\) e" by auto + with XA have eA: "e \ A" by auto + from fXfA have fX: "f ` X = fX" by (auto simp: X_def intro!: equalityI) + show "\fe. extreme fX (well_image f A (\))\<^sup>- fe" + proof (safe intro!: exI extremeI elim!: subset_imageE) + from Xe fX show fefX: "f e \ fX" by auto + fix fx assume fxfX: "fx \ fX" + show "well_image f A (\) (f e) fx" + proof (intro well_imageI) + fix a b + assume fea: "extreme {a \ A. f e = f a} (\) a" + and feb: "extreme {b \ A. fx = f b} (\) b" + from fea eA have aA: "a \ A" and ae: "a \ e" by auto + from feb fxfX have bA: "b \ A" and bX: "b \ X" by (auto simp: X_def) + with Xe have eb: "e \ b" by auto + from trans[OF ae eb aA eA bA] + show "a \ b". + qed + qed + qed + show ?thesis by unfold_locales +qed + + +end \ No newline at end of file diff --git a/thys/Complete_Non_Orders/document/root.bib b/thys/Complete_Non_Orders/document/root.bib --- a/thys/Complete_Non_Orders/document/root.bib +++ b/thys/Complete_Non_Orders/document/root.bib @@ -1,319 +1,426 @@ %% This BibTeX bibliography file was created using BibDesk. -%% http://bibdesk.sourceforge.net/ +%% https://bibdesk.sourceforge.io/ -%% Created for Jeremy Dubut at 2019-03-29 18:38:21 +0900 +%% Created for Jeremy Dubut at 2023-05-15 08:45:25 +0900 %% Saved with string encoding Unicode (UTF-8) -@InProceedings{YamadaD2019, - author = {Akihisa Yamada and J{\'e}r{\'e}my Dubut}, - title = {{Complete Non-Orders and Fixed Points}}, - booktitle = {10th International Conference on Interactive Theorem Proving (ITP 2019)}, - pages = {30:1--30:16}, - series = {Leibniz International Proceedings in Informatics (LIPIcs)}, - ISBN = {978-3-95977-122-1}, - ISSN = {1868-8969}, - year = {2019}, - volume = {141}, - editor = {John Harrison and John O'Leary and Andrew Tolmach}, - publisher = {Schloss Dagstuhl--Leibniz-Zentrum fuer Informatik}, - address = {Dagstuhl, Germany}, - URL = {http://drops.dagstuhl.de/opus/volltexte/2019/11085}, - URN = {urn:nbn:de:0030-drops-110852}, - doi = {10.4230/LIPIcs.ITP.2019.30}, - annote = {Keywords: Order Theory, Lattice Theory, Fixed-Points, Isabelle/HOL} -} -@Inbook{Bergman2015, -author="Bergman, George M.", -title="Lattices, Closure Operators, and Galois Connections", -bookTitle="An Invitation to General Algebra and Universal Constructions", -year="2015", -publisher="Springer International Publishing", -address="Cham", -pages="173--212", -isbn="978-3-319-11478-1", -doi="10.1007/978-3-319-11478-1_6", -url="https://doi.org/10.1007/978-3-319-11478-1_6" -} + + +@inproceedings{yamada23, + author = {Akihisa Yamada and J\'{e}r\'{e}my Dubut}, + booktitle = {Proceedings of the fourteenth conference on Interactive Theorem Proving (ITP'23)}, + date-added = {2023-05-15 08:42:28 +0900}, + date-modified = {2023-05-15 08:45:25 +0900}, + title = {{F}ormalizing {R}esults on {D}irected {S}ets in {I}sabelle/{HOL}}, + year = {2023}} + +@article{DubutY21, + author = {J\'{e}r\'{e}my Dubut and Akihisa Yamada}, + date-added = {2021-01-14 11:30:49 +0900}, + date-modified = {2023-05-15 08:41:50 +0900}, + journal = {Logical Methods in Computer Science}, + number = {1}, + title = {Fixed {P}oints {T}heorems for {N}on-{T}ransitive {R}elations}, + volume = {18}, + year = {2021}} + +@inbook{granas03, + author = {Andrzej Granas and James Dugundji}, + booktitle = {Fixed Point Theory}, + doi = {10.1007/978-0-387-21593-8_2}, + publisher = {Springer}, + series = {Monographs in Mathematics}, + title = {Elementary Fixed Point Theorems}, + year = 2003, + bdsk-url-1 = {https://doi.org/10.1007/978-0-387-21593-8_2}} + +@article{kantorovitch39, + author = {L. Kantorovitch}, + doi = {doi:10.1007/BF02547750}, + journal = {Acta Math.}, + pages = {63-97}, + title = {The method of successive approximations for functional equations}, + volume = {71}, + year = {1939}, + bdsk-url-1 = {https://doi.org/10.1007/BF02547750}} + +@article{jachymski00, + author = {Jacek Jachymski and Leslaw Gajek and Piotr Pokarowski}, + date-added = {2020-09-11 15:07:37 +0900}, + date-modified = {2020-09-11 15:09:39 +0900}, + journal = {Bulletin of the Australian Mathematical Society}, + number = {2}, + pages = {247--261}, + title = {The {T}arski--{K}antorovitch prinicple and the theory of iterated function systems}, + volume = {61}, + year = {2000}} + +@techreport{scott71, + author = {Dana Scott and Christopher Strachey}, + date-added = {2020-09-09 15:50:30 +0900}, + date-modified = {2020-09-09 15:53:17 +0900}, + institution = {Oxford Programming Research Group}, + number = {PRG-6}, + title = {Toward a mathematical semantics for computer languages}, + type = {Technical Monograph}, + year = {1971}} + +@inproceedings{cousot77, + author = {Patrick Cousot and Radhia Cousot}, + booktitle = {Proceedings of the 4th ACM Symposium on Principles of Programming Languages (POPL'77)}, + date-added = {2020-09-09 15:42:54 +0900}, + date-modified = {2020-09-09 15:45:38 +0900}, + pages = {238--252}, + publisher = {ACM Press}, + title = {{A}bstract {I}nterpretation: {A} {U}nified {L}attice {M}odel for {S}tatic {A}nalysis of {P}rograms by {C}onstruction or {A}pproximation of {F}ixpoints}, + year = {1977}} + +@inproceedings{yamada19, + author = {Akihisa Yamada and J\'{e}r\'{e}my Dubut}, + booktitle = {Proceedings of the 10th International Conference on Interactive Theorem Proving (ITP 2019)}, + date-added = {2020-08-31 14:39:28 +0900}, + date-modified = {2020-08-31 14:58:11 +0900}, + pages = {30:1--30:16}, + publisher = {Leibniz-Zentrum f{\"u}r Informatik}, + series = {Leibniz International Proceedings in Informatics}, + title = {Complete {N}on-{O}rders and {F}ixed {P}oints}, + volume = {141}, + year = {2019}} + +@inproceedings{grall10, + author = {Herv\'e Grall}, + booktitle = {Fixed Points in Computer Science 2010}, + pages = {41-46}, + title = {Proving fixed points}, + year = 2010} + +@article{bourbaki49, + author = {Nicolas Bourbaki}, + journal = {Archiv der Mathematik}, + number = 6, + pages = {434--437}, + title = {Sur le th\'eor\`eme de Zorn}, + volume = 2, + year = 1949} + +@misc{pataraia97, + author = {Dito Pataraia}, + date-added = {2020-08-19 11:52:38 +0900}, + date-modified = {2020-08-19 11:54:15 +0900}, + howpublished = {Presented in the 65th Peripatetic Seminar on Sheaves and Logic, in Aarhus, Denmark}, + title = {A constructive proof of {T}arski's fixed-point theorem for dcpo's}, + year = {1997}} + +@article{markowsky76, + author = {George Markowsky}, + date-added = {2020-08-19 11:47:04 +0900}, + date-modified = {2020-08-19 11:49:10 +0900}, + journal = {Algebra Universalis}, + pages = {53--68}, + title = {Chain{-}complete posets and directed sets with applications}, + volume = {6}, + year = {1976}} + +@book{Bergman2015, + address = {Cham}, + author = {Bergman, George M.}, + doi = {10.1007/978-3-319-11478-1}, + isbn = {978-3-319-11478-1}, + publisher = {Springer International Publishing}, + title = {An Invitation to General Algebra and Universal Constructions}, + year = {2015}, + bdsk-url-1 = {https://doi.org/10.1007/978-3-319-11478-1}} @book{Schmidt1993, -author="Schmidt, Gunther -and Str{\"o}hlein, Thomas", -title="Relations and Graphs: Discrete Mathematics for Computer Scientists", -year="1993", -publisher="Springer Berlin Heidelberg", -address="Berlin, Heidelberg", -pages="5--27", -} + address = {Berlin, Heidelberg}, + author = {Schmidt, Gunther and Str{\"o}hlein, Thomas}, + pages = {5--27}, + publisher = {Springer Berlin Heidelberg}, + title = {Relations and Graphs: Discrete Mathematics for Computer Scientists}, + year = {1993}} @book{gierz03, - Author = {G. Gierz and K. H. Hofmann and K. Keimel and J. D. Lawson and M. W. Mislove and D. S. Scott}, - Date-Added = {2019-03-29 18:36:56 +0900}, - Date-Modified = {2019-03-29 18:38:15 +0900}, - Publisher = {Cambridge University Press}, - Title = {Continuous Lattices and Domains}, - Year = {2003}} + author = {G. Gierz and K. H. Hofmann and K. Keimel and J. D. Lawson and M. W. Mislove and D. S. Scott}, + date-added = {2019-03-29 18:36:56 +0900}, + date-modified = {2019-03-29 18:38:15 +0900}, + doi = {10.1017/CBO9780511542725}, + publisher = {Cambridge University Press}, + title = {Continuous Lattices and Domains}, + year = {2003}, + bdsk-url-1 = {https://doi.org/10.1017/CBO9780511542725}} @book{abramski94, - Author = {Samson Abramsky and Achim Jung}, - Date-Added = {2019-03-29 18:27:24 +0900}, - Date-Modified = {2019-03-29 18:30:21 +0900}, - Number = {III}, - Publisher = {Oxford University Press}, - Series = {Handbook of Logic in Computer Science}, - Title = {Domain Theory}, - Year = {1994}} + author = {Samson Abramsky and Achim Jung}, + date-added = {2019-03-29 18:27:24 +0900}, + date-modified = {2019-03-29 18:30:21 +0900}, + number = {III}, + publisher = {Oxford University Press}, + series = {Handbook of Logic in Computer Science}, + title = {Domain Theory}, + year = {1994}} @book{davey02, - Author = {B. A. Davey and H. A. Priestley}, - Date-Added = {2019-03-29 18:24:39 +0900}, - Date-Modified = {2019-03-29 18:25:37 +0900}, - Publisher = {Cambridge University Press}, - Title = {Introduction to Lattices and Order}, - Year = {2002}} + author = {B. A. Davey and H. A. Priestley}, + date-added = {2019-03-29 18:24:39 +0900}, + date-modified = {2019-03-29 18:25:37 +0900}, + doi = {10.1017/CBO9780511809088}, + publisher = {Cambridge University Press}, + title = {Introduction to Lattices and Order}, + year = {2002}, + bdsk-url-1 = {https://doi.org/10.1017/CBO9780511809088}} @article{mashburn83, - Author = {J. D. Mashburn}, - Date-Added = {2019-03-27 22:36:00 +0900}, - Date-Modified = {2019-03-27 22:38:19 +0900}, - Journal = {Houston Journal of Mathematics}, - Number = {2}, - Pages = {231--244}, - Title = {The least fixed point property for omega-chain continuous functions}, - Volume = {9}, - Year = {1983}} + author = {J. D. Mashburn}, + date-added = {2019-03-27 22:36:00 +0900}, + date-modified = {2019-03-27 22:38:19 +0900}, + journal = {Houston Journal of Mathematics}, + number = {2}, + pages = {231--244}, + title = {The least fixed point property for omega-chain continuous functions}, + volume = {9}, + year = {1983}} @article{tarski55, - Author = {Alfred Tarski}, - Date-Added = {2019-03-27 11:16:30 +0900}, - Date-Modified = {2019-03-27 11:17:45 +0900}, - Journal = {Pacific Journal of Mathematics}, - Number = {2}, - Pages = {285--309}, - Title = {A lattice-theoretical fixpoint theorem and its applications}, - Volume = {5}, - Year = {1955}} + author = {Alfred Tarski}, + date-added = {2019-03-27 11:16:30 +0900}, + date-modified = {2019-03-27 11:17:45 +0900}, + journal = {Pacific Journal of Mathematics}, + number = {2}, + pages = {285--309}, + title = {A lattice-theoretical fixpoint theorem and its applications}, + volume = {5}, + year = {1955}} @article{trellis, - Author = {Skala, H.L.}, - Doi = {10.1007/BF02944982}, - Issue = 1, - Journal = {Algebra Univ.}, - Pages = {218--233}, - Title = {Trellis theory}, - Volume = 1, - Year = 1971, - Bdsk-Url-1 = {https://doi.org/10.1007/BF02944982}} + author = {Skala, H.L.}, + doi = {10.1007/BF02944982}, + issue = 1, + journal = {Algebra Univ.}, + pages = {218--233}, + publisher = {Birkh{\"a}user-Verlag}, + title = {Trellis theory}, + volume = 1, + year = 1971, + bdsk-url-1 = {https://doi.org/10.1007/BF02944982}} @article{SM13, - Author = {Stouti, Abdelkader and Maaden, Abdelhakim}, - Doi = {10.4067/S0716-09172013000400008}, - Journal = {Proyecciones. Revista de Matem{\'a}tica}, - Number = {4}, - Pages = {409-418}, - Title = {Fixed points and common fixed points theorems in pseudo-ordered sets}, - Volume = {32}, - Year = {2013}, - Bdsk-Url-1 = {https://doi.org/10.4067/S0716-09172013000400008}} + author = {Stouti, Abdelkader and Maaden, Abdelhakim}, + doi = {10.4067/S0716-09172013000400008}, + journal = {Proyecciones}, + number = {4}, + pages = {409-418}, + publisher = {Universidad Cat{\'o}lica del Norte, Departamento de Matem{\'a}ticas}, + title = {Fixed points and common fixed points theorems in pseudo-ordered sets}, + volume = {32}, + year = {2013}, + bdsk-url-1 = {https://doi.org/10.4067/S0716-09172013000400008}} @article{LN83, - Author = {Leutola, K and Nieminen, J}, - Journal = {Algebra Universalis}, - Number = {1}, - Pages = {344--354}, - Publisher = {Springer}, - Title = {Posets and generalized lattices}, - Volume = {16}, - Year = {1983}} + author = {Leutola, K. and Nieminen, J.}, + journal = {Algebra Universalis}, + number = {1}, + pages = {344--354}, + publisher = {Springer}, + title = {Posets and generalized lattices}, + volume = {16}, + year = {1983}} @article{Bhatta05, - Author = {Bhatta, S Parameshwara}, - Journal = {Czechoslovak Mathematical Journal}, - Number = {2}, - Pages = {365--369}, - Publisher = {Springer}, - Title = {Weak chain-completeness and fixed point property for pseudo-ordered sets}, - Volume = {55}, - Year = {2005}} + author = {Bhatta, S. Parameshwara}, + journal = {Czechoslovak Mathematical Journal}, + number = {2}, + pages = {365--369}, + publisher = {Springer}, + title = {Weak chain-completeness and fixed point property for pseudo-ordered sets}, + volume = {55}, + year = {2005}} -@article{PG11, - Author = {Parameshwara Bhatta, S and George, Shiju}, - Journal = {Algebra and Discrete Mathematics}, - Number = {1}, - Pages = {17--22}, - Publisher = {Луганский национальный университет им. Т. Шевченко}, - Title = {Some fixed point theorems for pseudo ordered sets}, - Volume = {11}, - Year = {2011}} +@article{BG11, + author = {Bhatta, S. Parameshwara and George, Shiju}, + date-modified = {2020-09-09 14:15:07 +0900}, + journal = {Algebra and Discrete Mathematics}, + number = {1}, + pages = {17--22}, + publisher = {Луганский национальный университет им. Т. Шевченко}, + title = {Some fixed point theorems for pseudo ordered sets}, + volume = {11}, + year = {2011}} @article{flyspeck, - Author = {Hales, Thomas and Adams, Mark and Bauer, Gertrud and Dang, Tat Dat and Harrison, John and Le Truong, Hoang and Kaliszyk, Cezary and Magron, Victor and McLaughlin, Sean and Nguyen, Tat Thang and others}, - Journal = {Forum of Mathematics, Pi}, - Publisher = {Cambridge University Press}, - Title = {A formal proof of the {K}epler conjecture}, - Volume = {5}, - Year = {2017}} + author = {Hales, Thomas and Adams, Mark and Bauer, Gertrud and Dang, Tat Dat and Harrison, John and Le Truong, Hoang and Kaliszyk, Cezary and Magron, Victor and McLaughlin, Sean and Nguyen, Tat Thang and others}, + doi = {10.1017/fmp.2017.1}, + journal = {Forum of Mathematics, Pi}, + pages = {e2}, + publisher = {Cambridge University Press}, + title = {A formal proof of the {K}epler conjecture}, + volume = {5}, + year = {2017}, + bdsk-url-1 = {https://doi.org/10.1017/fmp.2017.1}} @article{4color, - Author = {Gonthier, Georges}, - Journal = {Notices of the AMS}, - Number = {11}, - Pages = {1382--1393}, - Title = {Formal proof -- the four-color theorem}, - Volume = {55}, - Year = {2008}} + author = {Gonthier, Georges}, + journal = {Notices of the AMS}, + number = {11}, + pages = {1382--1393}, + title = {Formal proof -- the four-color theorem}, + volume = {55}, + year = {2008}} @inproceedings{sel4, - Author = {Klein, Gerwin and Elphinstone, Kevin and Heiser, Gernot and Andronick, June and Cock, David and Derrin, Philip and Elkaduwe, Dhammika and Engelhardt, Kai and Kolanski, Rafal and Norrish, Michael and others}, - Booktitle = {Proceedings of the ACM SIGOPS 22nd Symposium on Operating Systems principles}, - Organization = {ACM}, - Pages = {207--220}, - Title = {se{L}4: Formal verification of an {OS} kernel}, - Year = {2009}} + author = {Klein, Gerwin and Elphinstone, Kevin and Heiser, Gernot and Andronick, June and Cock, David and Derrin, Philip and Elkaduwe, Dhammika and Engelhardt, Kai and Kolanski, Rafal and Norrish, Michael and Sewell, Thomas and Tuch, Harvey and Wiwood, Simon}, + booktitle = {Proceedings of the ACM SIGOPS 22nd Symposium on Operating Systems Principles (SOSP 2009)}, + doi = {10.1145/1629575.1629596}, + organization = {ACM}, + pages = {207--220}, + title = {se{L}4: Formal verification of an {OS} kernel}, + year = {2009}, + bdsk-url-1 = {https://doi.org/10.1145/1629575.1629596}} @inproceedings{isabelle/jedit, - Author = {Wenzel, Makarius}, - Booktitle = {International Conference on Intelligent Computer Mathematics}, - Organization = {Springer}, - Pages = {468--471}, - Title = {{I}sabelle/j{E}dit--a Prover {IDE} within the {PIDE} framework}, - Year = {2012}} + author = {Wenzel, Makarius}, + booktitle = {Proceedings of the 5th Conferences on Intelligent Computer Mathematics (CICM 2012)}, + doi = {10.1007/978-3-642-31374-5_38}, + pages = {468--471}, + publisher = {Springer Berlin Heidelberg}, + series = {LNCS}, + title = {{I}sabelle/j{E}dit -- a Prover {IDE} within the {PIDE} framework}, + volume = 7362, + year = {2012}, + bdsk-url-1 = {https://doi.org/10.1007/978-3-642-31374-5_38}} @inproceedings{quickcheck, - Acmid = {1030056}, - Address = {Washington, DC, USA}, - Author = {Berghofer, Stefan and Nipkow, Tobias}, - Booktitle = {Proceedings of the Software Engineering and Formal Methods, Second International Conference}, - Doi = {10.1109/SEFM.2004.36}, - OPIsbn = {0-7695-2222-X}, - Numpages = {10}, - Pages = {230--239}, - Publisher = {IEEE Computer Society}, - Series = {SEFM '04}, - Title = {Random Testing in {I}sabelle/{HOL}}, - Year = {2004}, - Bdsk-Url-1 = {http://dx.doi.org/10.1109/SEFM.2004.36}} + acmid = {1030056}, + author = {Berghofer, Stefan and Nipkow, Tobias}, + booktitle = {Proceedings of the 2nd International Conference on Software Engineering and Formal Methods (SEFM 2004)}, + doi = {10.1109/SEFM.2004.36}, + numpages = {10}, + opaddress = {Washington, DC, USA}, + opisbn = {0-7695-2222-X}, + pages = {230--239}, + publisher = {IEEE Computer Society}, + title = {Random Testing in {I}sabelle/{HOL}}, + year = {2004}, + bdsk-url-1 = {https://doi.org/10.1109/SEFM.2004.36}} @inproceedings{nitpick, - Address = {Berlin, Heidelberg}, - Author = {Blanchette, Jasmin Christian and Nipkow, Tobias}, - Booktitle = {Interactive Theorem Proving}, - Editor = {Kaufmann, Matt and Paulson, Lawrence C.}, - OPIsbn = {978-3-642-14052-5}, - Pages = {131--146}, - Publisher = {Springer Berlin Heidelberg}, - Title = {Nitpick: A Counterexample Generator for Higher-Order Logic Based on a Relational Model Finder}, - Year = {2010}} + author = {Blanchette, Jasmin Christian and Nipkow, Tobias}, + booktitle = {Proceedings of the 1st International Conference on Interactive Theorem Proving (ITP 2010)}, + doi = {10.1007/978-3-642-14052-5_11}, + editor = {Kaufmann, Matt and Paulson, Lawrence C.}, + opaddress = {Berlin, Heidelberg}, + pages = {131--146}, + publisher = {Springer Berlin Heidelberg}, + series = {LNCS}, + title = {Nitpick: A Counterexample Generator for Higher-Order Logic Based on a Relational Model Finder}, + volume = 6172, + year = {2010}, + bdsk-url-1 = {https://doi.org/10.1007/978-3-642-14052-5_11}} @inproceedings{codegen, - Author = {Haftmann, Florian and Nipkow, Tobias}, - Booktitle = {Theorem Proving in Higher Order Logics (TPHOLs 2007)}, - Pages = {128--143}, - Series = {Lecture Notes in Computer Science}, - Title = {A code generator framework for {I}sabelle/{HOL}}, - Volume = {4732}, - Year = {2007}} + author = {Haftmann, Florian and Nipkow, Tobias}, + booktitle = {Theorem Proving in Higher Order Logics: Emerging Trends}, + editor = {Klaus Schneider and Jens Brandt}, + opnumber = {364/07}, + pages = {128--143}, + publisher = {Department of Computer Science, University of Kaiserslautern}, + title = {A code generator framework for {I}sabelle/{HOL}}, + year = {2007}} @inproceedings{locale, - Address = {Berlin, Heidelberg}, - Author = {Ballarin, Clemens}, - Booktitle = {Mathematical Knowledge Management}, - Doi = {10.1007/11812289_4}, - Editor = {Borwein, Jonathan M. and Farmer, William M.}, - OPIsbn = {978-3-540-37106-9}, - Pages = {31--43}, - Publisher = {Springer Berlin Heidelberg}, - Title = {Interpretation of Locales in {I}sabelle: Theories and Proof Contexts}, - Year = {2006}, - Bdsk-Url-1 = {https://doi.org/10.1007/11812289_4}} + author = {Ballarin, Clemens}, + booktitle = {Proceedings of the 5th International Conference on Mathematical Knowledge Management (MKM 2006)}, + doi = {10.1007/11812289_4}, + editor = {Borwein, Jonathan M. and Farmer, William M.}, + opaddress = {Berlin, Heidelberg}, + opisbn = {978-3-540-37106-9}, + pages = {31--43}, + publisher = {Springer Berlin Heidelberg}, + series = {LNCS}, + title = {Interpretation of Locales in {I}sabelle: Theories and Proof Contexts}, + volume = 4108, + year = {2006}, + bdsk-url-1 = {https://doi.org/10.1007/11812289_4}} @inproceedings{Kammuller00, - Address = {Berlin, Heidelberg}, - Author = {Kamm{\"u}ller, Florian}, - Booktitle = {Automated Deduction - CADE-17}, - Doi = {10.1007/10721959_7}, - Editor = {McAllester, David}, - OPIsbn = {978-3-540-45101-3}, - Pages = {99--114}, - Publisher = {Springer Berlin Heidelberg}, - Title = {Modular Reasoning in {I}sabelle}, - Year = {2000}, - Bdsk-Url-1 = {https://doi.org/10.1007/10721959_7}} + author = {Kamm{\"u}ller, Florian}, + booktitle = {Proceedings of the 17th International Conference on Automated Deduction (CADE-17)}, + doi = {10.1007/10721959_7}, + editor = {McAllester, David}, + opaddress = {Berlin, Heidelberg}, + pages = {99--114}, + publisher = {Springer Berlin Heidelberg}, + series = {LNCS}, + title = {Modular Reasoning in {I}sabelle}, + volume = 1831, + year = {2000}, + bdsk-url-1 = {https://doi.org/10.1007/10721959_7}} + @inproceedings{isafor, - title={Certification of termination proofs using {CeTA}}, - author={Thiemann, Ren{\'e} and Sternagel, Christian}, - booktitle={International Conference on Theorem Proving in Higher Order Logics}, - pages={452--468}, - year={2009}, - organization={Springer} -} + author = {Thiemann, Ren{\'e} and Sternagel, Christian}, + booktitle = {International Conference on Theorem Proving in Higher Order Logics}, + organization = {Springer}, + pages = {452--468}, + title = {Certification of termination proofs using {CeTA}}, + year = {2009}} + @book{Isabelle, - Author = {T.~Nipkow and L.C.~Paulson and M.~Wenzel}, - Publisher = {Springer}, - Series = {LNCS}, - Title = {{Isabelle/HOL} -- A Proof Assistant for Higher-Order Logic}, - Volume = 2283, - Year = 2002 -} + author = {T.~Nipkow and L.C.~Paulson and M.~Wenzel}, + doi = {10.1007/3-540-45949-9}, + publisher = {Springer}, + series = {LNCS}, + title = {{Isabelle/HOL} -- A Proof Assistant for Higher-Order Logic}, + volume = 2283, + year = 2002, + bdsk-url-1 = {https://doi.org/10.1007/3-540-45949-9}} + @inproceedings{hol4, - title={A brief overview of {HOL4}}, - author={Slind, Konrad and Norrish, Michael}, - booktitle={International Conference on Theorem Proving in Higher Order Logics}, - pages={28--32}, - year={2008}, - organization={Springer} -} + author = {Slind, Konrad and Norrish, Michael}, + booktitle = {International Conference on Theorem Proving in Higher Order Logics}, + organization = {Springer}, + pages = {28--32}, + title = {A brief overview of {HOL4}}, + year = {2008}} + @inproceedings{hol-light, - title={{HOL} light: An overview}, - author={Harrison, John}, - booktitle={International Conference on Theorem Proving in Higher Order Logics}, - pages={60--66}, - year={2009}, - organization={Springer} -} + author = {Harrison, John}, + booktitle = {International Conference on Theorem Proving in Higher Order Logics}, + organization = {Springer}, + pages = {60--66}, + title = {{HOL} light: An overview}, + year = {2009}} + @inproceedings{agda, - title={A brief overview of {A}gda -- a functional language with dependent types}, - author={Bove, Ana and Dybjer, Peter and Norell, Ulf}, - booktitle={International Conference on Theorem Proving in Higher Order Logics}, - pages={73--78}, - year={2009}, - organization={Springer} -} + author = {Bove, Ana and Dybjer, Peter and Norell, Ulf}, + booktitle = {International Conference on Theorem Proving in Higher Order Logics}, + organization = {Springer}, + pages = {73--78}, + title = {A brief overview of {A}gda -- a functional language with dependent types}, + year = {2009}} + @book{coq, - title={Certified programming with dependent types: a pragmatic introduction to the Coq proof assistant}, - author={Chlipala, Adam}, - year={2013}, - publisher={MIT Press} -} + author = {Chlipala, Adam}, + publisher = {MIT Press}, + title = {Certified programming with dependent types: a pragmatic introduction to the {C}oq proof assistant}, + year = {2013}} + @inproceedings{compcert, - title={Formal {C} semantics: {C}omp{C}ert and the {C} standard}, - author={Krebbers, Robbert and Leroy, Xavier and Wiedijk, Freek}, - booktitle={International Conference on Interactive Theorem Proving}, - pages={543--548}, - year={2014}, - organization={Springer} -} + author = {Krebbers, Robbert and Leroy, Xavier and Wiedijk, Freek}, + booktitle = {International Conference on Interactive Theorem Proving}, + organization = {Springer}, + pages = {543--548}, + title = {Formal {C} semantics: {C}omp{C}ert and the {C} standard}, + year = {2014}} + @inproceedings{sledgehammer, - title={Sledgehammer: judgement day}, - author={B{\"o}hme, Sascha and Nipkow, Tobias}, - booktitle={International Joint Conference on Automated Reasoning}, - pages={107--121}, - year={2010}, - organization={Springer} -} - -@inproceedings{grall10, - author = {Herv{\'{e}} Grall}, - editor = {Luigi Santocanale}, - title = {Proving Fixed Points}, - booktitle = {7th Workshop on Fixed Points in Computer Science, {FICS} 2010, Brno, - Czech Republic, August 21-22, 2010}, - pages = {41--46}, - publisher = {Laboratoire d'Informatique Fondamentale de Marseille}, - year = {2010}, - url = {https://hal.archives-ouvertes.fr/hal-00512377/document\#page=42}, - timestamp = {Tue, 21 Jul 2020 00:40:32 +0200}, - biburl = {https://dblp.org/rec/conf/fics/Grall10.bib}, - bibsource = {dblp computer science bibliography, https://dblp.org} -} + author = {B{\"o}hme, Sascha and Nipkow, Tobias}, + booktitle = {Proceedings of the 5th International Joint Conference on Automated Reasoning (IJCAR 2010)}, + doi = {10.1007/978-3-642-14203-1_9}, + pages = {107--121}, + publisher = {Springer Berlin Heidelberg}, + series = {LNCS}, + title = {Sledgehammer: Judgement Day}, + volume = 6173, + year = {2010}, + bdsk-url-1 = {https://doi.org/10.1007/978-3-642-14203-1_9}} diff --git a/thys/Complete_Non_Orders/document/root.tex b/thys/Complete_Non_Orders/document/root.tex --- a/thys/Complete_Non_Orders/document/root.tex +++ b/thys/Complete_Non_Orders/document/root.tex @@ -1,196 +1,210 @@ \documentclass[11pt,a4paper]{article} -\usepackage[T1]{fontenc} \usepackage{isabelle,isabellesym} \usepackage{amssymb,amsmath,stmaryrd} \usepackage{tikz} \usetikzlibrary{backgrounds} \usetikzlibrary{positioning} \usetikzlibrary{shapes} % this should be the last package used \usepackage{pdfsetup} % urls in roman style, theory text in math-similar italics \urlstyle{rm} \isabellestyle{it} \newcommand\SLE{\sqsubseteq} \newcommand\Nat{\mathbb{N}} +\newcommand\SLT{\sqsubset} +\newcommand\SUP{\bigsqcup} \makeatletter \def\tp@#1#2{\@ifnextchar[{\tp@@{#1}{#2}}{\tp@@@{#1}{#2}}} \def\tp@@#1#2[#3]#4{#3#1\def\mid{\mathrel{#3|}}#4#3#2} \def\tp@@@#1#2#3{\bgroup\left#1\def\mid{\;\middle|\;}#3\right#2\egroup} \def\pa{\tp@()} \def\tp{\tp@\langle\rangle} \def\set{\tp@\{\}} \makeatother \begin{document} \title{Complete Non-Orders and Fixed Points} -\author{Akihisa Yamada and Jérémy Dubut} +\author{Akihisa Yamada and J\'er\'emy Dubut} \maketitle \begin{abstract} - We develop an Isabelle/HOL library of order-theoretic concepts, such as various completeness conditions and - fixed-point theorems. We keep our formalization as general as possible: we reprove several well-known results - about complete orders, often without any properties of ordering, thus complete non-orders. In particular, we - generalize the Knaster--Tarski theorem so that we ensure the existence of a quasi-fixed point of monotone maps - over complete non-orders, and show that the set of quasi-fixed points is complete under a mild condition---% - attractivity---which is implied by either antisymmetry or transitivity. This result generalizes and strengthens a result - by Stauti and Maaden. Finally, we recover Kleene's fixed-point theorem for omega-complete non-orders, again - using attractivity to prove that Kleene's fixed points are least quasi-fixed points. +We develop an Isabelle/HOL library +of order-theoretic concepts, such as various completeness conditions and +fixed-point theorems. +We keep our formalization as general as possible: +we reprove several well-known results about complete orders, +often with only antisymmetry or attractivity, a mild condition implied by either +antisymmetry or transitivity. +In particular, we generalize various theorems ensuring +the existence of a quasi-fixed point of monotone maps over complete relations, +and show that the set of (quasi-)fixed points is itself complete. +This result generalizes and strengthens theorems of Knaster--Tarski, Bourbaki--Witt, Kleene, Markowsky, Pataraia, Mashburn, Bhatta--George, and Stouti--Maaden. \end{abstract} \tableofcontents \section{Introduction} The main driving force towards mechanizing mathematics using proof assistants +% For a journal version: +%such as Coq~\cite{coq}, Agda~\cite{agda}, HOL-Light~\cite{hol-light}, and Isabelle/HOL~\cite{isabelle}, has been the reliability they offer, +%Once a theorem is formally proved using a proof assistant, +%we can certainly trust the theorem without understanding how it is proved, +%provided the claim itself is correctly formalized. +%The most prominent of such examples are +%the four-color theorem proved using Coq~\cite{4color}, and +%the Kepler theorem proved with HOL-Light and Isabelle/HOL~\cite{flyspeck}. +%Moreover the reliability of proof assistants is utilized for +%developing bug-free software, for instance, +%the CompCert C compliler~\cite{compcert}, +%the seL4 OS kernel~\cite{sel4}, +%the IsaFoR/CeTA certifier for validating program analyzers~\cite{isafor}, +%... +% exemplified prominently by~\cite{4color},~\cite{flyspeck},~\cite{sel4}, etc. -In this work, we utilize another aspect of Isabelle/JEdit~\cite{isabelle/jedit} -as engineering tools for developing mathematical theories. -We formalize order-theoretic concepts and results, -adhering to an \emph{as-general-as-possible} approach: -most results concerning order-theoretic completeness and fixed-point theorems -are proved without assuming the underlying relations to be orders (non-orders). +In this work, we utilize another aspect of proof assistants: +they are also engineering tools for developing mathematical theories. + +\emph{Fixed-point theorems} are important in computer science, such as +in denotational semantics \cite{scott71} and +in abstract interpretation \cite{cousot77}, as they allow the definition +of semantics of loops and recursive functions. +The Knaster--Tarski theorem~\cite{tarski55} +shows that +any monotone map $f : A \to A$ over complete lattice $(A,\SLE)$ has a fixed point, +and the set of fixed points forms also a complete lattice. +The result was generalized in various ways: +Markowsky~\cite{markowsky76} +showed a corresponding result for \emph{chain-complete} posets. +The proof uses +the Bourbaki--Witt theorem~\cite{bourbaki49}, +stating that any inflationary map over a chain-complete poset has a fixed point. +The original proof of the latter is non-elementary in +the sense that it relies on ordinals and Hartogs' theorem. +Pataraia~\cite{pataraia97} +gave an elementary proof that +monotone maps over \emph{pointed directed-complete} poset +has a fixed point. +Fixed points are studied also for \emph{pseudo-orders}~\cite{trellis}, +relaxing transitivity. +Stouti and Maaden~\cite{SM13} showed that every monotone map over a complete pseudo-order has a (least) fixed point. +Markowsky's result was also generalized to +\emph{weak} chain-complete pseudo-orders by Bhatta and George~\cite{Bhatta05,BG11}. + +Another line of order-theoretic fixed points is the \emph{iterative} approach. +Kantorovitch showed that +for \emph{$\omega$-continuous} map $f$ +over a complete lattice,\footnote{ +More precisely, +he assumes a conditionally complete lattice defined over vectors and that +$\bot \SLE f\:\bot$ and $f\:v' \SLE v'$. +Hence $f$, which is monotone, +is a map over the complete lattice $\{v \mid \bot \SLE v \SLE v'\}$. +} +the +iteration $\bot, f\:\bot,f^2\:\bot,\dots$ converges to a fixed point +\cite[Theorem I]{kantorovitch39}. +Tarski~\cite{tarski55} also claimed a similar result for a +\emph{countably distributive} map over a +\emph{countably complete Boolean algebra}. +Kleene's fixed-point theorem states that, +for \emph{Scott-continuous} maps over pointed directed-complete posets, +the iteration converges to the least fixed point. +Finally, Mashburn~\cite{mashburn83} proved a version for +$\omega$-continuous maps over $\omega$-complete posets, +which covers Kantorovitch's, Tarski's and Kleene's claims. + + + +%In this paper, +%we formalize these fixed-point theorems in Isabelle/HOL. +%Here we adopt an \emph{as-general-as-possible} approach: +%all theorems +%are proved without assuming the underlying relations to be orders. +%One can easily find several formalizations of complete partial orders or lattices in Isabelle's standard library. +%They are, however, defined on partial orders +%and thus not directly reusable for general relations. + In particular, we provide the following: \begin{itemize} -\item -A locale-based library for binary relations, -as partly depicted in Figure~\ref{fig:non-orders}. -\item -Various completeness results that generalize known theorems in order theory: -Actually most relationships and duality of completeness conditions are proved without -\emph{any} properties of the underlying relations. +\item Several \emph{locales} that help organizing the different order-theoretic conditions, +such as reflexivity, transitivity, antisymmetry, and their combination, as well as concepts such as connex and well-related sets, analogues of chains +and well-ordered sets in a non-ordered context. +%(\prettyref{sec:prelim}). \item Existence of fixed points: -We show that a relation-preserving mapping $f : A \to A$ -over a complete non-order $\tp{A,\SLE}$ -admits a \emph{quasi-fixed point} $f(x) \sim x$, -meaning $x \SLE f(x) \wedge f(x) \SLE x$. -Clearly if $\SLE$ is antisymmetric then this implies the existence of fixed points $f(x) = x$. +We provide two proof schemes to prove that monotone or inflationary mapping +$f : A \to A$ over a complete related set $\tp{A,\SLE}$ +has a \emph{quasi-fixed point} $f\:x \sim x$, +meaning $x \SLE f\:x \mathrel\land f\:x \SLE x$, for various notions of completeness. +The first one, +%(\prettyref{sec:knaster-tarski}) +similar to the original proof by Tarski \cite{tarski55}, +does not require any ordering assumptions, +but relies on completeness with respect to all subsets. +The second one, +%(\prettyref{sec:weak_chain}), +inspired by a \emph{constructive} approach by Grall \cite{grall10}, +is a proof scheme based on the notion of derivations. +Here we demand antisymmetry (to avoid the necessity of the axiom of choice), +but can be instantiated to \emph{well-complete} sets, +a generalization of weak chain-completeness. +This also allows us to generalize Bourbaki--Witt theorem \cite{bourbaki49} to pseudo-orders. \item Completeness of the set of fixed points: -We further show that -if $\SLE$ satisfies a mild condition, which we call \emph{attractivity} and +%(\prettyref{sec:completeness}): We further show that +if $(A,\SLE)$ satisfies a mild condition, which we call \emph{attractivity} and which is implied by either transitivity or antisymmetry, -then the set of quasi-fixed points is complete. -Furthermore, we also show that if $\SLE$ is antisymmetric, -then the set of \emph{strict} fixed points $f(x) = x$ is complete. -\item Kleene-style fixed-point theorems: -For an $\omega$-complete non-order $\tp{A,\SLE}$ with a bottom element $\bot \in A$ (not necessarily unique) -and for every $\omega$-continuous map $f : A \to A$, -a supremum exists for the set $\set{ f^n(\bot) \mid n \in \Nat}$, -and it is a quasi-fixed point. -If $\SLE$ is attractive, then -the quasi-fixed points obtained this way are precisely the least quasi-fixed points. +then the set of quasi-fixed points inherits the completeness class from +$(A,\SLE)$, if it is at least well-complete. +The result instantiates to +the full completeness (generalizing Knaster--Tarski and \cite{SM13}), +directed-completeness \cite{pataraia97}, +chain-completeness \cite{markowsky76}, +and weak chain-completeness \cite{BG11}. +\item Iterative construction: +% (\prettyref{sec:kleene}): +For an $\omega$-continuous map over an $\omega$-complete related set, +we show that suprema of $\set{f^n\:\bot \mid n\in\Nat}$ are quasi-fixed points. +Under attractivity, the quasi-fixed points obtained from this scheme +are precisely the least quasi-fixed points of $f$. +This generalizes Mashburn's result, and thus ones by +Kantorovitch, Tarski and Kleene. \end{itemize} We remark that all these results would have required much more effort than we spent (if possible at all), -if we were not with the smart assistance by Isabelle. -Our workflow was often the following: first we formalize existing proofs, try relaxing assumptions, see where proof breaks, -and at some point ask for a counterexample. +if we were not with the aforementioned smart assistance by Isabelle. +Our workflow was often the following: first we formalize existing proofs, try relaxing assumptions, +see where proof breaks, and at some point ask for a counterexample. -\begin{figure} -\small -\centering -\def\isa#1{\textsf{#1}} -\def\t{-1.8} -\def\a{3.6} -\def\at{1.8} -\def\s{-3.6} -\def\st{-5.4} -\begin{tikzpicture} -\tikzstyle{every node}=[draw,ellipse] -\tikzstyle{every edge}=[draw] -\draw - (0,0) node[fill] (rel) {} - (\t,1) node (trans) {\isa{transitive}} - (0,-2) node (refl) {\isa{reflexive}} - (0,2) node (irr) {\isa{irreflexive}} - (\s,0) node (sym) {\isa{symmetric}} - (\a,0) node (anti) {\isa{antisymmetric}} - (\at,1) node (near) {\isa{near\_order}} - (\a,2) node (asym) {\isa{asymmetric}} - (\a,-2) node (pso) {\isa{pseudo\_order}} - (\at,-1) node (po) {\isa{partial\_order}} - (\t,-1) node (qo) {\isa{quasi\_order}} - (0,3) node (str) {\hspace{3em}\isa{strict\_order}\mbox{\hspace{3em}}} - (\st,-1) node (equiv) {\isa{equivalence}} - (\st,1) node (peq) {\hspace{-.8em}\isa{partial\_equivalence}\mbox{\hspace{-.8em}}} - (\st,3) node (emp) {$\emptyset$} - (\s,-2) node (tol) {\isa{tolerance}} - (\s,2) node (ntol) {$\neg$\isa{tolerance}} - ; -\draw[->] - (near) edge[color=blue] ([xshift=51,yshift=-7]str) - (irr) edge[color=red] ([xshift=-46,yshift=-8]str) - (trans) edge[color=blue] ([xshift=-51,yshift=-7]str) - (asym) edge[color=red] ([xshift=55,yshift=-7]str) - (trans) edge[color=green] (near) - (anti) edge[color=red] (near) - (irr) edge[color=green] (asym) - (anti) edge[color=blue] (asym) - (anti) edge[color=blue] (pso) - (near) edge[color=blue] (po) - (pso) edge[color=red] (po) - (refl) edge[color=green] (pso) - (trans) edge[color=blue] (qo) - (refl) edge[color=red] (qo) - (qo) edge[color=green] (po) - (qo) edge[color=green] (equiv) - (peq) edge[color=blue] (equiv) - (trans) edge[color=green] (peq) - (peq) edge[color=blue] (emp) - (str) edge[color=green] (emp) - (sym) edge[color=red] (peq) - (sym) edge[color=blue] (tol) - (sym) edge[color=blue] (ntol) - (irr) edge[color=green] (ntol) - (refl) edge[color=green] (tol) - (ntol) edge[color=red] (emp) - (tol) edge[color=red] (equiv) - (rel) edge[color=red, line width=1.5pt] (trans) - (rel) edge[color=blue, line width=1.5pt] (irr) - (rel) edge[color=blue, line width=1.5pt] (refl) - (rel) edge[color=green, line width=1.5pt] (sym) - (rel) edge[color=green, line width=1.5pt] (anti) - ; -\end{tikzpicture} -\caption{\label{fig:non-orders} -Combinations of basic properties. -The black dot around the center represents arbitrary binary relations, -and the five outgoing arrows indicate atomic assumptions. -We do not present the combination -of \isa{reflexive} and \isa{irreflexive}, which is empty, -and one of \isa{symmetric} and \isa{antisymmetric}, -which is a subset of equality. -Node ``$\neg$\isa{tolerance}'' indicates the negated relation is \isa{tolerance}, -and ``$\emptyset$'' is the empty relation. -} -\end{figure} +Concerning Isabelle formalization, one can easily find several formalizations +of complete partial orders or lattices in Isabelle?s standard library. They are, +however, defined on partial orders, either in form of classes or locales, and +thus not directly reusable for non-orders. Nevertheless we tried to make our +formalization compatible with the existing ones, and various +correspondences are ensured. -\paragraph*{Related Work} -Many attempts have been made to generalize the notion of completeness for lattices, conducted in different directions: by relaxing the notion of order itself, removing transitivity (pseudo-orders \cite{trellis}); by relaxing the notion of lattice, considering minimal upper bounds instead of least upper bounds ($\chi$-posets \cite{LN83}); by relaxing the notion of completeness, requiring the existence of least upper bounds for restricted classes of subsets (e.g., directed complete and $\omega$-complete, see \cite{davey02} for a textbook). Considering those generalizations, it was natural to prove new versions of classical fixed-point theorems for maps preserving those structures, e.g., existence of least fixed points for monotone maps on (weak chain) complete pseudo-orders \cite{Bhatta05, SM13}, construction of least fixed points for $\omega$-continuous functions for $\omega$-complete lattices \cite{mashburn83}, (weak chain) completeness of the set of fixed points for monotone functions on (weak chain) complete pseudo-orders \cite{PG11}. - -Concerning Isabelle formalization, -one can easily find several formalizations of complete partial orders or lattices in Isabelle's standard library. -They are, however, defined on partial orders, either in form of classes or locales, -and thus not directly reusable for non-orders. -Nevertheless we tried to make our formalization compatible with the existing ones, -and various correspondences are ensured. - -This work has been published in the conference paper \cite{YamadaD2019}. +This archive is the third version of this work. +The first version has been published in +the conference paper \cite{yamada19}. +The second version has been published in the +journal paper \cite{DubutY21}. +The third version is a restructuration of the second version +for future formalizations, including \cite{yamada23}. % include generated text of all theories \input{session} \bibliographystyle{abbrv} \bibliography{root} \end{document}