| author | haftmann | 
| Thu, 30 Sep 2010 08:50:45 +0200 | |
| changeset 39794 | 51451d73c3d4 | 
| parent 39380 | 5a2662c1e44a | 
| child 40612 | 7ae5b89d8913 | 
| 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: 
36111 
diff
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: 
36111 
diff
changeset
 | 
5  | 
(*<*)  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
6  | 
theory RBT  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
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: 
36111 
diff
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: 
36111 
diff
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: 
36111 
diff
changeset
 | 
15  | 
have "RBT_Impl.Empty \<in> ?rbt" by simp  | 
| 35617 | 16  | 
then show ?thesis ..  | 
17  | 
qed  | 
|
18  | 
||
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
19  | 
lemma rbt_eq_iff:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
20  | 
"t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
21  | 
by (simp add: impl_of_inject)  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
22  | 
|
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
23  | 
lemma rbt_eqI:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
24  | 
"impl_of t1 = impl_of t2 \<Longrightarrow> t1 = t2"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
25  | 
by (simp add: rbt_eq_iff)  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
26  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
27  | 
lemma is_rbt_impl_of [simp, intro]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
28  | 
"is_rbt (impl_of t)"  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
29  | 
using impl_of [of t] by simp  | 
| 35617 | 30  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
31  | 
lemma RBT_impl_of [simp, code abstype]:  | 
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
32  | 
"RBT (impl_of t) = t"  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
33  | 
by (simp add: impl_of_inverse)  | 
| 35617 | 34  | 
|
35  | 
||
36  | 
subsection {* Primitive operations *}
 | 
|
37  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
38  | 
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: 
36111 
diff
changeset
 | 
39  | 
[code]: "lookup t = RBT_Impl.lookup (impl_of t)"  | 
| 35617 | 40  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
41  | 
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: 
36111 
diff
changeset
 | 
42  | 
"empty = RBT RBT_Impl.Empty"  | 
| 35617 | 43  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
44  | 
lemma impl_of_empty [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
45  | 
"impl_of empty = RBT_Impl.Empty"  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
46  | 
by (simp add: empty_def RBT_inverse)  | 
| 35617 | 47  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
48  | 
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: 
36111 
diff
changeset
 | 
49  | 
"insert k v t = RBT (RBT_Impl.insert k v (impl_of t))"  | 
| 35617 | 50  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
51  | 
lemma impl_of_insert [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
52  | 
"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: 
36111 
diff
changeset
 | 
53  | 
by (simp add: insert_def RBT_inverse)  | 
| 35617 | 54  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
55  | 
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: 
36111 
diff
changeset
 | 
56  | 
"delete k t = RBT (RBT_Impl.delete k (impl_of t))"  | 
| 35617 | 57  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
58  | 
lemma impl_of_delete [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
59  | 
"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: 
36111 
diff
changeset
 | 
60  | 
by (simp add: delete_def RBT_inverse)  | 
| 35617 | 61  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
62  | 
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: 
36111 
diff
changeset
 | 
63  | 
[code]: "entries t = RBT_Impl.entries (impl_of t)"  | 
| 35617 | 64  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
65  | 
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: 
36111 
diff
changeset
 | 
66  | 
[code]: "keys t = RBT_Impl.keys (impl_of t)"  | 
| 36111 | 67  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
68  | 
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: 
36111 
diff
changeset
 | 
69  | 
"bulkload xs = RBT (RBT_Impl.bulkload xs)"  | 
| 35617 | 70  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
71  | 
lemma impl_of_bulkload [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
72  | 
"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: 
36111 
diff
changeset
 | 
73  | 
by (simp add: bulkload_def RBT_inverse)  | 
| 35617 | 74  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
75  | 
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: 
36111 
diff
changeset
 | 
76  | 
"map_entry k f t = RBT (RBT_Impl.map_entry k f (impl_of t))"  | 
| 35617 | 77  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
78  | 
lemma impl_of_map_entry [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
79  | 
"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: 
36111 
diff
changeset
 | 
80  | 
by (simp add: map_entry_def RBT_inverse)  | 
| 35617 | 81  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
82  | 
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: 
36111 
diff
changeset
 | 
83  | 
"map f t = RBT (RBT_Impl.map f (impl_of t))"  | 
| 35617 | 84  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
85  | 
lemma impl_of_map [code abstract]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
86  | 
"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: 
36111 
diff
changeset
 | 
87  | 
by (simp add: map_def RBT_inverse)  | 
| 35617 | 88  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
89  | 
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: 
36111 
diff
changeset
 | 
90  | 
[code]: "fold f t = RBT_Impl.fold f (impl_of t)"  | 
| 35617 | 91  | 
|
92  | 
||
93  | 
subsection {* Derived operations *}
 | 
|
94  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
95  | 
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: 
36111 
diff
changeset
 | 
96  | 
[code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 35617 | 97  | 
|
98  | 
||
99  | 
subsection {* Abstract lookup properties *}
 | 
|
100  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
101  | 
lemma lookup_RBT:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
102  | 
"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: 
36111 
diff
changeset
 | 
103  | 
by (simp add: lookup_def RBT_inverse)  | 
| 35617 | 104  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
105  | 
lemma lookup_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
106  | 
"RBT_Impl.lookup (impl_of t) = lookup t"  | 
| 35617 | 107  | 
by (simp add: lookup_def)  | 
108  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
109  | 
lemma entries_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
110  | 
"RBT_Impl.entries (impl_of t) = entries t"  | 
| 35617 | 111  | 
by (simp add: entries_def)  | 
112  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
113  | 
lemma keys_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
114  | 
"RBT_Impl.keys (impl_of t) = keys t"  | 
| 36111 | 115  | 
by (simp add: keys_def)  | 
116  | 
||
| 35617 | 117  | 
lemma lookup_empty [simp]:  | 
118  | 
"lookup empty = Map.empty"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
119  | 
by (simp add: empty_def lookup_RBT fun_eq_iff)  | 
| 35617 | 120  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
121  | 
lemma lookup_insert [simp]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
122  | 
"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: 
36111 
diff
changeset
 | 
123  | 
by (simp add: insert_def lookup_RBT lookup_insert lookup_impl_of)  | 
| 35617 | 124  | 
|
125  | 
lemma lookup_delete [simp]:  | 
|
126  | 
"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: 
36111 
diff
changeset
 | 
127  | 
by (simp add: delete_def lookup_RBT RBT_Impl.lookup_delete lookup_impl_of restrict_complement_singleton_eq)  | 
| 35617 | 128  | 
|
129  | 
lemma map_of_entries [simp]:  | 
|
130  | 
"map_of (entries t) = lookup t"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
131  | 
by (simp add: entries_def map_of_entries lookup_impl_of)  | 
| 35617 | 132  | 
|
| 36111 | 133  | 
lemma entries_lookup:  | 
134  | 
"entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"  | 
|
135  | 
by (simp add: entries_def lookup_def entries_lookup)  | 
|
136  | 
||
| 35617 | 137  | 
lemma lookup_bulkload [simp]:  | 
138  | 
"lookup (bulkload xs) = map_of xs"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
139  | 
by (simp add: bulkload_def lookup_RBT RBT_Impl.lookup_bulkload)  | 
| 35617 | 140  | 
|
141  | 
lemma lookup_map_entry [simp]:  | 
|
142  | 
"lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"  | 
|
| 37027 | 143  | 
by (simp add: map_entry_def lookup_RBT RBT_Impl.lookup_map_entry lookup_impl_of)  | 
| 35617 | 144  | 
|
145  | 
lemma lookup_map [simp]:  | 
|
146  | 
"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: 
36111 
diff
changeset
 | 
147  | 
by (simp add: map_def lookup_RBT lookup_map lookup_impl_of)  | 
| 35617 | 148  | 
|
149  | 
lemma fold_fold:  | 
|
| 37462 | 150  | 
"fold f t = More_List.fold (prod_case f) (entries t)"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
151  | 
by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of)  | 
| 35617 | 152  | 
|
| 36111 | 153  | 
lemma is_empty_empty [simp]:  | 
154  | 
"is_empty t \<longleftrightarrow> t = empty"  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
155  | 
by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split)  | 
| 36111 | 156  | 
|
157  | 
lemma RBT_lookup_empty [simp]: (*FIXME*)  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
158  | 
"RBT_Impl.lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
159  | 
by (cases t) (auto simp add: fun_eq_iff)  | 
| 36111 | 160  | 
|
161  | 
lemma lookup_empty_empty [simp]:  | 
|
162  | 
"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: 
36111 
diff
changeset
 | 
163  | 
by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse)  | 
| 36111 | 164  | 
|
165  | 
lemma sorted_keys [iff]:  | 
|
166  | 
"sorted (keys t)"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
167  | 
by (simp add: keys_def RBT_Impl.keys_def sorted_entries)  | 
| 36111 | 168  | 
|
169  | 
lemma distinct_keys [iff]:  | 
|
170  | 
"distinct (keys t)"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
171  | 
by (simp add: keys_def RBT_Impl.keys_def distinct_entries)  | 
| 36111 | 172  | 
|
173  | 
||
174  | 
subsection {* Implementation of mappings *}
 | 
|
175  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
176  | 
definition Mapping :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) mapping" where
 | 
| 36111 | 177  | 
"Mapping t = Mapping.Mapping (lookup t)"  | 
178  | 
||
179  | 
code_datatype Mapping  | 
|
180  | 
||
181  | 
lemma lookup_Mapping [simp, code]:  | 
|
182  | 
"Mapping.lookup (Mapping t) = lookup t"  | 
|
183  | 
by (simp add: Mapping_def)  | 
|
184  | 
||
185  | 
lemma empty_Mapping [code]:  | 
|
186  | 
"Mapping.empty = Mapping empty"  | 
|
187  | 
by (rule mapping_eqI) simp  | 
|
188  | 
||
189  | 
lemma is_empty_Mapping [code]:  | 
|
190  | 
"Mapping.is_empty (Mapping t) \<longleftrightarrow> is_empty t"  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
191  | 
by (simp add: rbt_eq_iff Mapping.is_empty_empty Mapping_def)  | 
| 36111 | 192  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
193  | 
lemma insert_Mapping [code]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
194  | 
"Mapping.update k v (Mapping t) = Mapping (insert k v t)"  | 
| 36111 | 195  | 
by (rule mapping_eqI) simp  | 
196  | 
||
197  | 
lemma delete_Mapping [code]:  | 
|
| 37027 | 198  | 
"Mapping.delete k (Mapping t) = Mapping (delete k t)"  | 
199  | 
by (rule mapping_eqI) simp  | 
|
200  | 
||
201  | 
lemma map_entry_Mapping [code]:  | 
|
202  | 
"Mapping.map_entry k f (Mapping t) = Mapping (map_entry k f t)"  | 
|
| 36111 | 203  | 
by (rule mapping_eqI) simp  | 
204  | 
||
205  | 
lemma keys_Mapping [code]:  | 
|
206  | 
"Mapping.keys (Mapping t) = set (keys t)"  | 
|
207  | 
by (simp add: keys_def Mapping_def Mapping.keys_def lookup_def lookup_keys)  | 
|
208  | 
||
209  | 
lemma ordered_keys_Mapping [code]:  | 
|
210  | 
"Mapping.ordered_keys (Mapping t) = keys t"  | 
|
211  | 
by (rule sorted_distinct_set_unique) (simp_all add: ordered_keys_def keys_Mapping)  | 
|
212  | 
||
213  | 
lemma Mapping_size_card_keys: (*FIXME*)  | 
|
214  | 
"Mapping.size m = card (Mapping.keys m)"  | 
|
215  | 
by (simp add: Mapping.size_def Mapping.keys_def)  | 
|
216  | 
||
217  | 
lemma size_Mapping [code]:  | 
|
218  | 
"Mapping.size (Mapping t) = length (keys t)"  | 
|
219  | 
by (simp add: Mapping_size_card_keys keys_Mapping distinct_card)  | 
|
220  | 
||
221  | 
lemma tabulate_Mapping [code]:  | 
|
222  | 
"Mapping.tabulate ks f = Mapping (bulkload (List.map (\<lambda>k. (k, f k)) ks))"  | 
|
223  | 
by (rule mapping_eqI) (simp add: map_of_map_restrict)  | 
|
224  | 
||
225  | 
lemma bulkload_Mapping [code]:  | 
|
226  | 
"Mapping.bulkload vs = Mapping (bulkload (List.map (\<lambda>n. (n, vs ! n)) [0..<length vs]))"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
227  | 
by (rule mapping_eqI) (simp add: map_of_map_restrict fun_eq_iff)  | 
| 36111 | 228  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
229  | 
lemma equal_Mapping [code]:  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
230  | 
"HOL.equal (Mapping t1) (Mapping t2) \<longleftrightarrow> entries t1 = entries t2"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
231  | 
by (simp add: equal Mapping_def entries_lookup)  | 
| 36111 | 232  | 
|
| 
38857
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
233  | 
lemma [code nbe]:  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
234  | 
"HOL.equal (x :: (_, _) mapping) x \<longleftrightarrow> True"  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
235  | 
by (fact equal_refl)  | 
| 
 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 
haftmann 
parents: 
37462 
diff
changeset
 | 
236  | 
|
| 36111 | 237  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36147 
diff
changeset
 | 
238  | 
hide_const (open) impl_of lookup empty insert delete  | 
| 36111 | 239  | 
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: 
36111 
diff
changeset
 | 
240  | 
(*>*)  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
241  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
242  | 
text {* 
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
243  | 
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: 
36111 
diff
changeset
 | 
244  | 
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: 
36111 
diff
changeset
 | 
245  | 
  in @{theory RBT_Impl}.
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
246  | 
*}  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
247  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
248  | 
subsection {* Data type and invariant *}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
249  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
250  | 
text {*
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
251  | 
  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: 
36111 
diff
changeset
 | 
252  | 
  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: 
36111 
diff
changeset
 | 
253  | 
  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: 
36111 
diff
changeset
 | 
254  | 
class.  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
255  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
256  | 
  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: 
36111 
diff
changeset
 | 
257  | 
  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: 
36111 
diff
changeset
 | 
258  | 
  "('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: 
36111 
diff
changeset
 | 
259  | 
  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: 
36111 
diff
changeset
 | 
260  | 
'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: 
36111 
diff
changeset
 | 
261  | 
properties about the operations must be established.  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
262  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
263  | 
  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: 
36111 
diff
changeset
 | 
264  | 
map represented by a red-black tree:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
265  | 
  @{term_type[display] "RBT.lookup"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
266  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
267  | 
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: 
36111 
diff
changeset
 | 
268  | 
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: 
36111 
diff
changeset
 | 
269  | 
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: 
36111 
diff
changeset
 | 
270  | 
$O(\log n)$.  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
271  | 
*}  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
272  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
273  | 
subsection {* Operations *}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
274  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
275  | 
text {*
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
276  | 
Currently, the following operations are supported:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
277  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
278  | 
  @{term_type [display] "RBT.empty"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
279  | 
Returns the empty tree. $O(1)$  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
280  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
281  | 
  @{term_type [display] "RBT.insert"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
282  | 
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: 
36111 
diff
changeset
 | 
283  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
284  | 
  @{term_type [display] "RBT.delete"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
285  | 
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: 
36111 
diff
changeset
 | 
286  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
287  | 
  @{term_type [display] "RBT.entries"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
288  | 
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: 
36111 
diff
changeset
 | 
289  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
290  | 
  @{term_type [display] "RBT.bulkload"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
291  | 
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: 
36111 
diff
changeset
 | 
292  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
293  | 
  @{term_type [display] "RBT.map_entry"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
294  | 
Maps a single entry in a tree.  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
295  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
296  | 
  @{term_type [display] "RBT.map"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
297  | 
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: 
36111 
diff
changeset
 | 
298  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
299  | 
  @{term_type [display] "RBT.fold"}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
300  | 
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: 
36111 
diff
changeset
 | 
301  | 
*}  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
302  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
303  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
304  | 
subsection {* Invariant preservation *}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
305  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
306  | 
text {*
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
307  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
308  | 
  @{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: 
36111 
diff
changeset
 | 
309  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
310  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
311  | 
  @{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: 
36111 
diff
changeset
 | 
312  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
313  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
314  | 
  @{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: 
36111 
diff
changeset
 | 
315  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
316  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
317  | 
  @{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: 
36111 
diff
changeset
 | 
318  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
319  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
320  | 
  @{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: 
36111 
diff
changeset
 | 
321  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
322  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
323  | 
  @{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: 
36111 
diff
changeset
 | 
324  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
325  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
326  | 
  @{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: 
36111 
diff
changeset
 | 
327  | 
*}  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
328  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
329  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
330  | 
subsection {* Map Semantics *}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
331  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
332  | 
text {*
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
333  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
334  | 
  \underline{@{text "lookup_empty"}}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
335  | 
  @{thm [display] lookup_empty}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
336  | 
  \vspace{1ex}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
337  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
338  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
339  | 
  \underline{@{text "lookup_insert"}}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
340  | 
  @{thm [display] lookup_insert}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
341  | 
  \vspace{1ex}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
342  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
343  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
344  | 
  \underline{@{text "lookup_delete"}}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
345  | 
  @{thm [display] lookup_delete}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
346  | 
  \vspace{1ex}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
347  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
348  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
349  | 
  \underline{@{text "lookup_bulkload"}}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
350  | 
  @{thm [display] lookup_bulkload}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
351  | 
  \vspace{1ex}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
352  | 
|
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
353  | 
\noindent  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
354  | 
  \underline{@{text "lookup_map"}}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
355  | 
  @{thm [display] lookup_map}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
356  | 
  \vspace{1ex}
 | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
357  | 
*}  | 
| 35617 | 358  | 
|
359  | 
end  |