author | haftmann |
Fri, 03 Dec 2010 22:34:20 +0100 | |
changeset 40949 | 1d46d893d682 |
parent 39921 | 45f95e4de831 |
child 40951 | 6c35a88d8f61 |
permissions | -rw-r--r-- |
37025 | 1 |
(* Author: Florian Haftmann, TU Muenchen *) |
2 |
||
3 |
header {* Operations on lists beyond the standard List theory *} |
|
4 |
||
5 |
theory More_List |
|
40949
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
6 |
imports Main Multiset |
37025 | 7 |
begin |
8 |
||
9 |
hide_const (open) Finite_Set.fold |
|
10 |
||
11 |
text {* Repairing code generator setup *} |
|
12 |
||
13 |
declare (in lattice) Inf_fin_set_fold [code_unfold del] |
|
14 |
declare (in lattice) Sup_fin_set_fold [code_unfold del] |
|
15 |
declare (in linorder) Min_fin_set_fold [code_unfold del] |
|
16 |
declare (in linorder) Max_fin_set_fold [code_unfold del] |
|
17 |
declare (in complete_lattice) Inf_set_fold [code_unfold del] |
|
18 |
declare (in complete_lattice) Sup_set_fold [code_unfold del] |
|
19 |
||
20 |
text {* Fold combinator with canonical argument order *} |
|
21 |
||
22 |
primrec fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b \<Rightarrow> 'b" where |
|
23 |
"fold f [] = id" |
|
24 |
| "fold f (x # xs) = fold f xs \<circ> f x" |
|
25 |
||
26 |
lemma foldl_fold: |
|
27 |
"foldl f s xs = fold (\<lambda>x s. f s x) xs s" |
|
28 |
by (induct xs arbitrary: s) simp_all |
|
29 |
||
30 |
lemma foldr_fold_rev: |
|
31 |
"foldr f xs = fold f (rev xs)" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
32 |
by (simp add: foldr_foldl foldl_fold fun_eq_iff) |
37025 | 33 |
|
34 |
lemma fold_rev_conv [code_unfold]: |
|
35 |
"fold f (rev xs) = foldr f xs" |
|
36 |
by (simp add: foldr_fold_rev) |
|
37 |
||
38 |
lemma fold_cong [fundef_cong, recdef_cong]: |
|
39 |
"a = b \<Longrightarrow> xs = ys \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> f x = g x) |
|
40 |
\<Longrightarrow> fold f xs a = fold g ys b" |
|
41 |
by (induct ys arbitrary: a b xs) simp_all |
|
42 |
||
43 |
lemma fold_id: |
|
44 |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x = id" |
|
45 |
shows "fold f xs = id" |
|
46 |
using assms by (induct xs) simp_all |
|
47 |
||
39921 | 48 |
lemma fold_commute: |
37025 | 49 |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h" |
50 |
shows "h \<circ> fold g xs = fold f xs \<circ> h" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
51 |
using assms by (induct xs) (simp_all add: fun_eq_iff) |
37025 | 52 |
|
39921 | 53 |
lemma fold_commute_apply: |
54 |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<circ> g x = f x \<circ> h" |
|
55 |
shows "h (fold g xs s) = fold f xs (h s)" |
|
56 |
proof - |
|
57 |
from assms have "h \<circ> fold g xs = fold f xs \<circ> h" by (rule fold_commute) |
|
58 |
then show ?thesis by (simp add: fun_eq_iff) |
|
59 |
qed |
|
60 |
||
37025 | 61 |
lemma fold_invariant: |
62 |
assumes "\<And>x. x \<in> set xs \<Longrightarrow> Q x" and "P s" |
|
63 |
and "\<And>x s. Q x \<Longrightarrow> P s \<Longrightarrow> P (f x s)" |
|
64 |
shows "P (fold f xs s)" |
|
65 |
using assms by (induct xs arbitrary: s) simp_all |
|
66 |
||
67 |
lemma fold_weak_invariant: |
|
68 |
assumes "P s" |
|
69 |
and "\<And>s x. x \<in> set xs \<Longrightarrow> P s \<Longrightarrow> P (f x s)" |
|
70 |
shows "P (fold f xs s)" |
|
71 |
using assms by (induct xs arbitrary: s) simp_all |
|
72 |
||
73 |
lemma fold_append [simp]: |
|
74 |
"fold f (xs @ ys) = fold f ys \<circ> fold f xs" |
|
75 |
by (induct xs) simp_all |
|
76 |
||
77 |
lemma fold_map [code_unfold]: |
|
78 |
"fold g (map f xs) = fold (g o f) xs" |
|
79 |
by (induct xs) simp_all |
|
80 |
||
40949
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
81 |
lemma fold_remove1_split: |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
82 |
assumes f: "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
83 |
and x: "x \<in> set xs" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
84 |
shows "fold f xs = fold f (remove1 x xs) \<circ> f x" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
85 |
using assms by (induct xs) (auto simp add: o_assoc [symmetric]) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
86 |
|
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
87 |
lemma fold_multiset_equiv: |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
88 |
assumes f: "\<And>x y. x \<in># multiset_of xs \<Longrightarrow> y \<in># multiset_of xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
89 |
and equiv: "multiset_of xs = multiset_of ys" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
90 |
shows "fold f xs = fold f ys" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
91 |
using f equiv [symmetric] proof (induct xs arbitrary: ys) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
92 |
case Nil then show ?case by simp |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
93 |
next |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
94 |
case (Cons x xs) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
95 |
have "\<And>x y. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
96 |
by (rule Cons.prems(1)) (simp_all add: mem_set_multiset_eq Cons.prems(2)) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
97 |
moreover from Cons.prems have "x \<in> set ys" by (simp add: mem_set_multiset_eq) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
98 |
ultimately have "fold f ys = fold f (remove1 x ys) \<circ> f x" by (fact fold_remove1_split) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
99 |
moreover from Cons.prems have "fold f xs = fold f (remove1 x ys)" by (auto intro: Cons.hyps) |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
100 |
ultimately show ?case by simp |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
101 |
qed |
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
102 |
|
37025 | 103 |
lemma fold_rev: |
104 |
assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y" |
|
105 |
shows "fold f (rev xs) = fold f xs" |
|
40949
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
haftmann
parents:
39921
diff
changeset
|
106 |
by (rule fold_multiset_equiv, rule assms) (simp_all add: in_multiset_in_set) |
37025 | 107 |
|
108 |
lemma foldr_fold: |
|
109 |
assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y" |
|
110 |
shows "foldr f xs = fold f xs" |
|
111 |
using assms unfolding foldr_fold_rev by (rule fold_rev) |
|
112 |
||
113 |
lemma fold_Cons_rev: |
|
114 |
"fold Cons xs = append (rev xs)" |
|
115 |
by (induct xs) simp_all |
|
116 |
||
117 |
lemma rev_conv_fold [code]: |
|
118 |
"rev xs = fold Cons xs []" |
|
119 |
by (simp add: fold_Cons_rev) |
|
120 |
||
121 |
lemma fold_append_concat_rev: |
|
122 |
"fold append xss = append (concat (rev xss))" |
|
123 |
by (induct xss) simp_all |
|
124 |
||
125 |
lemma concat_conv_foldr [code]: |
|
126 |
"concat xss = foldr append xss []" |
|
127 |
by (simp add: fold_append_concat_rev foldr_fold_rev) |
|
128 |
||
129 |
lemma fold_plus_listsum_rev: |
|
130 |
"fold plus xs = plus (listsum (rev xs))" |
|
131 |
by (induct xs) (simp_all add: add.assoc) |
|
132 |
||
39773 | 133 |
lemma (in monoid_add) listsum_conv_fold [code]: |
134 |
"listsum xs = fold (\<lambda>x y. y + x) xs 0" |
|
135 |
by (auto simp add: listsum_foldl foldl_fold fun_eq_iff) |
|
37025 | 136 |
|
39773 | 137 |
lemma (in linorder) sort_key_conv_fold: |
37025 | 138 |
assumes "inj_on f (set xs)" |
139 |
shows "sort_key f xs = fold (insort_key f) xs []" |
|
140 |
proof - |
|
141 |
have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" |
|
142 |
proof (rule fold_rev, rule ext) |
|
143 |
fix zs |
|
144 |
fix x y |
|
145 |
assume "x \<in> set xs" "y \<in> set xs" |
|
146 |
with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD) |
|
39773 | 147 |
have **: "x = y \<longleftrightarrow> y = x" by auto |
37025 | 148 |
show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs" |
39773 | 149 |
by (induct zs) (auto intro: * simp add: **) |
37025 | 150 |
qed |
151 |
then show ?thesis by (simp add: sort_key_def foldr_fold_rev) |
|
152 |
qed |
|
153 |
||
39773 | 154 |
lemma (in linorder) sort_conv_fold: |
37025 | 155 |
"sort xs = fold insort xs []" |
156 |
by (rule sort_key_conv_fold) simp |
|
157 |
||
158 |
text {* @{const Finite_Set.fold} and @{const fold} *} |
|
159 |
||
160 |
lemma (in fun_left_comm) fold_set_remdups: |
|
161 |
"Finite_Set.fold f y (set xs) = fold f (remdups xs) y" |
|
162 |
by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb) |
|
163 |
||
164 |
lemma (in fun_left_comm_idem) fold_set: |
|
165 |
"Finite_Set.fold f y (set xs) = fold f xs y" |
|
166 |
by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm) |
|
167 |
||
168 |
lemma (in ab_semigroup_idem_mult) fold1_set: |
|
169 |
assumes "xs \<noteq> []" |
|
170 |
shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)" |
|
171 |
proof - |
|
172 |
interpret fun_left_comm_idem times by (fact fun_left_comm_idem) |
|
173 |
from assms obtain y ys where xs: "xs = y # ys" |
|
174 |
by (cases xs) auto |
|
175 |
show ?thesis |
|
176 |
proof (cases "set ys = {}") |
|
177 |
case True with xs show ?thesis by simp |
|
178 |
next |
|
179 |
case False |
|
180 |
then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)" |
|
181 |
by (simp only: finite_set fold1_eq_fold_idem) |
|
182 |
with xs show ?thesis by (simp add: fold_set mult_commute) |
|
183 |
qed |
|
184 |
qed |
|
185 |
||
186 |
lemma (in lattice) Inf_fin_set_fold: |
|
187 |
"Inf_fin (set (x # xs)) = fold inf xs x" |
|
188 |
proof - |
|
189 |
interpret ab_semigroup_idem_mult "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
190 |
by (fact ab_semigroup_idem_mult_inf) |
|
191 |
show ?thesis |
|
192 |
by (simp add: Inf_fin_def fold1_set del: set.simps) |
|
193 |
qed |
|
194 |
||
195 |
lemma (in lattice) Inf_fin_set_foldr [code_unfold]: |
|
196 |
"Inf_fin (set (x # xs)) = foldr inf xs x" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
197 |
by (simp add: Inf_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) |
37025 | 198 |
|
199 |
lemma (in lattice) Sup_fin_set_fold: |
|
200 |
"Sup_fin (set (x # xs)) = fold sup xs x" |
|
201 |
proof - |
|
202 |
interpret ab_semigroup_idem_mult "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
203 |
by (fact ab_semigroup_idem_mult_sup) |
|
204 |
show ?thesis |
|
205 |
by (simp add: Sup_fin_def fold1_set del: set.simps) |
|
206 |
qed |
|
207 |
||
208 |
lemma (in lattice) Sup_fin_set_foldr [code_unfold]: |
|
209 |
"Sup_fin (set (x # xs)) = foldr sup xs x" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
210 |
by (simp add: Sup_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) |
37025 | 211 |
|
212 |
lemma (in linorder) Min_fin_set_fold: |
|
213 |
"Min (set (x # xs)) = fold min xs x" |
|
214 |
proof - |
|
215 |
interpret ab_semigroup_idem_mult "min :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
216 |
by (fact ab_semigroup_idem_mult_min) |
|
217 |
show ?thesis |
|
218 |
by (simp add: Min_def fold1_set del: set.simps) |
|
219 |
qed |
|
220 |
||
221 |
lemma (in linorder) Min_fin_set_foldr [code_unfold]: |
|
222 |
"Min (set (x # xs)) = foldr min xs x" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
223 |
by (simp add: Min_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) |
37025 | 224 |
|
225 |
lemma (in linorder) Max_fin_set_fold: |
|
226 |
"Max (set (x # xs)) = fold max xs x" |
|
227 |
proof - |
|
228 |
interpret ab_semigroup_idem_mult "max :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
229 |
by (fact ab_semigroup_idem_mult_max) |
|
230 |
show ?thesis |
|
231 |
by (simp add: Max_def fold1_set del: set.simps) |
|
232 |
qed |
|
233 |
||
234 |
lemma (in linorder) Max_fin_set_foldr [code_unfold]: |
|
235 |
"Max (set (x # xs)) = foldr max xs x" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
236 |
by (simp add: Max_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) |
37025 | 237 |
|
238 |
lemma (in complete_lattice) Inf_set_fold: |
|
239 |
"Inf (set xs) = fold inf xs top" |
|
240 |
proof - |
|
241 |
interpret fun_left_comm_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
242 |
by (fact fun_left_comm_idem_inf) |
|
243 |
show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute) |
|
244 |
qed |
|
245 |
||
246 |
lemma (in complete_lattice) Inf_set_foldr [code_unfold]: |
|
247 |
"Inf (set xs) = foldr inf xs top" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
248 |
by (simp add: Inf_set_fold ac_simps foldr_fold fun_eq_iff) |
37025 | 249 |
|
250 |
lemma (in complete_lattice) Sup_set_fold: |
|
251 |
"Sup (set xs) = fold sup xs bot" |
|
252 |
proof - |
|
253 |
interpret fun_left_comm_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" |
|
254 |
by (fact fun_left_comm_idem_sup) |
|
255 |
show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute) |
|
256 |
qed |
|
257 |
||
258 |
lemma (in complete_lattice) Sup_set_foldr [code_unfold]: |
|
259 |
"Sup (set xs) = foldr sup xs bot" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
260 |
by (simp add: Sup_set_fold ac_simps foldr_fold fun_eq_iff) |
37025 | 261 |
|
262 |
lemma (in complete_lattice) INFI_set_fold: |
|
263 |
"INFI (set xs) f = fold (inf \<circ> f) xs top" |
|
264 |
unfolding INFI_def set_map [symmetric] Inf_set_fold fold_map .. |
|
265 |
||
266 |
lemma (in complete_lattice) SUPR_set_fold: |
|
267 |
"SUPR (set xs) f = fold (sup \<circ> f) xs bot" |
|
268 |
unfolding SUPR_def set_map [symmetric] Sup_set_fold fold_map .. |
|
269 |
||
37028 | 270 |
text {* @{text nth_map} *} |
37025 | 271 |
|
272 |
definition nth_map :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a list \<Rightarrow> 'a list" where |
|
273 |
"nth_map n f xs = (if n < length xs then |
|
274 |
take n xs @ [f (xs ! n)] @ drop (Suc n) xs |
|
275 |
else xs)" |
|
276 |
||
277 |
lemma nth_map_id: |
|
278 |
"n \<ge> length xs \<Longrightarrow> nth_map n f xs = xs" |
|
279 |
by (simp add: nth_map_def) |
|
280 |
||
281 |
lemma nth_map_unfold: |
|
282 |
"n < length xs \<Longrightarrow> nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs" |
|
283 |
by (simp add: nth_map_def) |
|
284 |
||
285 |
lemma nth_map_Nil [simp]: |
|
286 |
"nth_map n f [] = []" |
|
287 |
by (simp add: nth_map_def) |
|
288 |
||
289 |
lemma nth_map_zero [simp]: |
|
290 |
"nth_map 0 f (x # xs) = f x # xs" |
|
291 |
by (simp add: nth_map_def) |
|
292 |
||
293 |
lemma nth_map_Suc [simp]: |
|
294 |
"nth_map (Suc n) f (x # xs) = x # nth_map n f xs" |
|
295 |
by (simp add: nth_map_def) |
|
296 |
||
297 |
end |