58101
|
1 |
|
|
2 |
(* Author: Tobias Nipkow, TU Muenchen *)
|
|
3 |
|
58320
|
4 |
header {* Sum over lists *}
|
58101
|
5 |
|
|
6 |
theory Groups_List
|
|
7 |
imports List
|
|
8 |
begin
|
|
9 |
|
58320
|
10 |
no_notation times (infixl "*" 70)
|
|
11 |
no_notation Groups.one ("1")
|
|
12 |
|
|
13 |
locale monoid_list = monoid
|
|
14 |
begin
|
|
15 |
|
|
16 |
definition F :: "'a list \<Rightarrow> 'a"
|
|
17 |
where
|
|
18 |
eq_foldr [code]: "F xs = foldr f xs 1"
|
|
19 |
|
|
20 |
lemma Nil [simp]:
|
|
21 |
"F [] = 1"
|
|
22 |
by (simp add: eq_foldr)
|
|
23 |
|
|
24 |
lemma Cons [simp]:
|
|
25 |
"F (x # xs) = x * F xs"
|
|
26 |
by (simp add: eq_foldr)
|
|
27 |
|
|
28 |
lemma append [simp]:
|
|
29 |
"F (xs @ ys) = F xs * F ys"
|
|
30 |
by (induct xs) (simp_all add: assoc)
|
|
31 |
|
|
32 |
end
|
58101
|
33 |
|
58320
|
34 |
locale comm_monoid_list = comm_monoid + monoid_list
|
|
35 |
begin
|
|
36 |
|
|
37 |
lemma rev [simp]:
|
|
38 |
"F (rev xs) = F xs"
|
|
39 |
by (simp add: eq_foldr foldr_fold fold_rev fun_eq_iff assoc left_commute)
|
|
40 |
|
|
41 |
end
|
|
42 |
|
|
43 |
locale comm_monoid_list_set = list: comm_monoid_list + set: comm_monoid_set
|
|
44 |
begin
|
58101
|
45 |
|
58320
|
46 |
lemma distinct_set_conv_list:
|
|
47 |
"distinct xs \<Longrightarrow> set.F g (set xs) = list.F (map g xs)"
|
|
48 |
by (induct xs) simp_all
|
58101
|
49 |
|
58320
|
50 |
lemma set_conv_list [code]:
|
|
51 |
"set.F g (set xs) = list.F (map g (remdups xs))"
|
|
52 |
by (simp add: distinct_set_conv_list [symmetric])
|
|
53 |
|
|
54 |
end
|
|
55 |
|
|
56 |
notation times (infixl "*" 70)
|
|
57 |
notation Groups.one ("1")
|
|
58 |
|
|
59 |
|
|
60 |
subsection {* List summation *}
|
|
61 |
|
|
62 |
context monoid_add
|
|
63 |
begin
|
|
64 |
|
|
65 |
definition listsum :: "'a list \<Rightarrow> 'a"
|
|
66 |
where
|
|
67 |
"listsum = monoid_list.F plus 0"
|
58101
|
68 |
|
58320
|
69 |
sublocale listsum!: monoid_list plus 0
|
|
70 |
where
|
|
71 |
"monoid_list.F plus 0 = listsum"
|
|
72 |
proof -
|
|
73 |
show "monoid_list plus 0" ..
|
|
74 |
then interpret listsum!: monoid_list plus 0 .
|
|
75 |
from listsum_def show "monoid_list.F plus 0 = listsum" by rule
|
|
76 |
qed
|
|
77 |
|
|
78 |
end
|
|
79 |
|
|
80 |
context comm_monoid_add
|
|
81 |
begin
|
|
82 |
|
|
83 |
sublocale listsum!: comm_monoid_list plus 0
|
|
84 |
where
|
|
85 |
"monoid_list.F plus 0 = listsum"
|
|
86 |
proof -
|
|
87 |
show "comm_monoid_list plus 0" ..
|
|
88 |
then interpret listsum!: comm_monoid_list plus 0 .
|
|
89 |
from listsum_def show "monoid_list.F plus 0 = listsum" by rule
|
58101
|
90 |
qed
|
|
91 |
|
58320
|
92 |
sublocale setsum!: comm_monoid_list_set plus 0
|
|
93 |
where
|
|
94 |
"monoid_list.F plus 0 = listsum"
|
|
95 |
and "comm_monoid_set.F plus 0 = setsum"
|
|
96 |
proof -
|
|
97 |
show "comm_monoid_list_set plus 0" ..
|
|
98 |
then interpret setsum!: comm_monoid_list_set plus 0 .
|
|
99 |
from listsum_def show "monoid_list.F plus 0 = listsum" by rule
|
|
100 |
from setsum_def show "comm_monoid_set.F plus 0 = setsum" by rule
|
|
101 |
qed
|
|
102 |
|
|
103 |
end
|
|
104 |
|
|
105 |
text {* Some syntactic sugar for summing a function over a list: *}
|
58101
|
106 |
|
|
107 |
syntax
|
|
108 |
"_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3SUM _<-_. _)" [0, 51, 10] 10)
|
|
109 |
syntax (xsymbols)
|
|
110 |
"_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
|
|
111 |
syntax (HTML output)
|
|
112 |
"_listsum" :: "pttrn => 'a list => 'b => 'b" ("(3\<Sum>_\<leftarrow>_. _)" [0, 51, 10] 10)
|
|
113 |
|
|
114 |
translations -- {* Beware of argument permutation! *}
|
|
115 |
"SUM x<-xs. b" == "CONST listsum (CONST map (%x. b) xs)"
|
|
116 |
"\<Sum>x\<leftarrow>xs. b" == "CONST listsum (CONST map (%x. b) xs)"
|
|
117 |
|
58320
|
118 |
text {* TODO duplicates *}
|
|
119 |
lemmas listsum_simps = listsum.Nil listsum.Cons
|
|
120 |
lemmas listsum_append = listsum.append
|
|
121 |
lemmas listsum_rev = listsum.rev
|
|
122 |
|
|
123 |
lemma (in monoid_add) fold_plus_listsum_rev:
|
|
124 |
"fold plus xs = plus (listsum (rev xs))"
|
|
125 |
proof
|
|
126 |
fix x
|
|
127 |
have "fold plus xs x = listsum (rev xs @ [x])"
|
|
128 |
by (simp add: foldr_conv_fold listsum.eq_foldr)
|
|
129 |
also have "\<dots> = listsum (rev xs) + x"
|
|
130 |
by simp
|
|
131 |
finally show "fold plus xs x = listsum (rev xs) + x"
|
|
132 |
.
|
|
133 |
qed
|
|
134 |
|
58101
|
135 |
lemma (in comm_monoid_add) listsum_map_remove1:
|
|
136 |
"x \<in> set xs \<Longrightarrow> listsum (map f xs) = f x + listsum (map f (remove1 x xs))"
|
|
137 |
by (induct xs) (auto simp add: ac_simps)
|
|
138 |
|
|
139 |
lemma (in monoid_add) size_list_conv_listsum:
|
|
140 |
"size_list f xs = listsum (map f xs) + size xs"
|
|
141 |
by (induct xs) auto
|
|
142 |
|
|
143 |
lemma (in monoid_add) length_concat:
|
|
144 |
"length (concat xss) = listsum (map length xss)"
|
|
145 |
by (induct xss) simp_all
|
|
146 |
|
|
147 |
lemma (in monoid_add) length_product_lists:
|
|
148 |
"length (product_lists xss) = foldr op * (map length xss) 1"
|
|
149 |
proof (induct xss)
|
|
150 |
case (Cons xs xss) then show ?case by (induct xs) (auto simp: length_concat o_def)
|
|
151 |
qed simp
|
|
152 |
|
|
153 |
lemma (in monoid_add) listsum_map_filter:
|
|
154 |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<not> P x \<Longrightarrow> f x = 0"
|
|
155 |
shows "listsum (map f (filter P xs)) = listsum (map f xs)"
|
|
156 |
using assms by (induct xs) auto
|
|
157 |
|
|
158 |
lemma (in comm_monoid_add) distinct_listsum_conv_Setsum:
|
|
159 |
"distinct xs \<Longrightarrow> listsum xs = Setsum (set xs)"
|
|
160 |
by (induct xs) simp_all
|
|
161 |
|
|
162 |
lemma listsum_eq_0_nat_iff_nat [simp]:
|
|
163 |
"listsum ns = (0::nat) \<longleftrightarrow> (\<forall>n \<in> set ns. n = 0)"
|
|
164 |
by (induct ns) simp_all
|
|
165 |
|
|
166 |
lemma member_le_listsum_nat:
|
|
167 |
"(n :: nat) \<in> set ns \<Longrightarrow> n \<le> listsum ns"
|
|
168 |
by (induct ns) auto
|
|
169 |
|
|
170 |
lemma elem_le_listsum_nat:
|
|
171 |
"k < size ns \<Longrightarrow> ns ! k \<le> listsum (ns::nat list)"
|
|
172 |
by (rule member_le_listsum_nat) simp
|
|
173 |
|
|
174 |
lemma listsum_update_nat:
|
|
175 |
"k < size ns \<Longrightarrow> listsum (ns[k := (n::nat)]) = listsum ns + n - ns ! k"
|
|
176 |
apply(induct ns arbitrary:k)
|
|
177 |
apply (auto split:nat.split)
|
|
178 |
apply(drule elem_le_listsum_nat)
|
|
179 |
apply arith
|
|
180 |
done
|
|
181 |
|
|
182 |
lemma (in monoid_add) listsum_triv:
|
|
183 |
"(\<Sum>x\<leftarrow>xs. r) = of_nat (length xs) * r"
|
|
184 |
by (induct xs) (simp_all add: distrib_right)
|
|
185 |
|
|
186 |
lemma (in monoid_add) listsum_0 [simp]:
|
|
187 |
"(\<Sum>x\<leftarrow>xs. 0) = 0"
|
|
188 |
by (induct xs) (simp_all add: distrib_right)
|
|
189 |
|
|
190 |
text{* For non-Abelian groups @{text xs} needs to be reversed on one side: *}
|
|
191 |
lemma (in ab_group_add) uminus_listsum_map:
|
|
192 |
"- listsum (map f xs) = listsum (map (uminus \<circ> f) xs)"
|
|
193 |
by (induct xs) simp_all
|
|
194 |
|
|
195 |
lemma (in comm_monoid_add) listsum_addf:
|
|
196 |
"(\<Sum>x\<leftarrow>xs. f x + g x) = listsum (map f xs) + listsum (map g xs)"
|
|
197 |
by (induct xs) (simp_all add: algebra_simps)
|
|
198 |
|
|
199 |
lemma (in ab_group_add) listsum_subtractf:
|
|
200 |
"(\<Sum>x\<leftarrow>xs. f x - g x) = listsum (map f xs) - listsum (map g xs)"
|
|
201 |
by (induct xs) (simp_all add: algebra_simps)
|
|
202 |
|
|
203 |
lemma (in semiring_0) listsum_const_mult:
|
|
204 |
"(\<Sum>x\<leftarrow>xs. c * f x) = c * (\<Sum>x\<leftarrow>xs. f x)"
|
|
205 |
by (induct xs) (simp_all add: algebra_simps)
|
|
206 |
|
|
207 |
lemma (in semiring_0) listsum_mult_const:
|
|
208 |
"(\<Sum>x\<leftarrow>xs. f x * c) = (\<Sum>x\<leftarrow>xs. f x) * c"
|
|
209 |
by (induct xs) (simp_all add: algebra_simps)
|
|
210 |
|
|
211 |
lemma (in ordered_ab_group_add_abs) listsum_abs:
|
|
212 |
"\<bar>listsum xs\<bar> \<le> listsum (map abs xs)"
|
|
213 |
by (induct xs) (simp_all add: order_trans [OF abs_triangle_ineq])
|
|
214 |
|
|
215 |
lemma listsum_mono:
|
|
216 |
fixes f g :: "'a \<Rightarrow> 'b::{monoid_add, ordered_ab_semigroup_add}"
|
|
217 |
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)"
|
|
218 |
by (induct xs) (simp, simp add: add_mono)
|
|
219 |
|
|
220 |
lemma (in monoid_add) listsum_distinct_conv_setsum_set:
|
|
221 |
"distinct xs \<Longrightarrow> listsum (map f xs) = setsum f (set xs)"
|
|
222 |
by (induct xs) simp_all
|
|
223 |
|
|
224 |
lemma (in monoid_add) interv_listsum_conv_setsum_set_nat:
|
|
225 |
"listsum (map f [m..<n]) = setsum f (set [m..<n])"
|
|
226 |
by (simp add: listsum_distinct_conv_setsum_set)
|
|
227 |
|
|
228 |
lemma (in monoid_add) interv_listsum_conv_setsum_set_int:
|
|
229 |
"listsum (map f [k..l]) = setsum f (set [k..l])"
|
|
230 |
by (simp add: listsum_distinct_conv_setsum_set)
|
|
231 |
|
|
232 |
text {* General equivalence between @{const listsum} and @{const setsum} *}
|
|
233 |
lemma (in monoid_add) listsum_setsum_nth:
|
|
234 |
"listsum xs = (\<Sum> i = 0 ..< length xs. xs ! i)"
|
|
235 |
using interv_listsum_conv_setsum_set_nat [of "op ! xs" 0 "length xs"] by (simp add: map_nth)
|
|
236 |
|
|
237 |
|
|
238 |
subsection {* Further facts about @{const List.n_lists} *}
|
|
239 |
|
|
240 |
lemma length_n_lists: "length (List.n_lists n xs) = length xs ^ n"
|
|
241 |
by (induct n) (auto simp add: comp_def length_concat listsum_triv)
|
|
242 |
|
|
243 |
lemma distinct_n_lists:
|
|
244 |
assumes "distinct xs"
|
|
245 |
shows "distinct (List.n_lists n xs)"
|
|
246 |
proof (rule card_distinct)
|
|
247 |
from assms have card_length: "card (set xs) = length xs" by (rule distinct_card)
|
|
248 |
have "card (set (List.n_lists n xs)) = card (set xs) ^ n"
|
|
249 |
proof (induct n)
|
|
250 |
case 0 then show ?case by simp
|
|
251 |
next
|
|
252 |
case (Suc n)
|
|
253 |
moreover have "card (\<Union>ys\<in>set (List.n_lists n xs). (\<lambda>y. y # ys) ` set xs)
|
|
254 |
= (\<Sum>ys\<in>set (List.n_lists n xs). card ((\<lambda>y. y # ys) ` set xs))"
|
|
255 |
by (rule card_UN_disjoint) auto
|
|
256 |
moreover have "\<And>ys. card ((\<lambda>y. y # ys) ` set xs) = card (set xs)"
|
|
257 |
by (rule card_image) (simp add: inj_on_def)
|
|
258 |
ultimately show ?case by auto
|
|
259 |
qed
|
|
260 |
also have "\<dots> = length xs ^ n" by (simp add: card_length)
|
|
261 |
finally show "card (set (List.n_lists n xs)) = length (List.n_lists n xs)"
|
|
262 |
by (simp add: length_n_lists)
|
|
263 |
qed
|
|
264 |
|
|
265 |
|
|
266 |
subsection {* Tools setup *}
|
|
267 |
|
58320
|
268 |
lemmas setsum_code = setsum.set_conv_list
|
|
269 |
|
58101
|
270 |
lemma setsum_set_upto_conv_listsum_int [code_unfold]:
|
|
271 |
"setsum f (set [i..j::int]) = listsum (map f [i..j])"
|
|
272 |
by (simp add: interv_listsum_conv_setsum_set_int)
|
|
273 |
|
|
274 |
lemma setsum_set_upt_conv_listsum_nat [code_unfold]:
|
|
275 |
"setsum f (set [m..<n]) = listsum (map f [m..<n])"
|
|
276 |
by (simp add: interv_listsum_conv_setsum_set_nat)
|
|
277 |
|
|
278 |
context
|
|
279 |
begin
|
|
280 |
|
|
281 |
interpretation lifting_syntax .
|
|
282 |
|
|
283 |
lemma listsum_transfer[transfer_rule]:
|
|
284 |
assumes [transfer_rule]: "A 0 0"
|
|
285 |
assumes [transfer_rule]: "(A ===> A ===> A) op + op +"
|
|
286 |
shows "(list_all2 A ===> A) listsum listsum"
|
58320
|
287 |
unfolding listsum.eq_foldr [abs_def]
|
58101
|
288 |
by transfer_prover
|
|
289 |
|
|
290 |
end
|
|
291 |
|
58320
|
292 |
end
|