changeset 82237 | 96cca71aa212 |
parent 82236 | d60c3f1ba86f |
child 82238 | c8ed5e759d22 |
--- a/src/HOL/Library/Multiset.thy Tue Mar 04 15:19:08 2025 +0100 +++ b/src/HOL/Library/Multiset.thy Tue Mar 04 16:07:55 2025 +0100 @@ -1337,6 +1337,9 @@ qed qed +lemma filter_mset_eq_mempty_iff: "filter_mset P A = {#} \<longleftrightarrow> (\<forall>x. x \<in># A \<longrightarrow> \<not>P x)" + by (auto simp: multiset_eq_iff count_eq_zero_iff) + lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \<in># M. Q x \<and> P x#}" by (auto simp: multiset_eq_iff)