# HG changeset patch # User bulwahn # Date 1326793542 -3600 # Node ID 9ace9e5b79beabc243abc00f17b584fa0098d148 # Parent 99c80c2f841a53a480d382d30c97a125a11ae69d renaming theory AList_Impl back to AList (reverting 1fec5b365f9b; AList with distinct key invariant is called DAList) diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Tue Jan 17 09:38:30 2012 +0100 +++ b/src/HOL/IsaMakefile Tue Jan 17 10:45:42 2012 +0100 @@ -436,7 +436,7 @@ $(OUT)/HOL-Library: $(OUT)/HOL Library/ROOT.ML \ $(SRC)/HOL/Tools/float_arith.ML $(SRC)/Tools/float.ML \ Library/Abstract_Rat.thy $(SRC)/Tools/Adhoc_Overloading.thy \ - Library/AList_Impl.thy Library/AList_Mapping.thy \ + Library/AList.thy Library/AList_Mapping.thy \ Library/BigO.thy Library/Binomial.thy \ Library/Bit.thy Library/Boolean_Algebra.thy Library/Cardinality.thy \ Library/Char_nat.thy Library/Code_Char.thy Library/Code_Char_chr.thy \ @@ -446,8 +446,8 @@ Tools/Predicate_Compile/code_prolog.ML Library/ContNotDenum.thy \ Library/Cset.thy Library/Cset_Monad.thy Library/Continuity.thy \ Library/Convex.thy Library/Countable.thy \ - Library/DAList.thy Library/Dlist.thy Library/Dlist_Cset.thy \ - Library/Efficient_Nat.thy Library/Eval_Witness.thy \ + Library/DAList.thy Library/Dlist.thy Library/Dlist_Cset.thy \ + Library/Efficient_Nat.thy Library/Eval_Witness.thy \ Library/Extended_Real.thy Library/Extended_Nat.thy Library/Float.thy \ Library/Formal_Power_Series.thy Library/Fraction_Field.thy \ Library/FrechetDeriv.thy Library/Cset.thy Library/FuncSet.thy \ diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/Library/AList.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/AList.thy Tue Jan 17 10:45:42 2012 +0100 @@ -0,0 +1,698 @@ +(* Title: HOL/Library/AList.thy + Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser, TU Muenchen +*) + +header {* Implementation of Association Lists *} + +theory AList +imports Main +begin + +text {* + The operations preserve distinctness of keys and + function @{term "clearjunk"} distributes over them. Since + @{term clearjunk} enforces distinctness of keys it can be used + to establish the invariant, e.g. for inductive proofs. +*} + +subsection {* @{text update} and @{text updates} *} + +primrec update :: "'key \ 'val \ ('key \ 'val) list \ ('key \ 'val) list" where + "update k v [] = [(k, v)]" + | "update k v (p#ps) = (if fst p = k then (k, v) # ps else p # update k v ps)" + +lemma update_conv': "map_of (update k v al) = (map_of al)(k\v)" + by (induct al) (auto simp add: fun_eq_iff) + +corollary update_conv: "map_of (update k v al) k' = ((map_of al)(k\v)) k'" + by (simp add: update_conv') + +lemma dom_update: "fst ` set (update k v al) = {k} \ fst ` set al" + by (induct al) auto + +lemma update_keys: + "map fst (update k v al) = + (if k \ set (map fst al) then map fst al else map fst al @ [k])" + by (induct al) simp_all + +lemma distinct_update: + assumes "distinct (map fst al)" + shows "distinct (map fst (update k v al))" + using assms by (simp add: update_keys) + +lemma update_filter: + "a\k \ update k v [q\ps . fst q \ a] = [q\update k v ps . fst q \ a]" + by (induct ps) auto + +lemma update_triv: "map_of al k = Some v \ update k v al = al" + by (induct al) auto + +lemma update_nonempty [simp]: "update k v al \ []" + by (induct al) auto + +lemma update_eqD: "update k v al = update k v' al' \ v = v'" +proof (induct al arbitrary: al') + case Nil thus ?case + by (cases al') (auto split: split_if_asm) +next + case Cons thus ?case + by (cases al') (auto split: split_if_asm) +qed + +lemma update_last [simp]: "update k v (update k v' al) = update k v al" + by (induct al) auto + +text {* Note that the lists are not necessarily the same: + @{term "update k v (update k' v' []) = [(k', v'), (k, v)]"} and + @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.*} +lemma update_swap: "k\k' + \ map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))" + by (simp add: update_conv' fun_eq_iff) + +lemma update_Some_unfold: + "map_of (update k v al) x = Some y \ + x = k \ v = y \ x \ k \ map_of al x = Some y" + by (simp add: update_conv' map_upd_Some_unfold) + +lemma image_update [simp]: + "x \ A \ map_of (update x y al) ` A = map_of al ` A" + by (simp add: update_conv') + +definition updates :: "'key list \ 'val list \ ('key \ 'val) list \ ('key \ 'val) list" where + "updates ks vs = fold (prod_case update) (zip ks vs)" + +lemma updates_simps [simp]: + "updates [] vs ps = ps" + "updates ks [] ps = ps" + "updates (k#ks) (v#vs) ps = updates ks vs (update k v ps)" + by (simp_all add: updates_def) + +lemma updates_key_simp [simp]: + "updates (k # ks) vs ps = + (case vs of [] \ ps | v # vs \ updates ks vs (update k v ps))" + by (cases vs) simp_all + +lemma updates_conv': "map_of (updates ks vs al) = (map_of al)(ks[\]vs)" +proof - + have "map_of \ fold (prod_case update) (zip ks vs) = + fold (\(k, v) f. f(k \ v)) (zip ks vs) \ map_of" + by (rule fold_commute) (auto simp add: fun_eq_iff update_conv') + then show ?thesis by (auto simp add: updates_def fun_eq_iff map_upds_fold_map_upd foldl_def split_def) +qed + +lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\]vs)) k" + by (simp add: updates_conv') + +lemma distinct_updates: + assumes "distinct (map fst al)" + shows "distinct (map fst (updates ks vs al))" +proof - + have "distinct (fold + (\(k, v) al. if k \ set al then al else al @ [k]) + (zip ks vs) (map fst al))" + by (rule fold_invariant [of "zip ks vs" "\_. True"]) (auto intro: assms) + moreover have "map fst \ fold (prod_case update) (zip ks vs) = + fold (\(k, v) al. if k \ set al then al else al @ [k]) (zip ks vs) \ map fst" + by (rule fold_commute) (simp add: update_keys split_def prod_case_beta comp_def) + ultimately show ?thesis by (simp add: updates_def fun_eq_iff) +qed + +lemma updates_append1[simp]: "size ks < size vs \ + updates (ks@[k]) vs al = update k (vs!size ks) (updates ks vs al)" + by (induct ks arbitrary: vs al) (auto split: list.splits) + +lemma updates_list_update_drop[simp]: + "\size ks \ i; i < size vs\ + \ updates ks (vs[i:=v]) al = updates ks vs al" + by (induct ks arbitrary: al vs i) (auto split:list.splits nat.splits) + +lemma update_updates_conv_if: " + map_of (updates xs ys (update x y al)) = + map_of (if x \ set(take (length ys) xs) then updates xs ys al + else (update x y (updates xs ys al)))" + by (simp add: updates_conv' update_conv' map_upd_upds_conv_if) + +lemma updates_twist [simp]: + "k \ set ks \ + map_of (updates ks vs (update k v al)) = map_of (update k v (updates ks vs al))" + by (simp add: updates_conv' update_conv' map_upds_twist) + +lemma updates_apply_notin[simp]: + "k \ set ks ==> map_of (updates ks vs al) k = map_of al k" + by (simp add: updates_conv) + +lemma updates_append_drop[simp]: + "size xs = size ys \ updates (xs@zs) ys al = updates xs ys al" + by (induct xs arbitrary: ys al) (auto split: list.splits) + +lemma updates_append2_drop[simp]: + "size xs = size ys \ updates xs (ys@zs) al = updates xs ys al" + by (induct xs arbitrary: ys al) (auto split: list.splits) + + +subsection {* @{text delete} *} + +definition delete :: "'key \ ('key \ 'val) list \ ('key \ 'val) list" where + delete_eq: "delete k = filter (\(k', _). k \ k')" + +lemma delete_simps [simp]: + "delete k [] = []" + "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)" + by (auto simp add: delete_eq) + +lemma delete_conv': "map_of (delete k al) = (map_of al)(k := None)" + by (induct al) (auto simp add: fun_eq_iff) + +corollary delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'" + by (simp add: delete_conv') + +lemma delete_keys: + "map fst (delete k al) = removeAll k (map fst al)" + by (simp add: delete_eq removeAll_filter_not_eq filter_map split_def comp_def) + +lemma distinct_delete: + assumes "distinct (map fst al)" + shows "distinct (map fst (delete k al))" + using assms by (simp add: delete_keys distinct_removeAll) + +lemma delete_id [simp]: "k \ fst ` set al \ delete k al = al" + by (auto simp add: image_iff delete_eq filter_id_conv) + +lemma delete_idem: "delete k (delete k al) = delete k al" + by (simp add: delete_eq) + +lemma map_of_delete [simp]: + "k' \ k \ map_of (delete k al) k' = map_of al k'" + by (simp add: delete_conv') + +lemma delete_notin_dom: "k \ fst ` set (delete k al)" + by (auto simp add: delete_eq) + +lemma dom_delete_subset: "fst ` set (delete k al) \ fst ` set al" + by (auto simp add: delete_eq) + +lemma delete_update_same: + "delete k (update k v al) = delete k al" + by (induct al) simp_all + +lemma delete_update: + "k \ l \ delete l (update k v al) = update k v (delete l al)" + by (induct al) simp_all + +lemma delete_twist: "delete x (delete y al) = delete y (delete x al)" + by (simp add: delete_eq conj_commute) + +lemma length_delete_le: "length (delete k al) \ length al" + by (simp add: delete_eq) + + +subsection {* @{text restrict} *} + +definition restrict :: "'key set \ ('key \ 'val) list \ ('key \ 'val) list" where + restrict_eq: "restrict A = filter (\(k, v). k \ A)" + +lemma restr_simps [simp]: + "restrict A [] = []" + "restrict A (p#ps) = (if fst p \ A then p # restrict A ps else restrict A ps)" + by (auto simp add: restrict_eq) + +lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)" +proof + fix k + show "map_of (restrict A al) k = ((map_of al)|` A) k" + by (induct al) (simp, cases "k \ A", auto) +qed + +corollary restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k" + by (simp add: restr_conv') + +lemma distinct_restr: + "distinct (map fst al) \ distinct (map fst (restrict A al))" + by (induct al) (auto simp add: restrict_eq) + +lemma restr_empty [simp]: + "restrict {} al = []" + "restrict A [] = []" + by (induct al) (auto simp add: restrict_eq) + +lemma restr_in [simp]: "x \ A \ map_of (restrict A al) x = map_of al x" + by (simp add: restr_conv') + +lemma restr_out [simp]: "x \ A \ map_of (restrict A al) x = None" + by (simp add: restr_conv') + +lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \ A" + by (induct al) (auto simp add: restrict_eq) + +lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al" + by (induct al) (auto simp add: restrict_eq) + +lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\B) al" + by (induct al) (auto simp add: restrict_eq) + +lemma restr_update[simp]: + "map_of (restrict D (update x y al)) = + map_of ((if x \ D then (update x y (restrict (D-{x}) al)) else restrict D al))" + by (simp add: restr_conv' update_conv') + +lemma restr_delete [simp]: + "(delete x (restrict D al)) = + (if x \ D then restrict (D - {x}) al else restrict D al)" +apply (simp add: delete_eq restrict_eq) +apply (auto simp add: split_def) +proof - + have "\y. y \ x \ x \ y" by auto + then show "[p \ al. fst p \ D \ x \ fst p] = [p \ al. fst p \ D \ fst p \ x]" + by simp + assume "x \ D" + then have "\y. y \ D \ y \ D \ x \ y" by auto + then show "[p \ al . fst p \ D \ x \ fst p] = [p \ al . fst p \ D]" + by simp +qed + +lemma update_restr: + "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))" + by (simp add: update_conv' restr_conv') (rule fun_upd_restrict) + +lemma update_restr_conv [simp]: + "x \ D \ + map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))" + by (simp add: update_conv' restr_conv') + +lemma restr_updates [simp]: " + \ length xs = length ys; set xs \ D \ + \ map_of (restrict D (updates xs ys al)) = + map_of (updates xs ys (restrict (D - set xs) al))" + by (simp add: updates_conv' restr_conv') + +lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)" + by (induct ps) auto + + +subsection {* @{text clearjunk} *} + +function clearjunk :: "('key \ 'val) list \ ('key \ 'val) list" where + "clearjunk [] = []" + | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)" + by pat_completeness auto +termination by (relation "measure length") + (simp_all add: less_Suc_eq_le length_delete_le) + +lemma map_of_clearjunk: + "map_of (clearjunk al) = map_of al" + by (induct al rule: clearjunk.induct) + (simp_all add: fun_eq_iff) + +lemma clearjunk_keys_set: + "set (map fst (clearjunk al)) = set (map fst al)" + by (induct al rule: clearjunk.induct) + (simp_all add: delete_keys) + +lemma dom_clearjunk: + "fst ` set (clearjunk al) = fst ` set al" + using clearjunk_keys_set by simp + +lemma distinct_clearjunk [simp]: + "distinct (map fst (clearjunk al))" + by (induct al rule: clearjunk.induct) + (simp_all del: set_map add: clearjunk_keys_set delete_keys) + +lemma ran_clearjunk: + "ran (map_of (clearjunk al)) = ran (map_of al)" + by (simp add: map_of_clearjunk) + +lemma ran_map_of: + "ran (map_of al) = snd ` set (clearjunk al)" +proof - + have "ran (map_of al) = ran (map_of (clearjunk al))" + by (simp add: ran_clearjunk) + also have "\ = snd ` set (clearjunk al)" + by (simp add: ran_distinct) + finally show ?thesis . +qed + +lemma clearjunk_update: + "clearjunk (update k v al) = update k v (clearjunk al)" + by (induct al rule: clearjunk.induct) + (simp_all add: delete_update) + +lemma clearjunk_updates: + "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)" +proof - + have "clearjunk \ fold (prod_case update) (zip ks vs) = + fold (prod_case update) (zip ks vs) \ clearjunk" + by (rule fold_commute) (simp add: clearjunk_update prod_case_beta o_def) + then show ?thesis by (simp add: updates_def fun_eq_iff) +qed + +lemma clearjunk_delete: + "clearjunk (delete x al) = delete x (clearjunk al)" + by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist) + +lemma clearjunk_restrict: + "clearjunk (restrict A al) = restrict A (clearjunk al)" + by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist) + +lemma distinct_clearjunk_id [simp]: + "distinct (map fst al) \ clearjunk al = al" + by (induct al rule: clearjunk.induct) auto + +lemma clearjunk_idem: + "clearjunk (clearjunk al) = clearjunk al" + by simp + +lemma length_clearjunk: + "length (clearjunk al) \ length al" +proof (induct al rule: clearjunk.induct [case_names Nil Cons]) + case Nil then show ?case by simp +next + case (Cons kv al) + moreover have "length (delete (fst kv) al) \ length al" by (fact length_delete_le) + ultimately have "length (clearjunk (delete (fst kv) al)) \ length al" by (rule order_trans) + then show ?case by simp +qed + +lemma delete_map: + assumes "\kv. fst (f kv) = fst kv" + shows "delete k (map f ps) = map f (delete k ps)" + by (simp add: delete_eq filter_map comp_def split_def assms) + +lemma clearjunk_map: + assumes "\kv. fst (f kv) = fst kv" + shows "clearjunk (map f ps) = map f (clearjunk ps)" + by (induct ps rule: clearjunk.induct [case_names Nil Cons]) + (simp_all add: clearjunk_delete delete_map assms) + + +subsection {* @{text map_ran} *} + +definition map_ran :: "('key \ 'val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" where + "map_ran f = map (\(k, v). (k, f k v))" + +lemma map_ran_simps [simp]: + "map_ran f [] = []" + "map_ran f ((k, v) # ps) = (k, f k v) # map_ran f ps" + by (simp_all add: map_ran_def) + +lemma dom_map_ran: + "fst ` set (map_ran f al) = fst ` set al" + by (simp add: map_ran_def image_image split_def) + +lemma map_ran_conv: + "map_of (map_ran f al) k = Option.map (f k) (map_of al k)" + by (induct al) auto + +lemma distinct_map_ran: + "distinct (map fst al) \ distinct (map fst (map_ran f al))" + by (simp add: map_ran_def split_def comp_def) + +lemma map_ran_filter: + "map_ran f [p\ps. fst p \ a] = [p\map_ran f ps. fst p \ a]" + by (simp add: map_ran_def filter_map split_def comp_def) + +lemma clearjunk_map_ran: + "clearjunk (map_ran f al) = map_ran f (clearjunk al)" + by (simp add: map_ran_def split_def clearjunk_map) + + +subsection {* @{text merge} *} + +definition merge :: "('key \ 'val) list \ ('key \ 'val) list \ ('key \ 'val) list" where + "merge qs ps = foldr (\(k, v). update k v) ps qs" + +lemma merge_simps [simp]: + "merge qs [] = qs" + "merge qs (p#ps) = update (fst p) (snd p) (merge qs ps)" + by (simp_all add: merge_def split_def) + +lemma merge_updates: + "merge qs ps = updates (rev (map fst ps)) (rev (map snd ps)) qs" + by (simp add: merge_def updates_def foldr_def zip_rev zip_map_fst_snd) + +lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \ fst ` set ys" + by (induct ys arbitrary: xs) (auto simp add: dom_update) + +lemma distinct_merge: + assumes "distinct (map fst xs)" + shows "distinct (map fst (merge xs ys))" +using assms by (simp add: merge_updates distinct_updates) + +lemma clearjunk_merge: + "clearjunk (merge xs ys) = merge (clearjunk xs) ys" + by (simp add: merge_updates clearjunk_updates) + +lemma merge_conv': + "map_of (merge xs ys) = map_of xs ++ map_of ys" +proof - + have "map_of \ fold (prod_case update) (rev ys) = + fold (\(k, v) m. m(k \ v)) (rev ys) \ map_of" + by (rule fold_commute) (simp add: update_conv' prod_case_beta split_def fun_eq_iff) + then show ?thesis + by (simp add: merge_def map_add_map_of_foldr foldr_def fun_eq_iff) +qed + +corollary merge_conv: + "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k" + by (simp add: merge_conv') + +lemma merge_empty: "map_of (merge [] ys) = map_of ys" + by (simp add: merge_conv') + +lemma merge_assoc[simp]: "map_of (merge m1 (merge m2 m3)) = + map_of (merge (merge m1 m2) m3)" + by (simp add: merge_conv') + +lemma merge_Some_iff: + "(map_of (merge m n) k = Some x) = + (map_of n k = Some x \ map_of n k = None \ map_of m k = Some x)" + by (simp add: merge_conv' map_add_Some_iff) + +lemmas merge_SomeD [dest!] = merge_Some_iff [THEN iffD1] + +lemma merge_find_right[simp]: "map_of n k = Some v \ map_of (merge m n) k = Some v" + by (simp add: merge_conv') + +lemma merge_None [iff]: + "(map_of (merge m n) k = None) = (map_of n k = None \ map_of m k = None)" + by (simp add: merge_conv') + +lemma merge_upd[simp]: + "map_of (merge m (update k v n)) = map_of (update k v (merge m n))" + by (simp add: update_conv' merge_conv') + +lemma merge_updatess[simp]: + "map_of (merge m (updates xs ys n)) = map_of (updates xs ys (merge m n))" + by (simp add: updates_conv' merge_conv') + +lemma merge_append: "map_of (xs@ys) = map_of (merge ys xs)" + by (simp add: merge_conv') + + +subsection {* @{text compose} *} + +function compose :: "('key \ 'a) list \ ('a \ 'b) list \ ('key \ 'b) list" where + "compose [] ys = []" + | "compose (x#xs) ys = (case map_of ys (snd x) + of None \ compose (delete (fst x) xs) ys + | Some v \ (fst x, v) # compose xs ys)" + by pat_completeness auto +termination by (relation "measure (length \ fst)") + (simp_all add: less_Suc_eq_le length_delete_le) + +lemma compose_first_None [simp]: + assumes "map_of xs k = None" + shows "map_of (compose xs ys) k = None" +using assms by (induct xs ys rule: compose.induct) + (auto split: option.splits split_if_asm) + +lemma compose_conv: + shows "map_of (compose xs ys) k = (map_of ys \\<^sub>m map_of xs) k" +proof (induct xs ys rule: compose.induct) + case 1 then show ?case by simp +next + case (2 x xs ys) show ?case + proof (cases "map_of ys (snd x)") + case None with 2 + have hyp: "map_of (compose (delete (fst x) xs) ys) k = + (map_of ys \\<^sub>m map_of (delete (fst x) xs)) k" + by simp + show ?thesis + proof (cases "fst x = k") + case True + from True delete_notin_dom [of k xs] + have "map_of (delete (fst x) xs) k = None" + by (simp add: map_of_eq_None_iff) + with hyp show ?thesis + using True None + by simp + next + case False + from False have "map_of (delete (fst x) xs) k = map_of xs k" + by simp + with hyp show ?thesis + using False None + by (simp add: map_comp_def) + qed + next + case (Some v) + with 2 + have "map_of (compose xs ys) k = (map_of ys \\<^sub>m map_of xs) k" + by simp + with Some show ?thesis + by (auto simp add: map_comp_def) + qed +qed + +lemma compose_conv': + shows "map_of (compose xs ys) = (map_of ys \\<^sub>m map_of xs)" + by (rule ext) (rule compose_conv) + +lemma compose_first_Some [simp]: + assumes "map_of xs k = Some v" + shows "map_of (compose xs ys) k = map_of ys v" +using assms by (simp add: compose_conv) + +lemma dom_compose: "fst ` set (compose xs ys) \ fst ` set xs" +proof (induct xs ys rule: compose.induct) + case 1 thus ?case by simp +next + case (2 x xs ys) + show ?case + proof (cases "map_of ys (snd x)") + case None + with "2.hyps" + have "fst ` set (compose (delete (fst x) xs) ys) \ fst ` set (delete (fst x) xs)" + by simp + also + have "\ \ fst ` set xs" + by (rule dom_delete_subset) + finally show ?thesis + using None + by auto + next + case (Some v) + with "2.hyps" + have "fst ` set (compose xs ys) \ fst ` set xs" + by simp + with Some show ?thesis + by auto + qed +qed + +lemma distinct_compose: + assumes "distinct (map fst xs)" + shows "distinct (map fst (compose xs ys))" +using assms +proof (induct xs ys rule: compose.induct) + case 1 thus ?case by simp +next + case (2 x xs ys) + show ?case + proof (cases "map_of ys (snd x)") + case None + with 2 show ?thesis by simp + next + case (Some v) + with 2 dom_compose [of xs ys] show ?thesis + by (auto) + qed +qed + +lemma compose_delete_twist: "(compose (delete k xs) ys) = delete k (compose xs ys)" +proof (induct xs ys rule: compose.induct) + case 1 thus ?case by simp +next + case (2 x xs ys) + show ?case + proof (cases "map_of ys (snd x)") + case None + with 2 have + hyp: "compose (delete k (delete (fst x) xs)) ys = + delete k (compose (delete (fst x) xs) ys)" + by simp + show ?thesis + proof (cases "fst x = k") + case True + with None hyp + show ?thesis + by (simp add: delete_idem) + next + case False + from None False hyp + show ?thesis + by (simp add: delete_twist) + qed + next + case (Some v) + with 2 have hyp: "compose (delete k xs) ys = delete k (compose xs ys)" by simp + with Some show ?thesis + by simp + qed +qed + +lemma compose_clearjunk: "compose xs (clearjunk ys) = compose xs ys" + by (induct xs ys rule: compose.induct) + (auto simp add: map_of_clearjunk split: option.splits) + +lemma clearjunk_compose: "clearjunk (compose xs ys) = compose (clearjunk xs) ys" + by (induct xs rule: clearjunk.induct) + (auto split: option.splits simp add: clearjunk_delete delete_idem + compose_delete_twist) + +lemma compose_empty [simp]: + "compose xs [] = []" + by (induct xs) (auto simp add: compose_delete_twist) + +lemma compose_Some_iff: + "(map_of (compose xs ys) k = Some v) = + (\k'. map_of xs k = Some k' \ map_of ys k' = Some v)" + by (simp add: compose_conv map_comp_Some_iff) + +lemma map_comp_None_iff: + "(map_of (compose xs ys) k = None) = + (map_of xs k = None \ (\k'. map_of xs k = Some k' \ map_of ys k' = None)) " + by (simp add: compose_conv map_comp_None_iff) + +subsection {* @{text map_entry} *} + +fun map_entry :: "'key \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" +where + "map_entry k f [] = []" +| "map_entry k f (p # ps) = (if fst p = k then (k, f (snd p)) # ps else p # map_entry k f ps)" + +lemma map_of_map_entry: + "map_of (map_entry k f xs) = (map_of xs)(k := case map_of xs k of None => None | Some v' => Some (f v'))" +by (induct xs) auto + +lemma dom_map_entry: + "fst ` set (map_entry k f xs) = fst ` set xs" +by (induct xs) auto + +lemma distinct_map_entry: + assumes "distinct (map fst xs)" + shows "distinct (map fst (map_entry k f xs))" +using assms by (induct xs) (auto simp add: dom_map_entry) + +subsection {* @{text map_default} *} + +fun map_default :: "'key \ 'val \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" +where + "map_default k v f [] = [(k, v)]" +| "map_default k v f (p # ps) = (if fst p = k then (k, f (snd p)) # ps else p # map_default k v f ps)" + +lemma map_of_map_default: + "map_of (map_default k v f xs) = (map_of xs)(k := case map_of xs k of None => Some v | Some v' => Some (f v'))" +by (induct xs) auto + +lemma dom_map_default: + "fst ` set (map_default k v f xs) = insert k (fst ` set xs)" +by (induct xs) auto + +lemma distinct_map_default: + assumes "distinct (map fst xs)" + shows "distinct (map fst (map_default k v f xs))" +using assms by (induct xs) (auto simp add: dom_map_default) + +hide_const (open) update updates delete restrict clearjunk merge compose map_entry + +end diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/Library/AList_Impl.thy --- a/src/HOL/Library/AList_Impl.thy Tue Jan 17 09:38:30 2012 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,698 +0,0 @@ -(* Title: HOL/Library/AList_Impl.thy - Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser, TU Muenchen -*) - -header {* Implementation of Association Lists *} - -theory AList_Impl -imports Main -begin - -text {* - The operations preserve distinctness of keys and - function @{term "clearjunk"} distributes over them. Since - @{term clearjunk} enforces distinctness of keys it can be used - to establish the invariant, e.g. for inductive proofs. -*} - -subsection {* @{text update} and @{text updates} *} - -primrec update :: "'key \ 'val \ ('key \ 'val) list \ ('key \ 'val) list" where - "update k v [] = [(k, v)]" - | "update k v (p#ps) = (if fst p = k then (k, v) # ps else p # update k v ps)" - -lemma update_conv': "map_of (update k v al) = (map_of al)(k\v)" - by (induct al) (auto simp add: fun_eq_iff) - -corollary update_conv: "map_of (update k v al) k' = ((map_of al)(k\v)) k'" - by (simp add: update_conv') - -lemma dom_update: "fst ` set (update k v al) = {k} \ fst ` set al" - by (induct al) auto - -lemma update_keys: - "map fst (update k v al) = - (if k \ set (map fst al) then map fst al else map fst al @ [k])" - by (induct al) simp_all - -lemma distinct_update: - assumes "distinct (map fst al)" - shows "distinct (map fst (update k v al))" - using assms by (simp add: update_keys) - -lemma update_filter: - "a\k \ update k v [q\ps . fst q \ a] = [q\update k v ps . fst q \ a]" - by (induct ps) auto - -lemma update_triv: "map_of al k = Some v \ update k v al = al" - by (induct al) auto - -lemma update_nonempty [simp]: "update k v al \ []" - by (induct al) auto - -lemma update_eqD: "update k v al = update k v' al' \ v = v'" -proof (induct al arbitrary: al') - case Nil thus ?case - by (cases al') (auto split: split_if_asm) -next - case Cons thus ?case - by (cases al') (auto split: split_if_asm) -qed - -lemma update_last [simp]: "update k v (update k v' al) = update k v al" - by (induct al) auto - -text {* Note that the lists are not necessarily the same: - @{term "update k v (update k' v' []) = [(k', v'), (k, v)]"} and - @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.*} -lemma update_swap: "k\k' - \ map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))" - by (simp add: update_conv' fun_eq_iff) - -lemma update_Some_unfold: - "map_of (update k v al) x = Some y \ - x = k \ v = y \ x \ k \ map_of al x = Some y" - by (simp add: update_conv' map_upd_Some_unfold) - -lemma image_update [simp]: - "x \ A \ map_of (update x y al) ` A = map_of al ` A" - by (simp add: update_conv') - -definition updates :: "'key list \ 'val list \ ('key \ 'val) list \ ('key \ 'val) list" where - "updates ks vs = fold (prod_case update) (zip ks vs)" - -lemma updates_simps [simp]: - "updates [] vs ps = ps" - "updates ks [] ps = ps" - "updates (k#ks) (v#vs) ps = updates ks vs (update k v ps)" - by (simp_all add: updates_def) - -lemma updates_key_simp [simp]: - "updates (k # ks) vs ps = - (case vs of [] \ ps | v # vs \ updates ks vs (update k v ps))" - by (cases vs) simp_all - -lemma updates_conv': "map_of (updates ks vs al) = (map_of al)(ks[\]vs)" -proof - - have "map_of \ fold (prod_case update) (zip ks vs) = - fold (\(k, v) f. f(k \ v)) (zip ks vs) \ map_of" - by (rule fold_commute) (auto simp add: fun_eq_iff update_conv') - then show ?thesis by (auto simp add: updates_def fun_eq_iff map_upds_fold_map_upd foldl_def split_def) -qed - -lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\]vs)) k" - by (simp add: updates_conv') - -lemma distinct_updates: - assumes "distinct (map fst al)" - shows "distinct (map fst (updates ks vs al))" -proof - - have "distinct (fold - (\(k, v) al. if k \ set al then al else al @ [k]) - (zip ks vs) (map fst al))" - by (rule fold_invariant [of "zip ks vs" "\_. True"]) (auto intro: assms) - moreover have "map fst \ fold (prod_case update) (zip ks vs) = - fold (\(k, v) al. if k \ set al then al else al @ [k]) (zip ks vs) \ map fst" - by (rule fold_commute) (simp add: update_keys split_def prod_case_beta comp_def) - ultimately show ?thesis by (simp add: updates_def fun_eq_iff) -qed - -lemma updates_append1[simp]: "size ks < size vs \ - updates (ks@[k]) vs al = update k (vs!size ks) (updates ks vs al)" - by (induct ks arbitrary: vs al) (auto split: list.splits) - -lemma updates_list_update_drop[simp]: - "\size ks \ i; i < size vs\ - \ updates ks (vs[i:=v]) al = updates ks vs al" - by (induct ks arbitrary: al vs i) (auto split:list.splits nat.splits) - -lemma update_updates_conv_if: " - map_of (updates xs ys (update x y al)) = - map_of (if x \ set(take (length ys) xs) then updates xs ys al - else (update x y (updates xs ys al)))" - by (simp add: updates_conv' update_conv' map_upd_upds_conv_if) - -lemma updates_twist [simp]: - "k \ set ks \ - map_of (updates ks vs (update k v al)) = map_of (update k v (updates ks vs al))" - by (simp add: updates_conv' update_conv' map_upds_twist) - -lemma updates_apply_notin[simp]: - "k \ set ks ==> map_of (updates ks vs al) k = map_of al k" - by (simp add: updates_conv) - -lemma updates_append_drop[simp]: - "size xs = size ys \ updates (xs@zs) ys al = updates xs ys al" - by (induct xs arbitrary: ys al) (auto split: list.splits) - -lemma updates_append2_drop[simp]: - "size xs = size ys \ updates xs (ys@zs) al = updates xs ys al" - by (induct xs arbitrary: ys al) (auto split: list.splits) - - -subsection {* @{text delete} *} - -definition delete :: "'key \ ('key \ 'val) list \ ('key \ 'val) list" where - delete_eq: "delete k = filter (\(k', _). k \ k')" - -lemma delete_simps [simp]: - "delete k [] = []" - "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)" - by (auto simp add: delete_eq) - -lemma delete_conv': "map_of (delete k al) = (map_of al)(k := None)" - by (induct al) (auto simp add: fun_eq_iff) - -corollary delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'" - by (simp add: delete_conv') - -lemma delete_keys: - "map fst (delete k al) = removeAll k (map fst al)" - by (simp add: delete_eq removeAll_filter_not_eq filter_map split_def comp_def) - -lemma distinct_delete: - assumes "distinct (map fst al)" - shows "distinct (map fst (delete k al))" - using assms by (simp add: delete_keys distinct_removeAll) - -lemma delete_id [simp]: "k \ fst ` set al \ delete k al = al" - by (auto simp add: image_iff delete_eq filter_id_conv) - -lemma delete_idem: "delete k (delete k al) = delete k al" - by (simp add: delete_eq) - -lemma map_of_delete [simp]: - "k' \ k \ map_of (delete k al) k' = map_of al k'" - by (simp add: delete_conv') - -lemma delete_notin_dom: "k \ fst ` set (delete k al)" - by (auto simp add: delete_eq) - -lemma dom_delete_subset: "fst ` set (delete k al) \ fst ` set al" - by (auto simp add: delete_eq) - -lemma delete_update_same: - "delete k (update k v al) = delete k al" - by (induct al) simp_all - -lemma delete_update: - "k \ l \ delete l (update k v al) = update k v (delete l al)" - by (induct al) simp_all - -lemma delete_twist: "delete x (delete y al) = delete y (delete x al)" - by (simp add: delete_eq conj_commute) - -lemma length_delete_le: "length (delete k al) \ length al" - by (simp add: delete_eq) - - -subsection {* @{text restrict} *} - -definition restrict :: "'key set \ ('key \ 'val) list \ ('key \ 'val) list" where - restrict_eq: "restrict A = filter (\(k, v). k \ A)" - -lemma restr_simps [simp]: - "restrict A [] = []" - "restrict A (p#ps) = (if fst p \ A then p # restrict A ps else restrict A ps)" - by (auto simp add: restrict_eq) - -lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)" -proof - fix k - show "map_of (restrict A al) k = ((map_of al)|` A) k" - by (induct al) (simp, cases "k \ A", auto) -qed - -corollary restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k" - by (simp add: restr_conv') - -lemma distinct_restr: - "distinct (map fst al) \ distinct (map fst (restrict A al))" - by (induct al) (auto simp add: restrict_eq) - -lemma restr_empty [simp]: - "restrict {} al = []" - "restrict A [] = []" - by (induct al) (auto simp add: restrict_eq) - -lemma restr_in [simp]: "x \ A \ map_of (restrict A al) x = map_of al x" - by (simp add: restr_conv') - -lemma restr_out [simp]: "x \ A \ map_of (restrict A al) x = None" - by (simp add: restr_conv') - -lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \ A" - by (induct al) (auto simp add: restrict_eq) - -lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al" - by (induct al) (auto simp add: restrict_eq) - -lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\B) al" - by (induct al) (auto simp add: restrict_eq) - -lemma restr_update[simp]: - "map_of (restrict D (update x y al)) = - map_of ((if x \ D then (update x y (restrict (D-{x}) al)) else restrict D al))" - by (simp add: restr_conv' update_conv') - -lemma restr_delete [simp]: - "(delete x (restrict D al)) = - (if x \ D then restrict (D - {x}) al else restrict D al)" -apply (simp add: delete_eq restrict_eq) -apply (auto simp add: split_def) -proof - - have "\y. y \ x \ x \ y" by auto - then show "[p \ al. fst p \ D \ x \ fst p] = [p \ al. fst p \ D \ fst p \ x]" - by simp - assume "x \ D" - then have "\y. y \ D \ y \ D \ x \ y" by auto - then show "[p \ al . fst p \ D \ x \ fst p] = [p \ al . fst p \ D]" - by simp -qed - -lemma update_restr: - "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))" - by (simp add: update_conv' restr_conv') (rule fun_upd_restrict) - -lemma update_restr_conv [simp]: - "x \ D \ - map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))" - by (simp add: update_conv' restr_conv') - -lemma restr_updates [simp]: " - \ length xs = length ys; set xs \ D \ - \ map_of (restrict D (updates xs ys al)) = - map_of (updates xs ys (restrict (D - set xs) al))" - by (simp add: updates_conv' restr_conv') - -lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)" - by (induct ps) auto - - -subsection {* @{text clearjunk} *} - -function clearjunk :: "('key \ 'val) list \ ('key \ 'val) list" where - "clearjunk [] = []" - | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)" - by pat_completeness auto -termination by (relation "measure length") - (simp_all add: less_Suc_eq_le length_delete_le) - -lemma map_of_clearjunk: - "map_of (clearjunk al) = map_of al" - by (induct al rule: clearjunk.induct) - (simp_all add: fun_eq_iff) - -lemma clearjunk_keys_set: - "set (map fst (clearjunk al)) = set (map fst al)" - by (induct al rule: clearjunk.induct) - (simp_all add: delete_keys) - -lemma dom_clearjunk: - "fst ` set (clearjunk al) = fst ` set al" - using clearjunk_keys_set by simp - -lemma distinct_clearjunk [simp]: - "distinct (map fst (clearjunk al))" - by (induct al rule: clearjunk.induct) - (simp_all del: set_map add: clearjunk_keys_set delete_keys) - -lemma ran_clearjunk: - "ran (map_of (clearjunk al)) = ran (map_of al)" - by (simp add: map_of_clearjunk) - -lemma ran_map_of: - "ran (map_of al) = snd ` set (clearjunk al)" -proof - - have "ran (map_of al) = ran (map_of (clearjunk al))" - by (simp add: ran_clearjunk) - also have "\ = snd ` set (clearjunk al)" - by (simp add: ran_distinct) - finally show ?thesis . -qed - -lemma clearjunk_update: - "clearjunk (update k v al) = update k v (clearjunk al)" - by (induct al rule: clearjunk.induct) - (simp_all add: delete_update) - -lemma clearjunk_updates: - "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)" -proof - - have "clearjunk \ fold (prod_case update) (zip ks vs) = - fold (prod_case update) (zip ks vs) \ clearjunk" - by (rule fold_commute) (simp add: clearjunk_update prod_case_beta o_def) - then show ?thesis by (simp add: updates_def fun_eq_iff) -qed - -lemma clearjunk_delete: - "clearjunk (delete x al) = delete x (clearjunk al)" - by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist) - -lemma clearjunk_restrict: - "clearjunk (restrict A al) = restrict A (clearjunk al)" - by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist) - -lemma distinct_clearjunk_id [simp]: - "distinct (map fst al) \ clearjunk al = al" - by (induct al rule: clearjunk.induct) auto - -lemma clearjunk_idem: - "clearjunk (clearjunk al) = clearjunk al" - by simp - -lemma length_clearjunk: - "length (clearjunk al) \ length al" -proof (induct al rule: clearjunk.induct [case_names Nil Cons]) - case Nil then show ?case by simp -next - case (Cons kv al) - moreover have "length (delete (fst kv) al) \ length al" by (fact length_delete_le) - ultimately have "length (clearjunk (delete (fst kv) al)) \ length al" by (rule order_trans) - then show ?case by simp -qed - -lemma delete_map: - assumes "\kv. fst (f kv) = fst kv" - shows "delete k (map f ps) = map f (delete k ps)" - by (simp add: delete_eq filter_map comp_def split_def assms) - -lemma clearjunk_map: - assumes "\kv. fst (f kv) = fst kv" - shows "clearjunk (map f ps) = map f (clearjunk ps)" - by (induct ps rule: clearjunk.induct [case_names Nil Cons]) - (simp_all add: clearjunk_delete delete_map assms) - - -subsection {* @{text map_ran} *} - -definition map_ran :: "('key \ 'val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" where - "map_ran f = map (\(k, v). (k, f k v))" - -lemma map_ran_simps [simp]: - "map_ran f [] = []" - "map_ran f ((k, v) # ps) = (k, f k v) # map_ran f ps" - by (simp_all add: map_ran_def) - -lemma dom_map_ran: - "fst ` set (map_ran f al) = fst ` set al" - by (simp add: map_ran_def image_image split_def) - -lemma map_ran_conv: - "map_of (map_ran f al) k = Option.map (f k) (map_of al k)" - by (induct al) auto - -lemma distinct_map_ran: - "distinct (map fst al) \ distinct (map fst (map_ran f al))" - by (simp add: map_ran_def split_def comp_def) - -lemma map_ran_filter: - "map_ran f [p\ps. fst p \ a] = [p\map_ran f ps. fst p \ a]" - by (simp add: map_ran_def filter_map split_def comp_def) - -lemma clearjunk_map_ran: - "clearjunk (map_ran f al) = map_ran f (clearjunk al)" - by (simp add: map_ran_def split_def clearjunk_map) - - -subsection {* @{text merge} *} - -definition merge :: "('key \ 'val) list \ ('key \ 'val) list \ ('key \ 'val) list" where - "merge qs ps = foldr (\(k, v). update k v) ps qs" - -lemma merge_simps [simp]: - "merge qs [] = qs" - "merge qs (p#ps) = update (fst p) (snd p) (merge qs ps)" - by (simp_all add: merge_def split_def) - -lemma merge_updates: - "merge qs ps = updates (rev (map fst ps)) (rev (map snd ps)) qs" - by (simp add: merge_def updates_def foldr_def zip_rev zip_map_fst_snd) - -lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \ fst ` set ys" - by (induct ys arbitrary: xs) (auto simp add: dom_update) - -lemma distinct_merge: - assumes "distinct (map fst xs)" - shows "distinct (map fst (merge xs ys))" -using assms by (simp add: merge_updates distinct_updates) - -lemma clearjunk_merge: - "clearjunk (merge xs ys) = merge (clearjunk xs) ys" - by (simp add: merge_updates clearjunk_updates) - -lemma merge_conv': - "map_of (merge xs ys) = map_of xs ++ map_of ys" -proof - - have "map_of \ fold (prod_case update) (rev ys) = - fold (\(k, v) m. m(k \ v)) (rev ys) \ map_of" - by (rule fold_commute) (simp add: update_conv' prod_case_beta split_def fun_eq_iff) - then show ?thesis - by (simp add: merge_def map_add_map_of_foldr foldr_def fun_eq_iff) -qed - -corollary merge_conv: - "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k" - by (simp add: merge_conv') - -lemma merge_empty: "map_of (merge [] ys) = map_of ys" - by (simp add: merge_conv') - -lemma merge_assoc[simp]: "map_of (merge m1 (merge m2 m3)) = - map_of (merge (merge m1 m2) m3)" - by (simp add: merge_conv') - -lemma merge_Some_iff: - "(map_of (merge m n) k = Some x) = - (map_of n k = Some x \ map_of n k = None \ map_of m k = Some x)" - by (simp add: merge_conv' map_add_Some_iff) - -lemmas merge_SomeD [dest!] = merge_Some_iff [THEN iffD1] - -lemma merge_find_right[simp]: "map_of n k = Some v \ map_of (merge m n) k = Some v" - by (simp add: merge_conv') - -lemma merge_None [iff]: - "(map_of (merge m n) k = None) = (map_of n k = None \ map_of m k = None)" - by (simp add: merge_conv') - -lemma merge_upd[simp]: - "map_of (merge m (update k v n)) = map_of (update k v (merge m n))" - by (simp add: update_conv' merge_conv') - -lemma merge_updatess[simp]: - "map_of (merge m (updates xs ys n)) = map_of (updates xs ys (merge m n))" - by (simp add: updates_conv' merge_conv') - -lemma merge_append: "map_of (xs@ys) = map_of (merge ys xs)" - by (simp add: merge_conv') - - -subsection {* @{text compose} *} - -function compose :: "('key \ 'a) list \ ('a \ 'b) list \ ('key \ 'b) list" where - "compose [] ys = []" - | "compose (x#xs) ys = (case map_of ys (snd x) - of None \ compose (delete (fst x) xs) ys - | Some v \ (fst x, v) # compose xs ys)" - by pat_completeness auto -termination by (relation "measure (length \ fst)") - (simp_all add: less_Suc_eq_le length_delete_le) - -lemma compose_first_None [simp]: - assumes "map_of xs k = None" - shows "map_of (compose xs ys) k = None" -using assms by (induct xs ys rule: compose.induct) - (auto split: option.splits split_if_asm) - -lemma compose_conv: - shows "map_of (compose xs ys) k = (map_of ys \\<^sub>m map_of xs) k" -proof (induct xs ys rule: compose.induct) - case 1 then show ?case by simp -next - case (2 x xs ys) show ?case - proof (cases "map_of ys (snd x)") - case None with 2 - have hyp: "map_of (compose (delete (fst x) xs) ys) k = - (map_of ys \\<^sub>m map_of (delete (fst x) xs)) k" - by simp - show ?thesis - proof (cases "fst x = k") - case True - from True delete_notin_dom [of k xs] - have "map_of (delete (fst x) xs) k = None" - by (simp add: map_of_eq_None_iff) - with hyp show ?thesis - using True None - by simp - next - case False - from False have "map_of (delete (fst x) xs) k = map_of xs k" - by simp - with hyp show ?thesis - using False None - by (simp add: map_comp_def) - qed - next - case (Some v) - with 2 - have "map_of (compose xs ys) k = (map_of ys \\<^sub>m map_of xs) k" - by simp - with Some show ?thesis - by (auto simp add: map_comp_def) - qed -qed - -lemma compose_conv': - shows "map_of (compose xs ys) = (map_of ys \\<^sub>m map_of xs)" - by (rule ext) (rule compose_conv) - -lemma compose_first_Some [simp]: - assumes "map_of xs k = Some v" - shows "map_of (compose xs ys) k = map_of ys v" -using assms by (simp add: compose_conv) - -lemma dom_compose: "fst ` set (compose xs ys) \ fst ` set xs" -proof (induct xs ys rule: compose.induct) - case 1 thus ?case by simp -next - case (2 x xs ys) - show ?case - proof (cases "map_of ys (snd x)") - case None - with "2.hyps" - have "fst ` set (compose (delete (fst x) xs) ys) \ fst ` set (delete (fst x) xs)" - by simp - also - have "\ \ fst ` set xs" - by (rule dom_delete_subset) - finally show ?thesis - using None - by auto - next - case (Some v) - with "2.hyps" - have "fst ` set (compose xs ys) \ fst ` set xs" - by simp - with Some show ?thesis - by auto - qed -qed - -lemma distinct_compose: - assumes "distinct (map fst xs)" - shows "distinct (map fst (compose xs ys))" -using assms -proof (induct xs ys rule: compose.induct) - case 1 thus ?case by simp -next - case (2 x xs ys) - show ?case - proof (cases "map_of ys (snd x)") - case None - with 2 show ?thesis by simp - next - case (Some v) - with 2 dom_compose [of xs ys] show ?thesis - by (auto) - qed -qed - -lemma compose_delete_twist: "(compose (delete k xs) ys) = delete k (compose xs ys)" -proof (induct xs ys rule: compose.induct) - case 1 thus ?case by simp -next - case (2 x xs ys) - show ?case - proof (cases "map_of ys (snd x)") - case None - with 2 have - hyp: "compose (delete k (delete (fst x) xs)) ys = - delete k (compose (delete (fst x) xs) ys)" - by simp - show ?thesis - proof (cases "fst x = k") - case True - with None hyp - show ?thesis - by (simp add: delete_idem) - next - case False - from None False hyp - show ?thesis - by (simp add: delete_twist) - qed - next - case (Some v) - with 2 have hyp: "compose (delete k xs) ys = delete k (compose xs ys)" by simp - with Some show ?thesis - by simp - qed -qed - -lemma compose_clearjunk: "compose xs (clearjunk ys) = compose xs ys" - by (induct xs ys rule: compose.induct) - (auto simp add: map_of_clearjunk split: option.splits) - -lemma clearjunk_compose: "clearjunk (compose xs ys) = compose (clearjunk xs) ys" - by (induct xs rule: clearjunk.induct) - (auto split: option.splits simp add: clearjunk_delete delete_idem - compose_delete_twist) - -lemma compose_empty [simp]: - "compose xs [] = []" - by (induct xs) (auto simp add: compose_delete_twist) - -lemma compose_Some_iff: - "(map_of (compose xs ys) k = Some v) = - (\k'. map_of xs k = Some k' \ map_of ys k' = Some v)" - by (simp add: compose_conv map_comp_Some_iff) - -lemma map_comp_None_iff: - "(map_of (compose xs ys) k = None) = - (map_of xs k = None \ (\k'. map_of xs k = Some k' \ map_of ys k' = None)) " - by (simp add: compose_conv map_comp_None_iff) - -subsection {* @{text map_entry} *} - -fun map_entry :: "'key \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" -where - "map_entry k f [] = []" -| "map_entry k f (p # ps) = (if fst p = k then (k, f (snd p)) # ps else p # map_entry k f ps)" - -lemma map_of_map_entry: - "map_of (map_entry k f xs) = (map_of xs)(k := case map_of xs k of None => None | Some v' => Some (f v'))" -by (induct xs) auto - -lemma dom_map_entry: - "fst ` set (map_entry k f xs) = fst ` set xs" -by (induct xs) auto - -lemma distinct_map_entry: - assumes "distinct (map fst xs)" - shows "distinct (map fst (map_entry k f xs))" -using assms by (induct xs) (auto simp add: dom_map_entry) - -subsection {* @{text map_default} *} - -fun map_default :: "'key \ 'val \ ('val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list" -where - "map_default k v f [] = [(k, v)]" -| "map_default k v f (p # ps) = (if fst p = k then (k, f (snd p)) # ps else p # map_default k v f ps)" - -lemma map_of_map_default: - "map_of (map_default k v f xs) = (map_of xs)(k := case map_of xs k of None => Some v | Some v' => Some (f v'))" -by (induct xs) auto - -lemma dom_map_default: - "fst ` set (map_default k v f xs) = insert k (fst ` set xs)" -by (induct xs) auto - -lemma distinct_map_default: - assumes "distinct (map fst xs)" - shows "distinct (map fst (map_default k v f xs))" -using assms by (induct xs) (auto simp add: dom_map_default) - -hide_const (open) update updates delete restrict clearjunk merge compose map_entry - -end diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/Library/AList_Mapping.thy --- a/src/HOL/Library/AList_Mapping.thy Tue Jan 17 09:38:30 2012 +0100 +++ b/src/HOL/Library/AList_Mapping.thy Tue Jan 17 10:45:42 2012 +0100 @@ -5,7 +5,7 @@ header {* Implementation of mappings with Association Lists *} theory AList_Mapping -imports AList_Impl Mapping +imports AList Mapping begin definition Mapping :: "('a \ 'b) list \ ('a, 'b) mapping" where @@ -30,11 +30,11 @@ by (cases xs) (simp_all add: is_empty_def null_def) lemma update_Mapping [code]: - "Mapping.update k v (Mapping xs) = Mapping (AList_Impl.update k v xs)" + "Mapping.update k v (Mapping xs) = Mapping (AList.update k v xs)" by (rule mapping_eqI) (simp add: update_conv') lemma delete_Mapping [code]: - "Mapping.delete k (Mapping xs) = Mapping (AList_Impl.delete k xs)" + "Mapping.delete k (Mapping xs) = Mapping (AList.delete k xs)" by (rule mapping_eqI) (simp add: delete_conv') lemma ordered_keys_Mapping [code]: diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/Library/DAList.thy --- a/src/HOL/Library/DAList.thy Tue Jan 17 09:38:30 2012 +0100 +++ b/src/HOL/Library/DAList.thy Tue Jan 17 10:45:42 2012 +0100 @@ -4,7 +4,7 @@ header {* Abstract type of association lists with unique keys *} theory DAList -imports AList_Impl +imports AList begin text {* This was based on some existing fragments in the AFP-Collection framework. *} @@ -36,34 +36,34 @@ where [code del]: "empty = Alist []" definition update :: "'key \ 'value \ ('key, 'value) alist \ ('key, 'value) alist" -where [code del]: "update k v xs = Alist (AList_Impl.update k v (impl_of xs))" +where [code del]: "update k v xs = Alist (AList.update k v (impl_of xs))" (* FIXME: we use an unoptimised delete operation. *) definition delete :: "'key \ ('key, 'value) alist \ ('key, 'value) alist" -where [code del]: "delete k xs = Alist (AList_Impl.delete k (impl_of xs))" +where [code del]: "delete k xs = Alist (AList.delete k (impl_of xs))" definition map_entry :: "'key \ ('value \ 'value) \ ('key, 'value) alist \ ('key, 'value) alist" -where [code del]: "map_entry k f xs = Alist (AList_Impl.map_entry k f (impl_of xs))" +where [code del]: "map_entry k f xs = Alist (AList.map_entry k f (impl_of xs))" definition filter :: "('key \ 'value \ bool) \ ('key, 'value) alist \ ('key, 'value) alist" where [code del]: "filter P xs = Alist (List.filter P (impl_of xs))" definition map_default :: "'key => 'value => ('value => 'value) => ('key, 'value) alist => ('key, 'value) alist" where - "map_default k v f xs = Alist (AList_Impl.map_default k v f (impl_of xs))" + "map_default k v f xs = Alist (AList.map_default k v f (impl_of xs))" lemma impl_of_empty [code abstract]: "impl_of empty = []" by (simp add: empty_def Alist_inverse) -lemma impl_of_update [code abstract]: "impl_of (update k v xs) = AList_Impl.update k v (impl_of xs)" +lemma impl_of_update [code abstract]: "impl_of (update k v xs) = AList.update k v (impl_of xs)" by (simp add: update_def Alist_inverse distinct_update) lemma impl_of_delete [code abstract]: - "impl_of (delete k al) = AList_Impl.delete k (impl_of al)" + "impl_of (delete k al) = AList.delete k (impl_of al)" unfolding delete_def by (simp add: Alist_inverse distinct_delete) lemma impl_of_map_entry [code abstract]: - "impl_of (map_entry k f xs) = AList_Impl.map_entry k f (impl_of xs)" + "impl_of (map_entry k f xs) = AList.map_entry k f (impl_of xs)" unfolding map_entry_def by (simp add: Alist_inverse distinct_map_entry) lemma distinct_map_fst_filter: @@ -75,7 +75,7 @@ unfolding filter_def by (simp add: Alist_inverse distinct_map_fst_filter) lemma impl_of_map_default [code abstract]: - "impl_of (map_default k v f xs) = AList_Impl.map_default k v f (impl_of xs)" + "impl_of (map_default k v f xs) = AList.map_default k v f (impl_of xs)" by (auto simp add: map_default_def Alist_inverse distinct_map_default) subsection {* Abstract operation properties *} @@ -176,4 +176,4 @@ hide_fact (open) lookup_def empty_def update_def delete_def map_entry_def filter_def map_default_def hide_const (open) impl_of lookup empty update delete map_entry filter map_default -end \ No newline at end of file +end diff -r 99c80c2f841a -r 9ace9e5b79be src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Tue Jan 17 09:38:30 2012 +0100 +++ b/src/HOL/Library/Multiset.thy Tue Jan 17 10:45:42 2012 +0100 @@ -1078,7 +1078,7 @@ qed auto definition - "subtract_entries_raw xs ys = foldr (%(k, v). AList_Impl.map_entry k (%v'. v' - v)) ys xs" + "subtract_entries_raw xs ys = foldr (%(k, v). AList.map_entry k (%v'. v' - v)) ys xs" lemma map_of_subtract_entries_raw: "distinct (map fst ys) ==> map_of (subtract_entries_raw xs ys) x = (case map_of xs x of None => None | Some v => (case map_of ys x of None => Some v | Some v' => Some (v - v')))"