# HG changeset patch # User desharna # Date 1653727545 -7200 # Node ID 9e34819a7ca18e637025e2bb87ed5d3d90058ef8 # Parent 5f2a1efd05602097bca68d46cdd0f694a956a744 added lemmas Multiset.bex_{least,greatest}_element diff -r 5f2a1efd0560 -r 9e34819a7ca1 NEWS --- a/NEWS Fri May 27 16:16:45 2022 +0200 +++ b/NEWS Sat May 28 10:45:45 2022 +0200 @@ -59,6 +59,8 @@ - Lifted multiple lemmas from mult to multp. - Redefined less_multiset to be based on multp. INCOMPATIBILITY. - Added lemmas. + Multiset.bex_greatest_element + Multiset.bex_least_element filter_mset_cong filter_mset_cong0 image_mset_filter_mset_swap diff -r 5f2a1efd0560 -r 9e34819a7ca1 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Fri May 27 16:16:45 2022 +0200 +++ b/src/HOL/Library/Multiset.thy Sat May 28 10:45:45 2022 +0200 @@ -1605,6 +1605,64 @@ qed +subsection \Least and greatest elements\ + +context begin + +qualified lemma + assumes + (* FIXME: Replace by transp_on (set_mset M) R if it gets introduced. *) + "\x \ set_mset M. \y \ set_mset M. \z \ set_mset M. R x y \ R y z \ R x z" and + "totalp_on (set_mset M) R" and + "M \ {#}" + shows + bex_least_element: "(\least \# M. \x \# M. least \ x \ R least x)" and + bex_greatest_element: "(\greatest \# M. \x \# M. greatest \ x \ R x greatest)" + unfolding atomize_conj + using assms +proof (induction M rule: multiset_induct) + case empty + hence False by simp + thus ?case .. +next + case (add x M) + from add.prems(1) have transp_on_x_M_raw: "\y\#M. \z\#M. R x y \ R y z \ R x z" + by (metis insert_iff set_mset_add_mset_insert) + + from add.prems(1) have transp_on_R_M: + "\x \ set_mset M. \y \ set_mset M. \z \ set_mset M. R x y \ R y z \ R x z" + by (meson mset_subsetD multi_psub_of_add_self) + + from add.prems(2) have + totalp_on_x_M_raw: "\y \# M. x \ y \ R x y \ R y x" and + totalp_on_M_R: "totalp_on (set_mset M) R" + by (simp_all add: totalp_on_def) + + show ?case + proof (cases "M = {#}") + case True + thus ?thesis by simp + next + case False + then obtain least greatest where + least_of_M: "least \# M" "\y\#M. least \ y \ R least y" and + greatest_of_M: "greatest\#M" "\y\#M. greatest \ y \ R y greatest" + using add.IH[OF transp_on_R_M totalp_on_M_R] by blast + + show ?thesis + proof (rule conjI) + from least_of_M show "\y\#add_mset x M. \z\#add_mset x M. y \ z \ R y z" + by (metis insert_iff set_mset_add_mset_insert totalp_on_x_M_raw transp_on_x_M_raw) + next + from greatest_of_M show "\y\#add_mset x M. \z\#add_mset x M. y \ z \ R z y" + by (metis insert_iff set_mset_add_mset_insert totalp_on_x_M_raw transp_on_x_M_raw) + qed + qed +qed + +end + + subsection \The fold combinator\ definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b"