src/HOL/Library/DAList_Multiset.thy
changeset 63310 caaacf37943f
parent 63195 f3f08c0d4aaf
child 63793 e68a0b651eb5
--- a/src/HOL/Library/DAList_Multiset.thy	Thu Jun 16 17:11:00 2016 +0200
+++ b/src/HOL/Library/DAList_Multiset.thy	Fri Jun 17 12:37:43 2016 +0200
@@ -244,7 +244,7 @@
 next
   assume ?rhs
   show ?lhs
-  proof (rule mset_less_eqI)
+  proof (rule mset_subset_eqI)
     fix x
     from \<open>?rhs\<close> have "count_of (DAList.impl_of xs) x \<le> count A x"
       by (cases "x \<in> fst ` set (DAList.impl_of xs)") (auto simp add: count_of_empty)