| author | blanchet | 
| Wed, 16 Oct 2013 19:55:23 +0200 | |
| changeset 54119 | 2c13cb4a057d | 
| parent 53026 | e1a548c11845 | 
| child 54853 | a435932a9f12 | 
| permissions | -rw-r--r-- | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 1 | (* Title: HOL/Library/Mapping.thy | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 2 | Author: Florian Haftmann and Ondrej Kuncar | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 3 | *) | 
| 29708 | 4 | |
| 5 | header {* An abstract view on maps for code generation. *}
 | |
| 6 | ||
| 7 | theory Mapping | |
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 8 | imports Main | 
| 29708 | 9 | begin | 
| 10 | ||
| 51379 | 11 | subsection {* Parametricity transfer rules *}
 | 
| 12 | ||
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 13 | context | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 14 | begin | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 15 | interpretation lifting_syntax . | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 16 | |
| 51379 | 17 | lemma empty_transfer: "(A ===> option_rel B) Map.empty Map.empty" by transfer_prover | 
| 18 | ||
| 19 | lemma lookup_transfer: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)" by transfer_prover | |
| 20 | ||
| 21 | lemma update_transfer: | |
| 22 | assumes [transfer_rule]: "bi_unique A" | |
| 23 | shows "(A ===> B ===> (A ===> option_rel B) ===> A ===> option_rel B) | |
| 24 | (\<lambda>k v m. m(k \<mapsto> v)) (\<lambda>k v m. m(k \<mapsto> v))" | |
| 25 | by transfer_prover | |
| 26 | ||
| 27 | lemma delete_transfer: | |
| 28 | assumes [transfer_rule]: "bi_unique A" | |
| 29 | shows "(A ===> (A ===> option_rel B) ===> A ===> option_rel B) | |
| 30 | (\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))" | |
| 31 | by transfer_prover | |
| 32 | ||
| 33 | definition equal_None :: "'a option \<Rightarrow> bool" where "equal_None x \<equiv> x = None" | |
| 34 | ||
| 35 | lemma [transfer_rule]: "(option_rel A ===> op=) equal_None equal_None" | |
| 53026 
e1a548c11845
got rid of the dependency of Lifting_* on the function package; use the original rel constants for basic BNFs;
 traytel parents: 
53013diff
changeset | 36 | unfolding fun_rel_def option_rel_def equal_None_def by (auto split: option.split) | 
| 51379 | 37 | |
| 38 | lemma dom_transfer: | |
| 39 | assumes [transfer_rule]: "bi_total A" | |
| 40 | shows "((A ===> option_rel B) ===> set_rel A) dom dom" | |
| 41 | unfolding dom_def[abs_def] equal_None_def[symmetric] | |
| 42 | by transfer_prover | |
| 43 | ||
| 44 | lemma map_of_transfer [transfer_rule]: | |
| 45 | assumes [transfer_rule]: "bi_unique R1" | |
| 46 | shows "(list_all2 (prod_rel R1 R2) ===> R1 ===> option_rel R2) map_of map_of" | |
| 47 | unfolding map_of_def by transfer_prover | |
| 48 | ||
| 49 | lemma tabulate_transfer: | |
| 50 | assumes [transfer_rule]: "bi_unique A" | |
| 51 | shows "(list_all2 A ===> (A ===> B) ===> A ===> option_rel B) | |
| 52 | (\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks)))" | |
| 53 | by transfer_prover | |
| 54 | ||
| 55 | lemma bulkload_transfer: | |
| 56 | "(list_all2 A ===> op= ===> option_rel A) | |
| 57 | (\<lambda>xs k. if k < length xs then Some (xs ! k) else None) (\<lambda>xs k. if k < length xs then Some (xs ! k) else None)" | |
| 58 | unfolding fun_rel_def | |
| 59 | apply clarsimp | |
| 60 | apply (erule list_all2_induct) | |
| 61 | apply simp | |
| 62 | apply (case_tac xa) | |
| 63 | apply simp | |
| 64 | by (auto dest: list_all2_lengthD list_all2_nthD) | |
| 65 | ||
| 66 | lemma map_transfer: | |
| 67 | "((A ===> B) ===> (C ===> D) ===> (B ===> option_rel C) ===> A ===> option_rel D) | |
| 68 | (\<lambda>f g m. (Option.map g \<circ> m \<circ> f)) (\<lambda>f g m. (Option.map g \<circ> m \<circ> f))" | |
| 69 | by transfer_prover | |
| 70 | ||
| 71 | lemma map_entry_transfer: | |
| 72 | assumes [transfer_rule]: "bi_unique A" | |
| 73 | shows "(A ===> (B ===> B) ===> (A ===> option_rel B) ===> A ===> option_rel B) | |
| 74 | (\<lambda>k f m. (case m k of None \<Rightarrow> m | |
| 75 | | Some v \<Rightarrow> m (k \<mapsto> (f v)))) (\<lambda>k f m. (case m k of None \<Rightarrow> m | |
| 76 | | Some v \<Rightarrow> m (k \<mapsto> (f v))))" | |
| 77 | by transfer_prover | |
| 78 | ||
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 79 | end | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 80 | |
| 29708 | 81 | subsection {* Type definition and primitive operations *}
 | 
| 82 | ||
| 49834 | 83 | typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set"
 | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 84 | morphisms rep Mapping .. | 
| 37700 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 haftmann parents: 
37299diff
changeset | 85 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 86 | setup_lifting(no_code) type_definition_mapping | 
| 37700 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 haftmann parents: 
37299diff
changeset | 87 | |
| 51379 | 88 | lift_definition empty :: "('a, 'b) mapping" is Map.empty parametric empty_transfer .
 | 
| 37700 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 haftmann parents: 
37299diff
changeset | 89 | |
| 51379 | 90 | lift_definition lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<Rightarrow> 'b option" is "\<lambda>m k. m k" 
 | 
| 91 | parametric lookup_transfer . | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 92 | |
| 51379 | 93 | lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is "\<lambda>k v m. m(k \<mapsto> v)" 
 | 
| 94 | parametric update_transfer . | |
| 37700 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 haftmann parents: 
37299diff
changeset | 95 | |
| 51379 | 96 | lift_definition delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is "\<lambda>k m. m(k := None)" 
 | 
| 97 | parametric delete_transfer . | |
| 39380 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 haftmann parents: 
39302diff
changeset | 98 | |
| 51379 | 99 | lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set" is dom parametric dom_transfer .
 | 
| 29708 | 100 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 101 | lift_definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping" is
 | 
| 51379 | 102 | "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" parametric tabulate_transfer . | 
| 29708 | 103 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 104 | lift_definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping" is | 
| 51379 | 105 | "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" parametric bulkload_transfer . | 
| 29708 | 106 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 107 | lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping" is
 | 
| 51379 | 108 | "\<lambda>f g m. (Option.map g \<circ> m \<circ> f)" parametric map_transfer . | 
| 29708 | 109 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 110 | |
| 40605 | 111 | subsection {* Functorial structure *}
 | 
| 112 | ||
| 41505 
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
 haftmann parents: 
41372diff
changeset | 113 | enriched_type map: map | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 114 | by (transfer, auto simp add: fun_eq_iff Option.map.compositionality Option.map.id)+ | 
| 40605 | 115 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 116 | |
| 29708 | 117 | subsection {* Derived operations *}
 | 
| 118 | ||
| 35194 | 119 | definition ordered_keys :: "('a\<Colon>linorder, 'b) mapping \<Rightarrow> 'a list" where
 | 
| 37052 | 120 | "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])" | 
| 35194 | 121 | |
| 35157 | 122 | definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool" where
 | 
| 37052 | 123 |   "is_empty m \<longleftrightarrow> keys m = {}"
 | 
| 35157 | 124 | |
| 125 | definition size :: "('a, 'b) mapping \<Rightarrow> nat" where
 | |
| 37052 | 126 | "size m = (if finite (keys m) then card (keys m) else 0)" | 
| 35157 | 127 | |
| 128 | definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
 | |
| 37052 | 129 | "replace k v m = (if k \<in> keys m then update k v m else m)" | 
| 29814 | 130 | |
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 131 | definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
 | 
| 37052 | 132 | "default k v m = (if k \<in> keys m then m else update k v m)" | 
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 133 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 134 | lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is
 | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 135 | "\<lambda>k f m. (case m k of None \<Rightarrow> m | 
| 51379 | 136 | | Some v \<Rightarrow> m (k \<mapsto> (f v)))" parametric map_entry_transfer . | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 137 | |
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 138 | lemma map_entry_code [code]: "map_entry k f m = (case lookup m k of None \<Rightarrow> m | 
| 49975 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
49973diff
changeset | 139 | | Some v \<Rightarrow> update k (f v) m)" | 
| 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
49973diff
changeset | 140 | by transfer rule | 
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 141 | |
| 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 142 | definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" where
 | 
| 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 143 | "map_default k v f m = map_entry k f (default k v m)" | 
| 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 144 | |
| 51379 | 145 | lift_definition assoc_list_to_mapping :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping"
 | 
| 146 | is map_of parametric map_of_transfer . | |
| 147 | ||
| 148 | lemma assoc_list_to_mapping_code [code]: | |
| 149 | "assoc_list_to_mapping xs = foldr (\<lambda>(k, v) m. update k v m) xs empty" | |
| 150 | by transfer(simp add: map_add_map_of_foldr[symmetric]) | |
| 151 | ||
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 152 | instantiation mapping :: (type, type) equal | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 153 | begin | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 154 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 155 | definition | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 156 | "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)" | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 157 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 158 | instance proof | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 159 | qed (unfold equal_mapping_def, transfer, auto) | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 160 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 161 | end | 
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 162 | |
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 163 | context | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 164 | begin | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 165 | interpretation lifting_syntax . | 
| 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 166 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 167 | lemma [transfer_rule]: | 
| 51379 | 168 | assumes [transfer_rule]: "bi_total A" | 
| 169 | assumes [transfer_rule]: "bi_unique B" | |
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 170 | shows "(pcr_mapping A B ===> pcr_mapping A B ===> op=) HOL.eq HOL.equal" | 
| 51379 | 171 | by (unfold equal) transfer_prover | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 172 | |
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
51379diff
changeset | 173 | end | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 174 | |
| 29708 | 175 | subsection {* Properties *}
 | 
| 176 | ||
| 49973 | 177 | lemma lookup_update: "lookup (update k v m) k = Some v" | 
| 178 | by transfer simp | |
| 179 | ||
| 180 | lemma lookup_update_neq: "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'" | |
| 181 | by transfer simp | |
| 182 | ||
| 183 | lemma lookup_empty: "lookup empty k = None" | |
| 184 | by transfer simp | |
| 185 | ||
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 186 | lemma keys_is_none_rep [code_unfold]: | 
| 37052 | 187 | "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 188 | by transfer (auto simp add: is_none_def) | 
| 29708 | 189 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 190 | lemma tabulate_alt_def: | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 191 | "map_of (List.map (\<lambda>k. (k, f k)) ks) = (Some o f) |` set ks" | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 192 | by (induct ks) (auto simp add: tabulate_def restrict_map_def) | 
| 29826 | 193 | |
| 29708 | 194 | lemma update_update: | 
| 195 | "update k v (update k w m) = update k v m" | |
| 196 | "k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 197 | by (transfer, simp add: fun_upd_twist)+ | 
| 29708 | 198 | |
| 35157 | 199 | lemma update_delete [simp]: | 
| 200 | "update k v (delete k m) = update k v m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 201 | by transfer simp | 
| 29708 | 202 | |
| 203 | lemma delete_update: | |
| 204 | "delete k (update k v m) = delete k m" | |
| 205 | "k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 206 | by (transfer, simp add: fun_upd_twist)+ | 
| 29708 | 207 | |
| 35157 | 208 | lemma delete_empty [simp]: | 
| 209 | "delete k empty = empty" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 210 | by transfer simp | 
| 29708 | 211 | |
| 35157 | 212 | lemma replace_update: | 
| 37052 | 213 | "k \<notin> keys m \<Longrightarrow> replace k v m = m" | 
| 214 | "k \<in> keys m \<Longrightarrow> replace k v m = update k v m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 215 | by (transfer, auto simp add: replace_def fun_upd_twist)+ | 
| 29708 | 216 | |
| 217 | lemma size_empty [simp]: | |
| 218 | "size empty = 0" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 219 | unfolding size_def by transfer simp | 
| 29708 | 220 | |
| 221 | lemma size_update: | |
| 37052 | 222 | "finite (keys m) \<Longrightarrow> size (update k v m) = | 
| 223 | (if k \<in> keys m then size m else Suc (size m))" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 224 | unfolding size_def by transfer (auto simp add: insert_dom) | 
| 29708 | 225 | |
| 226 | lemma size_delete: | |
| 37052 | 227 | "size (delete k m) = (if k \<in> keys m then size m - 1 else size m)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 228 | unfolding size_def by transfer simp | 
| 29708 | 229 | |
| 37052 | 230 | lemma size_tabulate [simp]: | 
| 29708 | 231 | "size (tabulate ks f) = length (remdups ks)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 232 | unfolding size_def by transfer (auto simp add: tabulate_alt_def card_set comp_def) | 
| 29708 | 233 | |
| 29831 | 234 | lemma bulkload_tabulate: | 
| 29826 | 235 | "bulkload xs = tabulate [0..<length xs] (nth xs)" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 236 | by transfer (auto simp add: tabulate_alt_def) | 
| 29826 | 237 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 238 | lemma is_empty_empty [simp]: | 
| 37052 | 239 | "is_empty empty" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 240 | unfolding is_empty_def by transfer simp | 
| 37052 | 241 | |
| 242 | lemma is_empty_update [simp]: | |
| 243 | "\<not> is_empty (update k v m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 244 | unfolding is_empty_def by transfer simp | 
| 37052 | 245 | |
| 246 | lemma is_empty_delete: | |
| 247 |   "is_empty (delete k m) \<longleftrightarrow> is_empty m \<or> keys m = {k}"
 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 248 | unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv) | 
| 37052 | 249 | |
| 250 | lemma is_empty_replace [simp]: | |
| 251 | "is_empty (replace k v m) \<longleftrightarrow> is_empty m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 252 | unfolding is_empty_def replace_def by transfer auto | 
| 37052 | 253 | |
| 254 | lemma is_empty_default [simp]: | |
| 255 | "\<not> is_empty (default k v m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 256 | unfolding is_empty_def default_def by transfer auto | 
| 37052 | 257 | |
| 258 | lemma is_empty_map_entry [simp]: | |
| 259 | "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 260 | unfolding is_empty_def | 
| 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 261 | apply transfer by (case_tac "m k") auto | 
| 37052 | 262 | |
| 263 | lemma is_empty_map_default [simp]: | |
| 264 | "\<not> is_empty (map_default k v f m)" | |
| 265 | by (simp add: map_default_def) | |
| 266 | ||
| 267 | lemma keys_empty [simp]: | |
| 268 |   "keys empty = {}"
 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 269 | by transfer simp | 
| 37052 | 270 | |
| 271 | lemma keys_update [simp]: | |
| 272 | "keys (update k v m) = insert k (keys m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 273 | by transfer simp | 
| 37052 | 274 | |
| 275 | lemma keys_delete [simp]: | |
| 276 |   "keys (delete k m) = keys m - {k}"
 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 277 | by transfer simp | 
| 37052 | 278 | |
| 279 | lemma keys_replace [simp]: | |
| 280 | "keys (replace k v m) = keys m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 281 | unfolding replace_def by transfer (simp add: insert_absorb) | 
| 37052 | 282 | |
| 283 | lemma keys_default [simp]: | |
| 284 | "keys (default k v m) = insert k (keys m)" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 285 | unfolding default_def by transfer (simp add: insert_absorb) | 
| 37052 | 286 | |
| 287 | lemma keys_map_entry [simp]: | |
| 288 | "keys (map_entry k f m) = keys m" | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 289 | apply transfer by (case_tac "m k") auto | 
| 37052 | 290 | |
| 291 | lemma keys_map_default [simp]: | |
| 292 | "keys (map_default k v f m) = insert k (keys m)" | |
| 293 | by (simp add: map_default_def) | |
| 294 | ||
| 295 | lemma keys_tabulate [simp]: | |
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 296 | "keys (tabulate ks f) = set ks" | 
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 297 | by transfer (simp add: map_of_map_restrict o_def) | 
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 298 | |
| 37052 | 299 | lemma keys_bulkload [simp]: | 
| 37026 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 300 |   "keys (bulkload xs) = {0..<length xs}"
 | 
| 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 301 | by (simp add: keys_tabulate bulkload_tabulate) | 
| 
7e8979a155ae
operations default, map_entry, map_default; more lemmas
 haftmann parents: 
36176diff
changeset | 302 | |
| 37052 | 303 | lemma distinct_ordered_keys [simp]: | 
| 304 | "distinct (ordered_keys m)" | |
| 305 | by (simp add: ordered_keys_def) | |
| 306 | ||
| 307 | lemma ordered_keys_infinite [simp]: | |
| 308 | "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []" | |
| 309 | by (simp add: ordered_keys_def) | |
| 310 | ||
| 311 | lemma ordered_keys_empty [simp]: | |
| 312 | "ordered_keys empty = []" | |
| 313 | by (simp add: ordered_keys_def) | |
| 314 | ||
| 315 | lemma ordered_keys_update [simp]: | |
| 316 | "k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m" | |
| 317 | "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (update k v m) = insort k (ordered_keys m)" | |
| 318 | by (simp_all add: ordered_keys_def) (auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb) | |
| 319 | ||
| 320 | lemma ordered_keys_delete [simp]: | |
| 321 | "ordered_keys (delete k m) = remove1 k (ordered_keys m)" | |
| 322 | proof (cases "finite (keys m)") | |
| 323 | case False then show ?thesis by simp | |
| 324 | next | |
| 325 | case True note fin = True | |
| 326 | show ?thesis | |
| 327 | proof (cases "k \<in> keys m") | |
| 328 | case False with fin have "k \<notin> set (sorted_list_of_set (keys m))" by simp | |
| 329 | with False show ?thesis by (simp add: ordered_keys_def remove1_idem) | |
| 330 | next | |
| 331 | case True with fin show ?thesis by (simp add: ordered_keys_def sorted_list_of_set_remove) | |
| 332 | qed | |
| 333 | qed | |
| 334 | ||
| 335 | lemma ordered_keys_replace [simp]: | |
| 336 | "ordered_keys (replace k v m) = ordered_keys m" | |
| 337 | by (simp add: replace_def) | |
| 338 | ||
| 339 | lemma ordered_keys_default [simp]: | |
| 340 | "k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m" | |
| 341 | "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)" | |
| 342 | by (simp_all add: default_def) | |
| 343 | ||
| 344 | lemma ordered_keys_map_entry [simp]: | |
| 345 | "ordered_keys (map_entry k f m) = ordered_keys m" | |
| 346 | by (simp add: ordered_keys_def) | |
| 347 | ||
| 348 | lemma ordered_keys_map_default [simp]: | |
| 349 | "k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m" | |
| 350 | "finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)" | |
| 351 | by (simp_all add: map_default_def) | |
| 352 | ||
| 353 | lemma ordered_keys_tabulate [simp]: | |
| 354 | "ordered_keys (tabulate ks f) = sort (remdups ks)" | |
| 355 | by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups) | |
| 356 | ||
| 357 | lemma ordered_keys_bulkload [simp]: | |
| 358 | "ordered_keys (bulkload ks) = [0..<length ks]" | |
| 359 | by (simp add: ordered_keys_def) | |
| 36110 | 360 | |
| 31459 | 361 | |
| 37700 
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
 haftmann parents: 
37299diff
changeset | 362 | subsection {* Code generator setup *}
 | 
| 31459 | 363 | |
| 37701 
411717732710
explicit code_datatype declaration prevents multiple instantiations later on
 haftmann parents: 
37700diff
changeset | 364 | code_datatype empty update | 
| 
411717732710
explicit code_datatype declaration prevents multiple instantiations later on
 haftmann parents: 
37700diff
changeset | 365 | |
| 49929 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 kuncar parents: 
49834diff
changeset | 366 | hide_const (open) empty is_empty rep lookup update delete ordered_keys keys size | 
| 40605 | 367 | replace default map_entry map_default tabulate bulkload map | 
| 35157 | 368 | |
| 49975 
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
 huffman parents: 
49973diff
changeset | 369 | end | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
49975diff
changeset | 370 | |
| 51379 | 371 |