| author | huffman | 
| Sat, 22 May 2010 13:27:36 -0700 | |
| changeset 37082 | a1fb7947dc2b | 
| parent 37053 | a89b47a94b19 | 
| child 37462 | 802619d7576d | 
| permissions | -rw-r--r-- | 
| 35617 | 1 | (* Author: Florian Haftmann, TU Muenchen *) | 
| 2 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 3 | header {* Abstract type of Red-Black Trees *}
 | 
| 35617 | 4 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 5 | (*<*) | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 6 | theory RBT | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 7 | imports Main RBT_Impl Mapping | 
| 35617 | 8 | begin | 
| 9 | ||
| 10 | subsection {* Type definition *}
 | |
| 11 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 12 | typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}"
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 13 | morphisms impl_of RBT | 
| 35617 | 14 | proof - | 
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 15 | have "RBT_Impl.Empty \<in> ?rbt" by simp | 
| 35617 | 16 | then show ?thesis .. | 
| 17 | qed | |
| 18 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 19 | lemma is_rbt_impl_of [simp, intro]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 20 | "is_rbt (impl_of t)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 21 | using impl_of [of t] by simp | 
| 35617 | 22 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 23 | lemma rbt_eq: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 24 | "t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 25 | by (simp add: impl_of_inject) | 
| 35617 | 26 | |
| 36111 | 27 | lemma [code abstype]: | 
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 28 | "RBT (impl_of t) = t" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 29 | by (simp add: impl_of_inverse) | 
| 35617 | 30 | |
| 31 | ||
| 32 | subsection {* Primitive operations *}
 | |
| 33 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 34 | definition lookup :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 35 | [code]: "lookup t = RBT_Impl.lookup (impl_of t)" | 
| 35617 | 36 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 37 | definition empty :: "('a\<Colon>linorder, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 38 | "empty = RBT RBT_Impl.Empty" | 
| 35617 | 39 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 40 | lemma impl_of_empty [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 41 | "impl_of empty = RBT_Impl.Empty" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 42 | by (simp add: empty_def RBT_inverse) | 
| 35617 | 43 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 44 | definition insert :: "'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 45 | "insert k v t = RBT (RBT_Impl.insert k v (impl_of t))" | 
| 35617 | 46 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 47 | lemma impl_of_insert [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 48 | "impl_of (insert k v t) = RBT_Impl.insert k v (impl_of t)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 49 | by (simp add: insert_def RBT_inverse) | 
| 35617 | 50 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 51 | definition delete :: "'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 52 | "delete k t = RBT (RBT_Impl.delete k (impl_of t))" | 
| 35617 | 53 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 54 | lemma impl_of_delete [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 55 | "impl_of (delete k t) = RBT_Impl.delete k (impl_of t)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 56 | by (simp add: delete_def RBT_inverse) | 
| 35617 | 57 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 58 | definition entries :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 59 | [code]: "entries t = RBT_Impl.entries (impl_of t)" | 
| 35617 | 60 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 61 | definition keys :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 62 | [code]: "keys t = RBT_Impl.keys (impl_of t)" | 
| 36111 | 63 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 64 | definition bulkload :: "('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 65 | "bulkload xs = RBT (RBT_Impl.bulkload xs)" | 
| 35617 | 66 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 67 | lemma impl_of_bulkload [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 68 | "impl_of (bulkload xs) = RBT_Impl.bulkload xs" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 69 | by (simp add: bulkload_def RBT_inverse) | 
| 35617 | 70 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 71 | definition map_entry :: "'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 72 | "map_entry k f t = RBT (RBT_Impl.map_entry k f (impl_of t))" | 
| 35617 | 73 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 74 | lemma impl_of_map_entry [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 75 | "impl_of (map_entry k f t) = RBT_Impl.map_entry k f (impl_of t)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 76 | by (simp add: map_entry_def RBT_inverse) | 
| 35617 | 77 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 78 | definition map :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 79 | "map f t = RBT (RBT_Impl.map f (impl_of t))" | 
| 35617 | 80 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 81 | lemma impl_of_map [code abstract]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 82 | "impl_of (map f t) = RBT_Impl.map f (impl_of t)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 83 | by (simp add: map_def RBT_inverse) | 
| 35617 | 84 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 85 | definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 86 | [code]: "fold f t = RBT_Impl.fold f (impl_of t)" | 
| 35617 | 87 | |
| 88 | ||
| 89 | subsection {* Derived operations *}
 | |
| 90 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 91 | definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 92 | [code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)" | 
| 35617 | 93 | |
| 94 | ||
| 95 | subsection {* Abstract lookup properties *}
 | |
| 96 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 97 | lemma lookup_RBT: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 98 | "is_rbt t \<Longrightarrow> lookup (RBT t) = RBT_Impl.lookup t" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 99 | by (simp add: lookup_def RBT_inverse) | 
| 35617 | 100 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 101 | lemma lookup_impl_of: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 102 | "RBT_Impl.lookup (impl_of t) = lookup t" | 
| 35617 | 103 | by (simp add: lookup_def) | 
| 104 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 105 | lemma entries_impl_of: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 106 | "RBT_Impl.entries (impl_of t) = entries t" | 
| 35617 | 107 | by (simp add: entries_def) | 
| 108 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 109 | lemma keys_impl_of: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 110 | "RBT_Impl.keys (impl_of t) = keys t" | 
| 36111 | 111 | by (simp add: keys_def) | 
| 112 | ||
| 35617 | 113 | lemma lookup_empty [simp]: | 
| 114 | "lookup empty = Map.empty" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 115 | by (simp add: empty_def lookup_RBT expand_fun_eq) | 
| 35617 | 116 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 117 | lemma lookup_insert [simp]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 118 | "lookup (insert k v t) = (lookup t)(k \<mapsto> v)" | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 119 | by (simp add: insert_def lookup_RBT lookup_insert lookup_impl_of) | 
| 35617 | 120 | |
| 121 | lemma lookup_delete [simp]: | |
| 122 | "lookup (delete k t) = (lookup t)(k := None)" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 123 | by (simp add: delete_def lookup_RBT RBT_Impl.lookup_delete lookup_impl_of restrict_complement_singleton_eq) | 
| 35617 | 124 | |
| 125 | lemma map_of_entries [simp]: | |
| 126 | "map_of (entries t) = lookup t" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 127 | by (simp add: entries_def map_of_entries lookup_impl_of) | 
| 35617 | 128 | |
| 36111 | 129 | lemma entries_lookup: | 
| 130 | "entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2" | |
| 131 | by (simp add: entries_def lookup_def entries_lookup) | |
| 132 | ||
| 35617 | 133 | lemma lookup_bulkload [simp]: | 
| 134 | "lookup (bulkload xs) = map_of xs" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 135 | by (simp add: bulkload_def lookup_RBT RBT_Impl.lookup_bulkload) | 
| 35617 | 136 | |
| 137 | lemma lookup_map_entry [simp]: | |
| 138 | "lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))" | |
| 37027 | 139 | by (simp add: map_entry_def lookup_RBT RBT_Impl.lookup_map_entry lookup_impl_of) | 
| 35617 | 140 | |
| 141 | lemma lookup_map [simp]: | |
| 142 | "lookup (map f t) k = Option.map (f k) (lookup t k)" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 143 | by (simp add: map_def lookup_RBT lookup_map lookup_impl_of) | 
| 35617 | 144 | |
| 145 | lemma fold_fold: | |
| 146 | "fold f t = (\<lambda>s. foldl (\<lambda>s (k, v). f k v s) s (entries t))" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 147 | by (simp add: fold_def expand_fun_eq RBT_Impl.fold_def entries_impl_of) | 
| 35617 | 148 | |
| 36111 | 149 | lemma is_empty_empty [simp]: | 
| 150 | "is_empty t \<longleftrightarrow> t = empty" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 151 | by (simp add: rbt_eq is_empty_def impl_of_empty split: rbt.split) | 
| 36111 | 152 | |
| 153 | lemma RBT_lookup_empty [simp]: (*FIXME*) | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 154 | "RBT_Impl.lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty" | 
| 36111 | 155 | by (cases t) (auto simp add: expand_fun_eq) | 
| 156 | ||
| 157 | lemma lookup_empty_empty [simp]: | |
| 158 | "lookup t = Map.empty \<longleftrightarrow> t = empty" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 159 | by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse) | 
| 36111 | 160 | |
| 161 | lemma sorted_keys [iff]: | |
| 162 | "sorted (keys t)" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 163 | by (simp add: keys_def RBT_Impl.keys_def sorted_entries) | 
| 36111 | 164 | |
| 165 | lemma distinct_keys [iff]: | |
| 166 | "distinct (keys t)" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 167 | by (simp add: keys_def RBT_Impl.keys_def distinct_entries) | 
| 36111 | 168 | |
| 169 | ||
| 170 | subsection {* Implementation of mappings *}
 | |
| 171 | ||
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 172 | definition Mapping :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" where
 | 
| 36111 | 173 | "Mapping t = Mapping.Mapping (lookup t)" | 
| 174 | ||
| 175 | code_datatype Mapping | |
| 176 | ||
| 177 | lemma lookup_Mapping [simp, code]: | |
| 178 | "Mapping.lookup (Mapping t) = lookup t" | |
| 179 | by (simp add: Mapping_def) | |
| 180 | ||
| 181 | lemma empty_Mapping [code]: | |
| 182 | "Mapping.empty = Mapping empty" | |
| 183 | by (rule mapping_eqI) simp | |
| 184 | ||
| 185 | lemma is_empty_Mapping [code]: | |
| 186 | "Mapping.is_empty (Mapping t) \<longleftrightarrow> is_empty t" | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 187 | by (simp add: rbt_eq Mapping.is_empty_empty Mapping_def) | 
| 36111 | 188 | |
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 189 | lemma insert_Mapping [code]: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 190 | "Mapping.update k v (Mapping t) = Mapping (insert k v t)" | 
| 36111 | 191 | by (rule mapping_eqI) simp | 
| 192 | ||
| 193 | lemma delete_Mapping [code]: | |
| 37027 | 194 | "Mapping.delete k (Mapping t) = Mapping (delete k t)" | 
| 195 | by (rule mapping_eqI) simp | |
| 196 | ||
| 197 | lemma map_entry_Mapping [code]: | |
| 198 | "Mapping.map_entry k f (Mapping t) = Mapping (map_entry k f t)" | |
| 36111 | 199 | by (rule mapping_eqI) simp | 
| 200 | ||
| 201 | lemma keys_Mapping [code]: | |
| 202 | "Mapping.keys (Mapping t) = set (keys t)" | |
| 203 | by (simp add: keys_def Mapping_def Mapping.keys_def lookup_def lookup_keys) | |
| 204 | ||
| 205 | lemma ordered_keys_Mapping [code]: | |
| 206 | "Mapping.ordered_keys (Mapping t) = keys t" | |
| 207 | by (rule sorted_distinct_set_unique) (simp_all add: ordered_keys_def keys_Mapping) | |
| 208 | ||
| 209 | lemma Mapping_size_card_keys: (*FIXME*) | |
| 210 | "Mapping.size m = card (Mapping.keys m)" | |
| 211 | by (simp add: Mapping.size_def Mapping.keys_def) | |
| 212 | ||
| 213 | lemma size_Mapping [code]: | |
| 214 | "Mapping.size (Mapping t) = length (keys t)" | |
| 215 | by (simp add: Mapping_size_card_keys keys_Mapping distinct_card) | |
| 216 | ||
| 217 | lemma tabulate_Mapping [code]: | |
| 218 | "Mapping.tabulate ks f = Mapping (bulkload (List.map (\<lambda>k. (k, f k)) ks))" | |
| 219 | by (rule mapping_eqI) (simp add: map_of_map_restrict) | |
| 220 | ||
| 221 | lemma bulkload_Mapping [code]: | |
| 222 | "Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))" | |
| 223 | by (rule mapping_eqI) (simp add: map_of_map_restrict expand_fun_eq) | |
| 224 | ||
| 37053 | 225 | lemma [code, code del]: | 
| 226 | "HOL.eq (x :: (_, _) mapping) y \<longleftrightarrow> x = y" by (fact eq_equals) (*FIXME*) | |
| 36111 | 227 | |
| 228 | lemma eq_Mapping [code]: | |
| 229 | "HOL.eq (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2" | |
| 230 | by (simp add: eq Mapping_def entries_lookup) | |
| 231 | ||
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36147diff
changeset | 232 | hide_const (open) impl_of lookup empty insert delete | 
| 36111 | 233 | entries keys bulkload map_entry map fold | 
| 36147 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 234 | (*>*) | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 235 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 236 | text {* 
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 237 | This theory defines abstract red-black trees as an efficient | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 238 | representation of finite maps, backed by the implementation | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 239 |   in @{theory RBT_Impl}.
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 240 | *} | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 241 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 242 | subsection {* Data type and invariant *}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 243 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 244 | text {*
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 245 |   The type @{typ "('k, 'v) RBT_Impl.rbt"} denotes red-black trees with
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 246 |   keys of type @{typ "'k"} and values of type @{typ "'v"}. To function
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 247 |   properly, the key type musorted belong to the @{text "linorder"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 248 | class. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 249 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 250 |   A value @{term t} of this type is a valid red-black tree if it
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 251 |   satisfies the invariant @{text "is_rbt t"}.  The abstract type @{typ
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 252 |   "('k, 'v) rbt"} always obeys this invariant, and for this reason you
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 253 |   should only use this in our application.  Going back to @{typ "('k,
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 254 | 'v) RBT_Impl.rbt"} may be necessary in proofs if not yet proven | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 255 | properties about the operations must be established. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 256 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 257 |   The interpretation function @{const "RBT.lookup"} returns the partial
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 258 | map represented by a red-black tree: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 259 |   @{term_type[display] "RBT.lookup"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 260 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 261 | This function should be used for reasoning about the semantics of the RBT | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 262 | operations. Furthermore, it implements the lookup functionality for | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 263 | the data structure: It is executable and the lookup is performed in | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 264 | $O(\log n)$. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 265 | *} | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 266 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 267 | subsection {* Operations *}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 268 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 269 | text {*
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 270 | Currently, the following operations are supported: | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 271 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 272 |   @{term_type [display] "RBT.empty"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 273 | Returns the empty tree. $O(1)$ | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 274 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 275 |   @{term_type [display] "RBT.insert"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 276 | Updates the map at a given position. $O(\log n)$ | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 277 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 278 |   @{term_type [display] "RBT.delete"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 279 | Deletes a map entry at a given position. $O(\log n)$ | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 280 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 281 |   @{term_type [display] "RBT.entries"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 282 | Return a corresponding key-value list for a tree. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 283 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 284 |   @{term_type [display] "RBT.bulkload"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 285 | Builds a tree from a key-value list. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 286 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 287 |   @{term_type [display] "RBT.map_entry"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 288 | Maps a single entry in a tree. | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 289 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 290 |   @{term_type [display] "RBT.map"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 291 | Maps all values in a tree. $O(n)$ | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 292 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 293 |   @{term_type [display] "RBT.fold"}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 294 | Folds over all entries in a tree. $O(n)$ | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 295 | *} | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 296 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 297 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 298 | subsection {* Invariant preservation *}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 299 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 300 | text {*
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 301 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 302 |   @{thm Empty_is_rbt}\hfill(@{text "Empty_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 303 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 304 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 305 |   @{thm insert_is_rbt}\hfill(@{text "insert_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 306 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 307 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 308 |   @{thm delete_is_rbt}\hfill(@{text "delete_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 309 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 310 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 311 |   @{thm bulkload_is_rbt}\hfill(@{text "bulkload_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 312 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 313 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 314 |   @{thm map_entry_is_rbt}\hfill(@{text "map_entry_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 315 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 316 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 317 |   @{thm map_is_rbt}\hfill(@{text "map_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 318 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 319 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 320 |   @{thm union_is_rbt}\hfill(@{text "union_is_rbt"})
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 321 | *} | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 322 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 323 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 324 | subsection {* Map Semantics *}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 325 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 326 | text {*
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 327 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 328 |   \underline{@{text "lookup_empty"}}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 329 |   @{thm [display] lookup_empty}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 330 |   \vspace{1ex}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 331 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 332 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 333 |   \underline{@{text "lookup_insert"}}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 334 |   @{thm [display] lookup_insert}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 335 |   \vspace{1ex}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 336 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 337 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 338 |   \underline{@{text "lookup_delete"}}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 339 |   @{thm [display] lookup_delete}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 340 |   \vspace{1ex}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 341 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 342 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 343 |   \underline{@{text "lookup_bulkload"}}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 344 |   @{thm [display] lookup_bulkload}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 345 |   \vspace{1ex}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 346 | |
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 347 | \noindent | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 348 |   \underline{@{text "lookup_map"}}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 349 |   @{thm [display] lookup_map}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 350 |   \vspace{1ex}
 | 
| 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 haftmann parents: 
36111diff
changeset | 351 | *} | 
| 35617 | 352 | |
| 353 | end |