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