| author | wenzelm | 
| Wed, 07 Aug 2024 14:44:51 +0200 | |
| changeset 80661 | 231d58c412b5 | 
| parent 80061 | 4c1347e172b1 | 
| child 80760 | be8c0e039a5e | 
| permissions | -rw-r--r-- | 
| 58101 | 1 | (* Author: Tobias Nipkow, TU Muenchen *) | 
| 2 | ||
| 60758 | 3 | section \<open>Sum and product over lists\<close> | 
| 58101 | 4 | |
| 5 | theory Groups_List | |
| 6 | imports List | |
| 7 | begin | |
| 8 | ||
| 58320 | 9 | locale monoid_list = monoid | 
| 10 | begin | |
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 11 | |
| 58320 | 12 | definition F :: "'a list \<Rightarrow> 'a" | 
| 13 | where | |
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63101diff
changeset | 14 | eq_foldr [code]: "F xs = foldr f xs \<^bold>1" | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 15 | |
| 58320 | 16 | lemma Nil [simp]: | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63101diff
changeset | 17 | "F [] = \<^bold>1" | 
| 58320 | 18 | by (simp add: eq_foldr) | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 19 | |
| 58320 | 20 | lemma Cons [simp]: | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63101diff
changeset | 21 | "F (x # xs) = x \<^bold>* F xs" | 
| 58320 | 22 | by (simp add: eq_foldr) | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 23 | |
| 58320 | 24 | lemma append [simp]: | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63101diff
changeset | 25 | "F (xs @ ys) = F xs \<^bold>* F ys" | 
| 58320 | 26 | by (induct xs) (simp_all add: assoc) | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 27 | |
| 58320 | 28 | end | 
| 58101 | 29 | |
| 58320 | 30 | locale comm_monoid_list = comm_monoid + monoid_list | 
| 31 | begin | |
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 32 | |
| 58320 | 33 | lemma rev [simp]: | 
| 34 | "F (rev xs) = F xs" | |
| 35 | by (simp add: eq_foldr foldr_fold fold_rev fun_eq_iff assoc left_commute) | |
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 36 | |
| 58320 | 37 | end | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 38 | |
| 58320 | 39 | locale comm_monoid_list_set = list: comm_monoid_list + set: comm_monoid_set | 
| 40 | begin | |
| 58101 | 41 | |
| 58320 | 42 | lemma distinct_set_conv_list: | 
| 43 | "distinct xs \<Longrightarrow> set.F g (set xs) = list.F (map g xs)" | |
| 44 | by (induct xs) simp_all | |
| 58101 | 45 | |
| 58320 | 46 | lemma set_conv_list [code]: | 
| 47 | "set.F g (set xs) = list.F (map g (remdups xs))" | |
| 48 | by (simp add: distinct_set_conv_list [symmetric]) | |
| 49 | ||
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 50 | lemma list_conv_set_nth: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 51 |   "list.F xs = set.F (\<lambda>i. xs ! i) {0..<length xs}"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 52 | proof - | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 53 | have "xs = map (\<lambda>i. xs ! i) [0..<length xs]" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 54 | by (simp add: map_nth) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 55 |   also have "list.F \<dots> = set.F (\<lambda>i. xs ! i) {0..<length xs}"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 56 | by (subst distinct_set_conv_list [symmetric]) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 57 | finally show ?thesis . | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 58 | qed | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79017diff
changeset | 59 | |
| 58320 | 60 | end | 
| 61 | ||
| 62 | ||
| 60758 | 63 | subsection \<open>List summation\<close> | 
| 58320 | 64 | |
| 65 | context monoid_add | |
| 66 | begin | |
| 67 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 68 | sublocale sum_list: monoid_list plus 0 | 
| 61776 | 69 | defines | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 70 | sum_list = sum_list.F .. | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 71 | |
| 58320 | 72 | end | 
| 73 | ||
| 74 | context comm_monoid_add | |
| 75 | begin | |
| 76 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 77 | sublocale sum_list: comm_monoid_list plus 0 | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
61378diff
changeset | 78 | rewrites | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 79 | "monoid_list.F plus 0 = sum_list" | 
| 58320 | 80 | proof - | 
| 81 | show "comm_monoid_list plus 0" .. | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 82 | then interpret sum_list: comm_monoid_list plus 0 . | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 83 | from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp | 
| 58101 | 84 | qed | 
| 85 | ||
| 64267 | 86 | sublocale sum: comm_monoid_list_set plus 0 | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
61378diff
changeset | 87 | rewrites | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 88 | "monoid_list.F plus 0 = sum_list" | 
| 64267 | 89 | and "comm_monoid_set.F plus 0 = sum" | 
| 58320 | 90 | proof - | 
| 91 | show "comm_monoid_list_set plus 0" .. | |
| 64267 | 92 | then interpret sum: comm_monoid_list_set plus 0 . | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 93 | from sum_list_def show "monoid_list.F plus 0 = sum_list" by simp | 
| 64267 | 94 | from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym) | 
| 58320 | 95 | qed | 
| 96 | ||
| 97 | end | |
| 98 | ||
| 60758 | 99 | text \<open>Some syntactic sugar for summing a function over a list:\<close> | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 100 | syntax (ASCII) | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 101 |   "_sum_list" :: "pttrn => 'a list => 'b => 'b"    ("(3SUM _<-_. _)" [0, 51, 10] 10)
 | 
| 58101 | 102 | syntax | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 103 |   "_sum_list" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
 | 
| 61799 | 104 | translations \<comment> \<open>Beware of argument permutation!\<close> | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 105 | "\<Sum>x\<leftarrow>xs. b" == "CONST sum_list (CONST map (\<lambda>x. b) xs)" | 
| 58101 | 106 | |
| 70928 | 107 | context | 
| 108 | includes lifting_syntax | |
| 109 | begin | |
| 110 | ||
| 111 | lemma sum_list_transfer [transfer_rule]: | |
| 112 | "(list_all2 A ===> A) sum_list sum_list" | |
| 113 | if [transfer_rule]: "A 0 0" "(A ===> A ===> A) (+) (+)" | |
| 114 | unfolding sum_list.eq_foldr [abs_def] | |
| 115 | by transfer_prover | |
| 116 | ||
| 117 | end | |
| 118 | ||
| 60758 | 119 | text \<open>TODO duplicates\<close> | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 120 | lemmas sum_list_simps = sum_list.Nil sum_list.Cons | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 121 | lemmas sum_list_append = sum_list.append | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 122 | lemmas sum_list_rev = sum_list.rev | 
| 58320 | 123 | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 124 | lemma (in monoid_add) fold_plus_sum_list_rev: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 125 | "fold plus xs = plus (sum_list (rev xs))" | 
| 58320 | 126 | proof | 
| 127 | fix x | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 128 | have "fold plus xs x = sum_list (rev xs @ [x])" | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 129 | by (simp add: foldr_conv_fold sum_list.eq_foldr) | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 130 | also have "\<dots> = sum_list (rev xs) + x" | 
| 58320 | 131 | by simp | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 132 | finally show "fold plus xs x = sum_list (rev xs) + x" | 
| 58320 | 133 | . | 
| 134 | qed | |
| 135 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 136 | lemma (in comm_monoid_add) sum_list_map_remove1: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 137 | "x \<in> set xs \<Longrightarrow> sum_list (map f xs) = f x + sum_list (map f (remove1 x xs))" | 
| 58101 | 138 | by (induct xs) (auto simp add: ac_simps) | 
| 139 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 140 | lemma (in monoid_add) size_list_conv_sum_list: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 141 | "size_list f xs = sum_list (map f xs) + size xs" | 
| 58101 | 142 | by (induct xs) auto | 
| 143 | ||
| 144 | lemma (in monoid_add) length_concat: | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 145 | "length (concat xss) = sum_list (map length xss)" | 
| 58101 | 146 | by (induct xss) simp_all | 
| 147 | ||
| 148 | lemma (in monoid_add) length_product_lists: | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
67489diff
changeset | 149 | "length (product_lists xss) = foldr (*) (map length xss) 1" | 
| 58101 | 150 | proof (induct xss) | 
| 151 | case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def) | |
| 152 | qed simp | |
| 153 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 154 | lemma (in monoid_add) sum_list_map_filter: | 
| 58101 | 155 | assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0" | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 156 | shows "sum_list (map f (filter P xs)) = sum_list (map f xs)" | 
| 58101 | 157 | using assms by (induct xs) auto | 
| 158 | ||
| 69231 | 159 | lemma sum_list_filter_le_nat: | 
| 160 | fixes f :: "'a \<Rightarrow> nat" | |
| 161 | shows "sum_list (map f (filter P xs)) \<le> sum_list (map f xs)" | |
| 162 | by(induction xs; simp) | |
| 163 | ||
| 64267 | 164 | lemma (in comm_monoid_add) distinct_sum_list_conv_Sum: | 
| 165 | "distinct xs \<Longrightarrow> sum_list xs = Sum (set xs)" | |
| 58101 | 166 | by (induct xs) simp_all | 
| 167 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 168 | lemma sum_list_upt[simp]: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 169 |   "m \<le> n \<Longrightarrow> sum_list [m..<n] = \<Sum> {m..<n}"
 | 
| 64267 | 170 | by(simp add: distinct_sum_list_conv_Sum) | 
| 58995 | 171 | |
| 66311 | 172 | context ordered_comm_monoid_add | 
| 173 | begin | |
| 174 | ||
| 175 | lemma sum_list_nonneg: "(\<And>x. x \<in> set xs \<Longrightarrow> 0 \<le> x) \<Longrightarrow> 0 \<le> sum_list xs" | |
| 176 | by (induction xs) auto | |
| 177 | ||
| 178 | lemma sum_list_nonpos: "(\<And>x. x \<in> set xs \<Longrightarrow> x \<le> 0) \<Longrightarrow> sum_list xs \<le> 0" | |
| 179 | by (induction xs) (auto simp: add_nonpos_nonpos) | |
| 58101 | 180 | |
| 66311 | 181 | lemma sum_list_nonneg_eq_0_iff: | 
| 182 | "(\<And>x. x \<in> set xs \<Longrightarrow> 0 \<le> x) \<Longrightarrow> sum_list xs = 0 \<longleftrightarrow> (\<forall>x\<in> set xs. x = 0)" | |
| 183 | by (induction xs) (simp_all add: add_nonneg_eq_0_iff sum_list_nonneg) | |
| 184 | ||
| 185 | end | |
| 186 | ||
| 187 | context canonically_ordered_monoid_add | |
| 188 | begin | |
| 58101 | 189 | |
| 66311 | 190 | lemma sum_list_eq_0_iff [simp]: | 
| 191 | "sum_list ns = 0 \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)" | |
| 192 | by (simp add: sum_list_nonneg_eq_0_iff) | |
| 193 | ||
| 194 | lemma member_le_sum_list: | |
| 195 | "x \<in> set xs \<Longrightarrow> x \<le> sum_list xs" | |
| 196 | by (induction xs) (auto simp: add_increasing add_increasing2) | |
| 58101 | 197 | |
| 66311 | 198 | lemma elem_le_sum_list: | 
| 199 | "k < size ns \<Longrightarrow> ns ! k \<le> sum_list (ns)" | |
| 200 | by (rule member_le_sum_list) simp | |
| 201 | ||
| 202 | end | |
| 203 | ||
| 204 | lemma (in ordered_cancel_comm_monoid_diff) sum_list_update: | |
| 205 | "k < size xs \<Longrightarrow> sum_list (xs[k := x]) = sum_list xs + x - xs ! k" | |
| 206 | apply(induction xs arbitrary:k) | |
| 207 | apply (auto simp: add_ac split: nat.split) | |
| 208 | apply(drule elem_le_sum_list) | |
| 209 | by (simp add: local.add_diff_assoc local.add_increasing) | |
| 58101 | 210 | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 211 | lemma (in monoid_add) sum_list_triv: | 
| 58101 | 212 | "(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r" | 
| 213 | by (induct xs) (simp_all add: distrib_right) | |
| 214 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 215 | lemma (in monoid_add) sum_list_0 [simp]: | 
| 58101 | 216 | "(\<Sum>x\<leftarrow>xs. 0) = 0" | 
| 217 | by (induct xs) (simp_all add: distrib_right) | |
| 218 | ||
| 61799 | 219 | text\<open>For non-Abelian groups \<open>xs\<close> needs to be reversed on one side:\<close> | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 220 | lemma (in ab_group_add) uminus_sum_list_map: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 221 | "- sum_list (map f xs) = sum_list (map (uminus \<circ> f) xs)" | 
| 58101 | 222 | by (induct xs) simp_all | 
| 223 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 224 | lemma (in comm_monoid_add) sum_list_addf: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 225 | "(\<Sum>x\<leftarrow>xs. f x + g x) = sum_list (map f xs) + sum_list (map g xs)" | 
| 58101 | 226 | by (induct xs) (simp_all add: algebra_simps) | 
| 227 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 228 | lemma (in ab_group_add) sum_list_subtractf: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 229 | "(\<Sum>x\<leftarrow>xs. f x - g x) = sum_list (map f xs) - sum_list (map g xs)" | 
| 58101 | 230 | by (induct xs) (simp_all add: algebra_simps) | 
| 231 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 232 | lemma (in semiring_0) sum_list_const_mult: | 
| 58101 | 233 | "(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)" | 
| 234 | by (induct xs) (simp_all add: algebra_simps) | |
| 235 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 236 | lemma (in semiring_0) sum_list_mult_const: | 
| 58101 | 237 | "(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c" | 
| 238 | by (induct xs) (simp_all add: algebra_simps) | |
| 239 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 240 | lemma (in ordered_ab_group_add_abs) sum_list_abs: | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 241 | "\<bar>sum_list xs\<bar> \<le> sum_list (map abs xs)" | 
| 58101 | 242 | by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq]) | 
| 243 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 244 | lemma sum_list_mono: | 
| 58101 | 245 |   fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
 | 
| 246 | shows "(\<And>x. x \<in> set xs \<Longrightarrow> f x \<le> g x) \<Longrightarrow> (\<Sum>x\<leftarrow>xs. f x) \<le> (\<Sum>x\<leftarrow>xs. g x)" | |
| 69231 | 247 | by (induct xs) (simp, simp add: add_mono) | 
| 248 | ||
| 249 | lemma sum_list_strict_mono: | |
| 250 |   fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, strict_ordered_ab_semigroup_add}"
 | |
| 251 | shows "\<lbrakk> xs \<noteq> []; \<And>x. x \<in> set xs \<Longrightarrow> f x < g x \<rbrakk> | |
| 252 | \<Longrightarrow> sum_list (map f xs) < sum_list (map g xs)" | |
| 253 | proof (induction xs) | |
| 254 | case Nil thus ?case by simp | |
| 255 | next | |
| 256 | case C: (Cons _ xs) | |
| 257 | show ?case | |
| 258 | proof (cases xs) | |
| 259 | case Nil thus ?thesis using C.prems by simp | |
| 260 | next | |
| 261 | case Cons thus ?thesis using C by(simp add: add_strict_mono) | |
| 262 | qed | |
| 263 | qed | |
| 58101 | 264 | |
| 75693 | 265 | text \<open>A much more general version of this monotonicity lemma | 
| 266 | can be formulated with multisets and the multiset order\<close> | |
| 267 | ||
| 268 | lemma sum_list_mono2: fixes xs :: "'a ::ordered_comm_monoid_add list" | |
| 269 | shows "\<lbrakk> length xs = length ys; \<And>i. i < length xs \<longrightarrow> xs!i \<le> ys!i \<rbrakk> | |
| 270 | \<Longrightarrow> sum_list xs \<le> sum_list ys" | |
| 271 | apply(induction xs ys rule: list_induct2) | |
| 272 | by(auto simp: nth_Cons' less_Suc_eq_0_disj imp_ex add_mono) | |
| 273 | ||
| 64267 | 274 | lemma (in monoid_add) sum_list_distinct_conv_sum_set: | 
| 275 | "distinct xs \<Longrightarrow> sum_list (map f xs) = sum f (set xs)" | |
| 58101 | 276 | by (induct xs) simp_all | 
| 277 | ||
| 64267 | 278 | lemma (in monoid_add) interv_sum_list_conv_sum_set_nat: | 
| 279 | "sum_list (map f [m..<n]) = sum f (set [m..<n])" | |
| 280 | by (simp add: sum_list_distinct_conv_sum_set) | |
| 58101 | 281 | |
| 64267 | 282 | lemma (in monoid_add) interv_sum_list_conv_sum_set_int: | 
| 283 | "sum_list (map f [k..l]) = sum f (set [k..l])" | |
| 284 | by (simp add: sum_list_distinct_conv_sum_set) | |
| 58101 | 285 | |
| 69593 | 286 | text \<open>General equivalence between \<^const>\<open>sum_list\<close> and \<^const>\<open>sum\<close>\<close> | 
| 64267 | 287 | lemma (in monoid_add) sum_list_sum_nth: | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 288 | "sum_list xs = (\<Sum> i = 0 ..< length xs. xs ! i)" | 
| 67399 | 289 | using interv_sum_list_conv_sum_set_nat [of "(!) xs" 0 "length xs"] by (simp add: map_nth) | 
| 58101 | 290 | |
| 64267 | 291 | lemma sum_list_map_eq_sum_count: | 
| 292 | "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) (set xs)" | |
| 59728 | 293 | proof(induction xs) | 
| 294 | case (Cons x xs) | |
| 295 | show ?case (is "?l = ?r") | |
| 296 | proof cases | |
| 297 | assume "x \<in> set xs" | |
| 60541 | 298 | have "?l = f x + (\<Sum>x\<in>set xs. count_list xs x * f x)" by (simp add: Cons.IH) | 
| 60758 | 299 |     also have "set xs = insert x (set xs - {x})" using \<open>x \<in> set xs\<close>by blast
 | 
| 60541 | 300 |     also have "f x + (\<Sum>x\<in>insert x (set xs - {x}). count_list xs x * f x) = ?r"
 | 
| 64267 | 301 | by (simp add: sum.insert_remove eq_commute) | 
| 59728 | 302 | finally show ?thesis . | 
| 303 | next | |
| 304 | assume "x \<notin> set xs" | |
| 305 | hence "\<And>xa. xa \<in> set xs \<Longrightarrow> x \<noteq> xa" by blast | |
| 60758 | 306 | thus ?thesis by (simp add: Cons.IH \<open>x \<notin> set xs\<close>) | 
| 59728 | 307 | qed | 
| 308 | qed simp | |
| 309 | ||
| 64267 | 310 | lemma sum_list_map_eq_sum_count2: | 
| 59728 | 311 | assumes "set xs \<subseteq> X" "finite X" | 
| 64267 | 312 | shows "sum_list (map f xs) = sum (\<lambda>x. count_list xs x * f x) X" | 
| 59728 | 313 | proof- | 
| 60541 | 314 | let ?F = "\<lambda>x. count_list xs x * f x" | 
| 64267 | 315 | have "sum ?F X = sum ?F (set xs \<union> (X - set xs))" | 
| 59728 | 316 | using Un_absorb1[OF assms(1)] by(simp) | 
| 64267 | 317 | also have "\<dots> = sum ?F (set xs)" | 
| 59728 | 318 | using assms(2) | 
| 64267 | 319 | by(simp add: sum.union_disjoint[OF _ _ Diff_disjoint] del: Un_Diff_cancel) | 
| 320 | finally show ?thesis by(simp add:sum_list_map_eq_sum_count) | |
| 59728 | 321 | qed | 
| 322 | ||
| 72545 | 323 | lemma sum_list_replicate: "sum_list (replicate n c) = of_nat n * c" | 
| 324 | by(induction n)(auto simp add: distrib_right) | |
| 325 | ||
| 326 | ||
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 327 | lemma sum_list_nonneg: | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 328 | "(\<And>x. x \<in> set xs \<Longrightarrow> (x :: 'a :: ordered_comm_monoid_add) \<ge> 0) \<Longrightarrow> sum_list xs \<ge> 0" | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
61955diff
changeset | 329 | by (induction xs) simp_all | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
61955diff
changeset | 330 | |
| 69231 | 331 | lemma sum_list_Suc: | 
| 332 | "sum_list (map (\<lambda>x. Suc(f x)) xs) = sum_list (map f xs) + length xs" | |
| 333 | by(induction xs; simp) | |
| 334 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 335 | lemma (in monoid_add) sum_list_map_filter': | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 336 | "sum_list (map f (filter P xs)) = sum_list (map (\<lambda>x. if P x then f x else 0) xs)" | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
61955diff
changeset | 337 | by (induction xs) simp_all | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
61955diff
changeset | 338 | |
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 339 | text \<open>Summation of a strictly ascending sequence with length \<open>n\<close> | 
| 66434 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 340 |   can be upper-bounded by summation over \<open>{0..<n}\<close>.\<close>
 | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 341 | |
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 342 | lemma sorted_wrt_less_sum_mono_lowerbound: | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 343 |   fixes f :: "nat \<Rightarrow> ('b::ordered_comm_monoid_add)"
 | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 344 | assumes mono: "\<And>x y. x\<le>y \<Longrightarrow> f x \<le> f y" | 
| 67399 | 345 | shows "sorted_wrt (<) ns \<Longrightarrow> | 
| 66434 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 346 |     (\<Sum>i\<in>{0..<length ns}. f i) \<le> (\<Sum>i\<leftarrow>ns. f i)"
 | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 347 | proof (induction ns rule: rev_induct) | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 348 | case Nil | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 349 | then show ?case by simp | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 350 | next | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 351 | case (snoc n ns) | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 352 |   have "sum f {0..<length (ns @ [n])}
 | 
| 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 353 |       = sum f {0..<length ns} + f (length ns)"
 | 
| 66434 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 354 | by simp | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 355 |   also have "sum f {0..<length ns} \<le> sum_list (map f ns)"
 | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 356 | using snoc by (auto simp: sorted_wrt_append) | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 357 | also have "length ns \<le> n" | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 358 | using sorted_wrt_less_idx[OF snoc.prems(1), of "length ns"] by auto | 
| 66434 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 359 |   finally have "sum f {0..<length (ns @ [n])} \<le> sum_list (map f ns) + f n"
 | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 360 | using mono add_mono by blast | 
| 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 361 | thus ?case by simp | 
| 67489 
f1ba59ddd9a6
drop redundant cong rules
 Lars Hupel <lars.hupel@mytum.de> parents: 
67399diff
changeset | 362 | qed | 
| 66434 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 nipkow parents: 
66311diff
changeset | 363 | |
| 58101 | 364 | |
| 72024 | 365 | subsection \<open>Horner sums\<close> | 
| 366 | ||
| 367 | context comm_semiring_0 | |
| 368 | begin | |
| 369 | ||
| 370 | definition horner_sum :: \<open>('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'a\<close>
 | |
| 371 | where horner_sum_foldr: \<open>horner_sum f a xs = foldr (\<lambda>x b. f x + a * b) xs 0\<close> | |
| 372 | ||
| 373 | lemma horner_sum_simps [simp]: | |
| 374 | \<open>horner_sum f a [] = 0\<close> | |
| 375 | \<open>horner_sum f a (x # xs) = f x + a * horner_sum f a xs\<close> | |
| 376 | by (simp_all add: horner_sum_foldr) | |
| 377 | ||
| 378 | lemma horner_sum_eq_sum_funpow: | |
| 379 | \<open>horner_sum f a xs = (\<Sum>n = 0..<length xs. ((*) a ^^ n) (f (xs ! n)))\<close> | |
| 380 | proof (induction xs) | |
| 381 | case Nil | |
| 382 | then show ?case | |
| 383 | by simp | |
| 384 | next | |
| 385 | case (Cons x xs) | |
| 386 | then show ?case | |
| 387 | by (simp add: sum.atLeast0_lessThan_Suc_shift sum_distrib_left del: sum.op_ivl_Suc) | |
| 388 | qed | |
| 389 | ||
| 390 | end | |
| 391 | ||
| 392 | context | |
| 393 | includes lifting_syntax | |
| 394 | begin | |
| 395 | ||
| 396 | lemma horner_sum_transfer [transfer_rule]: | |
| 397 | \<open>((B ===> A) ===> A ===> list_all2 B ===> A) horner_sum horner_sum\<close> | |
| 398 | if [transfer_rule]: \<open>A 0 0\<close> | |
| 399 | and [transfer_rule]: \<open>(A ===> A ===> A) (+) (+)\<close> | |
| 400 | and [transfer_rule]: \<open>(A ===> A ===> A) (*) (*)\<close> | |
| 401 | by (unfold horner_sum_foldr) transfer_prover | |
| 402 | ||
| 403 | end | |
| 404 | ||
| 405 | context comm_semiring_1 | |
| 406 | begin | |
| 407 | ||
| 408 | lemma horner_sum_eq_sum: | |
| 409 | \<open>horner_sum f a xs = (\<Sum>n = 0..<length xs. f (xs ! n) * a ^ n)\<close> | |
| 410 | proof - | |
| 411 | have \<open>(*) a ^^ n = (*) (a ^ n)\<close> for n | |
| 412 | by (induction n) (simp_all add: ac_simps) | |
| 413 | then show ?thesis | |
| 414 | by (simp add: horner_sum_eq_sum_funpow ac_simps) | |
| 415 | qed | |
| 416 | ||
| 72619 | 417 | lemma horner_sum_append: | 
| 418 | \<open>horner_sum f a (xs @ ys) = horner_sum f a xs + a ^ length xs * horner_sum f a ys\<close> | |
| 419 | using sum.atLeastLessThan_shift_bounds [of _ 0 \<open>length xs\<close> \<open>length ys\<close>] | |
| 420 | atLeastLessThan_add_Un [of 0 \<open>length xs\<close> \<open>length ys\<close>] | |
| 421 | by (simp add: horner_sum_eq_sum sum_distrib_left sum.union_disjoint ac_simps nth_append power_add) | |
| 422 | ||
| 72024 | 423 | end | 
| 424 | ||
| 75662 | 425 | context linordered_semidom | 
| 426 | begin | |
| 427 | ||
| 428 | lemma horner_sum_nonnegative: | |
| 429 | \<open>0 \<le> horner_sum of_bool 2 bs\<close> | |
| 430 | by (induction bs) simp_all | |
| 431 | ||
| 432 | end | |
| 433 | ||
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 434 | context discrete_linordered_semidom | 
| 75662 | 435 | begin | 
| 436 | ||
| 437 | lemma horner_sum_bound: | |
| 438 | \<open>horner_sum of_bool 2 bs < 2 ^ length bs\<close> | |
| 439 | proof (induction bs) | |
| 440 | case Nil | |
| 441 | then show ?case | |
| 442 | by simp | |
| 443 | next | |
| 444 | case (Cons b bs) | |
| 445 | moreover define a where \<open>a = 2 ^ length bs - horner_sum of_bool 2 bs\<close> | |
| 446 | ultimately have *: \<open>2 ^ length bs = horner_sum of_bool 2 bs + a\<close> | |
| 447 | by simp | |
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 448 | have \<open>0 < a\<close> | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 449 | using Cons * by simp | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 450 | moreover have \<open>1 \<le> a\<close> | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 451 | using \<open>0 < a\<close> by (simp add: less_eq_iff_succ_less) | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 452 | ultimately have \<open>0 + 1 < a + a\<close> | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 453 | by (rule add_less_le_mono) | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 454 | then have \<open>1 < a * 2\<close> | 
| 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 455 | by (simp add: mult_2_right) | 
| 75662 | 456 | with Cons show ?case | 
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 457 | by (simp add: * algebra_simps) | 
| 75662 | 458 | qed | 
| 459 | ||
| 79017 | 460 | lemma horner_sum_of_bool_2_less: | 
| 461 | \<open>(horner_sum of_bool 2 bs) < 2 ^ length bs\<close> | |
| 462 | by (fact horner_sum_bound) | |
| 463 | ||
| 75662 | 464 | end | 
| 465 | ||
| 466 | lemma nat_horner_sum [simp]: | |
| 467 | \<open>nat (horner_sum of_bool 2 bs) = horner_sum of_bool 2 bs\<close> | |
| 468 | by (induction bs) (auto simp add: nat_add_distrib horner_sum_nonnegative) | |
| 469 | ||
| 78935 
5e788ff7a489
explicit type class for discrete linordered semidoms
 haftmann parents: 
75693diff
changeset | 470 | context discrete_linordered_semidom | 
| 75662 | 471 | begin | 
| 472 | ||
| 473 | lemma horner_sum_less_eq_iff_lexordp_eq: | |
| 474 | \<open>horner_sum of_bool 2 bs \<le> horner_sum of_bool 2 cs \<longleftrightarrow> lexordp_eq (rev bs) (rev cs)\<close> | |
| 475 | if \<open>length bs = length cs\<close> | |
| 476 | proof - | |
| 477 | have \<open>horner_sum of_bool 2 (rev bs) \<le> horner_sum of_bool 2 (rev cs) \<longleftrightarrow> lexordp_eq bs cs\<close> | |
| 478 | if \<open>length bs = length cs\<close> for bs cs | |
| 479 | using that proof (induction bs cs rule: list_induct2) | |
| 480 | case Nil | |
| 481 | then show ?case | |
| 482 | by simp | |
| 483 | next | |
| 484 | case (Cons b bs c cs) | |
| 485 | with horner_sum_nonnegative [of \<open>rev bs\<close>] horner_sum_nonnegative [of \<open>rev cs\<close>] | |
| 486 | horner_sum_bound [of \<open>rev bs\<close>] horner_sum_bound [of \<open>rev cs\<close>] | |
| 487 | show ?case | |
| 488 | by (auto simp add: horner_sum_append not_le Cons intro: add_strict_increasing2 add_increasing) | |
| 489 | qed | |
| 490 | from that this [of \<open>rev bs\<close> \<open>rev cs\<close>] show ?thesis | |
| 491 | by simp | |
| 492 | qed | |
| 493 | ||
| 494 | lemma horner_sum_less_iff_lexordp: | |
| 495 | \<open>horner_sum of_bool 2 bs < horner_sum of_bool 2 cs \<longleftrightarrow> ord_class.lexordp (rev bs) (rev cs)\<close> | |
| 496 | if \<open>length bs = length cs\<close> | |
| 497 | proof - | |
| 498 | have \<open>horner_sum of_bool 2 (rev bs) < horner_sum of_bool 2 (rev cs) \<longleftrightarrow> ord_class.lexordp bs cs\<close> | |
| 499 | if \<open>length bs = length cs\<close> for bs cs | |
| 500 | using that proof (induction bs cs rule: list_induct2) | |
| 501 | case Nil | |
| 502 | then show ?case | |
| 503 | by simp | |
| 504 | next | |
| 505 | case (Cons b bs c cs) | |
| 506 | with horner_sum_nonnegative [of \<open>rev bs\<close>] horner_sum_nonnegative [of \<open>rev cs\<close>] | |
| 507 | horner_sum_bound [of \<open>rev bs\<close>] horner_sum_bound [of \<open>rev cs\<close>] | |
| 508 | show ?case | |
| 509 | by (auto simp add: horner_sum_append not_less Cons intro: add_strict_increasing2 add_increasing) | |
| 510 | qed | |
| 511 | from that this [of \<open>rev bs\<close> \<open>rev cs\<close>] show ?thesis | |
| 512 | by simp | |
| 513 | qed | |
| 514 | ||
| 515 | end | |
| 516 | ||
| 72024 | 517 | |
| 69593 | 518 | subsection \<open>Further facts about \<^const>\<open>List.n_lists\<close>\<close> | 
| 58101 | 519 | |
| 520 | lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n" | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 521 | by (induct n) (auto simp add: comp_def length_concat sum_list_triv) | 
| 58101 | 522 | |
| 523 | lemma distinct_n_lists: | |
| 524 | assumes "distinct xs" | |
| 525 | shows "distinct (List.n_lists n xs)" | |
| 526 | proof (rule card_distinct) | |
| 527 | from assms have card_length: "card (set xs) = length xs" by (rule distinct_card) | |
| 528 | have "card (set (List.n_lists n xs)) = card (set xs) ^ n" | |
| 529 | proof (induct n) | |
| 530 | case 0 then show ?case by simp | |
| 531 | next | |
| 532 | case (Suc n) | |
| 533 | moreover have "card (\<Union>ys\<in>set (List.n_lists n xs). (\<lambda>y. y # ys) ` set xs) | |
| 534 | = (\<Sum>ys\<in>set (List.n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))" | |
| 535 | by (rule card_UN_disjoint) auto | |
| 536 | moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)" | |
| 537 | by (rule card_image) (simp add: inj_on_def) | |
| 538 | ultimately show ?case by auto | |
| 539 | qed | |
| 540 | also have "\<dots> = length xs ^ n" by (simp add: card_length) | |
| 541 | finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)" | |
| 542 | by (simp add: length_n_lists) | |
| 543 | qed | |
| 544 | ||
| 545 | ||
| 60758 | 546 | subsection \<open>Tools setup\<close> | 
| 58101 | 547 | |
| 64267 | 548 | lemmas sum_code = sum.set_conv_list | 
| 58320 | 549 | |
| 64267 | 550 | lemma sum_set_upto_conv_sum_list_int [code_unfold]: | 
| 551 | "sum f (set [i..j::int]) = sum_list (map f [i..j])" | |
| 552 | by (simp add: interv_sum_list_conv_sum_set_int) | |
| 58101 | 553 | |
| 64267 | 554 | lemma sum_set_upt_conv_sum_list_nat [code_unfold]: | 
| 555 | "sum f (set [m..<n]) = sum_list (map f [m..<n])" | |
| 556 | by (simp add: interv_sum_list_conv_sum_set_nat) | |
| 58101 | 557 | |
| 58368 | 558 | |
| 60758 | 559 | subsection \<open>List product\<close> | 
| 58368 | 560 | |
| 561 | context monoid_mult | |
| 562 | begin | |
| 563 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 564 | sublocale prod_list: monoid_list times 1 | 
| 61776 | 565 | defines | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 566 | prod_list = prod_list.F .. | 
| 58368 | 567 | |
| 58320 | 568 | end | 
| 58368 | 569 | |
| 570 | context comm_monoid_mult | |
| 571 | begin | |
| 572 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 573 | sublocale prod_list: comm_monoid_list times 1 | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
61378diff
changeset | 574 | rewrites | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 575 | "monoid_list.F times 1 = prod_list" | 
| 58368 | 576 | proof - | 
| 577 | show "comm_monoid_list times 1" .. | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 578 | then interpret prod_list: comm_monoid_list times 1 . | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 579 | from prod_list_def show "monoid_list.F times 1 = prod_list" by simp | 
| 58368 | 580 | qed | 
| 581 | ||
| 64272 | 582 | sublocale prod: comm_monoid_list_set times 1 | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
61378diff
changeset | 583 | rewrites | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 584 | "monoid_list.F times 1 = prod_list" | 
| 64272 | 585 | and "comm_monoid_set.F times 1 = prod" | 
| 58368 | 586 | proof - | 
| 587 | show "comm_monoid_list_set times 1" .. | |
| 64272 | 588 | then interpret prod: comm_monoid_list_set times 1 . | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 589 | from prod_list_def show "monoid_list.F times 1 = prod_list" by simp | 
| 64272 | 590 | from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym) | 
| 58368 | 591 | qed | 
| 592 | ||
| 593 | end | |
| 594 | ||
| 60758 | 595 | text \<open>Some syntactic sugar:\<close> | 
| 58368 | 596 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 597 | syntax (ASCII) | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 598 |   "_prod_list" :: "pttrn => 'a list => 'b => 'b"    ("(3PROD _<-_. _)" [0, 51, 10] 10)
 | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61799diff
changeset | 599 | syntax | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 600 |   "_prod_list" :: "pttrn => 'a list => 'b => 'b"    ("(3\<Prod>_\<leftarrow>_. _)" [0, 51, 10] 10)
 | 
| 61799 | 601 | translations \<comment> \<open>Beware of argument permutation!\<close> | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63343diff
changeset | 602 | "\<Prod>x\<leftarrow>xs. b" \<rightleftharpoons> "CONST prod_list (CONST map (\<lambda>x. b) xs)" | 
| 58368 | 603 | |
| 70928 | 604 | context | 
| 605 | includes lifting_syntax | |
| 606 | begin | |
| 607 | ||
| 608 | lemma prod_list_transfer [transfer_rule]: | |
| 609 | "(list_all2 A ===> A) prod_list prod_list" | |
| 610 | if [transfer_rule]: "A 1 1" "(A ===> A ===> A) (*) (*)" | |
| 611 | unfolding prod_list.eq_foldr [abs_def] | |
| 612 | by transfer_prover | |
| 613 | ||
| 58368 | 614 | end | 
| 70928 | 615 | |
| 616 | lemma prod_list_zero_iff: | |
| 617 |   "prod_list xs = 0 \<longleftrightarrow> (0 :: 'a :: {semiring_no_zero_divisors, semiring_1}) \<in> set xs"
 | |
| 618 | by (induction xs) simp_all | |
| 619 | ||
| 620 | end |