# HG changeset patch # User nipkow # Date 1428920032 -7200 # Node ID 7c4a2e87e4d0596b1d206ac5338da5fa49b123d0 # Parent 7ff7fdfbb9a079134c21fbdff96d3b7824b554d4# Parent 177d740a0d48a243a6a66765bdedc0c3c6434305 merged diff -r 7ff7fdfbb9a0 -r 7c4a2e87e4d0 src/HOL/Library/AList.thy --- a/src/HOL/Library/AList.thy Mon Apr 13 00:59:17 2015 +0200 +++ b/src/HOL/Library/AList.thy Mon Apr 13 12:13:52 2015 +0200 @@ -215,6 +215,87 @@ by (simp add: delete_eq) +subsection {* @{text update_with_aux} and @{text delete_aux}*} + +qualified primrec update_with_aux :: "'val \ 'key \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" +where + "update_with_aux v k f [] = [(k, f v)]" +| "update_with_aux v k f (p # ps) = (if (fst p = k) then (k, f (snd p)) # ps else p # update_with_aux v k f ps)" + +text {* + The above @{term "delete"} traverses all the list even if it has found the key. + This one does not have to keep going because is assumes the invariant that keys are distinct. +*} +qualified fun delete_aux :: "'key \ ('key \ 'val) list \ ('key \ 'val) list" +where + "delete_aux k [] = []" +| "delete_aux k ((k', v) # xs) = (if k = k' then xs else (k', v) # delete_aux k xs)" + +lemma map_of_update_with_aux': + "map_of (update_with_aux v k f ps) k' = ((map_of ps)(k \ (case map_of ps k of None \ f v | Some v \ f v))) k'" +by(induct ps) auto + +lemma map_of_update_with_aux: + "map_of (update_with_aux v k f ps) = (map_of ps)(k \ (case map_of ps k of None \ f v | Some v \ f v))" +by(simp add: fun_eq_iff map_of_update_with_aux') + +lemma dom_update_with_aux: "fst ` set (update_with_aux v k f ps) = {k} \ fst ` set ps" + by (induct ps) auto + +lemma distinct_update_with_aux [simp]: + "distinct (map fst (update_with_aux v k f ps)) = distinct (map fst ps)" +by(induct ps)(auto simp add: dom_update_with_aux) + +lemma set_update_with_aux: + "distinct (map fst xs) + \ set (update_with_aux v k f xs) = (set xs - {k} \ UNIV \ {(k, f (case map_of xs k of None \ v | Some v \ v))})" +by(induct xs)(auto intro: rev_image_eqI) + +lemma set_delete_aux: "distinct (map fst xs) \ set (delete_aux k xs) = set xs - {k} \ UNIV" +apply(induct xs) +apply simp_all +apply clarsimp +apply(fastforce intro: rev_image_eqI) +done + +lemma dom_delete_aux: "distinct (map fst ps) \ fst ` set (delete_aux k ps) = fst ` set ps - {k}" +by(auto simp add: set_delete_aux) + +lemma distinct_delete_aux [simp]: + "distinct (map fst ps) \ distinct (map fst (delete_aux k ps))" +proof(induct ps) + case Nil thus ?case by simp +next + case (Cons a ps) + obtain k' v where a: "a = (k', v)" by(cases a) + show ?case + proof(cases "k' = k") + case True with Cons a show ?thesis by simp + next + case False + with Cons a have "k' \ fst ` set ps" "distinct (map fst ps)" by simp_all + with False a have "k' \ fst ` set (delete_aux k ps)" + by(auto dest!: dom_delete_aux[where k=k]) + with Cons a show ?thesis by simp + qed +qed + +lemma map_of_delete_aux': + "distinct (map fst xs) \ map_of (delete_aux k xs) = (map_of xs)(k := None)" + apply (induct xs) + apply (fastforce simp add: map_of_eq_None_iff fun_upd_twist) + apply (auto intro!: ext) + apply (simp add: map_of_eq_None_iff) + done + +lemma map_of_delete_aux: + "distinct (map fst xs) \ map_of (delete_aux k xs) k' = ((map_of xs)(k := None)) k'" +by(simp add: map_of_delete_aux') + +lemma delete_aux_eq_Nil_conv: "delete_aux k ts = [] \ ts = [] \ (\v. ts = [(k, v)])" +by(cases ts)(auto split: split_if_asm) + + subsection {* @{text restrict} *} qualified definition restrict :: "'key set \ ('key \ 'val) list \ ('key \ 'val) list"