| author | wenzelm | 
| Tue, 06 Jun 2023 11:07:49 +0200 | |
| changeset 78134 | a11ebc8c751a | 
| parent 74157 | 8e2355ddce1b | 
| permissions | -rw-r--r-- | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/RBT_Mapping.thy  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
2  | 
Author: Florian Haftmann and Ondrej Kuncar  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
3  | 
*)  | 
| 43124 | 4  | 
|
| 60500 | 5  | 
section \<open>Implementation of mappings with Red-Black Trees\<close>  | 
| 43124 | 6  | 
|
7  | 
(*<*)  | 
|
8  | 
theory RBT_Mapping  | 
|
9  | 
imports RBT Mapping  | 
|
10  | 
begin  | 
|
11  | 
||
| 60500 | 12  | 
subsection \<open>Implementation of mappings\<close>  | 
| 43124 | 13  | 
|
| 56019 | 14  | 
context includes rbt.lifting begin  | 
| 61076 | 15  | 
lift_definition Mapping :: "('a::linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" is RBT.lookup .
 | 
| 56019 | 16  | 
end  | 
| 43124 | 17  | 
|
18  | 
code_datatype Mapping  | 
|
19  | 
||
| 56019 | 20  | 
context includes rbt.lifting begin  | 
21  | 
||
| 43124 | 22  | 
lemma lookup_Mapping [simp, code]:  | 
| 56019 | 23  | 
"Mapping.lookup (Mapping t) = RBT.lookup t"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
24  | 
by (transfer fixing: t) rule  | 
| 43124 | 25  | 
|
| 56019 | 26  | 
lemma empty_Mapping [code]: "Mapping.empty = Mapping RBT.empty"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
27  | 
proof -  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
28  | 
note RBT.empty.transfer[transfer_rule del]  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
29  | 
show ?thesis by transfer simp  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
30  | 
qed  | 
| 43124 | 31  | 
|
32  | 
lemma is_empty_Mapping [code]:  | 
|
| 56019 | 33  | 
"Mapping.is_empty (Mapping t) \<longleftrightarrow> RBT.is_empty t"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
34  | 
unfolding is_empty_def by (transfer fixing: t) simp  | 
| 43124 | 35  | 
|
36  | 
lemma insert_Mapping [code]:  | 
|
| 56019 | 37  | 
"Mapping.update k v (Mapping t) = Mapping (RBT.insert k v t)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
38  | 
by (transfer fixing: t) simp  | 
| 43124 | 39  | 
|
40  | 
lemma delete_Mapping [code]:  | 
|
| 56019 | 41  | 
"Mapping.delete k (Mapping t) = Mapping (RBT.delete k t)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
42  | 
by (transfer fixing: t) simp  | 
| 43124 | 43  | 
|
44  | 
lemma map_entry_Mapping [code]:  | 
|
| 56019 | 45  | 
"Mapping.map_entry k f (Mapping t) = Mapping (RBT.map_entry k f t)"  | 
| 63649 | 46  | 
apply (transfer fixing: t)  | 
47  | 
apply (case_tac "RBT.lookup t k")  | 
|
48  | 
apply auto  | 
|
49  | 
done  | 
|
| 43124 | 50  | 
|
51  | 
lemma keys_Mapping [code]:  | 
|
| 56019 | 52  | 
"Mapping.keys (Mapping t) = set (RBT.keys t)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
53  | 
by (transfer fixing: t) (simp add: lookup_keys)  | 
| 43124 | 54  | 
|
55  | 
lemma ordered_keys_Mapping [code]:  | 
|
| 56019 | 56  | 
"Mapping.ordered_keys (Mapping t) = RBT.keys t"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
57  | 
unfolding ordered_keys_def  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
58  | 
by (transfer fixing: t) (auto simp add: lookup_keys intro: sorted_distinct_set_unique)  | 
| 43124 | 59  | 
|
| 73832 | 60  | 
lemma Map_graph_lookup: "Map.graph (RBT.lookup t) = set (RBT.entries t)"  | 
| 
74157
 
8e2355ddce1b
add/rename some theorems about Map(pings)
 
Lukas Stevens <mail@lukas-stevens.de> 
parents: 
73832 
diff
changeset
 | 
61  | 
by (metis RBT.distinct_entries RBT.map_of_entries graph_map_of_if_distinct_dom)  | 
| 73832 | 62  | 
|
63  | 
lemma entries_Mapping [code]: "Mapping.entries (Mapping t) = set (RBT.entries t)"  | 
|
64  | 
by (transfer fixing: t) (fact Map_graph_lookup)  | 
|
65  | 
||
66  | 
lemma ordered_entries_Mapping [code]: "Mapping.ordered_entries (Mapping t) = RBT.entries t"  | 
|
67  | 
proof -  | 
|
68  | 
note folding_Map_graph.idem_if_sorted_distinct[  | 
|
69  | 
where ?m="RBT.lookup t", OF _ _ folding_Map_graph.distinct_if_distinct_map]  | 
|
70  | 
then show ?thesis  | 
|
71  | 
unfolding ordered_entries_def  | 
|
72  | 
by (transfer fixing: t) (auto simp: Map_graph_lookup distinct_entries sorted_entries)  | 
|
73  | 
qed  | 
|
74  | 
||
75  | 
lemma fold_Mapping [code]: "Mapping.fold f (Mapping t) a = RBT.fold f t a"  | 
|
76  | 
by (simp add: Mapping.fold_def fold_fold ordered_entries_Mapping)  | 
|
77  | 
||
| 43124 | 78  | 
lemma Mapping_size_card_keys: (*FIXME*)  | 
79  | 
"Mapping.size m = card (Mapping.keys m)"  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
80  | 
unfolding size_def by transfer simp  | 
| 43124 | 81  | 
|
82  | 
lemma size_Mapping [code]:  | 
|
| 56019 | 83  | 
"Mapping.size (Mapping t) = length (RBT.keys t)"  | 
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
84  | 
unfolding size_def  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
85  | 
by (transfer fixing: t) (simp add: lookup_keys distinct_card)  | 
| 43124 | 86  | 
|
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
87  | 
context  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
88  | 
notes RBT.bulkload.transfer[transfer_rule del]  | 
| 
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
89  | 
begin  | 
| 60373 | 90  | 
|
91  | 
lemma tabulate_Mapping [code]:  | 
|
92  | 
"Mapping.tabulate ks f = Mapping (RBT.bulkload (List.map (\<lambda>k. (k, f k)) ks))"  | 
|
93  | 
by transfer (simp add: map_of_map_restrict)  | 
|
94  | 
||
95  | 
lemma bulkload_Mapping [code]:  | 
|
96  | 
"Mapping.bulkload vs = Mapping (RBT.bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"  | 
|
97  | 
by transfer (simp add: map_of_map_restrict fun_eq_iff)  | 
|
98  | 
||
| 
49929
 
70300f1b6835
update RBT_Mapping, AList_Mapping and Mapping to use lifting/transfer
 
kuncar 
parents: 
47450 
diff
changeset
 | 
99  | 
end  | 
| 43124 | 100  | 
|
| 63194 | 101  | 
lemma map_values_Mapping [code]:  | 
102  | 
"Mapping.map_values f (Mapping t) = Mapping (RBT.map f t)"  | 
|
103  | 
by (transfer fixing: t) (auto simp: fun_eq_iff)  | 
|
104  | 
||
105  | 
lemma filter_Mapping [code]:  | 
|
106  | 
"Mapping.filter P (Mapping t) = Mapping (RBT.filter P t)"  | 
|
107  | 
by (transfer' fixing: P t) (simp add: RBT.lookup_filter fun_eq_iff)  | 
|
108  | 
||
109  | 
lemma combine_with_key_Mapping [code]:  | 
|
110  | 
"Mapping.combine_with_key f (Mapping t1) (Mapping t2) =  | 
|
111  | 
Mapping (RBT.combine_with_key f t1 t2)"  | 
|
112  | 
by (transfer fixing: f t1 t2) (simp_all add: fun_eq_iff)  | 
|
113  | 
||
114  | 
lemma combine_Mapping [code]:  | 
|
115  | 
"Mapping.combine f (Mapping t1) (Mapping t2) =  | 
|
116  | 
Mapping (RBT.combine f t1 t2)"  | 
|
117  | 
by (transfer fixing: f t1 t2) (simp_all add: fun_eq_iff)  | 
|
118  | 
||
| 43124 | 119  | 
lemma equal_Mapping [code]:  | 
| 56019 | 120  | 
"HOL.equal (Mapping t1) (Mapping t2) \<longleftrightarrow> RBT.entries t1 = RBT.entries t2"  | 
| 73832 | 121  | 
by (transfer fixing: t1 t2) (simp add: RBT.entries_lookup)  | 
| 43124 | 122  | 
|
123  | 
lemma [code nbe]:  | 
|
124  | 
"HOL.equal (x :: (_, _) mapping) x \<longleftrightarrow> True"  | 
|
125  | 
by (fact equal_refl)  | 
|
126  | 
||
| 56019 | 127  | 
end  | 
| 43124 | 128  | 
|
129  | 
(*>*)  | 
|
130  | 
||
| 60500 | 131  | 
text \<open>  | 
| 43124 | 132  | 
This theory defines abstract red-black trees as an efficient  | 
133  | 
representation of finite maps, backed by the implementation  | 
|
| 69593 | 134  | 
in \<^theory>\<open>HOL-Library.RBT_Impl\<close>.  | 
| 60500 | 135  | 
\<close>  | 
| 43124 | 136  | 
|
| 60500 | 137  | 
subsection \<open>Data type and invariant\<close>  | 
| 43124 | 138  | 
|
| 60500 | 139  | 
text \<open>  | 
| 69593 | 140  | 
  The type \<^typ>\<open>('k, 'v) RBT_Impl.rbt\<close> denotes red-black trees with
 | 
141  | 
keys of type \<^typ>\<open>'k\<close> and values of type \<^typ>\<open>'v\<close>. To function  | 
|
| 61585 | 142  | 
properly, the key type musorted belong to the \<open>linorder\<close>  | 
| 43124 | 143  | 
class.  | 
144  | 
||
| 69593 | 145  | 
A value \<^term>\<open>t\<close> of this type is a valid red-black tree if it  | 
146  | 
  satisfies the invariant \<open>is_rbt t\<close>.  The abstract type \<^typ>\<open>('k, 'v) rbt\<close> always obeys this invariant, and for this reason you
 | 
|
147  | 
  should only use this in our application.  Going back to \<^typ>\<open>('k,
 | 
|
148  | 
'v) RBT_Impl.rbt\<close> may be necessary in proofs if not yet proven  | 
|
| 43124 | 149  | 
properties about the operations must be established.  | 
150  | 
||
| 69593 | 151  | 
The interpretation function \<^const>\<open>RBT.lookup\<close> returns the partial  | 
| 43124 | 152  | 
map represented by a red-black tree:  | 
153  | 
  @{term_type[display] "RBT.lookup"}
 | 
|
154  | 
||
155  | 
This function should be used for reasoning about the semantics of the RBT  | 
|
156  | 
operations. Furthermore, it implements the lookup functionality for  | 
|
157  | 
the data structure: It is executable and the lookup is performed in  | 
|
158  | 
$O(\log n)$.  | 
|
| 60500 | 159  | 
\<close>  | 
| 43124 | 160  | 
|
| 60500 | 161  | 
subsection \<open>Operations\<close>  | 
| 43124 | 162  | 
|
| 60500 | 163  | 
text \<open>  | 
| 43124 | 164  | 
Currently, the following operations are supported:  | 
165  | 
||
166  | 
  @{term_type [display] "RBT.empty"}
 | 
|
167  | 
Returns the empty tree. $O(1)$  | 
|
168  | 
||
169  | 
  @{term_type [display] "RBT.insert"}
 | 
|
170  | 
Updates the map at a given position. $O(\log n)$  | 
|
171  | 
||
172  | 
  @{term_type [display] "RBT.delete"}
 | 
|
173  | 
Deletes a map entry at a given position. $O(\log n)$  | 
|
174  | 
||
175  | 
  @{term_type [display] "RBT.entries"}
 | 
|
176  | 
Return a corresponding key-value list for a tree.  | 
|
177  | 
||
178  | 
  @{term_type [display] "RBT.bulkload"}
 | 
|
179  | 
Builds a tree from a key-value list.  | 
|
180  | 
||
181  | 
  @{term_type [display] "RBT.map_entry"}
 | 
|
182  | 
Maps a single entry in a tree.  | 
|
183  | 
||
184  | 
  @{term_type [display] "RBT.map"}
 | 
|
185  | 
Maps all values in a tree. $O(n)$  | 
|
186  | 
||
187  | 
  @{term_type [display] "RBT.fold"}
 | 
|
188  | 
Folds over all entries in a tree. $O(n)$  | 
|
| 60500 | 189  | 
\<close>  | 
| 43124 | 190  | 
|
191  | 
||
| 60500 | 192  | 
subsection \<open>Invariant preservation\<close>  | 
| 43124 | 193  | 
|
| 60500 | 194  | 
text \<open>  | 
| 43124 | 195  | 
\noindent  | 
| 61585 | 196  | 
  @{thm Empty_is_rbt}\hfill(\<open>Empty_is_rbt\<close>)
 | 
| 43124 | 197  | 
|
198  | 
\noindent  | 
|
| 61585 | 199  | 
  @{thm rbt_insert_is_rbt}\hfill(\<open>rbt_insert_is_rbt\<close>)
 | 
| 43124 | 200  | 
|
201  | 
\noindent  | 
|
| 61585 | 202  | 
  @{thm rbt_delete_is_rbt}\hfill(\<open>delete_is_rbt\<close>)
 | 
| 43124 | 203  | 
|
204  | 
\noindent  | 
|
| 61585 | 205  | 
  @{thm rbt_bulkload_is_rbt}\hfill(\<open>bulkload_is_rbt\<close>)
 | 
| 43124 | 206  | 
|
207  | 
\noindent  | 
|
| 61585 | 208  | 
  @{thm rbt_map_entry_is_rbt}\hfill(\<open>map_entry_is_rbt\<close>)
 | 
| 43124 | 209  | 
|
210  | 
\noindent  | 
|
| 61585 | 211  | 
  @{thm map_is_rbt}\hfill(\<open>map_is_rbt\<close>)
 | 
| 43124 | 212  | 
|
213  | 
\noindent  | 
|
| 61585 | 214  | 
  @{thm rbt_union_is_rbt}\hfill(\<open>union_is_rbt\<close>)
 | 
| 60500 | 215  | 
\<close>  | 
| 43124 | 216  | 
|
217  | 
||
| 60500 | 218  | 
subsection \<open>Map Semantics\<close>  | 
| 43124 | 219  | 
|
| 60500 | 220  | 
text \<open>  | 
| 43124 | 221  | 
\noindent  | 
| 61585 | 222  | 
  \underline{\<open>lookup_empty\<close>}
 | 
| 43124 | 223  | 
  @{thm [display] lookup_empty}
 | 
224  | 
  \vspace{1ex}
 | 
|
225  | 
||
226  | 
\noindent  | 
|
| 61585 | 227  | 
  \underline{\<open>lookup_insert\<close>}
 | 
| 43124 | 228  | 
  @{thm [display] lookup_insert}
 | 
229  | 
  \vspace{1ex}
 | 
|
230  | 
||
231  | 
\noindent  | 
|
| 61585 | 232  | 
  \underline{\<open>lookup_delete\<close>}
 | 
| 43124 | 233  | 
  @{thm [display] lookup_delete}
 | 
234  | 
  \vspace{1ex}
 | 
|
235  | 
||
236  | 
\noindent  | 
|
| 61585 | 237  | 
  \underline{\<open>lookup_bulkload\<close>}
 | 
| 43124 | 238  | 
  @{thm [display] lookup_bulkload}
 | 
239  | 
  \vspace{1ex}
 | 
|
240  | 
||
241  | 
\noindent  | 
|
| 61585 | 242  | 
  \underline{\<open>lookup_map\<close>}
 | 
| 43124 | 243  | 
  @{thm [display] lookup_map}
 | 
244  | 
  \vspace{1ex}
 | 
|
| 60500 | 245  | 
\<close>  | 
| 43124 | 246  | 
|
| 62390 | 247  | 
end  |