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