| author | wenzelm | 
| Tue, 06 Jun 2023 15:29:43 +0200 | |
| changeset 78139 | bb85bda12b8e | 
| parent 73832 | 9db620f007fa | 
| child 82691 | b69e4da2604b | 
| permissions | -rw-r--r-- | 
| 63462 | 1 | (* Title: HOL/Library/AList_Mapping.thy | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 3 | *) | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 4 | |
| 60500 | 5 | section \<open>Implementation of mappings with Association Lists\<close> | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 6 | |
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 7 | theory AList_Mapping | 
| 63462 | 8 | imports AList Mapping | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 9 | begin | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 10 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 11 | lift_definition Mapping :: "('a \<times> 'b) list \<Rightarrow> ('a, 'b) mapping" is map_of .
 | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 12 | |
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 13 | code_datatype Mapping | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 14 | |
| 63462 | 15 | lemma lookup_Mapping [simp, code]: "Mapping.lookup (Mapping xs) = map_of xs" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 16 | by transfer rule | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 17 | |
| 63462 | 18 | lemma keys_Mapping [simp, code]: "Mapping.keys (Mapping xs) = set (map fst xs)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 19 | by transfer (simp add: dom_map_of_conv_image_fst) | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 20 | |
| 63462 | 21 | lemma empty_Mapping [code]: "Mapping.empty = Mapping []" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 22 | by transfer simp | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 23 | |
| 63462 | 24 | lemma is_empty_Mapping [code]: "Mapping.is_empty (Mapping xs) \<longleftrightarrow> List.null xs" | 
| 63649 | 25 | by (cases xs) (simp_all add: is_empty_def null_def) | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 26 | |
| 63462 | 27 | lemma update_Mapping [code]: "Mapping.update k v (Mapping xs) = Mapping (AList.update k v xs)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 28 | by transfer (simp add: update_conv') | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 29 | |
| 63462 | 30 | lemma delete_Mapping [code]: "Mapping.delete k (Mapping xs) = Mapping (AList.delete k xs)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 31 | by transfer (simp add: delete_conv') | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 32 | |
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 33 | lemma ordered_keys_Mapping [code]: | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 34 | "Mapping.ordered_keys (Mapping xs) = sort (remdups (map fst xs))" | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 35 | by (simp only: ordered_keys_def keys_Mapping sorted_list_of_set_sort_remdups) simp | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 36 | |
| 73832 | 37 | lemma entries_Mapping [code]: | 
| 38 | "Mapping.entries (Mapping xs) = set (AList.clearjunk xs)" | |
| 39 | by transfer (fact graph_map_of) | |
| 40 | ||
| 41 | lemma ordered_entries_Mapping [code]: | |
| 42 | "Mapping.ordered_entries (Mapping xs) = sort_key fst (AList.clearjunk xs)" | |
| 43 | proof - | |
| 44 | have distinct: "distinct (sort_key fst (AList.clearjunk xs))" | |
| 45 | using distinct_clearjunk distinct_map distinct_sort by blast | |
| 46 | note folding_Map_graph.idem_if_sorted_distinct[where ?m="map_of xs", OF _ sorted_sort_key distinct] | |
| 47 | then show ?thesis | |
| 48 | unfolding ordered_entries_def | |
| 49 | by (transfer fixing: xs) (auto simp: graph_map_of) | |
| 50 | qed | |
| 51 | ||
| 52 | lemma fold_Mapping [code]: | |
| 53 | "Mapping.fold f (Mapping xs) a = List.fold (case_prod f) (sort_key fst (AList.clearjunk xs)) a" | |
| 54 | by (simp add: Mapping.fold_def ordered_entries_Mapping) | |
| 55 | ||
| 63462 | 56 | lemma size_Mapping [code]: "Mapping.size (Mapping xs) = length (remdups (map fst xs))" | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 57 | by (simp add: size_def length_remdups_card_conv dom_map_of_conv_image_fst) | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 58 | |
| 63462 | 59 | lemma tabulate_Mapping [code]: "Mapping.tabulate ks f = Mapping (map (\<lambda>k. (k, f k)) ks)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 60 | by transfer (simp add: map_of_map_restrict) | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 61 | |
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 62 | lemma bulkload_Mapping [code]: | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 63 | "Mapping.bulkload vs = Mapping (map (\<lambda>n. (n, vs ! n)) [0..<length vs])" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
46238diff
changeset | 64 | by transfer (simp add: map_of_map_restrict fun_eq_iff) | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 65 | |
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 66 | lemma equal_Mapping [code]: | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 67 | "HOL.equal (Mapping xs) (Mapping ys) \<longleftrightarrow> | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 68 | (let ks = map fst xs; ls = map fst ys | 
| 63462 | 69 | in (\<forall>l\<in>set ls. l \<in> set ks) \<and> (\<forall>k\<in>set ks. k \<in> set ls \<and> map_of xs k = map_of ys k))" | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 70 | proof - | 
| 63462 | 71 | have *: "(a, b) \<in> set xs \<Longrightarrow> a \<in> fst ` set xs" for a b xs | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 72 | by (auto simp add: image_def intro!: bexI) | 
| 63476 | 73 | show ?thesis | 
| 74 | apply transfer | |
| 75 | apply (auto intro!: map_of_eqI) | |
| 76 | apply (auto dest!: map_of_eq_dom intro: *) | |
| 77 | done | |
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 78 | qed | 
| 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 79 | |
| 63194 | 80 | lemma map_values_Mapping [code]: | 
| 63462 | 81 | "Mapping.map_values f (Mapping xs) = Mapping (map (\<lambda>(x,y). (x, f x y)) xs)" | 
| 82 |   for f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b" and xs :: "('c \<times> 'a) list"
 | |
| 83 | apply transfer | |
| 84 | apply (rule ext) | |
| 85 | subgoal for f xs x by (induct xs) auto | |
| 86 | done | |
| 63194 | 87 | |
| 63462 | 88 | lemma combine_with_key_code [code]: | 
| 63195 | 89 | "Mapping.combine_with_key f (Mapping xs) (Mapping ys) = | 
| 63462 | 90 | Mapping.tabulate (remdups (map fst xs @ map fst ys)) | 
| 63195 | 91 | (\<lambda>x. the (combine_options (f x) (map_of xs x) (map_of ys x)))" | 
| 63462 | 92 | apply transfer | 
| 93 | apply (rule ext) | |
| 94 | apply (rule sym) | |
| 95 | subgoal for f xs ys x | |
| 96 | apply (cases "map_of xs x"; cases "map_of ys x"; simp) | |
| 63476 | 97 | apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff | 
| 98 | dest: map_of_SomeD split: option.splits)+ | |
| 63462 | 99 | done | 
| 100 | done | |
| 63195 | 101 | |
| 63462 | 102 | lemma combine_code [code]: | 
| 63194 | 103 | "Mapping.combine f (Mapping xs) (Mapping ys) = | 
| 63462 | 104 | Mapping.tabulate (remdups (map fst xs @ map fst ys)) | 
| 63194 | 105 | (\<lambda>x. the (combine_options f (map_of xs x) (map_of ys x)))" | 
| 63462 | 106 | apply transfer | 
| 107 | apply (rule ext) | |
| 108 | apply (rule sym) | |
| 109 | subgoal for f xs ys x | |
| 110 | apply (cases "map_of xs x"; cases "map_of ys x"; simp) | |
| 63476 | 111 | apply (force simp: map_of_eq_None_iff combine_options_def option.the_def o_def image_iff | 
| 112 | dest: map_of_SomeD split: option.splits)+ | |
| 63462 | 113 | done | 
| 114 | done | |
| 63194 | 115 | |
| 63462 | 116 | lemma map_of_filter_distinct: (* TODO: move? *) | 
| 63194 | 117 | assumes "distinct (map fst xs)" | 
| 63462 | 118 | shows "map_of (filter P xs) x = | 
| 119 | (case map_of xs x of | |
| 120 | None \<Rightarrow> None | |
| 121 | | Some y \<Rightarrow> if P (x,y) then Some y else None)" | |
| 63194 | 122 | using assms | 
| 123 | by (auto simp: map_of_eq_None_iff filter_map distinct_map_filter dest: map_of_SomeD | |
| 63462 | 124 | simp del: map_of_eq_Some_iff intro!: map_of_is_SomeI split: option.splits) | 
| 125 | ||
| 63194 | 126 | lemma filter_Mapping [code]: | 
| 127 | "Mapping.filter P (Mapping xs) = Mapping (filter (\<lambda>(k,v). P k v) (AList.clearjunk xs))" | |
| 63462 | 128 | apply transfer | 
| 129 | apply (rule ext) | |
| 130 | apply (subst map_of_filter_distinct) | |
| 63476 | 131 | apply (simp_all add: map_of_clearjunk split: option.split) | 
| 63462 | 132 | done | 
| 63194 | 133 | |
| 63462 | 134 | lemma [code nbe]: "HOL.equal (x :: ('a, 'b) mapping) x \<longleftrightarrow> True"
 | 
| 44897 
787983a08bfb
moving connection of association lists to Mappings into a separate theory
 bulwahn parents: diff
changeset | 135 | by (fact equal_refl) | 
| 59487 
adaa430fc0f7
default abstypes and default abstract equations make technical (no_code) annotation superfluous
 haftmann parents: 
58881diff
changeset | 136 | |
| 44913 | 137 | end |