kuncar@49929
|
1 |
(* Title: HOL/Library/RBT_Mapping.thy
|
kuncar@49929
|
2 |
Author: Florian Haftmann and Ondrej Kuncar
|
kuncar@49929
|
3 |
*)
|
bulwahn@43124
|
4 |
|
bulwahn@43124
|
5 |
header {* Implementation of mappings with Red-Black Trees *}
|
bulwahn@43124
|
6 |
|
bulwahn@43124
|
7 |
(*<*)
|
bulwahn@43124
|
8 |
theory RBT_Mapping
|
bulwahn@43124
|
9 |
imports RBT Mapping
|
bulwahn@43124
|
10 |
begin
|
bulwahn@43124
|
11 |
|
bulwahn@43124
|
12 |
subsection {* Implementation of mappings *}
|
bulwahn@43124
|
13 |
|
kuncar@49929
|
14 |
lift_definition Mapping :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" is lookup .
|
bulwahn@43124
|
15 |
|
bulwahn@43124
|
16 |
code_datatype Mapping
|
bulwahn@43124
|
17 |
|
bulwahn@43124
|
18 |
lemma lookup_Mapping [simp, code]:
|
bulwahn@43124
|
19 |
"Mapping.lookup (Mapping t) = lookup t"
|
kuncar@49929
|
20 |
by (transfer fixing: t) rule
|
bulwahn@43124
|
21 |
|
kuncar@49929
|
22 |
lemma empty_Mapping [code]: "Mapping.empty = Mapping empty"
|
kuncar@49929
|
23 |
proof -
|
kuncar@49929
|
24 |
note RBT.empty.transfer[transfer_rule del]
|
kuncar@49929
|
25 |
show ?thesis by transfer simp
|
kuncar@49929
|
26 |
qed
|
bulwahn@43124
|
27 |
|
bulwahn@43124
|
28 |
lemma is_empty_Mapping [code]:
|
bulwahn@43124
|
29 |
"Mapping.is_empty (Mapping t) \<longleftrightarrow> is_empty t"
|
kuncar@49929
|
30 |
unfolding is_empty_def by (transfer fixing: t) simp
|
bulwahn@43124
|
31 |
|
bulwahn@43124
|
32 |
lemma insert_Mapping [code]:
|
bulwahn@43124
|
33 |
"Mapping.update k v (Mapping t) = Mapping (insert k v t)"
|
kuncar@49929
|
34 |
by (transfer fixing: t) simp
|
bulwahn@43124
|
35 |
|
bulwahn@43124
|
36 |
lemma delete_Mapping [code]:
|
bulwahn@43124
|
37 |
"Mapping.delete k (Mapping t) = Mapping (delete k t)"
|
kuncar@49929
|
38 |
by (transfer fixing: t) simp
|
bulwahn@43124
|
39 |
|
bulwahn@43124
|
40 |
lemma map_entry_Mapping [code]:
|
bulwahn@43124
|
41 |
"Mapping.map_entry k f (Mapping t) = Mapping (map_entry k f t)"
|
kuncar@51438
|
42 |
apply (transfer fixing: t) by (case_tac "lookup t k") auto
|
bulwahn@43124
|
43 |
|
bulwahn@43124
|
44 |
lemma keys_Mapping [code]:
|
bulwahn@43124
|
45 |
"Mapping.keys (Mapping t) = set (keys t)"
|
kuncar@49929
|
46 |
by (transfer fixing: t) (simp add: lookup_keys)
|
bulwahn@43124
|
47 |
|
bulwahn@43124
|
48 |
lemma ordered_keys_Mapping [code]:
|
bulwahn@43124
|
49 |
"Mapping.ordered_keys (Mapping t) = keys t"
|
kuncar@49929
|
50 |
unfolding ordered_keys_def
|
kuncar@49929
|
51 |
by (transfer fixing: t) (auto simp add: lookup_keys intro: sorted_distinct_set_unique)
|
bulwahn@43124
|
52 |
|
bulwahn@43124
|
53 |
lemma Mapping_size_card_keys: (*FIXME*)
|
bulwahn@43124
|
54 |
"Mapping.size m = card (Mapping.keys m)"
|
kuncar@49929
|
55 |
unfolding size_def by transfer simp
|
bulwahn@43124
|
56 |
|
bulwahn@43124
|
57 |
lemma size_Mapping [code]:
|
bulwahn@43124
|
58 |
"Mapping.size (Mapping t) = length (keys t)"
|
kuncar@49929
|
59 |
unfolding size_def
|
kuncar@49929
|
60 |
by (transfer fixing: t) (simp add: lookup_keys distinct_card)
|
bulwahn@43124
|
61 |
|
kuncar@49929
|
62 |
context
|
kuncar@49929
|
63 |
notes RBT.bulkload.transfer[transfer_rule del]
|
kuncar@49929
|
64 |
begin
|
kuncar@49929
|
65 |
lemma tabulate_Mapping [code]:
|
kuncar@49929
|
66 |
"Mapping.tabulate ks f = Mapping (bulkload (List.map (\<lambda>k. (k, f k)) ks))"
|
kuncar@49929
|
67 |
by transfer (simp add: map_of_map_restrict)
|
kuncar@49929
|
68 |
|
kuncar@49929
|
69 |
lemma bulkload_Mapping [code]:
|
kuncar@49929
|
70 |
"Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"
|
kuncar@49929
|
71 |
by transfer (simp add: map_of_map_restrict fun_eq_iff)
|
kuncar@49929
|
72 |
end
|
bulwahn@43124
|
73 |
|
bulwahn@43124
|
74 |
lemma equal_Mapping [code]:
|
bulwahn@43124
|
75 |
"HOL.equal (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"
|
haftmann@51161
|
76 |
by (transfer fixing: t1 t2) (simp add: entries_lookup)
|
bulwahn@43124
|
77 |
|
bulwahn@43124
|
78 |
lemma [code nbe]:
|
bulwahn@43124
|
79 |
"HOL.equal (x :: (_, _) mapping) x \<longleftrightarrow> True"
|
bulwahn@43124
|
80 |
by (fact equal_refl)
|
bulwahn@43124
|
81 |
|
bulwahn@43124
|
82 |
|
bulwahn@43124
|
83 |
hide_const (open) impl_of lookup empty insert delete
|
bulwahn@43124
|
84 |
entries keys bulkload map_entry map fold
|
bulwahn@43124
|
85 |
(*>*)
|
bulwahn@43124
|
86 |
|
bulwahn@43124
|
87 |
text {*
|
bulwahn@43124
|
88 |
This theory defines abstract red-black trees as an efficient
|
bulwahn@43124
|
89 |
representation of finite maps, backed by the implementation
|
bulwahn@43124
|
90 |
in @{theory RBT_Impl}.
|
bulwahn@43124
|
91 |
*}
|
bulwahn@43124
|
92 |
|
bulwahn@43124
|
93 |
subsection {* Data type and invariant *}
|
bulwahn@43124
|
94 |
|
bulwahn@43124
|
95 |
text {*
|
bulwahn@43124
|
96 |
The type @{typ "('k, 'v) RBT_Impl.rbt"} denotes red-black trees with
|
bulwahn@43124
|
97 |
keys of type @{typ "'k"} and values of type @{typ "'v"}. To function
|
bulwahn@43124
|
98 |
properly, the key type musorted belong to the @{text "linorder"}
|
bulwahn@43124
|
99 |
class.
|
bulwahn@43124
|
100 |
|
bulwahn@43124
|
101 |
A value @{term t} of this type is a valid red-black tree if it
|
bulwahn@43124
|
102 |
satisfies the invariant @{text "is_rbt t"}. The abstract type @{typ
|
bulwahn@43124
|
103 |
"('k, 'v) rbt"} always obeys this invariant, and for this reason you
|
bulwahn@43124
|
104 |
should only use this in our application. Going back to @{typ "('k,
|
bulwahn@43124
|
105 |
'v) RBT_Impl.rbt"} may be necessary in proofs if not yet proven
|
bulwahn@43124
|
106 |
properties about the operations must be established.
|
bulwahn@43124
|
107 |
|
bulwahn@43124
|
108 |
The interpretation function @{const "RBT.lookup"} returns the partial
|
bulwahn@43124
|
109 |
map represented by a red-black tree:
|
bulwahn@43124
|
110 |
@{term_type[display] "RBT.lookup"}
|
bulwahn@43124
|
111 |
|
bulwahn@43124
|
112 |
This function should be used for reasoning about the semantics of the RBT
|
bulwahn@43124
|
113 |
operations. Furthermore, it implements the lookup functionality for
|
bulwahn@43124
|
114 |
the data structure: It is executable and the lookup is performed in
|
bulwahn@43124
|
115 |
$O(\log n)$.
|
bulwahn@43124
|
116 |
*}
|
bulwahn@43124
|
117 |
|
bulwahn@43124
|
118 |
subsection {* Operations *}
|
bulwahn@43124
|
119 |
|
bulwahn@43124
|
120 |
text {*
|
bulwahn@43124
|
121 |
Currently, the following operations are supported:
|
bulwahn@43124
|
122 |
|
bulwahn@43124
|
123 |
@{term_type [display] "RBT.empty"}
|
bulwahn@43124
|
124 |
Returns the empty tree. $O(1)$
|
bulwahn@43124
|
125 |
|
bulwahn@43124
|
126 |
@{term_type [display] "RBT.insert"}
|
bulwahn@43124
|
127 |
Updates the map at a given position. $O(\log n)$
|
bulwahn@43124
|
128 |
|
bulwahn@43124
|
129 |
@{term_type [display] "RBT.delete"}
|
bulwahn@43124
|
130 |
Deletes a map entry at a given position. $O(\log n)$
|
bulwahn@43124
|
131 |
|
bulwahn@43124
|
132 |
@{term_type [display] "RBT.entries"}
|
bulwahn@43124
|
133 |
Return a corresponding key-value list for a tree.
|
bulwahn@43124
|
134 |
|
bulwahn@43124
|
135 |
@{term_type [display] "RBT.bulkload"}
|
bulwahn@43124
|
136 |
Builds a tree from a key-value list.
|
bulwahn@43124
|
137 |
|
bulwahn@43124
|
138 |
@{term_type [display] "RBT.map_entry"}
|
bulwahn@43124
|
139 |
Maps a single entry in a tree.
|
bulwahn@43124
|
140 |
|
bulwahn@43124
|
141 |
@{term_type [display] "RBT.map"}
|
bulwahn@43124
|
142 |
Maps all values in a tree. $O(n)$
|
bulwahn@43124
|
143 |
|
bulwahn@43124
|
144 |
@{term_type [display] "RBT.fold"}
|
bulwahn@43124
|
145 |
Folds over all entries in a tree. $O(n)$
|
bulwahn@43124
|
146 |
*}
|
bulwahn@43124
|
147 |
|
bulwahn@43124
|
148 |
|
bulwahn@43124
|
149 |
subsection {* Invariant preservation *}
|
bulwahn@43124
|
150 |
|
bulwahn@43124
|
151 |
text {*
|
bulwahn@43124
|
152 |
\noindent
|
bulwahn@43124
|
153 |
@{thm Empty_is_rbt}\hfill(@{text "Empty_is_rbt"})
|
bulwahn@43124
|
154 |
|
bulwahn@43124
|
155 |
\noindent
|
Andreas@47450
|
156 |
@{thm rbt_insert_is_rbt}\hfill(@{text "rbt_insert_is_rbt"})
|
bulwahn@43124
|
157 |
|
bulwahn@43124
|
158 |
\noindent
|
Andreas@47450
|
159 |
@{thm rbt_delete_is_rbt}\hfill(@{text "delete_is_rbt"})
|
bulwahn@43124
|
160 |
|
bulwahn@43124
|
161 |
\noindent
|
Andreas@47450
|
162 |
@{thm rbt_bulkload_is_rbt}\hfill(@{text "bulkload_is_rbt"})
|
bulwahn@43124
|
163 |
|
bulwahn@43124
|
164 |
\noindent
|
Andreas@47450
|
165 |
@{thm rbt_map_entry_is_rbt}\hfill(@{text "map_entry_is_rbt"})
|
bulwahn@43124
|
166 |
|
bulwahn@43124
|
167 |
\noindent
|
bulwahn@43124
|
168 |
@{thm map_is_rbt}\hfill(@{text "map_is_rbt"})
|
bulwahn@43124
|
169 |
|
bulwahn@43124
|
170 |
\noindent
|
Andreas@47450
|
171 |
@{thm rbt_union_is_rbt}\hfill(@{text "union_is_rbt"})
|
bulwahn@43124
|
172 |
*}
|
bulwahn@43124
|
173 |
|
bulwahn@43124
|
174 |
|
bulwahn@43124
|
175 |
subsection {* Map Semantics *}
|
bulwahn@43124
|
176 |
|
bulwahn@43124
|
177 |
text {*
|
bulwahn@43124
|
178 |
\noindent
|
bulwahn@43124
|
179 |
\underline{@{text "lookup_empty"}}
|
bulwahn@43124
|
180 |
@{thm [display] lookup_empty}
|
bulwahn@43124
|
181 |
\vspace{1ex}
|
bulwahn@43124
|
182 |
|
bulwahn@43124
|
183 |
\noindent
|
bulwahn@43124
|
184 |
\underline{@{text "lookup_insert"}}
|
bulwahn@43124
|
185 |
@{thm [display] lookup_insert}
|
bulwahn@43124
|
186 |
\vspace{1ex}
|
bulwahn@43124
|
187 |
|
bulwahn@43124
|
188 |
\noindent
|
bulwahn@43124
|
189 |
\underline{@{text "lookup_delete"}}
|
bulwahn@43124
|
190 |
@{thm [display] lookup_delete}
|
bulwahn@43124
|
191 |
\vspace{1ex}
|
bulwahn@43124
|
192 |
|
bulwahn@43124
|
193 |
\noindent
|
bulwahn@43124
|
194 |
\underline{@{text "lookup_bulkload"}}
|
bulwahn@43124
|
195 |
@{thm [display] lookup_bulkload}
|
bulwahn@43124
|
196 |
\vspace{1ex}
|
bulwahn@43124
|
197 |
|
bulwahn@43124
|
198 |
\noindent
|
bulwahn@43124
|
199 |
\underline{@{text "lookup_map"}}
|
bulwahn@43124
|
200 |
@{thm [display] lookup_map}
|
bulwahn@43124
|
201 |
\vspace{1ex}
|
bulwahn@43124
|
202 |
*}
|
bulwahn@43124
|
203 |
|
bulwahn@43124
|
204 |
end |