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