merged
authornipkow
Thu, 03 Aug 2017 07:31:25 +0200
changeset 66309 ca985e87c123
parent 66307 50ed697e97f1 (current diff)
parent 66308 b6a0d95b94be (diff)
child 66310 e8d2862ec203
child 66311 037aaa0b6daf
child 66315 ce50687a700e
merged
--- a/src/HOL/Groups_List.thy	Tue Aug 01 20:38:39 2017 +0200
+++ b/src/HOL/Groups_List.thy	Thu Aug 03 07:31:25 2017 +0200
@@ -142,9 +142,9 @@
   "m \<le> n \<Longrightarrow> sum_list [m..<n] = \<Sum> {m..<n}"
 by(simp add: distinct_sum_list_conv_Sum)
 
-lemma sum_list_eq_0_nat_iff_nat [simp]:
-  "sum_list ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
-  by (induct ns) simp_all
+lemma (in canonically_ordered_monoid_add) sum_list_eq_0_iff [simp]:
+  "sum_list ns = 0 \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
+by (induct ns) simp_all
 
 lemma member_le_sum_list_nat:
   "(n :: nat) \<in> set ns \<Longrightarrow> n \<le> sum_list ns"