| author | blanchet | 
| Sun, 01 May 2011 18:37:24 +0200 | |
| changeset 42535 | 3c1f302b3ee6 | 
| parent 40951 | 6c35a88d8f61 | 
| child 42871 | 1c0b99f950d9 | 
| 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: 
39921diff
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: 
39198diff
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: 
39198diff
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: 
39921diff
changeset | 81 | lemma fold_remove1_split: | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
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: 
39921diff
changeset | 83 | and x: "x \<in> set xs" | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
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: 
39921diff
changeset | 85 | using assms by (induct xs) (auto simp add: o_assoc [symmetric]) | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 86 | |
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 87 | lemma fold_multiset_equiv: | 
| 40951 | 88 | 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" | 
| 40949 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 89 | and equiv: "multiset_of xs = multiset_of ys" | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 90 | shows "fold f xs = fold f ys" | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 91 | using f equiv [symmetric] proof (induct xs arbitrary: ys) | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 92 | case Nil then show ?case by simp | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 93 | next | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 94 | case (Cons x xs) | 
| 40951 | 95 | then have *: "set ys = set (x # xs)" by (blast dest: multiset_of_eq_setD) | 
| 40949 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 96 | have "\<And>x y. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" | 
| 40951 | 97 | by (rule Cons.prems(1)) (simp_all add: *) | 
| 98 | moreover from * have "x \<in> set ys" by simp | |
| 40949 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 99 | 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: 
39921diff
changeset | 100 | 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: 
39921diff
changeset | 101 | ultimately show ?case by simp | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 102 | qed | 
| 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 103 | |
| 37025 | 104 | lemma fold_rev: | 
| 105 | assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y" | |
| 106 | shows "fold f (rev xs) = fold f xs" | |
| 40949 
1d46d893d682
lemmas fold_remove1_split and fold_multiset_equiv
 haftmann parents: 
39921diff
changeset | 107 | by (rule fold_multiset_equiv, rule assms) (simp_all add: in_multiset_in_set) | 
| 37025 | 108 | |
| 109 | lemma foldr_fold: | |
| 110 | assumes "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f y \<circ> f x = f x \<circ> f y" | |
| 111 | shows "foldr f xs = fold f xs" | |
| 112 | using assms unfolding foldr_fold_rev by (rule fold_rev) | |
| 113 | ||
| 114 | lemma fold_Cons_rev: | |
| 115 | "fold Cons xs = append (rev xs)" | |
| 116 | by (induct xs) simp_all | |
| 117 | ||
| 118 | lemma rev_conv_fold [code]: | |
| 119 | "rev xs = fold Cons xs []" | |
| 120 | by (simp add: fold_Cons_rev) | |
| 121 | ||
| 122 | lemma fold_append_concat_rev: | |
| 123 | "fold append xss = append (concat (rev xss))" | |
| 124 | by (induct xss) simp_all | |
| 125 | ||
| 126 | lemma concat_conv_foldr [code]: | |
| 127 | "concat xss = foldr append xss []" | |
| 128 | by (simp add: fold_append_concat_rev foldr_fold_rev) | |
| 129 | ||
| 130 | lemma fold_plus_listsum_rev: | |
| 131 | "fold plus xs = plus (listsum (rev xs))" | |
| 132 | by (induct xs) (simp_all add: add.assoc) | |
| 133 | ||
| 39773 | 134 | lemma (in monoid_add) listsum_conv_fold [code]: | 
| 135 | "listsum xs = fold (\<lambda>x y. y + x) xs 0" | |
| 136 | by (auto simp add: listsum_foldl foldl_fold fun_eq_iff) | |
| 37025 | 137 | |
| 39773 | 138 | lemma (in linorder) sort_key_conv_fold: | 
| 37025 | 139 | assumes "inj_on f (set xs)" | 
| 140 | shows "sort_key f xs = fold (insort_key f) xs []" | |
| 141 | proof - | |
| 142 | have "fold (insort_key f) (rev xs) = fold (insort_key f) xs" | |
| 143 | proof (rule fold_rev, rule ext) | |
| 144 | fix zs | |
| 145 | fix x y | |
| 146 | assume "x \<in> set xs" "y \<in> set xs" | |
| 147 | with assms have *: "f y = f x \<Longrightarrow> y = x" by (auto dest: inj_onD) | |
| 39773 | 148 | have **: "x = y \<longleftrightarrow> y = x" by auto | 
| 37025 | 149 | show "(insort_key f y \<circ> insort_key f x) zs = (insort_key f x \<circ> insort_key f y) zs" | 
| 39773 | 150 | by (induct zs) (auto intro: * simp add: **) | 
| 37025 | 151 | qed | 
| 152 | then show ?thesis by (simp add: sort_key_def foldr_fold_rev) | |
| 153 | qed | |
| 154 | ||
| 39773 | 155 | lemma (in linorder) sort_conv_fold: | 
| 37025 | 156 | "sort xs = fold insort xs []" | 
| 157 | by (rule sort_key_conv_fold) simp | |
| 158 | ||
| 159 | text {* @{const Finite_Set.fold} and @{const fold} *}
 | |
| 160 | ||
| 161 | lemma (in fun_left_comm) fold_set_remdups: | |
| 162 | "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" | |
| 163 | by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm insert_absorb) | |
| 164 | ||
| 165 | lemma (in fun_left_comm_idem) fold_set: | |
| 166 | "Finite_Set.fold f y (set xs) = fold f xs y" | |
| 167 | by (rule sym, induct xs arbitrary: y) (simp_all add: fold_fun_comm) | |
| 168 | ||
| 169 | lemma (in ab_semigroup_idem_mult) fold1_set: | |
| 170 | assumes "xs \<noteq> []" | |
| 171 | shows "Finite_Set.fold1 times (set xs) = fold times (tl xs) (hd xs)" | |
| 172 | proof - | |
| 173 | interpret fun_left_comm_idem times by (fact fun_left_comm_idem) | |
| 174 | from assms obtain y ys where xs: "xs = y # ys" | |
| 175 | by (cases xs) auto | |
| 176 | show ?thesis | |
| 177 |   proof (cases "set ys = {}")
 | |
| 178 | case True with xs show ?thesis by simp | |
| 179 | next | |
| 180 | case False | |
| 181 | then have "fold1 times (insert y (set ys)) = Finite_Set.fold times y (set ys)" | |
| 182 | by (simp only: finite_set fold1_eq_fold_idem) | |
| 183 | with xs show ?thesis by (simp add: fold_set mult_commute) | |
| 184 | qed | |
| 185 | qed | |
| 186 | ||
| 187 | lemma (in lattice) Inf_fin_set_fold: | |
| 188 | "Inf_fin (set (x # xs)) = fold inf xs x" | |
| 189 | proof - | |
| 190 | interpret ab_semigroup_idem_mult "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 191 | by (fact ab_semigroup_idem_mult_inf) | |
| 192 | show ?thesis | |
| 193 | by (simp add: Inf_fin_def fold1_set del: set.simps) | |
| 194 | qed | |
| 195 | ||
| 196 | lemma (in lattice) Inf_fin_set_foldr [code_unfold]: | |
| 197 | "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: 
39198diff
changeset | 198 | by (simp add: Inf_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) | 
| 37025 | 199 | |
| 200 | lemma (in lattice) Sup_fin_set_fold: | |
| 201 | "Sup_fin (set (x # xs)) = fold sup xs x" | |
| 202 | proof - | |
| 203 | interpret ab_semigroup_idem_mult "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 204 | by (fact ab_semigroup_idem_mult_sup) | |
| 205 | show ?thesis | |
| 206 | by (simp add: Sup_fin_def fold1_set del: set.simps) | |
| 207 | qed | |
| 208 | ||
| 209 | lemma (in lattice) Sup_fin_set_foldr [code_unfold]: | |
| 210 | "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: 
39198diff
changeset | 211 | by (simp add: Sup_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) | 
| 37025 | 212 | |
| 213 | lemma (in linorder) Min_fin_set_fold: | |
| 214 | "Min (set (x # xs)) = fold min xs x" | |
| 215 | proof - | |
| 216 | interpret ab_semigroup_idem_mult "min :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 217 | by (fact ab_semigroup_idem_mult_min) | |
| 218 | show ?thesis | |
| 219 | by (simp add: Min_def fold1_set del: set.simps) | |
| 220 | qed | |
| 221 | ||
| 222 | lemma (in linorder) Min_fin_set_foldr [code_unfold]: | |
| 223 | "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: 
39198diff
changeset | 224 | by (simp add: Min_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) | 
| 37025 | 225 | |
| 226 | lemma (in linorder) Max_fin_set_fold: | |
| 227 | "Max (set (x # xs)) = fold max xs x" | |
| 228 | proof - | |
| 229 | interpret ab_semigroup_idem_mult "max :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 230 | by (fact ab_semigroup_idem_mult_max) | |
| 231 | show ?thesis | |
| 232 | by (simp add: Max_def fold1_set del: set.simps) | |
| 233 | qed | |
| 234 | ||
| 235 | lemma (in linorder) Max_fin_set_foldr [code_unfold]: | |
| 236 | "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: 
39198diff
changeset | 237 | by (simp add: Max_fin_set_fold ac_simps foldr_fold fun_eq_iff del: set.simps) | 
| 37025 | 238 | |
| 239 | lemma (in complete_lattice) Inf_set_fold: | |
| 240 | "Inf (set xs) = fold inf xs top" | |
| 241 | proof - | |
| 242 | interpret fun_left_comm_idem "inf :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 243 | by (fact fun_left_comm_idem_inf) | |
| 244 | show ?thesis by (simp add: Inf_fold_inf fold_set inf_commute) | |
| 245 | qed | |
| 246 | ||
| 247 | lemma (in complete_lattice) Inf_set_foldr [code_unfold]: | |
| 248 | "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: 
39198diff
changeset | 249 | by (simp add: Inf_set_fold ac_simps foldr_fold fun_eq_iff) | 
| 37025 | 250 | |
| 251 | lemma (in complete_lattice) Sup_set_fold: | |
| 252 | "Sup (set xs) = fold sup xs bot" | |
| 253 | proof - | |
| 254 | interpret fun_left_comm_idem "sup :: 'a \<Rightarrow> 'a \<Rightarrow> 'a" | |
| 255 | by (fact fun_left_comm_idem_sup) | |
| 256 | show ?thesis by (simp add: Sup_fold_sup fold_set sup_commute) | |
| 257 | qed | |
| 258 | ||
| 259 | lemma (in complete_lattice) Sup_set_foldr [code_unfold]: | |
| 260 | "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: 
39198diff
changeset | 261 | by (simp add: Sup_set_fold ac_simps foldr_fold fun_eq_iff) | 
| 37025 | 262 | |
| 263 | lemma (in complete_lattice) INFI_set_fold: | |
| 264 | "INFI (set xs) f = fold (inf \<circ> f) xs top" | |
| 265 | unfolding INFI_def set_map [symmetric] Inf_set_fold fold_map .. | |
| 266 | ||
| 267 | lemma (in complete_lattice) SUPR_set_fold: | |
| 268 | "SUPR (set xs) f = fold (sup \<circ> f) xs bot" | |
| 269 | unfolding SUPR_def set_map [symmetric] Sup_set_fold fold_map .. | |
| 270 | ||
| 37028 | 271 | text {* @{text nth_map} *}
 | 
| 37025 | 272 | |
| 273 | definition nth_map :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a list \<Rightarrow> 'a list" where
 | |
| 274 | "nth_map n f xs = (if n < length xs then | |
| 275 | take n xs @ [f (xs ! n)] @ drop (Suc n) xs | |
| 276 | else xs)" | |
| 277 | ||
| 278 | lemma nth_map_id: | |
| 279 | "n \<ge> length xs \<Longrightarrow> nth_map n f xs = xs" | |
| 280 | by (simp add: nth_map_def) | |
| 281 | ||
| 282 | lemma nth_map_unfold: | |
| 283 | "n < length xs \<Longrightarrow> nth_map n f xs = take n xs @ [f (xs ! n)] @ drop (Suc n) xs" | |
| 284 | by (simp add: nth_map_def) | |
| 285 | ||
| 286 | lemma nth_map_Nil [simp]: | |
| 287 | "nth_map n f [] = []" | |
| 288 | by (simp add: nth_map_def) | |
| 289 | ||
| 290 | lemma nth_map_zero [simp]: | |
| 291 | "nth_map 0 f (x # xs) = f x # xs" | |
| 292 | by (simp add: nth_map_def) | |
| 293 | ||
| 294 | lemma nth_map_Suc [simp]: | |
| 295 | "nth_map (Suc n) f (x # xs) = x # nth_map n f xs" | |
| 296 | by (simp add: nth_map_def) | |
| 297 | ||
| 298 | end |