| author | huffman | 
| Wed, 16 Jan 2008 22:41:49 +0100 | |
| changeset 25920 | 8df5eabda5f6 | 
| parent 23373 | ead82c82da9e | 
| child 25966 | 74f6817870f9 | 
| permissions | -rw-r--r-- | 
| 22803 | 1 | (* Title: HOL/Library/AssocList.thy | 
| 19234 | 2 | ID: $Id$ | 
| 3 | Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser | |
| 4 | *) | |
| 5 | ||
| 6 | header {* Map operations implemented on association lists*}
 | |
| 7 | ||
| 8 | theory AssocList | |
| 9 | imports Map | |
| 10 | begin | |
| 11 | ||
| 22740 | 12 | text {*
 | 
| 13 | The operations preserve distinctness of keys and | |
| 14 |   function @{term "clearjunk"} distributes over them. Since 
 | |
| 15 |   @{term clearjunk} enforces distinctness of keys it can be used
 | |
| 16 | to establish the invariant, e.g. for inductive proofs. | |
| 17 | *} | |
| 19234 | 18 | |
| 22740 | 19 | fun | 
| 20 |   delete :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 21 | where | |
| 23373 | 22 | "delete k [] = []" | 
| 22740 | 23 | | "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)" | 
| 19323 | 24 | |
| 22740 | 25 | fun | 
| 26 |   update :: "'key \<Rightarrow> 'val \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 27 | where | |
| 28 | "update k v [] = [(k, v)]" | |
| 29 | | "update k v (p#ps) = (if fst p = k then (k, v) # ps else p # update k v ps)" | |
| 19234 | 30 | |
| 22740 | 31 | function | 
| 32 |   updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 33 | where | |
| 34 | "updates [] vs ps = ps" | |
| 35 | | "updates (k#ks) vs ps = (case vs | |
| 36 | of [] \<Rightarrow> ps | |
| 37 | | (v#vs') \<Rightarrow> updates ks vs' (update k v ps))" | |
| 38 | by pat_completeness auto | |
| 39 | termination by lexicographic_order | |
| 19323 | 40 | |
| 22740 | 41 | fun | 
| 42 |   merge :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 43 | where | |
| 44 | "merge qs [] = qs" | |
| 45 | | "merge qs (p#ps) = update (fst p) (snd p) (merge qs ps)" | |
| 19234 | 46 | |
| 47 | lemma length_delete_le: "length (delete k al) \<le> length al" | |
| 48 | proof (induct al) | |
| 22740 | 49 | case Nil thus ?case by simp | 
| 19234 | 50 | next | 
| 51 | case (Cons a al) | |
| 52 | note length_filter_le [of "\<lambda>p. fst p \<noteq> fst a" al] | |
| 53 | also have "\<And>n. n \<le> Suc n" | |
| 54 | by simp | |
| 23281 | 55 | finally have "length [p\<leftarrow>al . fst p \<noteq> fst a] \<le> Suc (length al)" . | 
| 19234 | 56 | with Cons show ?case | 
| 22740 | 57 | by auto | 
| 19234 | 58 | qed | 
| 59 | ||
| 22740 | 60 | lemma compose_hint [simp]: | 
| 61 | "length (delete k al) < Suc (length al)" | |
| 19234 | 62 | proof - | 
| 63 | note length_delete_le | |
| 64 | also have "\<And>n. n < Suc n" | |
| 65 | by simp | |
| 66 | finally show ?thesis . | |
| 67 | qed | |
| 68 | ||
| 22740 | 69 | function | 
| 70 |   compose :: "('key \<times> 'a) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('key \<times> 'b) list"
 | |
| 71 | where | |
| 72 | "compose [] ys = []" | |
| 73 | | "compose (x#xs) ys = (case map_of ys (snd x) | |
| 74 | of None \<Rightarrow> compose (delete (fst x) xs) ys | |
| 75 | | Some v \<Rightarrow> (fst x, v) # compose xs ys)" | |
| 76 | by pat_completeness auto | |
| 77 | termination by lexicographic_order | |
| 19234 | 78 | |
| 22740 | 79 | fun | 
| 80 |   restrict :: "'key set \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 81 | where | |
| 23373 | 82 | "restrict A [] = []" | 
| 22740 | 83 | | "restrict A (p#ps) = (if fst p \<in> A then p#restrict A ps else restrict A ps)" | 
| 19234 | 84 | |
| 22740 | 85 | fun | 
| 86 |   map_ran :: "('key \<Rightarrow> 'val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 87 | where | |
| 23373 | 88 | "map_ran f [] = []" | 
| 22740 | 89 | | "map_ran f (p#ps) = (fst p, f (fst p) (snd p)) # map_ran f ps" | 
| 90 | ||
| 91 | fun | |
| 92 |   clearjunk  :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
 | |
| 93 | where | |
| 23373 | 94 | "clearjunk [] = []" | 
| 22740 | 95 | | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)" | 
| 96 | ||
| 97 | lemmas [simp del] = compose_hint | |
| 19234 | 98 | |
| 99 | ||
| 100 | subsection {* Lookup *}
 | |
| 101 | ||
| 22740 | 102 | lemma lookup_simps [code func]: | 
| 19234 | 103 | "map_of [] k = None" | 
| 104 | "map_of (p#ps) k = (if fst p = k then Some (snd p) else map_of ps k)" | |
| 105 | by simp_all | |
| 106 | ||
| 22740 | 107 | |
| 19234 | 108 | subsection {* @{const delete} *}
 | 
| 109 | ||
| 22740 | 110 | lemma delete_def: | 
| 111 | "delete k xs = filter (\<lambda>p. fst p \<noteq> k) xs" | |
| 112 | by (induct xs) auto | |
| 19234 | 113 | |
| 22740 | 114 | lemma delete_id [simp]: "k \<notin> fst ` set al \<Longrightarrow> delete k al = al" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 115 | by (induct al) auto | 
| 19234 | 116 | |
| 117 | lemma delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'" | |
| 118 | by (induct al) auto | |
| 119 | ||
| 120 | lemma delete_conv': "map_of (delete k al) = ((map_of al)(k := None))" | |
| 121 | by (rule ext) (rule delete_conv) | |
| 122 | ||
| 123 | lemma delete_idem: "delete k (delete k al) = delete k al" | |
| 124 | by (induct al) auto | |
| 125 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 126 | lemma map_of_delete [simp]: | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 127 | "k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'" | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 128 | by (induct al) auto | 
| 19234 | 129 | |
| 130 | lemma delete_notin_dom: "k \<notin> fst ` set (delete k al)" | |
| 131 | by (induct al) auto | |
| 132 | ||
| 133 | lemma dom_delete_subset: "fst ` set (delete k al) \<subseteq> fst ` set al" | |
| 134 | by (induct al) auto | |
| 135 | ||
| 136 | lemma distinct_delete: | |
| 137 | assumes "distinct (map fst al)" | |
| 138 | shows "distinct (map fst (delete k al))" | |
| 23373 | 139 | using assms | 
| 19234 | 140 | proof (induct al) | 
| 141 | case Nil thus ?case by simp | |
| 142 | next | |
| 143 | case (Cons a al) | |
| 144 | from Cons.prems obtain | |
| 145 | a_notin_al: "fst a \<notin> fst ` set al" and | |
| 146 | dist_al: "distinct (map fst al)" | |
| 147 | by auto | |
| 148 | show ?case | |
| 149 | proof (cases "fst a = k") | |
| 150 | case True | |
| 23373 | 151 | with Cons dist_al show ?thesis by simp | 
| 19234 | 152 | next | 
| 153 | case False | |
| 154 | from dist_al | |
| 155 | have "distinct (map fst (delete k al))" | |
| 156 | by (rule Cons.hyps) | |
| 157 | moreover from a_notin_al dom_delete_subset [of k al] | |
| 158 | have "fst a \<notin> fst ` set (delete k al)" | |
| 159 | by blast | |
| 160 | ultimately show ?thesis using False by simp | |
| 161 | qed | |
| 162 | qed | |
| 163 | ||
| 164 | lemma delete_twist: "delete x (delete y al) = delete y (delete x al)" | |
| 165 | by (induct al) auto | |
| 166 | ||
| 167 | lemma clearjunk_delete: "clearjunk (delete x al) = delete x (clearjunk al)" | |
| 168 | by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist) | |
| 169 | ||
| 23373 | 170 | |
| 19234 | 171 | subsection {* @{const clearjunk} *}
 | 
| 172 | ||
| 173 | lemma insert_fst_filter: | |
| 174 |   "insert a(fst ` {x \<in> set ps. fst x \<noteq> a}) = insert a (fst ` set ps)"
 | |
| 175 | by (induct ps) auto | |
| 176 | ||
| 177 | lemma dom_clearjunk: "fst ` set (clearjunk al) = fst ` set al" | |
| 178 | by (induct al rule: clearjunk.induct) (simp_all add: insert_fst_filter delete_def) | |
| 179 | ||
| 180 | lemma notin_filter_fst: "a \<notin> fst ` {x \<in> set ps. fst x \<noteq> a}"
 | |
| 181 | by (induct ps) auto | |
| 182 | ||
| 183 | lemma distinct_clearjunk [simp]: "distinct (map fst (clearjunk al))" | |
| 184 | by (induct al rule: clearjunk.induct) | |
| 185 | (simp_all add: dom_clearjunk notin_filter_fst delete_def) | |
| 186 | ||
| 23281 | 187 | lemma map_of_filter: "k \<noteq> a \<Longrightarrow> map_of [q\<leftarrow>ps . fst q \<noteq> a] k = map_of ps k" | 
| 19234 | 188 | by (induct ps) auto | 
| 189 | ||
| 190 | lemma map_of_clearjunk: "map_of (clearjunk al) = map_of al" | |
| 191 | apply (rule ext) | |
| 192 | apply (induct al rule: clearjunk.induct) | |
| 193 | apply simp | |
| 194 | apply (simp add: map_of_filter) | |
| 195 | done | |
| 196 | ||
| 197 | lemma length_clearjunk: "length (clearjunk al) \<le> length al" | |
| 198 | proof (induct al rule: clearjunk.induct [case_names Nil Cons]) | |
| 199 | case Nil thus ?case by simp | |
| 200 | next | |
| 22740 | 201 | case (Cons p ps) | 
| 23281 | 202 | from Cons have "length (clearjunk [q\<leftarrow>ps . fst q \<noteq> fst p]) \<le> length [q\<leftarrow>ps . fst q \<noteq> fst p]" | 
| 19234 | 203 | by (simp add: delete_def) | 
| 204 | also have "\<dots> \<le> length ps" | |
| 205 | by simp | |
| 206 | finally show ?case | |
| 207 | by (simp add: delete_def) | |
| 208 | qed | |
| 209 | ||
| 23281 | 210 | lemma notin_fst_filter: "a \<notin> fst ` set ps \<Longrightarrow> [q\<leftarrow>ps . fst q \<noteq> a] = ps" | 
| 19234 | 211 | by (induct ps) auto | 
| 212 | ||
| 213 | lemma distinct_clearjunk_id [simp]: "distinct (map fst al) \<Longrightarrow> clearjunk al = al" | |
| 214 | by (induct al rule: clearjunk.induct) (auto simp add: notin_fst_filter) | |
| 215 | ||
| 216 | lemma clearjunk_idem: "clearjunk (clearjunk al) = clearjunk al" | |
| 217 | by simp | |
| 218 | ||
| 23373 | 219 | |
| 19234 | 220 | subsection {* @{const dom} and @{term "ran"} *}
 | 
| 221 | ||
| 222 | lemma dom_map_of': "fst ` set al = dom (map_of al)" | |
| 223 | by (induct al) auto | |
| 224 | ||
| 225 | lemmas dom_map_of = dom_map_of' [symmetric] | |
| 226 | ||
| 227 | lemma ran_clearjunk: "ran (map_of (clearjunk al)) = ran (map_of al)" | |
| 228 | by (simp add: map_of_clearjunk) | |
| 229 | ||
| 230 | lemma ran_distinct: | |
| 231 | assumes dist: "distinct (map fst al)" | |
| 232 | shows "ran (map_of al) = snd ` set al" | |
| 233 | using dist | |
| 234 | proof (induct al) | |
| 235 | case Nil | |
| 236 | thus ?case by simp | |
| 237 | next | |
| 238 | case (Cons a al) | |
| 239 | hence hyp: "snd ` set al = ran (map_of al)" | |
| 240 | by simp | |
| 241 | ||
| 242 |   have "ran (map_of (a # al)) = {snd a} \<union> ran (map_of al)"
 | |
| 243 | proof | |
| 244 |     show "ran (map_of (a # al)) \<subseteq> {snd a} \<union> ran (map_of al)"
 | |
| 245 | proof | |
| 246 | fix v | |
| 247 | assume "v \<in> ran (map_of (a#al))" | |
| 248 | then obtain x where "map_of (a#al) x = Some v" | |
| 249 | by (auto simp add: ran_def) | |
| 250 |       then show "v \<in> {snd a} \<union> ran (map_of al)"
 | |
| 251 | by (auto split: split_if_asm simp add: ran_def) | |
| 252 | qed | |
| 253 | next | |
| 254 |     show "{snd a} \<union> ran (map_of al) \<subseteq> ran (map_of (a # al))"
 | |
| 255 | proof | |
| 256 | fix v | |
| 257 |       assume v_in: "v \<in> {snd a} \<union> ran (map_of al)"
 | |
| 258 | show "v \<in> ran (map_of (a#al))" | |
| 259 | proof (cases "v=snd a") | |
| 260 | case True | |
| 261 | with v_in show ?thesis | |
| 262 | by (auto simp add: ran_def) | |
| 263 | next | |
| 264 | case False | |
| 265 | with v_in have "v \<in> ran (map_of al)" by auto | |
| 266 | then obtain x where al_x: "map_of al x = Some v" | |
| 267 | by (auto simp add: ran_def) | |
| 268 | from map_of_SomeD [OF this] | |
| 269 | have "x \<in> fst ` set al" | |
| 270 | by (force simp add: image_def) | |
| 271 | with Cons.prems have "x\<noteq>fst a" | |
| 272 | by - (rule ccontr,simp) | |
| 273 | with al_x | |
| 274 | show ?thesis | |
| 275 | by (auto simp add: ran_def) | |
| 276 | qed | |
| 277 | qed | |
| 278 | qed | |
| 279 | with hyp show ?case | |
| 280 | by (simp only:) auto | |
| 281 | qed | |
| 282 | ||
| 283 | lemma ran_map_of: "ran (map_of al) = snd ` set (clearjunk al)" | |
| 284 | proof - | |
| 285 | have "ran (map_of al) = ran (map_of (clearjunk al))" | |
| 286 | by (simp add: ran_clearjunk) | |
| 287 | also have "\<dots> = snd ` set (clearjunk al)" | |
| 288 | by (simp add: ran_distinct) | |
| 289 | finally show ?thesis . | |
| 290 | qed | |
| 291 | ||
| 23373 | 292 | |
| 19234 | 293 | subsection {* @{const update} *}
 | 
| 294 | ||
| 295 | lemma update_conv: "map_of (update k v al) k' = ((map_of al)(k\<mapsto>v)) k'" | |
| 296 | by (induct al) auto | |
| 297 | ||
| 298 | lemma update_conv': "map_of (update k v al) = ((map_of al)(k\<mapsto>v))" | |
| 299 | by (rule ext) (rule update_conv) | |
| 300 | ||
| 301 | lemma dom_update: "fst ` set (update k v al) = {k} \<union> fst ` set al"
 | |
| 302 | by (induct al) auto | |
| 303 | ||
| 304 | lemma distinct_update: | |
| 305 | assumes "distinct (map fst al)" | |
| 306 | shows "distinct (map fst (update k v al))" | |
| 23373 | 307 | using assms | 
| 19234 | 308 | proof (induct al) | 
| 309 | case Nil thus ?case by simp | |
| 310 | next | |
| 311 | case (Cons a al) | |
| 312 | from Cons.prems obtain | |
| 313 | a_notin_al: "fst a \<notin> fst ` set al" and | |
| 314 | dist_al: "distinct (map fst al)" | |
| 315 | by auto | |
| 316 | show ?case | |
| 317 | proof (cases "fst a = k") | |
| 318 | case True | |
| 319 | from True dist_al a_notin_al show ?thesis by simp | |
| 320 | next | |
| 321 | case False | |
| 322 | from dist_al | |
| 323 | have "distinct (map fst (update k v al))" | |
| 324 | by (rule Cons.hyps) | |
| 325 | with False a_notin_al show ?thesis by (simp add: dom_update) | |
| 326 | qed | |
| 327 | qed | |
| 328 | ||
| 329 | lemma update_filter: | |
| 23281 | 330 | "a\<noteq>k \<Longrightarrow> update k v [q\<leftarrow>ps . fst q \<noteq> a] = [q\<leftarrow>update k v ps . fst q \<noteq> a]" | 
| 19234 | 331 | by (induct ps) auto | 
| 332 | ||
| 333 | lemma clearjunk_update: "clearjunk (update k v al) = update k v (clearjunk al)" | |
| 334 | by (induct al rule: clearjunk.induct) (auto simp add: update_filter delete_def) | |
| 335 | ||
| 336 | lemma update_triv: "map_of al k = Some v \<Longrightarrow> update k v al = al" | |
| 337 | by (induct al) auto | |
| 338 | ||
| 339 | lemma update_nonempty [simp]: "update k v al \<noteq> []" | |
| 340 | by (induct al) auto | |
| 341 | ||
| 342 | lemma update_eqD: "update k v al = update k v' al' \<Longrightarrow> v=v'" | |
| 20503 | 343 | proof (induct al arbitrary: al') | 
| 19234 | 344 | case Nil thus ?case | 
| 345 | by (cases al') (auto split: split_if_asm) | |
| 346 | next | |
| 347 | case Cons thus ?case | |
| 348 | by (cases al') (auto split: split_if_asm) | |
| 349 | qed | |
| 350 | ||
| 351 | lemma update_last [simp]: "update k v (update k v' al) = update k v al" | |
| 352 | by (induct al) auto | |
| 353 | ||
| 354 | text {* Note that the lists are not necessarily the same:
 | |
| 355 |         @{term "update k v (update k' v' []) = [(k',v'),(k,v)]"} and 
 | |
| 356 |         @{term "update k' v' (update k v []) = [(k,v),(k',v')]"}.*}
 | |
| 357 | lemma update_swap: "k\<noteq>k' | |
| 358 | \<Longrightarrow> map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))" | |
| 359 | by (auto simp add: update_conv' intro: ext) | |
| 360 | ||
| 361 | lemma update_Some_unfold: | |
| 362 | "(map_of (update k v al) x = Some y) = | |
| 363 | (x = k \<and> v = y \<or> x \<noteq> k \<and> map_of al x = Some y)" | |
| 364 | by (simp add: update_conv' map_upd_Some_unfold) | |
| 365 | ||
| 366 | lemma image_update[simp]: "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A" | |
| 367 | by (simp add: update_conv' image_map_upd) | |
| 368 | ||
| 369 | ||
| 370 | subsection {* @{const updates} *}
 | |
| 371 | ||
| 372 | lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\<mapsto>]vs)) k" | |
| 20503 | 373 | proof (induct ks arbitrary: vs al) | 
| 19234 | 374 | case Nil | 
| 375 | thus ?case by simp | |
| 376 | next | |
| 377 | case (Cons k ks) | |
| 378 | show ?case | |
| 379 | proof (cases vs) | |
| 380 | case Nil | |
| 381 | with Cons show ?thesis by simp | |
| 382 | next | |
| 383 | case (Cons k ks') | |
| 384 | with Cons.hyps show ?thesis | |
| 385 | by (simp add: update_conv fun_upd_def) | |
| 386 | qed | |
| 387 | qed | |
| 388 | ||
| 389 | lemma updates_conv': "map_of (updates ks vs al) = ((map_of al)(ks[\<mapsto>]vs))" | |
| 390 | by (rule ext) (rule updates_conv) | |
| 391 | ||
| 392 | lemma distinct_updates: | |
| 393 | assumes "distinct (map fst al)" | |
| 394 | shows "distinct (map fst (updates ks vs al))" | |
| 23373 | 395 | using assms | 
| 22740 | 396 | by (induct ks arbitrary: vs al) | 
| 397 | (auto simp add: distinct_update split: list.splits) | |
| 19234 | 398 | |
| 399 | lemma clearjunk_updates: | |
| 400 | "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)" | |
| 20503 | 401 | by (induct ks arbitrary: vs al) (auto simp add: clearjunk_update split: list.splits) | 
| 19234 | 402 | |
| 403 | lemma updates_empty[simp]: "updates vs [] al = al" | |
| 404 | by (induct vs) auto | |
| 405 | ||
| 406 | lemma updates_Cons: "updates (k#ks) (v#vs) al = updates ks vs (update k v al)" | |
| 407 | by simp | |
| 408 | ||
| 409 | lemma updates_append1[simp]: "size ks < size vs \<Longrightarrow> | |
| 410 | updates (ks@[k]) vs al = update k (vs!size ks) (updates ks vs al)" | |
| 20503 | 411 | by (induct ks arbitrary: vs al) (auto split: list.splits) | 
| 19234 | 412 | |
| 413 | lemma updates_list_update_drop[simp]: | |
| 414 | "\<lbrakk>size ks \<le> i; i < size vs\<rbrakk> | |
| 415 | \<Longrightarrow> updates ks (vs[i:=v]) al = updates ks vs al" | |
| 20503 | 416 | by (induct ks arbitrary: al vs i) (auto split:list.splits nat.splits) | 
| 19234 | 417 | |
| 418 | lemma update_updates_conv_if: " | |
| 419 | map_of (updates xs ys (update x y al)) = | |
| 420 | map_of (if x \<in> set(take (length ys) xs) then updates xs ys al | |
| 421 | else (update x y (updates xs ys al)))" | |
| 422 | by (simp add: updates_conv' update_conv' map_upd_upds_conv_if) | |
| 423 | ||
| 424 | lemma updates_twist [simp]: | |
| 425 | "k \<notin> set ks \<Longrightarrow> | |
| 426 | map_of (updates ks vs (update k v al)) = map_of (update k v (updates ks vs al))" | |
| 427 | by (simp add: updates_conv' update_conv' map_upds_twist) | |
| 428 | ||
| 429 | lemma updates_apply_notin[simp]: | |
| 430 | "k \<notin> set ks ==> map_of (updates ks vs al) k = map_of al k" | |
| 431 | by (simp add: updates_conv) | |
| 432 | ||
| 433 | lemma updates_append_drop[simp]: | |
| 434 | "size xs = size ys \<Longrightarrow> updates (xs@zs) ys al = updates xs ys al" | |
| 20503 | 435 | by (induct xs arbitrary: ys al) (auto split: list.splits) | 
| 19234 | 436 | |
| 437 | lemma updates_append2_drop[simp]: | |
| 438 | "size xs = size ys \<Longrightarrow> updates xs (ys@zs) al = updates xs ys al" | |
| 20503 | 439 | by (induct xs arbitrary: ys al) (auto split: list.splits) | 
| 19234 | 440 | |
| 23373 | 441 | |
| 19333 | 442 | subsection {* @{const map_ran} *}
 | 
| 19234 | 443 | |
| 19333 | 444 | lemma map_ran_conv: "map_of (map_ran f al) k = option_map (f k) (map_of al k)" | 
| 19234 | 445 | by (induct al) auto | 
| 446 | ||
| 19333 | 447 | lemma dom_map_ran: "fst ` set (map_ran f al) = fst ` set al" | 
| 19234 | 448 | by (induct al) auto | 
| 449 | ||
| 19333 | 450 | lemma distinct_map_ran: "distinct (map fst al) \<Longrightarrow> distinct (map fst (map_ran f al))" | 
| 451 | by (induct al) (auto simp add: dom_map_ran) | |
| 19234 | 452 | |
| 23281 | 453 | lemma map_ran_filter: "map_ran f [p\<leftarrow>ps. fst p \<noteq> a] = [p\<leftarrow>map_ran f ps. fst p \<noteq> a]" | 
| 19234 | 454 | by (induct ps) auto | 
| 455 | ||
| 19333 | 456 | lemma clearjunk_map_ran: "clearjunk (map_ran f al) = map_ran f (clearjunk al)" | 
| 457 | by (induct al rule: clearjunk.induct) (auto simp add: delete_def map_ran_filter) | |
| 19234 | 458 | |
| 23373 | 459 | |
| 19234 | 460 | subsection {* @{const merge} *}
 | 
| 461 | ||
| 462 | lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \<union> fst ` set ys" | |
| 20503 | 463 | by (induct ys arbitrary: xs) (auto simp add: dom_update) | 
| 19234 | 464 | |
| 465 | lemma distinct_merge: | |
| 466 | assumes "distinct (map fst xs)" | |
| 467 | shows "distinct (map fst (merge xs ys))" | |
| 23373 | 468 | using assms | 
| 20503 | 469 | by (induct ys arbitrary: xs) (auto simp add: dom_merge distinct_update) | 
| 19234 | 470 | |
| 471 | lemma clearjunk_merge: | |
| 472 | "clearjunk (merge xs ys) = merge (clearjunk xs) ys" | |
| 473 | by (induct ys) (auto simp add: clearjunk_update) | |
| 474 | ||
| 475 | lemma merge_conv: "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k" | |
| 476 | proof (induct ys) | |
| 477 | case Nil thus ?case by simp | |
| 478 | next | |
| 479 | case (Cons y ys) | |
| 480 | show ?case | |
| 481 | proof (cases "k = fst y") | |
| 482 | case True | |
| 483 | from True show ?thesis | |
| 484 | by (simp add: update_conv) | |
| 485 | next | |
| 486 | case False | |
| 487 | from False show ?thesis | |
| 488 | by (auto simp add: update_conv Cons.hyps map_add_def) | |
| 489 | qed | |
| 490 | qed | |
| 491 | ||
| 492 | lemma merge_conv': "map_of (merge xs ys) = (map_of xs ++ map_of ys)" | |
| 493 | by (rule ext) (rule merge_conv) | |
| 494 | ||
| 495 | lemma merge_emty: "map_of (merge [] ys) = map_of ys" | |
| 496 | by (simp add: merge_conv') | |
| 497 | ||
| 498 | lemma merge_assoc[simp]: "map_of (merge m1 (merge m2 m3)) = | |
| 499 | map_of (merge (merge m1 m2) m3)" | |
| 500 | by (simp add: merge_conv') | |
| 501 | ||
| 502 | lemma merge_Some_iff: | |
| 503 | "(map_of (merge m n) k = Some x) = | |
| 504 | (map_of n k = Some x \<or> map_of n k = None \<and> map_of m k = Some x)" | |
| 505 | by (simp add: merge_conv' map_add_Some_iff) | |
| 506 | ||
| 507 | lemmas merge_SomeD = merge_Some_iff [THEN iffD1, standard] | |
| 508 | declare merge_SomeD [dest!] | |
| 509 | ||
| 510 | lemma merge_find_right[simp]: "map_of n k = Some v \<Longrightarrow> map_of (merge m n) k = Some v" | |
| 511 | by (simp add: merge_conv') | |
| 512 | ||
| 513 | lemma merge_None [iff]: | |
| 514 | "(map_of (merge m n) k = None) = (map_of n k = None \<and> map_of m k = None)" | |
| 515 | by (simp add: merge_conv') | |
| 516 | ||
| 517 | lemma merge_upd[simp]: | |
| 518 | "map_of (merge m (update k v n)) = map_of (update k v (merge m n))" | |
| 519 | by (simp add: update_conv' merge_conv') | |
| 520 | ||
| 521 | lemma merge_updatess[simp]: | |
| 522 | "map_of (merge m (updates xs ys n)) = map_of (updates xs ys (merge m n))" | |
| 523 | by (simp add: updates_conv' merge_conv') | |
| 524 | ||
| 525 | lemma merge_append: "map_of (xs@ys) = map_of (merge ys xs)" | |
| 526 | by (simp add: merge_conv') | |
| 527 | ||
| 23373 | 528 | |
| 19234 | 529 | subsection {* @{const compose} *}
 | 
| 530 | ||
| 531 | lemma compose_first_None [simp]: | |
| 532 | assumes "map_of xs k = None" | |
| 533 | shows "map_of (compose xs ys) k = None" | |
| 23373 | 534 | using assms by (induct xs ys rule: compose.induct) | 
| 22916 | 535 | (auto split: option.splits split_if_asm) | 
| 19234 | 536 | |
| 537 | lemma compose_conv: | |
| 538 | shows "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k" | |
| 22916 | 539 | proof (induct xs ys rule: compose.induct) | 
| 540 | case 1 then show ?case by simp | |
| 19234 | 541 | next | 
| 22916 | 542 | case (2 x xs ys) show ?case | 
| 19234 | 543 | proof (cases "map_of ys (snd x)") | 
| 22916 | 544 | case None with 2 | 
| 19234 | 545 | have hyp: "map_of (compose (delete (fst x) xs) ys) k = | 
| 546 | (map_of ys \<circ>\<^sub>m map_of (delete (fst x) xs)) k" | |
| 547 | by simp | |
| 548 | show ?thesis | |
| 549 | proof (cases "fst x = k") | |
| 550 | case True | |
| 551 | from True delete_notin_dom [of k xs] | |
| 552 | have "map_of (delete (fst x) xs) k = None" | |
| 553 | by (simp add: map_of_eq_None_iff) | |
| 554 | with hyp show ?thesis | |
| 555 | using True None | |
| 556 | by simp | |
| 557 | next | |
| 558 | case False | |
| 559 | from False have "map_of (delete (fst x) xs) k = map_of xs k" | |
| 560 | by simp | |
| 561 | with hyp show ?thesis | |
| 562 | using False None | |
| 563 | by (simp add: map_comp_def) | |
| 564 | qed | |
| 565 | next | |
| 566 | case (Some v) | |
| 22916 | 567 | with 2 | 
| 19234 | 568 | have "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k" | 
| 569 | by simp | |
| 570 | with Some show ?thesis | |
| 571 | by (auto simp add: map_comp_def) | |
| 572 | qed | |
| 573 | qed | |
| 574 | ||
| 575 | lemma compose_conv': | |
| 576 | shows "map_of (compose xs ys) = (map_of ys \<circ>\<^sub>m map_of xs)" | |
| 577 | by (rule ext) (rule compose_conv) | |
| 578 | ||
| 579 | lemma compose_first_Some [simp]: | |
| 580 | assumes "map_of xs k = Some v" | |
| 581 | shows "map_of (compose xs ys) k = map_of ys v" | |
| 23373 | 582 | using assms by (simp add: compose_conv) | 
| 19234 | 583 | |
| 584 | lemma dom_compose: "fst ` set (compose xs ys) \<subseteq> fst ` set xs" | |
| 22916 | 585 | proof (induct xs ys rule: compose.induct) | 
| 586 | case 1 thus ?case by simp | |
| 19234 | 587 | next | 
| 22916 | 588 | case (2 x xs ys) | 
| 19234 | 589 | show ?case | 
| 590 | proof (cases "map_of ys (snd x)") | |
| 591 | case None | |
| 22916 | 592 | with "2.hyps" | 
| 19234 | 593 | have "fst ` set (compose (delete (fst x) xs) ys) \<subseteq> fst ` set (delete (fst x) xs)" | 
| 594 | by simp | |
| 595 | also | |
| 596 | have "\<dots> \<subseteq> fst ` set xs" | |
| 597 | by (rule dom_delete_subset) | |
| 598 | finally show ?thesis | |
| 599 | using None | |
| 600 | by auto | |
| 601 | next | |
| 602 | case (Some v) | |
| 22916 | 603 | with "2.hyps" | 
| 19234 | 604 | have "fst ` set (compose xs ys) \<subseteq> fst ` set xs" | 
| 605 | by simp | |
| 606 | with Some show ?thesis | |
| 607 | by auto | |
| 608 | qed | |
| 609 | qed | |
| 610 | ||
| 611 | lemma distinct_compose: | |
| 612 | assumes "distinct (map fst xs)" | |
| 613 | shows "distinct (map fst (compose xs ys))" | |
| 23373 | 614 | using assms | 
| 22916 | 615 | proof (induct xs ys rule: compose.induct) | 
| 616 | case 1 thus ?case by simp | |
| 19234 | 617 | next | 
| 22916 | 618 | case (2 x xs ys) | 
| 19234 | 619 | show ?case | 
| 620 | proof (cases "map_of ys (snd x)") | |
| 621 | case None | |
| 22916 | 622 | with 2 show ?thesis by simp | 
| 19234 | 623 | next | 
| 624 | case (Some v) | |
| 22916 | 625 | with 2 dom_compose [of xs ys] show ?thesis | 
| 19234 | 626 | by (auto) | 
| 627 | qed | |
| 628 | qed | |
| 629 | ||
| 630 | lemma compose_delete_twist: "(compose (delete k xs) ys) = delete k (compose xs ys)" | |
| 22916 | 631 | proof (induct xs ys rule: compose.induct) | 
| 632 | case 1 thus ?case by simp | |
| 19234 | 633 | next | 
| 22916 | 634 | case (2 x xs ys) | 
| 19234 | 635 | show ?case | 
| 636 | proof (cases "map_of ys (snd x)") | |
| 637 | case None | |
| 22916 | 638 | with 2 have | 
| 19234 | 639 | hyp: "compose (delete k (delete (fst x) xs)) ys = | 
| 640 | delete k (compose (delete (fst x) xs) ys)" | |
| 641 | by simp | |
| 642 | show ?thesis | |
| 643 | proof (cases "fst x = k") | |
| 644 | case True | |
| 645 | with None hyp | |
| 646 | show ?thesis | |
| 647 | by (simp add: delete_idem) | |
| 648 | next | |
| 649 | case False | |
| 650 | from None False hyp | |
| 651 | show ?thesis | |
| 652 | by (simp add: delete_twist) | |
| 653 | qed | |
| 654 | next | |
| 655 | case (Some v) | |
| 22916 | 656 | with 2 have hyp: "compose (delete k xs) ys = delete k (compose xs ys)" by simp | 
| 19234 | 657 | with Some show ?thesis | 
| 658 | by simp | |
| 659 | qed | |
| 660 | qed | |
| 661 | ||
| 662 | lemma compose_clearjunk: "compose xs (clearjunk ys) = compose xs ys" | |
| 22916 | 663 | by (induct xs ys rule: compose.induct) | 
| 19234 | 664 | (auto simp add: map_of_clearjunk split: option.splits) | 
| 665 | ||
| 666 | lemma clearjunk_compose: "clearjunk (compose xs ys) = compose (clearjunk xs) ys" | |
| 667 | by (induct xs rule: clearjunk.induct) | |
| 668 | (auto split: option.splits simp add: clearjunk_delete delete_idem | |
| 669 | compose_delete_twist) | |
| 670 | ||
| 671 | lemma compose_empty [simp]: | |
| 672 | "compose xs [] = []" | |
| 22916 | 673 | by (induct xs) (auto simp add: compose_delete_twist) | 
| 19234 | 674 | |
| 675 | lemma compose_Some_iff: | |
| 676 | "(map_of (compose xs ys) k = Some v) = | |
| 677 | (\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = Some v)" | |
| 678 | by (simp add: compose_conv map_comp_Some_iff) | |
| 679 | ||
| 680 | lemma map_comp_None_iff: | |
| 681 | "(map_of (compose xs ys) k = None) = | |
| 682 | (map_of xs k = None \<or> (\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = None)) " | |
| 683 | by (simp add: compose_conv map_comp_None_iff) | |
| 684 | ||
| 685 | ||
| 686 | subsection {* @{const restrict} *}
 | |
| 687 | ||
| 22740 | 688 | lemma restrict_def: | 
| 689 | "restrict A = filter (\<lambda>p. fst p \<in> A)" | |
| 690 | proof | |
| 691 | fix xs | |
| 692 | show "restrict A xs = filter (\<lambda>p. fst p \<in> A) xs" | |
| 693 | by (induct xs) auto | |
| 694 | qed | |
| 19234 | 695 | |
| 696 | lemma distinct_restr: "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))" | |
| 697 | by (induct al) (auto simp add: restrict_def) | |
| 698 | ||
| 699 | lemma restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k" | |
| 700 | apply (induct al) | |
| 701 | apply (simp add: restrict_def) | |
| 702 | apply (cases "k\<in>A") | |
| 703 | apply (auto simp add: restrict_def) | |
| 704 | done | |
| 705 | ||
| 706 | lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)" | |
| 707 | by (rule ext) (rule restr_conv) | |
| 708 | ||
| 709 | lemma restr_empty [simp]: | |
| 710 |   "restrict {} al = []" 
 | |
| 711 | "restrict A [] = []" | |
| 712 | by (induct al) (auto simp add: restrict_def) | |
| 713 | ||
| 714 | lemma restr_in [simp]: "x \<in> A \<Longrightarrow> map_of (restrict A al) x = map_of al x" | |
| 715 | by (simp add: restr_conv') | |
| 716 | ||
| 717 | lemma restr_out [simp]: "x \<notin> A \<Longrightarrow> map_of (restrict A al) x = None" | |
| 718 | by (simp add: restr_conv') | |
| 719 | ||
| 720 | lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \<inter> A" | |
| 721 | by (induct al) (auto simp add: restrict_def) | |
| 722 | ||
| 723 | lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al"
 | |
| 724 | by (induct al) (auto simp add: restrict_def) | |
| 725 | ||
| 726 | lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\<inter>B) al" | |
| 727 | by (induct al) (auto simp add: restrict_def) | |
| 728 | ||
| 729 | lemma restr_update[simp]: | |
| 730 | "map_of (restrict D (update x y al)) = | |
| 731 |   map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
 | |
| 732 | by (simp add: restr_conv' update_conv') | |
| 733 | ||
| 734 | lemma restr_delete [simp]: | |
| 735 | "(delete x (restrict D al)) = | |
| 736 |     (if x\<in> D then restrict (D - {x}) al else restrict D al)"
 | |
| 737 | proof (induct al) | |
| 738 | case Nil thus ?case by simp | |
| 739 | next | |
| 740 | case (Cons a al) | |
| 741 | show ?case | |
| 742 | proof (cases "x \<in> D") | |
| 743 | case True | |
| 744 | note x_D = this | |
| 745 |     with Cons have hyp: "delete x (restrict D al) = restrict (D - {x}) al"
 | |
| 746 | by simp | |
| 747 | show ?thesis | |
| 748 | proof (cases "fst a = x") | |
| 749 | case True | |
| 750 | from Cons.hyps | |
| 751 | show ?thesis | |
| 752 | using x_D True | |
| 753 | by simp | |
| 754 | next | |
| 755 | case False | |
| 756 | note not_fst_a_x = this | |
| 757 | show ?thesis | |
| 758 | proof (cases "fst a \<in> D") | |
| 759 | case True | |
| 760 | with not_fst_a_x | |
| 761 | have "delete x (restrict D (a#al)) = a#(delete x (restrict D al))" | |
| 762 | by (cases a) (simp add: restrict_def) | |
| 763 | 	also from not_fst_a_x True hyp have "\<dots> = restrict (D - {x}) (a # al)"
 | |
| 764 | by (cases a) (simp add: restrict_def) | |
| 765 | finally show ?thesis | |
| 766 | using x_D by simp | |
| 767 | next | |
| 768 | case False | |
| 769 | hence "delete x (restrict D (a#al)) = delete x (restrict D al)" | |
| 770 | by (cases a) (simp add: restrict_def) | |
| 771 | moreover from False not_fst_a_x | |
| 772 | 	have "restrict (D - {x}) (a # al) = restrict (D - {x}) al"
 | |
| 773 | by (cases a) (simp add: restrict_def) | |
| 774 | ultimately | |
| 775 | show ?thesis using x_D hyp by simp | |
| 776 | qed | |
| 777 | qed | |
| 778 | next | |
| 779 | case False | |
| 780 | from False Cons show ?thesis | |
| 781 | by simp | |
| 782 | qed | |
| 783 | qed | |
| 784 | ||
| 785 | lemma update_restr: | |
| 786 |  "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
 | |
| 787 | by (simp add: update_conv' restr_conv') (rule fun_upd_restrict) | |
| 788 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 789 | lemma upate_restr_conv [simp]: | 
| 19234 | 790 | "x \<in> D \<Longrightarrow> | 
| 791 |  map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
 | |
| 792 | by (simp add: update_conv' restr_conv') | |
| 793 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20503diff
changeset | 794 | lemma restr_updates [simp]: " | 
| 19234 | 795 | \<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk> | 
| 796 | \<Longrightarrow> map_of (restrict D (updates xs ys al)) = | |
| 797 | map_of (updates xs ys (restrict (D - set xs) al))" | |
| 798 | by (simp add: updates_conv' restr_conv') | |
| 799 | ||
| 800 | lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)" | |
| 801 | by (induct ps) auto | |
| 802 | ||
| 803 | lemma clearjunk_restrict: | |
| 804 | "clearjunk (restrict A al) = restrict A (clearjunk al)" | |
| 805 | by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist) | |
| 806 | ||
| 807 | end |