| author | haftmann | 
| Thu, 29 Dec 2011 13:42:21 +0100 | |
| changeset 46034 | 773c0c4994df | 
| parent 45928 | 874845660119 | 
| child 46133 | d9fe85d3d2cd | 
| 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  | 
| 43124 | 7  | 
imports Main RBT_Impl  | 
| 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  | 
| 
45694
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
43124 
diff
changeset
 | 
14  | 
proof  | 
| 
 
4a8743618257
prefer typedef without extra definition and alternative name;
 
wenzelm 
parents: 
43124 
diff
changeset
 | 
15  | 
  show "RBT_Impl.Empty \<in> {t. is_rbt t}" by simp
 | 
| 35617 | 16  | 
qed  | 
17  | 
||
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
18  | 
lemma rbt_eq_iff:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
19  | 
"t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
20  | 
by (simp add: impl_of_inject)  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
21  | 
|
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
22  | 
lemma rbt_eqI:  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
23  | 
"impl_of t1 = impl_of t2 \<Longrightarrow> t1 = t2"  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
24  | 
by (simp add: rbt_eq_iff)  | 
| 
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
25  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
26  | 
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
 | 
27  | 
"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
 | 
28  | 
using impl_of [of t] by simp  | 
| 35617 | 29  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
30  | 
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
 | 
31  | 
"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
 | 
32  | 
by (simp add: impl_of_inverse)  | 
| 35617 | 33  | 
|
34  | 
||
35  | 
subsection {* Primitive operations *}
 | 
|
36  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
37  | 
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
 | 
38  | 
[code]: "lookup t = RBT_Impl.lookup (impl_of t)"  | 
| 35617 | 39  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
40  | 
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
 | 
41  | 
"empty = RBT RBT_Impl.Empty"  | 
| 35617 | 42  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
43  | 
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
 | 
44  | 
"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
 | 
45  | 
by (simp add: empty_def RBT_inverse)  | 
| 35617 | 46  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
47  | 
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
 | 
48  | 
"insert k v t = RBT (RBT_Impl.insert k v (impl_of t))"  | 
| 35617 | 49  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
50  | 
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
 | 
51  | 
"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
 | 
52  | 
by (simp add: insert_def RBT_inverse)  | 
| 35617 | 53  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
54  | 
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
 | 
55  | 
"delete k t = RBT (RBT_Impl.delete k (impl_of t))"  | 
| 35617 | 56  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
57  | 
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
 | 
58  | 
"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
 | 
59  | 
by (simp add: delete_def RBT_inverse)  | 
| 35617 | 60  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
61  | 
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
 | 
62  | 
[code]: "entries t = RBT_Impl.entries (impl_of t)"  | 
| 35617 | 63  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
64  | 
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
 | 
65  | 
[code]: "keys t = RBT_Impl.keys (impl_of t)"  | 
| 36111 | 66  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
67  | 
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
 | 
68  | 
"bulkload xs = RBT (RBT_Impl.bulkload xs)"  | 
| 35617 | 69  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
70  | 
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
 | 
71  | 
"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
 | 
72  | 
by (simp add: bulkload_def RBT_inverse)  | 
| 35617 | 73  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
74  | 
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
 | 
75  | 
"map_entry k f t = RBT (RBT_Impl.map_entry k f (impl_of t))"  | 
| 35617 | 76  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
77  | 
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
 | 
78  | 
"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
 | 
79  | 
by (simp add: map_entry_def RBT_inverse)  | 
| 35617 | 80  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
81  | 
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
 | 
82  | 
"map f t = RBT (RBT_Impl.map f (impl_of t))"  | 
| 35617 | 83  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
84  | 
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
 | 
85  | 
"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
 | 
86  | 
by (simp add: map_def RBT_inverse)  | 
| 35617 | 87  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
88  | 
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
 | 
89  | 
[code]: "fold f t = RBT_Impl.fold f (impl_of t)"  | 
| 35617 | 90  | 
|
91  | 
||
92  | 
subsection {* Derived operations *}
 | 
|
93  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
94  | 
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
 | 
95  | 
[code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)"  | 
| 35617 | 96  | 
|
97  | 
||
98  | 
subsection {* Abstract lookup properties *}
 | 
|
99  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
100  | 
lemma lookup_RBT:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
101  | 
"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
 | 
102  | 
by (simp add: lookup_def RBT_inverse)  | 
| 35617 | 103  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
104  | 
lemma lookup_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
105  | 
"RBT_Impl.lookup (impl_of t) = lookup t"  | 
| 35617 | 106  | 
by (simp add: lookup_def)  | 
107  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
108  | 
lemma entries_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
109  | 
"RBT_Impl.entries (impl_of t) = entries t"  | 
| 35617 | 110  | 
by (simp add: entries_def)  | 
111  | 
||
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
112  | 
lemma keys_impl_of:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
113  | 
"RBT_Impl.keys (impl_of t) = keys t"  | 
| 36111 | 114  | 
by (simp add: keys_def)  | 
115  | 
||
| 35617 | 116  | 
lemma lookup_empty [simp]:  | 
117  | 
"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
 | 
118  | 
by (simp add: empty_def lookup_RBT fun_eq_iff)  | 
| 35617 | 119  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
120  | 
lemma lookup_insert [simp]:  | 
| 
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
121  | 
"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
 | 
122  | 
by (simp add: insert_def lookup_RBT lookup_insert lookup_impl_of)  | 
| 35617 | 123  | 
|
124  | 
lemma lookup_delete [simp]:  | 
|
125  | 
"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
 | 
126  | 
by (simp add: delete_def lookup_RBT RBT_Impl.lookup_delete lookup_impl_of restrict_complement_singleton_eq)  | 
| 35617 | 127  | 
|
128  | 
lemma map_of_entries [simp]:  | 
|
129  | 
"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
 | 
130  | 
by (simp add: entries_def map_of_entries lookup_impl_of)  | 
| 35617 | 131  | 
|
| 36111 | 132  | 
lemma entries_lookup:  | 
133  | 
"entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2"  | 
|
134  | 
by (simp add: entries_def lookup_def entries_lookup)  | 
|
135  | 
||
| 35617 | 136  | 
lemma lookup_bulkload [simp]:  | 
137  | 
"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
 | 
138  | 
by (simp add: bulkload_def lookup_RBT RBT_Impl.lookup_bulkload)  | 
| 35617 | 139  | 
|
140  | 
lemma lookup_map_entry [simp]:  | 
|
141  | 
"lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))"  | 
|
| 37027 | 142  | 
by (simp add: map_entry_def lookup_RBT RBT_Impl.lookup_map_entry lookup_impl_of)  | 
| 35617 | 143  | 
|
144  | 
lemma lookup_map [simp]:  | 
|
145  | 
"lookup (map f t) k = Option.map (f k) (lookup t k)"  | 
|
| 
40612
 
7ae5b89d8913
proper qualification needed due to shadowing on theory merge
 
haftmann 
parents: 
39380 
diff
changeset
 | 
146  | 
by (simp add: map_def lookup_RBT RBT_Impl.lookup_map lookup_impl_of)  | 
| 35617 | 147  | 
|
148  | 
lemma fold_fold:  | 
|
| 37462 | 149  | 
"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
 | 
150  | 
by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of)  | 
| 35617 | 151  | 
|
| 36111 | 152  | 
lemma is_empty_empty [simp]:  | 
153  | 
"is_empty t \<longleftrightarrow> t = empty"  | 
|
| 
39380
 
5a2662c1e44a
established emerging canonical names *_eqI and *_eq_iff
 
haftmann 
parents: 
39302 
diff
changeset
 | 
154  | 
by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split)  | 
| 36111 | 155  | 
|
156  | 
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
 | 
157  | 
"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
 | 
158  | 
by (cases t) (auto simp add: fun_eq_iff)  | 
| 36111 | 159  | 
|
160  | 
lemma lookup_empty_empty [simp]:  | 
|
161  | 
"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
 | 
162  | 
by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse)  | 
| 36111 | 163  | 
|
164  | 
lemma sorted_keys [iff]:  | 
|
165  | 
"sorted (keys t)"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
166  | 
by (simp add: keys_def RBT_Impl.keys_def sorted_entries)  | 
| 36111 | 167  | 
|
168  | 
lemma distinct_keys [iff]:  | 
|
169  | 
"distinct (keys t)"  | 
|
| 
36147
 
b43b22f63665
theory RBT with abstract type of red-black trees backed by implementation RBT_Impl
 
haftmann 
parents: 
36111 
diff
changeset
 | 
170  | 
by (simp add: keys_def RBT_Impl.keys_def distinct_entries)  | 
| 36111 | 171  | 
|
| 
45928
 
874845660119
adding quickcheck generators in some HOL-Library theories
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
172  | 
subsection {* Quickcheck generators *}
 | 
| 
 
874845660119
adding quickcheck generators in some HOL-Library theories
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
173  | 
|
| 
 
874845660119
adding quickcheck generators in some HOL-Library theories
 
bulwahn 
parents: 
45694 
diff
changeset
 | 
174  | 
quickcheck_generator rbt constructors: empty, insert  | 
| 36111 | 175  | 
|
| 35617 | 176  | 
end  |