|
61515
|
1 |
(* Author: Tobias Nipkow *)
|
|
|
2 |
|
|
|
3 |
section \<open>A 2-3-4 Tree Implementation of Maps\<close>
|
|
|
4 |
|
|
|
5 |
theory Tree234_Map
|
|
|
6 |
imports
|
|
|
7 |
Tree234_Set
|
|
|
8 |
"../Data_Structures/Map_by_Ordered"
|
|
|
9 |
begin
|
|
|
10 |
|
|
|
11 |
subsection \<open>Map operations on 2-3-4 trees\<close>
|
|
|
12 |
|
|
|
13 |
fun lookup :: "('a::linorder * 'b) tree234 \<Rightarrow> 'a \<Rightarrow> 'b option" where
|
|
|
14 |
"lookup Leaf x = None" |
|
|
|
15 |
"lookup (Node2 l (a,b) r) x =
|
|
|
16 |
(if x < a then lookup l x else
|
|
|
17 |
if a < x then lookup r x else Some b)" |
|
|
|
18 |
"lookup (Node3 l (a1,b1) m (a2,b2) r) x =
|
|
|
19 |
(if x < a1 then lookup l x else
|
|
|
20 |
if x = a1 then Some b1 else
|
|
|
21 |
if x < a2 then lookup m x else
|
|
|
22 |
if x = a2 then Some b2
|
|
|
23 |
else lookup r x)" |
|
|
|
24 |
"lookup (Node4 l (a1,b1) m (a2,b2) n (a3,b3) r) x =
|
|
|
25 |
(if x < a2 then
|
|
|
26 |
if x = a1 then Some b1 else
|
|
|
27 |
if x < a1 then lookup l x else lookup m x
|
|
|
28 |
else
|
|
|
29 |
if x = a2 then Some b2 else
|
|
|
30 |
if x = a3 then Some b3 else
|
|
|
31 |
if x < a3 then lookup n x
|
|
|
32 |
else lookup r x)"
|
|
|
33 |
|
|
|
34 |
fun upd :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>i" where
|
|
|
35 |
"upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" |
|
|
|
36 |
"upd x y (Node2 l ab r) =
|
|
|
37 |
(if x < fst ab then
|
|
|
38 |
(case upd x y l of
|
|
|
39 |
T\<^sub>i l' => T\<^sub>i (Node2 l' ab r)
|
|
|
40 |
| Up\<^sub>i l1 q l2 => T\<^sub>i (Node3 l1 q l2 ab r))
|
|
|
41 |
else if x = fst ab then T\<^sub>i (Node2 l (x,y) r)
|
|
|
42 |
else
|
|
|
43 |
(case upd x y r of
|
|
|
44 |
T\<^sub>i r' => T\<^sub>i (Node2 l ab r')
|
|
|
45 |
| Up\<^sub>i r1 q r2 => T\<^sub>i (Node3 l ab r1 q r2)))" |
|
|
|
46 |
"upd x y (Node3 l ab1 m ab2 r) =
|
|
|
47 |
(if x < fst ab1 then
|
|
|
48 |
(case upd x y l of
|
|
|
49 |
T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r)
|
|
|
50 |
| Up\<^sub>i l1 q l2 => Up\<^sub>i (Node2 l1 q l2) ab1 (Node2 m ab2 r))
|
|
|
51 |
else if x = fst ab1 then T\<^sub>i (Node3 l (x,y) m ab2 r)
|
|
|
52 |
else if x < fst ab2 then
|
|
|
53 |
(case upd x y m of
|
|
|
54 |
T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r)
|
|
|
55 |
| Up\<^sub>i m1 q m2 => Up\<^sub>i (Node2 l ab1 m1) q (Node2 m2 ab2 r))
|
|
|
56 |
else if x = fst ab2 then T\<^sub>i (Node3 l ab1 m (x,y) r)
|
|
|
57 |
else
|
|
|
58 |
(case upd x y r of
|
|
|
59 |
T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r')
|
|
|
60 |
| Up\<^sub>i r1 q r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 q r2)))" |
|
|
|
61 |
"upd x y (Node4 l ab1 m ab2 n ab3 r) =
|
|
|
62 |
(if x < fst ab2 then
|
|
|
63 |
if x < fst ab1 then
|
|
|
64 |
(case upd x y l of
|
|
|
65 |
T\<^sub>i l' => T\<^sub>i (Node4 l' ab1 m ab2 n ab3 r)
|
|
|
66 |
| Up\<^sub>i l1 q l2 => Up\<^sub>i (Node2 l1 q l2) ab1 (Node3 m ab2 n ab3 r))
|
|
|
67 |
else
|
|
|
68 |
if x = fst ab1 then T\<^sub>i (Node4 l (x,y) m ab2 n ab3 r)
|
|
|
69 |
else
|
|
|
70 |
(case upd x y m of
|
|
|
71 |
T\<^sub>i m' => T\<^sub>i (Node4 l ab1 m' ab2 n ab3 r)
|
|
|
72 |
| Up\<^sub>i m1 q m2 => Up\<^sub>i (Node2 l ab1 m1) q (Node3 m2 ab2 n ab3 r))
|
|
|
73 |
else
|
|
|
74 |
if x = fst ab2 then T\<^sub>i (Node4 l ab1 m (x,y) n ab3 r) else
|
|
|
75 |
if x < fst ab3 then
|
|
|
76 |
(case upd x y n of
|
|
|
77 |
T\<^sub>i n' => T\<^sub>i (Node4 l ab1 m ab2 n' ab3 r)
|
|
|
78 |
| Up\<^sub>i n1 q n2 => Up\<^sub>i (Node2 l ab1 m) ab2(*q*) (Node3 n1 q n2 ab3 r))
|
|
|
79 |
else
|
|
|
80 |
if x = fst ab3 then T\<^sub>i (Node4 l ab1 m ab2 n (x,y) r)
|
|
|
81 |
else
|
|
|
82 |
(case upd x y r of
|
|
|
83 |
T\<^sub>i r' => T\<^sub>i (Node4 l ab1 m ab2 n ab3 r')
|
|
|
84 |
| Up\<^sub>i r1 q r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node3 n ab3 r1 q r2)))"
|
|
|
85 |
|
|
|
86 |
definition update :: "'a::linorder \<Rightarrow> 'b \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
|
|
|
87 |
"update a b t = tree\<^sub>i(upd a b t)"
|
|
|
88 |
|
|
|
89 |
fun del :: "'a::linorder \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) up\<^sub>d"
|
|
|
90 |
where
|
|
|
91 |
"del k Leaf = T\<^sub>d Leaf" |
|
|
|
92 |
"del k (Node2 Leaf p Leaf) = (if k=fst p then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf p Leaf))" |
|
|
|
93 |
"del k (Node3 Leaf p Leaf q Leaf) =
|
|
|
94 |
T\<^sub>d(if k=fst p then Node2 Leaf q Leaf else
|
|
|
95 |
if k=fst q then Node2 Leaf p Leaf
|
|
|
96 |
else Node3 Leaf p Leaf q Leaf)" |
|
|
|
97 |
"del k (Node4 Leaf ab1 Leaf ab2 Leaf ab3 Leaf) =
|
|
|
98 |
T\<^sub>d(if k=fst ab1 then Node3 Leaf ab2 Leaf ab3 Leaf else
|
|
|
99 |
if k=fst ab2 then Node3 Leaf ab1 Leaf ab3 Leaf else
|
|
|
100 |
if k=fst ab3 then Node3 Leaf ab1 Leaf ab2 Leaf
|
|
|
101 |
else Node4 Leaf ab1 Leaf ab2 Leaf ab3 Leaf)" |
|
|
|
102 |
"del k (Node2 l a r) =
|
|
|
103 |
(if k<fst a then node21 (del k l) a r else
|
|
|
104 |
if k > fst a then node22 l a (del k r)
|
|
|
105 |
else let (a',t) = del_min r in node22 l a' t)" |
|
|
|
106 |
"del k (Node3 l a m b r) =
|
|
|
107 |
(if k<fst a then node31 (del k l) a m b r else
|
|
|
108 |
if k = fst a then let (a',m') = del_min m in node32 l a' m' b r else
|
|
|
109 |
if k < fst b then node32 l a (del k m) b r else
|
|
|
110 |
if k = fst b then let (b',r') = del_min r in node33 l a m b' r'
|
|
|
111 |
else node33 l a m b (del k r))" |
|
|
|
112 |
"del x (Node4 l ab1 m ab2 n ab3 r) =
|
|
|
113 |
(if x < fst ab2 then
|
|
|
114 |
if x < fst ab1 then node41 (del x l) ab1 m ab2 n ab3 r else
|
|
|
115 |
if x = fst ab1 then let (ab',m') = del_min m in node42 l ab' m' ab2 n ab3 r
|
|
|
116 |
else node42 l ab1 (del x m) ab2 n ab3 r
|
|
|
117 |
else
|
|
|
118 |
if x = fst ab2 then let (ab',n') = del_min n in node43 l ab1 m ab' n' ab3 r else
|
|
|
119 |
if x < fst ab3 then node43 l ab1 m ab2 (del x n) ab3 r else
|
|
|
120 |
if x = fst ab3 then let (ab',r') = del_min r in node44 l ab1 m ab2 n ab' r'
|
|
|
121 |
else node44 l ab1 m ab2 n ab3 (del x r))"
|
|
|
122 |
|
|
|
123 |
definition delete :: "'a::linorder \<Rightarrow> ('a*'b) tree234 \<Rightarrow> ('a*'b) tree234" where
|
|
|
124 |
"delete k t = tree\<^sub>d(del k t)"
|
|
|
125 |
|
|
|
126 |
|
|
|
127 |
subsection "Functional correctness"
|
|
|
128 |
|
|
|
129 |
lemma lookup: "sorted1(inorder t) \<Longrightarrow> lookup t x = map_of (inorder t) x"
|
|
|
130 |
by (induction t) (auto simp: map_of_simps split: option.split)
|
|
|
131 |
|
|
|
132 |
|
|
|
133 |
lemma inorder_upd:
|
|
|
134 |
"sorted1(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(upd a b t)) = upd_list a b (inorder t)"
|
|
|
135 |
by(induction t)
|
|
|
136 |
(auto simp: upd_list_simps, auto simp: upd_list_simps split: up\<^sub>i.splits)
|
|
|
137 |
|
|
|
138 |
lemma inorder_update:
|
|
|
139 |
"sorted1(inorder t) \<Longrightarrow> inorder(update a b t) = upd_list a b (inorder t)"
|
|
|
140 |
by(simp add: update_def inorder_upd)
|
|
|
141 |
|
|
|
142 |
|
|
|
143 |
lemma inorder_del: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
|
|
|
144 |
inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
|
|
|
145 |
by(induction t rule: del.induct)
|
|
|
146 |
((auto simp: del_list_simps inorder_nodes del_minD split: prod.splits)[1])+
|
|
|
147 |
(* 290 secs (2015) *)
|
|
|
148 |
|
|
|
149 |
lemma inorder_delete: "\<lbrakk> bal t ; sorted1(inorder t) \<rbrakk> \<Longrightarrow>
|
|
|
150 |
inorder(delete x t) = del_list x (inorder t)"
|
|
|
151 |
by(simp add: delete_def inorder_del)
|
|
|
152 |
|
|
|
153 |
|
|
|
154 |
subsection \<open>Balancedness\<close>
|
|
|
155 |
|
|
|
156 |
lemma bal_upd: "bal t \<Longrightarrow> bal (tree\<^sub>i(upd x y t)) \<and> height(upd x y t) = height t"
|
|
|
157 |
by (induct t) (auto, auto split: up\<^sub>i.split) (* 33 secs (2015) *)
|
|
|
158 |
|
|
|
159 |
lemma bal_update: "bal t \<Longrightarrow> bal (update x y t)"
|
|
|
160 |
by (simp add: update_def bal_upd)
|
|
|
161 |
|
|
|
162 |
|
|
|
163 |
lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
|
|
|
164 |
by(induction x t rule: del.induct)
|
|
|
165 |
(auto simp add: heights height_del_min split: prod.split)
|
|
|
166 |
|
|
|
167 |
lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
|
|
|
168 |
by(induction x t rule: del.induct)
|
|
|
169 |
(auto simp: bals bal_del_min height_del height_del_min split: prod.split)
|
|
|
170 |
(* 110 secs (2015) *)
|
|
|
171 |
|
|
|
172 |
corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
|
|
|
173 |
by(simp add: delete_def bal_tree\<^sub>d_del)
|
|
|
174 |
|
|
|
175 |
|
|
|
176 |
subsection \<open>Overall Correctness\<close>
|
|
|
177 |
|
|
|
178 |
interpretation T234_Map: Map_by_Ordered
|
|
|
179 |
where empty = Leaf and lookup = lookup and update = update and delete = delete
|
|
|
180 |
and inorder = inorder and wf = bal
|
|
|
181 |
proof (standard, goal_cases)
|
|
|
182 |
case 2 thus ?case by(simp add: lookup)
|
|
|
183 |
next
|
|
|
184 |
case 3 thus ?case by(simp add: inorder_update)
|
|
|
185 |
next
|
|
|
186 |
case 4 thus ?case by(simp add: inorder_delete)
|
|
|
187 |
next
|
|
|
188 |
case 6 thus ?case by(simp add: bal_update)
|
|
|
189 |
next
|
|
|
190 |
case 7 thus ?case by(simp add: bal_delete)
|
|
|
191 |
qed simp+
|
|
|
192 |
|
|
|
193 |
end
|