author | wenzelm |
Wed, 15 Feb 2012 23:19:30 +0100 | |
changeset 46497 | 89ccf66aa73d |
parent 46133 | d9fe85d3d2cd |
child 47092 | fa3538d6004b |
permissions | -rw-r--r-- |
45577 | 1 |
(* Author: Lukas Bulwahn, TU Muenchen *) |
2 |
||
3 |
header {* Lifting operations of RBT trees *} |
|
4 |
||
5 |
theory Lift_RBT |
|
6 |
imports Main "~~/src/HOL/Library/RBT_Impl" |
|
7 |
begin |
|
8 |
||
9 |
subsection {* Type definition *} |
|
10 |
||
11 |
typedef (open) ('a, 'b) rbt = "{t :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt. is_rbt t}" |
|
12 |
morphisms impl_of RBT |
|
13 |
proof - |
|
14 |
have "RBT_Impl.Empty \<in> ?rbt" by simp |
|
15 |
then show ?thesis .. |
|
16 |
qed |
|
17 |
||
18 |
local_setup {* fn lthy => |
|
19 |
let |
|
20 |
val quotients = {qtyp = @{typ "('a, 'b) rbt"}, rtyp = @{typ "('a, 'b) RBT_Impl.rbt"}, |
|
21 |
equiv_rel = @{term "dummy"}, equiv_thm = @{thm refl}} |
|
22 |
val qty_full_name = @{type_name "rbt"} |
|
23 |
||
24 |
fun qinfo phi = Quotient_Info.transform_quotients phi quotients |
|
25 |
in lthy |
|
26 |
|> Local_Theory.declaration {syntax = false, pervasive = true} |
|
27 |
(fn phi => Quotient_Info.update_quotients qty_full_name (qinfo phi) |
|
28 |
#> Quotient_Info.update_abs_rep qty_full_name (Quotient_Info.transform_abs_rep phi |
|
29 |
{abs = @{term "RBT"}, rep = @{term "impl_of"}})) |
|
30 |
end |
|
31 |
*} |
|
32 |
||
33 |
lemma rbt_eq_iff: |
|
34 |
"t1 = t2 \<longleftrightarrow> impl_of t1 = impl_of t2" |
|
35 |
by (simp add: impl_of_inject) |
|
36 |
||
37 |
lemma rbt_eqI: |
|
38 |
"impl_of t1 = impl_of t2 \<Longrightarrow> t1 = t2" |
|
39 |
by (simp add: rbt_eq_iff) |
|
40 |
||
41 |
lemma is_rbt_impl_of [simp, intro]: |
|
42 |
"is_rbt (impl_of t)" |
|
43 |
using impl_of [of t] by simp |
|
44 |
||
45 |
lemma RBT_impl_of [simp, code abstype]: |
|
46 |
"RBT (impl_of t) = t" |
|
47 |
by (simp add: impl_of_inverse) |
|
48 |
||
49 |
||
50 |
subsection {* Primitive operations *} |
|
51 |
||
52 |
quotient_definition lookup where "lookup :: ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a \<rightharpoonup> 'b" is "RBT_Impl.lookup" |
|
53 |
||
54 |
declare lookup_def[unfolded map_fun_def comp_def id_def, code] |
|
55 |
||
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
56 |
(* FIXME: quotient_definition at the moment requires that types variables match exactly, |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
57 |
i.e., sort constraints must be annotated to the constant being lifted. |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
58 |
But, it should be possible to infer the necessary sort constraints, making the explicit |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
59 |
sort constraints unnecessary. |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
60 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
61 |
Hence this unannotated quotient_definition fails: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
62 |
|
45577 | 63 |
quotient_definition empty where "empty :: ('a\<Colon>linorder, 'b) rbt" |
64 |
is "RBT_Impl.Empty" |
|
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
65 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
66 |
Similar issue for quotient_definition for entries, keys, map, and fold. |
45577 | 67 |
*) |
68 |
||
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
69 |
quotient_definition empty where "empty :: ('a\<Colon>linorder, 'b) rbt" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
70 |
is "(RBT_Impl.Empty :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt)" |
45577 | 71 |
|
72 |
lemma impl_of_empty [code abstract]: |
|
73 |
"impl_of empty = RBT_Impl.Empty" |
|
74 |
by (simp add: empty_def RBT_inverse) |
|
75 |
||
76 |
quotient_definition insert where "insert :: 'a\<Colon>linorder \<Rightarrow> 'b \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" |
|
77 |
is "RBT_Impl.insert" |
|
78 |
||
79 |
lemma impl_of_insert [code abstract]: |
|
80 |
"impl_of (insert k v t) = RBT_Impl.insert k v (impl_of t)" |
|
81 |
by (simp add: insert_def RBT_inverse) |
|
82 |
||
83 |
quotient_definition delete where "delete :: 'a\<Colon>linorder \<Rightarrow> ('a, 'b) rbt \<Rightarrow> ('a, 'b) rbt" |
|
84 |
is "RBT_Impl.delete" |
|
85 |
||
86 |
lemma impl_of_delete [code abstract]: |
|
87 |
"impl_of (delete k t) = RBT_Impl.delete k (impl_of t)" |
|
88 |
by (simp add: delete_def RBT_inverse) |
|
89 |
||
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
90 |
(* FIXME: unnecessary type annotations *) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
91 |
quotient_definition "entries :: ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a \<times> 'b) list" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
92 |
is "RBT_Impl.entries :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt \<Rightarrow> ('a \<times> 'b) list" |
45577 | 93 |
|
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
94 |
lemma [code]: "entries t = RBT_Impl.entries (impl_of t)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
95 |
unfolding entries_def map_fun_def comp_def id_def .. |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
96 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
97 |
(* FIXME: unnecessary type annotations *) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
98 |
quotient_definition "keys :: ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'a list" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
99 |
is "RBT_Impl.keys :: ('a\<Colon>linorder, 'b) RBT_Impl.rbt \<Rightarrow> 'a list" |
45577 | 100 |
|
101 |
quotient_definition "bulkload :: ('a\<Colon>linorder \<times> 'b) list \<Rightarrow> ('a, 'b) rbt" |
|
102 |
is "RBT_Impl.bulkload" |
|
103 |
||
104 |
lemma impl_of_bulkload [code abstract]: |
|
105 |
"impl_of (bulkload xs) = RBT_Impl.bulkload xs" |
|
106 |
by (simp add: bulkload_def RBT_inverse) |
|
107 |
||
108 |
quotient_definition "map_entry :: 'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" |
|
109 |
is "RBT_Impl.map_entry" |
|
110 |
||
111 |
lemma impl_of_map_entry [code abstract]: |
|
112 |
"impl_of (map_entry k f t) = RBT_Impl.map_entry k f (impl_of t)" |
|
113 |
by (simp add: map_entry_def RBT_inverse) |
|
114 |
||
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
115 |
(* FIXME: unnecesary type annotations *) |
45577 | 116 |
quotient_definition map where "map :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> ('a, 'b) rbt" |
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
117 |
is "RBT_Impl.map :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('a\<Colon>linorder, 'b) RBT_Impl.rbt \<Rightarrow> ('a, 'b) RBT_Impl.rbt" |
45577 | 118 |
|
119 |
lemma impl_of_map [code abstract]: |
|
120 |
"impl_of (map f t) = RBT_Impl.map f (impl_of t)" |
|
121 |
by (simp add: map_def RBT_inverse) |
|
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
122 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
123 |
(* FIXME: unnecessary type annotations *) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
124 |
quotient_definition fold where "fold :: ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) rbt \<Rightarrow> 'c \<Rightarrow> 'c" is "RBT_Impl.fold :: ('a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> ('a\<Colon>linorder, 'b) RBT_Impl.rbt \<Rightarrow> 'c \<Rightarrow> 'c" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
125 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
126 |
lemma [code]: "fold f t = RBT_Impl.fold f (impl_of t)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
127 |
unfolding fold_def map_fun_def comp_def id_def .. |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
128 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
129 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
130 |
subsection {* Derived operations *} |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
131 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
132 |
definition is_empty :: "('a\<Colon>linorder, 'b) rbt \<Rightarrow> bool" where |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
133 |
[code]: "is_empty t = (case impl_of t of RBT_Impl.Empty \<Rightarrow> True | _ \<Rightarrow> False)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
134 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
135 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
136 |
subsection {* Abstract lookup properties *} |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
137 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
138 |
(* TODO: obtain the following lemmas by lifting existing theorems. *) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
139 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
140 |
lemma lookup_RBT: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
141 |
"is_rbt t \<Longrightarrow> lookup (RBT t) = RBT_Impl.lookup t" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
142 |
by (simp add: lookup_def RBT_inverse) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
143 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
144 |
lemma lookup_impl_of: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
145 |
"RBT_Impl.lookup (impl_of t) = lookup t" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
146 |
by (simp add: lookup_def) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
147 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
148 |
lemma entries_impl_of: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
149 |
"RBT_Impl.entries (impl_of t) = entries t" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
150 |
by (simp add: entries_def) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
151 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
152 |
lemma keys_impl_of: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
153 |
"RBT_Impl.keys (impl_of t) = keys t" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
154 |
by (simp add: keys_def) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
155 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
156 |
lemma lookup_empty [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
157 |
"lookup empty = Map.empty" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
158 |
by (simp add: empty_def lookup_RBT fun_eq_iff) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
159 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
160 |
lemma lookup_insert [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
161 |
"lookup (insert k v t) = (lookup t)(k \<mapsto> v)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
162 |
by (simp add: insert_def lookup_RBT lookup_insert lookup_impl_of) |
45577 | 163 |
|
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
164 |
lemma lookup_delete [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
165 |
"lookup (delete k t) = (lookup t)(k := None)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
166 |
by (simp add: delete_def lookup_RBT RBT_Impl.lookup_delete lookup_impl_of restrict_complement_singleton_eq) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
167 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
168 |
lemma map_of_entries [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
169 |
"map_of (entries t) = lookup t" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
170 |
by (simp add: entries_def map_of_entries lookup_impl_of) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
171 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
172 |
lemma entries_lookup: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
173 |
"entries t1 = entries t2 \<longleftrightarrow> lookup t1 = lookup t2" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
174 |
by (simp add: entries_def lookup_def entries_lookup) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
175 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
176 |
lemma lookup_bulkload [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
177 |
"lookup (bulkload xs) = map_of xs" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
178 |
by (simp add: bulkload_def lookup_RBT RBT_Impl.lookup_bulkload) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
179 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
180 |
lemma lookup_map_entry [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
181 |
"lookup (map_entry k f t) = (lookup t)(k := Option.map f (lookup t k))" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
182 |
by (simp add: map_entry_def lookup_RBT RBT_Impl.lookup_map_entry lookup_impl_of) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
183 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
184 |
lemma lookup_map [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
185 |
"lookup (map f t) k = Option.map (f k) (lookup t k)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
186 |
by (simp add: map_def lookup_RBT RBT_Impl.lookup_map lookup_impl_of) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
187 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
188 |
lemma fold_fold: |
46133
d9fe85d3d2cd
incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
haftmann
parents:
45629
diff
changeset
|
189 |
"fold f t = List.fold (prod_case f) (entries t)" |
45629
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
190 |
by (simp add: fold_def fun_eq_iff RBT_Impl.fold_def entries_impl_of) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
191 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
192 |
lemma is_empty_empty [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
193 |
"is_empty t \<longleftrightarrow> t = empty" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
194 |
by (simp add: rbt_eq_iff is_empty_def impl_of_empty split: rbt.split) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
195 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
196 |
lemma RBT_lookup_empty [simp]: (*FIXME*) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
197 |
"RBT_Impl.lookup t = Map.empty \<longleftrightarrow> t = RBT_Impl.Empty" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
198 |
by (cases t) (auto simp add: fun_eq_iff) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
199 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
200 |
lemma lookup_empty_empty [simp]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
201 |
"lookup t = Map.empty \<longleftrightarrow> t = empty" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
202 |
by (cases t) (simp add: empty_def lookup_def RBT_inject RBT_inverse) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
203 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
204 |
lemma sorted_keys [iff]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
205 |
"sorted (keys t)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
206 |
by (simp add: keys_def RBT_Impl.keys_def sorted_entries) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
207 |
|
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
208 |
lemma distinct_keys [iff]: |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
209 |
"distinct (keys t)" |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
210 |
by (simp add: keys_def RBT_Impl.keys_def distinct_entries) |
ef08425dd2d5
improved example to only use quotient_definitions; added lemmas to serve as possible replacement for the existing RBT theory
bulwahn
parents:
45577
diff
changeset
|
211 |
|
45577 | 212 |
|
213 |
||
214 |
end |