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