author | haftmann |
Fri, 14 Jun 2019 08:34:27 +0000 | |
changeset 70333 | 0f7edf0853df |
parent 68782 | 8ff34c1ad580 |
child 73832 | 9db620f007fa |
permissions | -rw-r--r-- |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
1 |
(* Title: HOL/Library/Mapping.thy |
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
2 |
Author: Florian Haftmann and Ondrej Kuncar |
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
3 |
*) |
29708 | 4 |
|
60500 | 5 |
section \<open>An abstract view on maps for code generation.\<close> |
29708 | 6 |
|
7 |
theory Mapping |
|
53013
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
kuncar
parents:
51379
diff
changeset
|
8 |
imports Main |
29708 | 9 |
begin |
10 |
||
60500 | 11 |
subsection \<open>Parametricity transfer rules\<close> |
51379 | 12 |
|
63462 | 13 |
lemma map_of_foldr: "map_of xs = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) xs Map.empty" (* FIXME move *) |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
14 |
using map_add_map_of_foldr [of Map.empty] by auto |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
15 |
|
63343 | 16 |
context includes lifting_syntax |
53013
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
kuncar
parents:
51379
diff
changeset
|
17 |
begin |
56528 | 18 |
|
63462 | 19 |
lemma empty_parametric: "(A ===> rel_option B) Map.empty Map.empty" |
56528 | 20 |
by transfer_prover |
51379 | 21 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
22 |
lemma lookup_parametric: "((A ===> B) ===> A ===> B) (\<lambda>m k. m k) (\<lambda>m k. m k)" |
56528 | 23 |
by transfer_prover |
51379 | 24 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
25 |
lemma update_parametric: |
51379 | 26 |
assumes [transfer_rule]: "bi_unique A" |
56528 | 27 |
shows "(A ===> B ===> (A ===> rel_option B) ===> A ===> rel_option B) |
28 |
(\<lambda>k v m. m(k \<mapsto> v)) (\<lambda>k v m. m(k \<mapsto> v))" |
|
29 |
by transfer_prover |
|
51379 | 30 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
31 |
lemma delete_parametric: |
51379 | 32 |
assumes [transfer_rule]: "bi_unique A" |
63462 | 33 |
shows "(A ===> (A ===> rel_option B) ===> A ===> rel_option B) |
56528 | 34 |
(\<lambda>k m. m(k := None)) (\<lambda>k m. m(k := None))" |
35 |
by transfer_prover |
|
51379 | 36 |
|
56528 | 37 |
lemma is_none_parametric [transfer_rule]: |
38 |
"(rel_option A ===> HOL.eq) Option.is_none Option.is_none" |
|
61068 | 39 |
by (auto simp add: Option.is_none_def rel_fun_def rel_option_iff split: option.split) |
51379 | 40 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
41 |
lemma dom_parametric: |
51379 | 42 |
assumes [transfer_rule]: "bi_total A" |
63462 | 43 |
shows "((A ===> rel_option B) ===> rel_set A) dom dom" |
61068 | 44 |
unfolding dom_def [abs_def] Option.is_none_def [symmetric] by transfer_prover |
51379 | 45 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
46 |
lemma map_of_parametric [transfer_rule]: |
51379 | 47 |
assumes [transfer_rule]: "bi_unique R1" |
55944 | 48 |
shows "(list_all2 (rel_prod R1 R2) ===> R1 ===> rel_option R2) map_of map_of" |
56528 | 49 |
unfolding map_of_def by transfer_prover |
51379 | 50 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
51 |
lemma map_entry_parametric [transfer_rule]: |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
52 |
assumes [transfer_rule]: "bi_unique A" |
63462 | 53 |
shows "(A ===> (B ===> B) ===> (A ===> rel_option B) ===> A ===> rel_option B) |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
54 |
(\<lambda>k f m. (case m k of None \<Rightarrow> m |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
55 |
| Some v \<Rightarrow> m (k \<mapsto> (f v)))) (\<lambda>k f m. (case m k of None \<Rightarrow> m |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
56 |
| Some v \<Rightarrow> m (k \<mapsto> (f v))))" |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
57 |
by transfer_prover |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
58 |
|
63462 | 59 |
lemma tabulate_parametric: |
51379 | 60 |
assumes [transfer_rule]: "bi_unique A" |
63462 | 61 |
shows "(list_all2 A ===> (A ===> B) ===> A ===> rel_option B) |
56528 | 62 |
(\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks))) (\<lambda>ks f. (map_of (map (\<lambda>k. (k, f k)) ks)))" |
63 |
by transfer_prover |
|
51379 | 64 |
|
63462 | 65 |
lemma bulkload_parametric: |
66 |
"(list_all2 A ===> HOL.eq ===> rel_option A) |
|
67 |
(\<lambda>xs k. if k < length xs then Some (xs ! k) else None) |
|
68 |
(\<lambda>xs k. if k < length xs then Some (xs ! k) else None)" |
|
56528 | 69 |
proof |
70 |
fix xs ys |
|
71 |
assume "list_all2 A xs ys" |
|
63462 | 72 |
then show |
73 |
"(HOL.eq ===> rel_option A) |
|
74 |
(\<lambda>k. if k < length xs then Some (xs ! k) else None) |
|
75 |
(\<lambda>k. if k < length ys then Some (ys ! k) else None)" |
|
56528 | 76 |
apply induct |
63476 | 77 |
apply auto |
56528 | 78 |
unfolding rel_fun_def |
63462 | 79 |
apply clarsimp |
80 |
apply (case_tac xa) |
|
63476 | 81 |
apply (auto dest: list_all2_lengthD list_all2_nthD) |
56528 | 82 |
done |
83 |
qed |
|
51379 | 84 |
|
63462 | 85 |
lemma map_parametric: |
86 |
"((A ===> B) ===> (C ===> D) ===> (B ===> rel_option C) ===> A ===> rel_option D) |
|
56528 | 87 |
(\<lambda>f g m. (map_option g \<circ> m \<circ> f)) (\<lambda>f g m. (map_option g \<circ> m \<circ> f))" |
88 |
by transfer_prover |
|
63462 | 89 |
|
90 |
lemma combine_with_key_parametric: |
|
91 |
"((A ===> B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===> |
|
92 |
(A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x)) |
|
93 |
(\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x))" |
|
63194 | 94 |
unfolding combine_options_def by transfer_prover |
63462 | 95 |
|
96 |
lemma combine_parametric: |
|
97 |
"((B ===> B ===> B) ===> (A ===> rel_option B) ===> (A ===> rel_option B) ===> |
|
98 |
(A ===> rel_option B)) (\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x)) |
|
99 |
(\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x))" |
|
63194 | 100 |
unfolding combine_options_def by transfer_prover |
51379 | 101 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
102 |
end |
51379 | 103 |
|
53013
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
kuncar
parents:
51379
diff
changeset
|
104 |
|
60500 | 105 |
subsection \<open>Type definition and primitive operations\<close> |
29708 | 106 |
|
49834 | 107 |
typedef ('a, 'b) mapping = "UNIV :: ('a \<rightharpoonup> 'b) set" |
63462 | 108 |
morphisms rep Mapping .. |
37700
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents:
37299
diff
changeset
|
109 |
|
59485 | 110 |
setup_lifting type_definition_mapping |
37700
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents:
37299
diff
changeset
|
111 |
|
56528 | 112 |
lift_definition empty :: "('a, 'b) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
113 |
is Map.empty parametric empty_parametric . |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
114 |
|
56528 | 115 |
lift_definition lookup :: "('a, 'b) mapping \<Rightarrow> 'a \<Rightarrow> 'b option" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
116 |
is "\<lambda>m k. m k" parametric lookup_parametric . |
56528 | 117 |
|
63194 | 118 |
definition "lookup_default d m k = (case Mapping.lookup m k of None \<Rightarrow> d | Some v \<Rightarrow> v)" |
119 |
||
56528 | 120 |
lift_definition update :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
121 |
is "\<lambda>k v m. m(k \<mapsto> v)" parametric update_parametric . |
37700
bd90378b8171
refrain from using datatype declaration -- opens chance for quickcheck later on
haftmann
parents:
37299
diff
changeset
|
122 |
|
56528 | 123 |
lift_definition delete :: "'a \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
124 |
is "\<lambda>k m. m(k := None)" parametric delete_parametric . |
39380
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
haftmann
parents:
39302
diff
changeset
|
125 |
|
63194 | 126 |
lift_definition filter :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
63462 | 127 |
is "\<lambda>P m k. case m k of None \<Rightarrow> None | Some v \<Rightarrow> if P k v then Some v else None" . |
63194 | 128 |
|
56528 | 129 |
lift_definition keys :: "('a, 'b) mapping \<Rightarrow> 'a set" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
130 |
is dom parametric dom_parametric . |
29708 | 131 |
|
56528 | 132 |
lift_definition tabulate :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
133 |
is "\<lambda>ks f. (map_of (List.map (\<lambda>k. (k, f k)) ks))" parametric tabulate_parametric . |
29708 | 134 |
|
56528 | 135 |
lift_definition bulkload :: "'a list \<Rightarrow> (nat, 'a) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
136 |
is "\<lambda>xs k. if k < length xs then Some (xs ! k) else None" parametric bulkload_parametric . |
29708 | 137 |
|
56528 | 138 |
lift_definition map :: "('c \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('c, 'd) mapping" |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
139 |
is "\<lambda>f g m. (map_option g \<circ> m \<circ> f)" parametric map_parametric . |
63462 | 140 |
|
63194 | 141 |
lift_definition map_values :: "('c \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> ('c, 'a) mapping \<Rightarrow> ('c, 'b) mapping" |
63462 | 142 |
is "\<lambda>f m x. map_option (f x) (m x)" . |
63194 | 143 |
|
63462 | 144 |
lift_definition combine_with_key :: |
63194 | 145 |
"('a \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping" |
146 |
is "\<lambda>f m1 m2 x. combine_options (f x) (m1 x) (m2 x)" parametric combine_with_key_parametric . |
|
147 |
||
63462 | 148 |
lift_definition combine :: |
63194 | 149 |
"('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping \<Rightarrow> ('a,'b) mapping" |
150 |
is "\<lambda>f m1 m2 x. combine_options f (m1 x) (m2 x)" parametric combine_parametric . |
|
151 |
||
63462 | 152 |
definition "All_mapping m P \<longleftrightarrow> |
153 |
(\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some y \<Rightarrow> P x y)" |
|
29708 | 154 |
|
59485 | 155 |
declare [[code drop: map]] |
156 |
||
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
157 |
|
60500 | 158 |
subsection \<open>Functorial structure\<close> |
40605 | 159 |
|
55467
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents:
55466
diff
changeset
|
160 |
functor map: map |
55466 | 161 |
by (transfer, auto simp add: fun_eq_iff option.map_comp option.map_id)+ |
40605 | 162 |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
163 |
|
60500 | 164 |
subsection \<open>Derived operations\<close> |
29708 | 165 |
|
61076 | 166 |
definition ordered_keys :: "('a::linorder, 'b) mapping \<Rightarrow> 'a list" |
63462 | 167 |
where "ordered_keys m = (if finite (keys m) then sorted_list_of_set (keys m) else [])" |
35194 | 168 |
|
56528 | 169 |
definition is_empty :: "('a, 'b) mapping \<Rightarrow> bool" |
63462 | 170 |
where "is_empty m \<longleftrightarrow> keys m = {}" |
35157 | 171 |
|
56528 | 172 |
definition size :: "('a, 'b) mapping \<Rightarrow> nat" |
63462 | 173 |
where "size m = (if finite (keys m) then card (keys m) else 0)" |
35157 | 174 |
|
56528 | 175 |
definition replace :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
63462 | 176 |
where "replace k v m = (if k \<in> keys m then update k v m else m)" |
29814 | 177 |
|
56528 | 178 |
definition default :: "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
63462 | 179 |
where "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:
36176
diff
changeset
|
180 |
|
60500 | 181 |
text \<open>Manual derivation of transfer rule is non-trivial\<close> |
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
182 |
|
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
183 |
lift_definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" is |
63462 | 184 |
"\<lambda>k f m. |
185 |
(case m k of |
|
186 |
None \<Rightarrow> m |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
187 |
| Some v \<Rightarrow> m (k \<mapsto> (f v)))" parametric map_entry_parametric . |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
188 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
189 |
lemma map_entry_code [code]: |
63462 | 190 |
"map_entry k f m = |
191 |
(case lookup m k of |
|
192 |
None \<Rightarrow> m |
|
49975
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
49973
diff
changeset
|
193 |
| Some v \<Rightarrow> update k (f v) m)" |
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
49973
diff
changeset
|
194 |
by transfer rule |
37026
7e8979a155ae
operations default, map_entry, map_default; more lemmas
haftmann
parents:
36176
diff
changeset
|
195 |
|
56528 | 196 |
definition map_default :: "'a \<Rightarrow> 'b \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a, 'b) mapping \<Rightarrow> ('a, 'b) mapping" |
63462 | 197 |
where "map_default k v f m = map_entry k f (default k v m)" |
37026
7e8979a155ae
operations default, map_entry, map_default; more lemmas
haftmann
parents:
36176
diff
changeset
|
198 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
199 |
definition of_alist :: "('k \<times> 'v) list \<Rightarrow> ('k, 'v) mapping" |
63462 | 200 |
where "of_alist xs = foldr (\<lambda>(k, v) m. update k v m) xs empty" |
51379 | 201 |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
202 |
instantiation mapping :: (type, type) equal |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
203 |
begin |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
204 |
|
63462 | 205 |
definition "HOL.equal m1 m2 \<longleftrightarrow> (\<forall>k. lookup m1 k = lookup m2 k)" |
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
206 |
|
60679 | 207 |
instance |
63462 | 208 |
apply standard |
209 |
unfolding equal_mapping_def |
|
210 |
apply transfer |
|
211 |
apply auto |
|
212 |
done |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
213 |
|
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
214 |
end |
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
215 |
|
63343 | 216 |
context includes lifting_syntax |
53013
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
kuncar
parents:
51379
diff
changeset
|
217 |
begin |
56528 | 218 |
|
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
219 |
lemma [transfer_rule]: |
51379 | 220 |
assumes [transfer_rule]: "bi_total A" |
63462 | 221 |
and [transfer_rule]: "bi_unique B" |
67399 | 222 |
shows "(pcr_mapping A B ===> pcr_mapping A B ===> (=)) HOL.eq HOL.equal" |
63462 | 223 |
unfolding equal by transfer_prover |
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
224 |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
225 |
lemma of_alist_transfer [transfer_rule]: |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
226 |
assumes [transfer_rule]: "bi_unique R1" |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
227 |
shows "(list_all2 (rel_prod R1 R2) ===> pcr_mapping R1 R2) map_of of_alist" |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
228 |
unfolding of_alist_def [abs_def] map_of_foldr [abs_def] by transfer_prover |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
229 |
|
53013
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
kuncar
parents:
51379
diff
changeset
|
230 |
end |
51161
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
haftmann
parents:
49975
diff
changeset
|
231 |
|
56528 | 232 |
|
60500 | 233 |
subsection \<open>Properties\<close> |
29708 | 234 |
|
63462 | 235 |
lemma mapping_eqI: "(\<And>x. lookup m x = lookup m' x) \<Longrightarrow> m = m'" |
63195 | 236 |
by transfer (simp add: fun_eq_iff) |
237 |
||
63462 | 238 |
lemma mapping_eqI': |
239 |
assumes "\<And>x. x \<in> Mapping.keys m \<Longrightarrow> Mapping.lookup_default d m x = Mapping.lookup_default d m' x" |
|
240 |
and "Mapping.keys m = Mapping.keys m'" |
|
241 |
shows "m = m'" |
|
63195 | 242 |
proof (intro mapping_eqI) |
63462 | 243 |
show "Mapping.lookup m x = Mapping.lookup m' x" for x |
63195 | 244 |
proof (cases "Mapping.lookup m x") |
245 |
case None |
|
63462 | 246 |
then have "x \<notin> Mapping.keys m" |
247 |
by transfer (simp add: dom_def) |
|
248 |
then have "x \<notin> Mapping.keys m'" |
|
249 |
by (simp add: assms) |
|
250 |
then have "Mapping.lookup m' x = None" |
|
251 |
by transfer (simp add: dom_def) |
|
252 |
with None show ?thesis |
|
253 |
by simp |
|
63195 | 254 |
next |
255 |
case (Some y) |
|
63462 | 256 |
then have A: "x \<in> Mapping.keys m" |
257 |
by transfer (simp add: dom_def) |
|
258 |
then have "x \<in> Mapping.keys m'" |
|
259 |
by (simp add: assms) |
|
260 |
then have "\<exists>y'. Mapping.lookup m' x = Some y'" |
|
261 |
by transfer (simp add: dom_def) |
|
262 |
with Some assms(1)[OF A] show ?thesis |
|
263 |
by (auto simp add: lookup_default_def) |
|
63195 | 264 |
qed |
265 |
qed |
|
266 |
||
63462 | 267 |
lemma lookup_update: "lookup (update k v m) k = Some v" |
49973 | 268 |
by transfer simp |
269 |
||
63462 | 270 |
lemma lookup_update_neq: "k \<noteq> k' \<Longrightarrow> lookup (update k v m) k' = lookup m k'" |
49973 | 271 |
by transfer simp |
272 |
||
63462 | 273 |
lemma lookup_update': "Mapping.lookup (update k v m) k' = (if k = k' then Some v else lookup m k')" |
63194 | 274 |
by (auto simp: lookup_update lookup_update_neq) |
275 |
||
63462 | 276 |
lemma lookup_empty: "lookup empty k = None" |
49973 | 277 |
by transfer simp |
278 |
||
63194 | 279 |
lemma lookup_filter: |
63462 | 280 |
"lookup (filter P m) k = |
281 |
(case lookup m k of |
|
282 |
None \<Rightarrow> None |
|
283 |
| Some v \<Rightarrow> if P k v then Some v else None)" |
|
63194 | 284 |
by transfer simp_all |
285 |
||
63462 | 286 |
lemma lookup_map_values: "lookup (map_values f m) k = map_option (f k) (lookup m k)" |
63194 | 287 |
by transfer simp_all |
288 |
||
289 |
lemma lookup_default_empty: "lookup_default d empty k = d" |
|
290 |
by (simp add: lookup_default_def lookup_empty) |
|
291 |
||
63462 | 292 |
lemma lookup_default_update: "lookup_default d (update k v m) k = v" |
63194 | 293 |
by (simp add: lookup_default_def lookup_update) |
294 |
||
295 |
lemma lookup_default_update_neq: |
|
63462 | 296 |
"k \<noteq> k' \<Longrightarrow> lookup_default d (update k v m) k' = lookup_default d m k'" |
63194 | 297 |
by (simp add: lookup_default_def lookup_update_neq) |
298 |
||
63462 | 299 |
lemma lookup_default_update': |
63194 | 300 |
"lookup_default d (update k v m) k' = (if k = k' then v else lookup_default d m k')" |
301 |
by (auto simp: lookup_default_update lookup_default_update_neq) |
|
302 |
||
303 |
lemma lookup_default_filter: |
|
63462 | 304 |
"lookup_default d (filter P m) k = |
63194 | 305 |
(if P k (lookup_default d m k) then lookup_default d m k else d)" |
306 |
by (simp add: lookup_default_def lookup_filter split: option.splits) |
|
307 |
||
308 |
lemma lookup_default_map_values: |
|
309 |
"lookup_default (f k d) (map_values f m) k = f k (lookup_default d m k)" |
|
63462 | 310 |
by (simp add: lookup_default_def lookup_map_values split: option.splits) |
63194 | 311 |
|
312 |
lemma lookup_combine_with_key: |
|
63462 | 313 |
"Mapping.lookup (combine_with_key f m1 m2) x = |
314 |
combine_options (f x) (Mapping.lookup m1 x) (Mapping.lookup m2 x)" |
|
63194 | 315 |
by transfer (auto split: option.splits) |
63462 | 316 |
|
63194 | 317 |
lemma combine_altdef: "combine f m1 m2 = combine_with_key (\<lambda>_. f) m1 m2" |
318 |
by transfer' (rule refl) |
|
319 |
||
320 |
lemma lookup_combine: |
|
63462 | 321 |
"Mapping.lookup (combine f m1 m2) x = |
63194 | 322 |
combine_options f (Mapping.lookup m1 x) (Mapping.lookup m2 x)" |
323 |
by transfer (auto split: option.splits) |
|
63462 | 324 |
|
325 |
lemma lookup_default_neutral_combine_with_key: |
|
63194 | 326 |
assumes "\<And>x. f k d x = x" "\<And>x. f k x d = x" |
63462 | 327 |
shows "Mapping.lookup_default d (combine_with_key f m1 m2) k = |
328 |
f k (Mapping.lookup_default d m1 k) (Mapping.lookup_default d m2 k)" |
|
63194 | 329 |
by (auto simp: lookup_default_def lookup_combine_with_key assms split: option.splits) |
63462 | 330 |
|
331 |
lemma lookup_default_neutral_combine: |
|
63194 | 332 |
assumes "\<And>x. f d x = x" "\<And>x. f x d = x" |
63462 | 333 |
shows "Mapping.lookup_default d (combine f m1 m2) x = |
334 |
f (Mapping.lookup_default d m1 x) (Mapping.lookup_default d m2 x)" |
|
63194 | 335 |
by (auto simp: lookup_default_def lookup_combine assms split: option.splits) |
336 |
||
63462 | 337 |
lemma lookup_map_entry: "lookup (map_entry x f m) x = map_option f (lookup m x)" |
63195 | 338 |
by transfer (auto split: option.splits) |
339 |
||
63462 | 340 |
lemma lookup_map_entry_neq: "x \<noteq> y \<Longrightarrow> lookup (map_entry x f m) y = lookup m y" |
63195 | 341 |
by transfer (auto split: option.splits) |
342 |
||
343 |
lemma lookup_map_entry': |
|
63462 | 344 |
"lookup (map_entry x f m) y = |
63195 | 345 |
(if x = y then map_option f (lookup m y) else lookup m y)" |
346 |
by transfer (auto split: option.splits) |
|
347 |
||
63462 | 348 |
lemma lookup_default: "lookup (default x d m) x = Some (lookup_default d m x)" |
349 |
unfolding lookup_default_def default_def |
|
350 |
by transfer (auto split: option.splits) |
|
351 |
||
352 |
lemma lookup_default_neq: "x \<noteq> y \<Longrightarrow> lookup (default x d m) y = lookup m y" |
|
353 |
unfolding lookup_default_def default_def |
|
354 |
by transfer (auto split: option.splits) |
|
63195 | 355 |
|
356 |
lemma lookup_default': |
|
63462 | 357 |
"lookup (default x d m) y = |
358 |
(if x = y then Some (lookup_default d m x) else lookup m y)" |
|
63195 | 359 |
unfolding lookup_default_def default_def |
360 |
by transfer (auto split: option.splits) |
|
361 |
||
63462 | 362 |
lemma lookup_map_default: "lookup (map_default x d f m) x = Some (f (lookup_default d m x))" |
363 |
unfolding lookup_default_def default_def |
|
364 |
by (simp add: map_default_def lookup_map_entry lookup_default lookup_default_def) |
|
365 |
||
366 |
lemma lookup_map_default_neq: "x \<noteq> y \<Longrightarrow> lookup (map_default x d f m) y = lookup m y" |
|
367 |
unfolding lookup_default_def default_def |
|
368 |
by (simp add: map_default_def lookup_map_entry_neq lookup_default_neq) |
|
63195 | 369 |
|
370 |
lemma lookup_map_default': |
|
63462 | 371 |
"lookup (map_default x d f m) y = |
372 |
(if x = y then Some (f (lookup_default d m x)) else lookup m y)" |
|
373 |
unfolding lookup_default_def default_def |
|
374 |
by (simp add: map_default_def lookup_map_entry' lookup_default' lookup_default_def) |
|
63195 | 375 |
|
63462 | 376 |
lemma lookup_tabulate: |
63194 | 377 |
assumes "distinct xs" |
63462 | 378 |
shows "Mapping.lookup (Mapping.tabulate xs f) x = (if x \<in> set xs then Some (f x) else None)" |
63194 | 379 |
using assms by transfer (auto simp: map_of_eq_None_iff o_def dest!: map_of_SomeD) |
380 |
||
381 |
lemma lookup_of_alist: "Mapping.lookup (Mapping.of_alist xs) k = map_of xs k" |
|
382 |
by transfer simp_all |
|
383 |
||
63462 | 384 |
lemma keys_is_none_rep [code_unfold]: "k \<in> keys m \<longleftrightarrow> \<not> (Option.is_none (lookup m k))" |
61068 | 385 |
by transfer (auto simp add: Option.is_none_def) |
29708 | 386 |
|
387 |
lemma update_update: |
|
388 |
"update k v (update k w m) = update k v m" |
|
389 |
"k \<noteq> l \<Longrightarrow> update k v (update l w m) = update l w (update k v m)" |
|
63462 | 390 |
by (transfer; simp add: fun_upd_twist)+ |
29708 | 391 |
|
63462 | 392 |
lemma update_delete [simp]: "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:
49834
diff
changeset
|
393 |
by transfer simp |
29708 | 394 |
|
395 |
lemma delete_update: |
|
396 |
"delete k (update k v m) = delete k m" |
|
397 |
"k \<noteq> l \<Longrightarrow> delete k (update l v m) = update l v (delete k m)" |
|
63462 | 398 |
by (transfer; simp add: fun_upd_twist)+ |
29708 | 399 |
|
63462 | 400 |
lemma delete_empty [simp]: "delete k empty = empty" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
401 |
by transfer simp |
29708 | 402 |
|
35157 | 403 |
lemma replace_update: |
37052 | 404 |
"k \<notin> keys m \<Longrightarrow> replace k v m = m" |
405 |
"k \<in> keys m \<Longrightarrow> replace k v m = update k v m" |
|
63462 | 406 |
by (transfer; auto simp add: replace_def fun_upd_twist)+ |
407 |
||
63194 | 408 |
lemma map_values_update: "map_values f (update k v m) = update k (f k v) (map_values f m)" |
409 |
by transfer (simp_all add: fun_eq_iff) |
|
63462 | 410 |
|
411 |
lemma size_mono: "finite (keys m') \<Longrightarrow> keys m \<subseteq> keys m' \<Longrightarrow> size m \<le> size m'" |
|
63194 | 412 |
unfolding size_def by (auto intro: card_mono) |
29708 | 413 |
|
63462 | 414 |
lemma size_empty [simp]: "size empty = 0" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
415 |
unfolding size_def by transfer simp |
29708 | 416 |
|
417 |
lemma size_update: |
|
37052 | 418 |
"finite (keys m) \<Longrightarrow> size (update k v m) = |
419 |
(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:
49834
diff
changeset
|
420 |
unfolding size_def by transfer (auto simp add: insert_dom) |
29708 | 421 |
|
63462 | 422 |
lemma size_delete: "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:
49834
diff
changeset
|
423 |
unfolding size_def by transfer simp |
29708 | 424 |
|
63462 | 425 |
lemma size_tabulate [simp]: "size (tabulate ks f) = length (remdups ks)" |
426 |
unfolding size_def by transfer (auto simp add: map_of_map_restrict card_set comp_def) |
|
29708 | 427 |
|
63194 | 428 |
lemma keys_filter: "keys (filter P m) \<subseteq> keys m" |
429 |
by transfer (auto split: option.splits) |
|
430 |
||
431 |
lemma size_filter: "finite (keys m) \<Longrightarrow> size (filter P m) \<le> size m" |
|
432 |
by (intro size_mono keys_filter) |
|
433 |
||
63462 | 434 |
lemma bulkload_tabulate: "bulkload xs = tabulate [0..<length xs] (nth xs)" |
56528 | 435 |
by transfer (auto simp add: map_of_map_restrict) |
29826 | 436 |
|
63462 | 437 |
lemma is_empty_empty [simp]: "is_empty empty" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
438 |
unfolding is_empty_def by transfer simp |
37052 | 439 |
|
63462 | 440 |
lemma is_empty_update [simp]: "\<not> is_empty (update k v m)" |
441 |
unfolding is_empty_def by transfer simp |
|
442 |
||
443 |
lemma is_empty_delete: "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:
49834
diff
changeset
|
444 |
unfolding is_empty_def by transfer (auto simp del: dom_eq_empty_conv) |
37052 | 445 |
|
63462 | 446 |
lemma is_empty_replace [simp]: "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:
49834
diff
changeset
|
447 |
unfolding is_empty_def replace_def by transfer auto |
37052 | 448 |
|
63462 | 449 |
lemma is_empty_default [simp]: "\<not> is_empty (default k v m)" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
450 |
unfolding is_empty_def default_def by transfer auto |
37052 | 451 |
|
63462 | 452 |
lemma is_empty_map_entry [simp]: "is_empty (map_entry k f m) \<longleftrightarrow> is_empty m" |
56528 | 453 |
unfolding is_empty_def by transfer (auto split: option.split) |
37052 | 454 |
|
63462 | 455 |
lemma is_empty_map_values [simp]: "is_empty (map_values f m) \<longleftrightarrow> is_empty m" |
63194 | 456 |
unfolding is_empty_def by transfer (auto simp: fun_eq_iff) |
457 |
||
63462 | 458 |
lemma is_empty_map_default [simp]: "\<not> is_empty (map_default k v f m)" |
37052 | 459 |
by (simp add: map_default_def) |
460 |
||
63462 | 461 |
lemma keys_dom_lookup: "keys m = dom (Mapping.lookup m)" |
56545 | 462 |
by transfer rule |
463 |
||
63462 | 464 |
lemma keys_empty [simp]: "keys empty = {}" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
465 |
by transfer simp |
37052 | 466 |
|
63462 | 467 |
lemma keys_update [simp]: "keys (update k v m) = insert k (keys m)" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
468 |
by transfer simp |
37052 | 469 |
|
63462 | 470 |
lemma keys_delete [simp]: "keys (delete k m) = keys m - {k}" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
471 |
by transfer simp |
37052 | 472 |
|
63462 | 473 |
lemma keys_replace [simp]: "keys (replace k v m) = keys m" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
474 |
unfolding replace_def by transfer (simp add: insert_absorb) |
37052 | 475 |
|
63462 | 476 |
lemma keys_default [simp]: "keys (default k v m) = insert k (keys m)" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
477 |
unfolding default_def by transfer (simp add: insert_absorb) |
37052 | 478 |
|
63462 | 479 |
lemma keys_map_entry [simp]: "keys (map_entry k f m) = keys m" |
56528 | 480 |
by transfer (auto split: option.split) |
37052 | 481 |
|
63462 | 482 |
lemma keys_map_default [simp]: "keys (map_default k v f m) = insert k (keys m)" |
37052 | 483 |
by (simp add: map_default_def) |
484 |
||
63462 | 485 |
lemma keys_map_values [simp]: "keys (map_values f m) = keys m" |
63194 | 486 |
by transfer (simp_all add: dom_def) |
487 |
||
63462 | 488 |
lemma keys_combine_with_key [simp]: |
63194 | 489 |
"Mapping.keys (combine_with_key f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2" |
63462 | 490 |
by transfer (auto simp: dom_def combine_options_def split: option.splits) |
63194 | 491 |
|
492 |
lemma keys_combine [simp]: "Mapping.keys (combine f m1 m2) = Mapping.keys m1 \<union> Mapping.keys m2" |
|
493 |
by (simp add: combine_altdef) |
|
494 |
||
63462 | 495 |
lemma keys_tabulate [simp]: "keys (tabulate ks f) = set ks" |
49929
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
kuncar
parents:
49834
diff
changeset
|
496 |
by transfer (simp add: map_of_map_restrict o_def) |
37026
7e8979a155ae
operations default, map_entry, map_default; more lemmas
haftmann
parents:
36176
diff
changeset
|
497 |
|
63194 | 498 |
lemma keys_of_alist [simp]: "keys (of_alist xs) = set (List.map fst xs)" |
499 |
by transfer (simp_all add: dom_map_of_conv_image_fst) |
|
500 |
||
63462 | 501 |
lemma keys_bulkload [simp]: "keys (bulkload xs) = {0..<length xs}" |
56528 | 502 |
by (simp add: bulkload_tabulate) |
37026
7e8979a155ae
operations default, map_entry, map_default; more lemmas
haftmann
parents:
36176
diff
changeset
|
503 |
|
63462 | 504 |
lemma distinct_ordered_keys [simp]: "distinct (ordered_keys m)" |
37052 | 505 |
by (simp add: ordered_keys_def) |
506 |
||
63462 | 507 |
lemma ordered_keys_infinite [simp]: "\<not> finite (keys m) \<Longrightarrow> ordered_keys m = []" |
37052 | 508 |
by (simp add: ordered_keys_def) |
509 |
||
63462 | 510 |
lemma ordered_keys_empty [simp]: "ordered_keys empty = []" |
37052 | 511 |
by (simp add: ordered_keys_def) |
512 |
||
513 |
lemma ordered_keys_update [simp]: |
|
514 |
"k \<in> keys m \<Longrightarrow> ordered_keys (update k v m) = ordered_keys m" |
|
63462 | 515 |
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> |
516 |
ordered_keys (update k v m) = insort k (ordered_keys m)" |
|
517 |
by (simp_all add: ordered_keys_def) |
|
518 |
(auto simp only: sorted_list_of_set_insert [symmetric] insert_absorb) |
|
37052 | 519 |
|
63462 | 520 |
lemma ordered_keys_delete [simp]: "ordered_keys (delete k m) = remove1 k (ordered_keys m)" |
37052 | 521 |
proof (cases "finite (keys m)") |
63462 | 522 |
case False |
523 |
then show ?thesis by simp |
|
37052 | 524 |
next |
63462 | 525 |
case fin: True |
37052 | 526 |
show ?thesis |
527 |
proof (cases "k \<in> keys m") |
|
63462 | 528 |
case False |
529 |
with fin have "k \<notin> set (sorted_list_of_set (keys m))" |
|
530 |
by simp |
|
531 |
with False show ?thesis |
|
532 |
by (simp add: ordered_keys_def remove1_idem) |
|
37052 | 533 |
next |
63462 | 534 |
case True |
535 |
with fin show ?thesis |
|
536 |
by (simp add: ordered_keys_def sorted_list_of_set_remove) |
|
37052 | 537 |
qed |
538 |
qed |
|
539 |
||
63462 | 540 |
lemma ordered_keys_replace [simp]: "ordered_keys (replace k v m) = ordered_keys m" |
37052 | 541 |
by (simp add: replace_def) |
542 |
||
543 |
lemma ordered_keys_default [simp]: |
|
544 |
"k \<in> keys m \<Longrightarrow> ordered_keys (default k v m) = ordered_keys m" |
|
545 |
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (default k v m) = insort k (ordered_keys m)" |
|
546 |
by (simp_all add: default_def) |
|
547 |
||
63462 | 548 |
lemma ordered_keys_map_entry [simp]: "ordered_keys (map_entry k f m) = ordered_keys m" |
37052 | 549 |
by (simp add: ordered_keys_def) |
550 |
||
551 |
lemma ordered_keys_map_default [simp]: |
|
552 |
"k \<in> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = ordered_keys m" |
|
553 |
"finite (keys m) \<Longrightarrow> k \<notin> keys m \<Longrightarrow> ordered_keys (map_default k v f m) = insort k (ordered_keys m)" |
|
554 |
by (simp_all add: map_default_def) |
|
555 |
||
63462 | 556 |
lemma ordered_keys_tabulate [simp]: "ordered_keys (tabulate ks f) = sort (remdups ks)" |
37052 | 557 |
by (simp add: ordered_keys_def sorted_list_of_set_sort_remdups) |
558 |
||
63462 | 559 |
lemma ordered_keys_bulkload [simp]: "ordered_keys (bulkload ks) = [0..<length ks]" |
37052 | 560 |
by (simp add: ordered_keys_def) |
36110 | 561 |
|
63462 | 562 |
lemma tabulate_fold: "tabulate xs f = fold (\<lambda>k m. update k (f k) m) xs empty" |
56528 | 563 |
proof transfer |
564 |
fix f :: "'a \<Rightarrow> 'b" and xs |
|
56529
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
565 |
have "map_of (List.map (\<lambda>k. (k, f k)) xs) = foldr (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty" |
aff193f53a64
restoring notion of primitive vs. derived operations in terms of generated code;
haftmann
parents:
56528
diff
changeset
|
566 |
by (simp add: foldr_map comp_def map_of_foldr) |
56528 | 567 |
also have "foldr (\<lambda>k m. m(k \<mapsto> f k)) xs = fold (\<lambda>k m. m(k \<mapsto> f k)) xs" |
568 |
by (rule foldr_fold) (simp add: fun_eq_iff) |
|
569 |
ultimately show "map_of (List.map (\<lambda>k. (k, f k)) xs) = fold (\<lambda>k m. m(k \<mapsto> f k)) xs Map.empty" |
|
570 |
by simp |
|
571 |
qed |
|
572 |
||
63194 | 573 |
lemma All_mapping_mono: |
574 |
"(\<And>k v. k \<in> keys m \<Longrightarrow> P k v \<Longrightarrow> Q k v) \<Longrightarrow> All_mapping m P \<Longrightarrow> All_mapping m Q" |
|
575 |
unfolding All_mapping_def by transfer (auto simp: All_mapping_def dom_def split: option.splits) |
|
31459 | 576 |
|
63194 | 577 |
lemma All_mapping_empty [simp]: "All_mapping Mapping.empty P" |
578 |
by (auto simp: All_mapping_def lookup_empty) |
|
63462 | 579 |
|
580 |
lemma All_mapping_update_iff: |
|
63194 | 581 |
"All_mapping (Mapping.update k v m) P \<longleftrightarrow> P k v \<and> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v')" |
63462 | 582 |
unfolding All_mapping_def |
63194 | 583 |
proof safe |
584 |
assume "\<forall>x. case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y" |
|
63462 | 585 |
then have *: "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some y \<Rightarrow> P x y" for x |
63194 | 586 |
by blast |
63462 | 587 |
from *[of k] show "P k v" |
588 |
by (simp add: lookup_update) |
|
63194 | 589 |
show "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x |
63462 | 590 |
using *[of x] by (auto simp add: lookup_update' split: if_splits option.splits) |
63194 | 591 |
next |
592 |
assume "P k v" |
|
593 |
assume "\<forall>x. case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" |
|
63462 | 594 |
then have A: "case Mapping.lookup m x of None \<Rightarrow> True | Some v' \<Rightarrow> k = x \<or> P x v'" for x |
595 |
by blast |
|
63194 | 596 |
show "case Mapping.lookup (Mapping.update k v m) x of None \<Rightarrow> True | Some xa \<Rightarrow> P x xa" for x |
597 |
using \<open>P k v\<close> A[of x] by (auto simp: lookup_update' split: option.splits) |
|
598 |
qed |
|
599 |
||
600 |
lemma All_mapping_update: |
|
601 |
"P k v \<Longrightarrow> All_mapping m (\<lambda>k' v'. k = k' \<or> P k' v') \<Longrightarrow> All_mapping (Mapping.update k v m) P" |
|
602 |
by (simp add: All_mapping_update_iff) |
|
603 |
||
63462 | 604 |
lemma All_mapping_filter_iff: "All_mapping (filter P m) Q \<longleftrightarrow> All_mapping m (\<lambda>k v. P k v \<longrightarrow> Q k v)" |
63194 | 605 |
by (auto simp: All_mapping_def lookup_filter split: option.splits) |
606 |
||
63462 | 607 |
lemma All_mapping_filter: "All_mapping m Q \<Longrightarrow> All_mapping (filter P m) Q" |
63194 | 608 |
by (auto simp: All_mapping_filter_iff intro: All_mapping_mono) |
31459 | 609 |
|
63462 | 610 |
lemma All_mapping_map_values: "All_mapping (map_values f m) P \<longleftrightarrow> All_mapping m (\<lambda>k v. P k (f k v))" |
63194 | 611 |
by (auto simp: All_mapping_def lookup_map_values split: option.splits) |
612 |
||
63462 | 613 |
lemma All_mapping_tabulate: "(\<forall>x\<in>set xs. P x (f x)) \<Longrightarrow> All_mapping (Mapping.tabulate xs f) P" |
614 |
unfolding All_mapping_def |
|
615 |
apply (intro allI) |
|
616 |
apply transfer |
|
617 |
apply (auto split: option.split dest!: map_of_SomeD) |
|
618 |
done |
|
63194 | 619 |
|
620 |
lemma All_mapping_alist: |
|
621 |
"(\<And>k v. (k, v) \<in> set xs \<Longrightarrow> P k v) \<Longrightarrow> All_mapping (Mapping.of_alist xs) P" |
|
622 |
by (auto simp: All_mapping_def lookup_of_alist dest!: map_of_SomeD split: option.splits) |
|
623 |
||
63462 | 624 |
lemma combine_empty [simp]: "combine f Mapping.empty y = y" "combine f y Mapping.empty = y" |
625 |
by (transfer; force)+ |
|
63194 | 626 |
|
627 |
lemma (in abel_semigroup) comm_monoid_set_combine: "comm_monoid_set (combine f) Mapping.empty" |
|
628 |
by standard (transfer fixing: f, simp add: combine_options_ac[of f] ac_simps)+ |
|
629 |
||
630 |
locale combine_mapping_abel_semigroup = abel_semigroup |
|
631 |
begin |
|
632 |
||
633 |
sublocale combine: comm_monoid_set "combine f" Mapping.empty |
|
634 |
by (rule comm_monoid_set_combine) |
|
635 |
||
636 |
lemma fold_combine_code: |
|
637 |
"combine.F g (set xs) = foldr (\<lambda>x. combine f (g x)) (remdups xs) Mapping.empty" |
|
638 |
proof - |
|
639 |
have "combine.F g (set xs) = foldr (\<lambda>x. combine f (g x)) xs Mapping.empty" |
|
640 |
if "distinct xs" for xs |
|
641 |
using that by (induction xs) simp_all |
|
642 |
from this[of "remdups xs"] show ?thesis by simp |
|
643 |
qed |
|
63462 | 644 |
|
645 |
lemma keys_fold_combine: "finite A \<Longrightarrow> Mapping.keys (combine.F g A) = (\<Union>x\<in>A. Mapping.keys (g x))" |
|
646 |
by (induct A rule: finite_induct) simp_all |
|
35157 | 647 |
|
49975
faf4afed009f
transfer package: more flexible handling of equality relations using is_equality predicate
huffman
parents:
49973
diff
changeset
|
648 |
end |
59485 | 649 |
|
63462 | 650 |
|
63194 | 651 |
subsection \<open>Code generator setup\<close> |
652 |
||
653 |
hide_const (open) empty is_empty rep lookup lookup_default filter update delete ordered_keys |
|
654 |
keys size replace default map_entry map_default tabulate bulkload map map_values combine of_alist |
|
655 |
||
656 |
end |