author | nipkow |
Sun, 08 Apr 2018 12:31:08 +0200 | |
changeset 67967 | 5a4280946a25 |
parent 67613 | ce654b0e6d69 |
child 68023 | 75130777ece4 |
permissions | -rw-r--r-- |
61793 | 1 |
(* |
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
2 |
Author: Tobias Nipkow, Daniel Stüwe |
61793 | 3 |
*) |
4 |
||
62130 | 5 |
section \<open>AA Tree Implementation of Sets\<close> |
61793 | 6 |
|
7 |
theory AA_Set |
|
8 |
imports |
|
9 |
Isin2 |
|
10 |
Cmp |
|
11 |
begin |
|
12 |
||
13 |
type_synonym 'a aa_tree = "('a,nat) tree" |
|
14 |
||
15 |
fun lvl :: "'a aa_tree \<Rightarrow> nat" where |
|
16 |
"lvl Leaf = 0" | |
|
17 |
"lvl (Node lv _ _ _) = lv" |
|
62496 | 18 |
|
61793 | 19 |
fun invar :: "'a aa_tree \<Rightarrow> bool" where |
20 |
"invar Leaf = True" | |
|
21 |
"invar (Node h l a r) = |
|
22 |
(invar l \<and> invar r \<and> |
|
23 |
h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node h lr b rr \<and> h = lvl rr + 1)))" |
|
62496 | 24 |
|
61793 | 25 |
fun skew :: "'a aa_tree \<Rightarrow> 'a aa_tree" where |
26 |
"skew (Node lva (Node lvb t1 b t2) a t3) = |
|
27 |
(if lva = lvb then Node lva t1 b (Node lva t2 a t3) else Node lva (Node lvb t1 b t2) a t3)" | |
|
28 |
"skew t = t" |
|
29 |
||
30 |
fun split :: "'a aa_tree \<Rightarrow> 'a aa_tree" where |
|
31 |
"split (Node lva t1 a (Node lvb t2 b (Node lvc t3 c t4))) = |
|
67369 | 32 |
(if lva = lvb \<and> lvb = lvc \<comment> \<open>\<open>lva = lvc\<close> suffices\<close> |
61793 | 33 |
then Node (lva+1) (Node lva t1 a t2) b (Node lva t3 c t4) |
34 |
else Node lva t1 a (Node lvb t2 b (Node lvc t3 c t4)))" | |
|
35 |
"split t = t" |
|
36 |
||
37 |
hide_const (open) insert |
|
38 |
||
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
39 |
fun insert :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where |
61793 | 40 |
"insert x Leaf = Node 1 Leaf x Leaf" | |
41 |
"insert x (Node lv t1 a t2) = |
|
42 |
(case cmp x a of |
|
43 |
LT \<Rightarrow> split (skew (Node lv (insert x t1) a t2)) | |
|
44 |
GT \<Rightarrow> split (skew (Node lv t1 a (insert x t2))) | |
|
45 |
EQ \<Rightarrow> Node lv t1 x t2)" |
|
46 |
||
47 |
fun sngl :: "'a aa_tree \<Rightarrow> bool" where |
|
48 |
"sngl Leaf = False" | |
|
49 |
"sngl (Node _ _ _ Leaf) = True" | |
|
50 |
"sngl (Node lva _ _ (Node lvb _ _ _)) = (lva > lvb)" |
|
51 |
||
52 |
definition adjust :: "'a aa_tree \<Rightarrow> 'a aa_tree" where |
|
53 |
"adjust t = |
|
54 |
(case t of |
|
55 |
Node lv l x r \<Rightarrow> |
|
56 |
(if lvl l >= lv-1 \<and> lvl r >= lv-1 then t else |
|
57 |
if lvl r < lv-1 \<and> sngl l then skew (Node (lv-1) l x r) else |
|
58 |
if lvl r < lv-1 |
|
59 |
then case l of |
|
60 |
Node lva t1 a (Node lvb t2 b t3) |
|
62496 | 61 |
\<Rightarrow> Node (lvb+1) (Node lva t1 a t2) b (Node (lv-1) t3 x r) |
61793 | 62 |
else |
63 |
if lvl r < lv then split (Node (lv-1) l x r) |
|
64 |
else |
|
65 |
case r of |
|
62160 | 66 |
Node lvb t1 b t4 \<Rightarrow> |
61793 | 67 |
(case t1 of |
68 |
Node lva t2 a t3 |
|
69 |
\<Rightarrow> Node (lva+1) (Node (lv-1) l x t2) a |
|
62496 | 70 |
(split (Node (if sngl t1 then lva else lva+1) t3 b t4)))))" |
71 |
||
67406 | 72 |
text\<open>In the paper, the last case of @{const adjust} is expressed with the help of an |
62496 | 73 |
incorrect auxiliary function \texttt{nlvl}. |
74 |
||
75 |
Function @{text del_max} below is called \texttt{dellrg} in the paper. |
|
76 |
The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest |
|
77 |
element but recurses on the left instead of the right subtree; the invariant |
|
67406 | 78 |
is not restored.\<close> |
62496 | 79 |
|
80 |
fun del_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where |
|
81 |
"del_max (Node lv l a Leaf) = (l,a)" | |
|
82 |
"del_max (Node lv l a r) = (let (r',b) = del_max r in (adjust(Node lv l a r'), b))" |
|
61793 | 83 |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
84 |
fun delete :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where |
61793 | 85 |
"delete _ Leaf = Leaf" | |
86 |
"delete x (Node lv l a r) = |
|
87 |
(case cmp x a of |
|
88 |
LT \<Rightarrow> adjust (Node lv (delete x l) a r) | |
|
89 |
GT \<Rightarrow> adjust (Node lv l a (delete x r)) | |
|
90 |
EQ \<Rightarrow> (if l = Leaf then r |
|
91 |
else let (l',b) = del_max l in adjust (Node lv l' b r)))" |
|
92 |
||
62496 | 93 |
fun pre_adjust where |
94 |
"pre_adjust (Node lv l a r) = (invar l \<and> invar r \<and> |
|
95 |
((lv = lvl l + 1 \<and> (lv = lvl r + 1 \<or> lv = lvl r + 2 \<or> lv = lvl r \<and> sngl r)) \<or> |
|
96 |
(lv = lvl l + 2 \<and> (lv = lvl r + 1 \<or> lv = lvl r \<and> sngl r))))" |
|
97 |
||
98 |
declare pre_adjust.simps [simp del] |
|
99 |
||
100 |
subsection "Auxiliary Proofs" |
|
101 |
||
102 |
lemma split_case: "split t = (case t of |
|
103 |
Node lvx a x (Node lvy b y (Node lvz c z d)) \<Rightarrow> |
|
104 |
(if lvx = lvy \<and> lvy = lvz |
|
105 |
then Node (lvx+1) (Node lvx a x b) y (Node lvx c z d) |
|
106 |
else t) |
|
107 |
| t \<Rightarrow> t)" |
|
108 |
by(auto split: tree.split) |
|
109 |
||
110 |
lemma skew_case: "skew t = (case t of |
|
111 |
Node lvx (Node lvy a y b) x c \<Rightarrow> |
|
112 |
(if lvx = lvy then Node lvx a y (Node lvx b x c) else t) |
|
113 |
| t \<Rightarrow> t)" |
|
114 |
by(auto split: tree.split) |
|
115 |
||
116 |
lemma lvl_0_iff: "invar t \<Longrightarrow> lvl t = 0 \<longleftrightarrow> t = Leaf" |
|
117 |
by(cases t) auto |
|
118 |
||
119 |
lemma lvl_Suc_iff: "lvl t = Suc n \<longleftrightarrow> (\<exists> l a r. t = Node (Suc n) l a r)" |
|
120 |
by(cases t) auto |
|
121 |
||
122 |
lemma lvl_skew: "lvl (skew t) = lvl t" |
|
62526 | 123 |
by(cases t rule: skew.cases) auto |
62496 | 124 |
|
125 |
lemma lvl_split: "lvl (split t) = lvl t \<or> lvl (split t) = lvl t + 1 \<and> sngl (split t)" |
|
62526 | 126 |
by(cases t rule: split.cases) auto |
62496 | 127 |
|
128 |
lemma invar_2Nodes:"invar (Node lv l x (Node rlv rl rx rr)) = |
|
129 |
(invar l \<and> invar \<langle>rlv, rl, rx, rr\<rangle> \<and> lv = Suc (lvl l) \<and> |
|
130 |
(lv = Suc rlv \<or> rlv = lv \<and> lv = Suc (lvl rr)))" |
|
131 |
by simp |
|
132 |
||
133 |
lemma invar_NodeLeaf[simp]: |
|
134 |
"invar (Node lv l x Leaf) = (invar l \<and> lv = Suc (lvl l) \<and> lv = Suc 0)" |
|
135 |
by simp |
|
136 |
||
137 |
lemma sngl_if_invar: "invar (Node n l a r) \<Longrightarrow> n = lvl r \<Longrightarrow> sngl r" |
|
138 |
by(cases r rule: sngl.cases) clarsimp+ |
|
139 |
||
140 |
||
141 |
subsection "Invariance" |
|
142 |
||
143 |
subsubsection "Proofs for insert" |
|
144 |
||
145 |
lemma lvl_insert_aux: |
|
146 |
"lvl (insert x t) = lvl t \<or> lvl (insert x t) = lvl t + 1 \<and> sngl (insert x t)" |
|
147 |
apply(induction t) |
|
148 |
apply (auto simp: lvl_skew) |
|
149 |
apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+ |
|
150 |
done |
|
151 |
||
152 |
lemma lvl_insert: obtains |
|
153 |
(Same) "lvl (insert x t) = lvl t" | |
|
154 |
(Incr) "lvl (insert x t) = lvl t + 1" "sngl (insert x t)" |
|
155 |
using lvl_insert_aux by blast |
|
156 |
||
157 |
lemma lvl_insert_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(insert x t) = lvl t" |
|
62526 | 158 |
proof (induction t rule: insert.induct) |
62496 | 159 |
case (2 x lv t1 a t2) |
160 |
consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" |
|
161 |
using less_linear by blast |
|
162 |
thus ?case proof cases |
|
163 |
case LT |
|
164 |
thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits) |
|
165 |
next |
|
166 |
case GT |
|
167 |
thus ?thesis using 2 proof (cases t1) |
|
168 |
case Node |
|
169 |
thus ?thesis using 2 GT |
|
170 |
apply (auto simp add: skew_case split_case split: tree.splits) |
|
171 |
by (metis less_not_refl2 lvl.simps(2) lvl_insert_aux n_not_Suc_n sngl.simps(3))+ |
|
172 |
qed (auto simp add: lvl_0_iff) |
|
173 |
qed simp |
|
174 |
qed simp |
|
175 |
||
176 |
lemma skew_invar: "invar t \<Longrightarrow> skew t = t" |
|
62526 | 177 |
by(cases t rule: skew.cases) auto |
62496 | 178 |
|
179 |
lemma split_invar: "invar t \<Longrightarrow> split t = t" |
|
62526 | 180 |
by(cases t rule: split.cases) clarsimp+ |
62496 | 181 |
|
182 |
lemma invar_NodeL: |
|
183 |
"\<lbrakk> invar(Node n l x r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node n l' x r)" |
|
184 |
by(auto) |
|
185 |
||
186 |
lemma invar_NodeR: |
|
187 |
"\<lbrakk> invar(Node n l x r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node n l x r')" |
|
188 |
by(auto) |
|
189 |
||
190 |
lemma invar_NodeR2: |
|
191 |
"\<lbrakk> invar(Node n l x r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> invar(Node n l x r')" |
|
192 |
by(cases r' rule: sngl.cases) clarsimp+ |
|
193 |
||
194 |
||
195 |
lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \<longleftrightarrow> |
|
67613 | 196 |
(\<exists>l x r. insert a t = Node (lvl t + 1) l x r \<and> lvl l = lvl r)" |
62496 | 197 |
apply(cases t) |
198 |
apply(auto simp add: skew_case split_case split: if_splits) |
|
199 |
apply(auto split: tree.splits if_splits) |
|
200 |
done |
|
201 |
||
202 |
lemma invar_insert: "invar t \<Longrightarrow> invar(insert a t)" |
|
203 |
proof(induction t) |
|
67040 | 204 |
case N: (Node n l x r) |
62496 | 205 |
hence il: "invar l" and ir: "invar r" by auto |
67040 | 206 |
note iil = N.IH(1)[OF il] |
207 |
note iir = N.IH(2)[OF ir] |
|
62496 | 208 |
let ?t = "Node n l x r" |
209 |
have "a < x \<or> a = x \<or> x < a" by auto |
|
210 |
moreover |
|
67040 | 211 |
have ?case if "a < x" |
212 |
proof (cases rule: lvl_insert[of a l]) |
|
213 |
case (Same) thus ?thesis |
|
214 |
using \<open>a<x\<close> invar_NodeL[OF N.prems iil Same] |
|
215 |
by (simp add: skew_invar split_invar del: invar.simps) |
|
216 |
next |
|
217 |
case (Incr) |
|
218 |
then obtain t1 w t2 where ial[simp]: "insert a l = Node n t1 w t2" |
|
219 |
using N.prems by (auto simp: lvl_Suc_iff) |
|
220 |
have l12: "lvl t1 = lvl t2" |
|
221 |
by (metis Incr(1) ial lvl_insert_incr_iff tree.inject) |
|
222 |
have "insert a ?t = split(skew(Node n (insert a l) x r))" |
|
223 |
by(simp add: \<open>a<x\<close>) |
|
224 |
also have "skew(Node n (insert a l) x r) = Node n t1 w (Node n t2 x r)" |
|
225 |
by(simp) |
|
226 |
also have "invar(split \<dots>)" |
|
227 |
proof (cases r) |
|
228 |
case Leaf |
|
229 |
hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff) |
|
230 |
thus ?thesis using Leaf ial by simp |
|
62496 | 231 |
next |
67040 | 232 |
case [simp]: (Node m t3 y t4) |
233 |
show ?thesis (*using N(3) iil l12 by(auto)*) |
|
234 |
proof cases |
|
235 |
assume "m = n" thus ?thesis using N(3) iil by(auto) |
|
62496 | 236 |
next |
67040 | 237 |
assume "m \<noteq> n" thus ?thesis using N(3) iil l12 by(auto) |
62496 | 238 |
qed |
239 |
qed |
|
67040 | 240 |
finally show ?thesis . |
241 |
qed |
|
62496 | 242 |
moreover |
67040 | 243 |
have ?case if "x < a" |
244 |
proof - |
|
62496 | 245 |
from \<open>invar ?t\<close> have "n = lvl r \<or> n = lvl r + 1" by auto |
67040 | 246 |
thus ?case |
62496 | 247 |
proof |
248 |
assume 0: "n = lvl r" |
|
249 |
have "insert a ?t = split(skew(Node n l x (insert a r)))" |
|
250 |
using \<open>a>x\<close> by(auto) |
|
251 |
also have "skew(Node n l x (insert a r)) = Node n l x (insert a r)" |
|
67040 | 252 |
using N.prems by(simp add: skew_case split: tree.split) |
62496 | 253 |
also have "invar(split \<dots>)" |
254 |
proof - |
|
255 |
from lvl_insert_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a] |
|
256 |
obtain t1 y t2 where iar: "insert a r = Node n t1 y t2" |
|
67040 | 257 |
using N.prems 0 by (auto simp: lvl_Suc_iff) |
258 |
from N.prems iar 0 iir |
|
62496 | 259 |
show ?thesis by (auto simp: split_case split: tree.splits) |
260 |
qed |
|
261 |
finally show ?thesis . |
|
262 |
next |
|
263 |
assume 1: "n = lvl r + 1" |
|
264 |
hence "sngl ?t" by(cases r) auto |
|
265 |
show ?thesis |
|
266 |
proof (cases rule: lvl_insert[of a r]) |
|
267 |
case (Same) |
|
67040 | 268 |
show ?thesis using \<open>x<a\<close> il ir invar_NodeR[OF N.prems 1 iir Same] |
62496 | 269 |
by (auto simp add: skew_invar split_invar) |
270 |
next |
|
271 |
case (Incr) |
|
67406 | 272 |
thus ?thesis using invar_NodeR2[OF \<open>invar ?t\<close> Incr(2) 1 iir] 1 \<open>x < a\<close> |
62496 | 273 |
by (auto simp add: skew_invar split_invar split: if_splits) |
274 |
qed |
|
275 |
qed |
|
67040 | 276 |
qed |
277 |
moreover |
|
278 |
have "a = x \<Longrightarrow> ?case" using N.prems by auto |
|
62496 | 279 |
ultimately show ?case by blast |
280 |
qed simp |
|
281 |
||
282 |
||
283 |
subsubsection "Proofs for delete" |
|
284 |
||
285 |
lemma invarL: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar l" |
|
286 |
by(simp add: ASSUMPTION_def) |
|
287 |
||
288 |
lemma invarR: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar r" |
|
289 |
by(simp add: ASSUMPTION_def) |
|
290 |
||
291 |
lemma sngl_NodeI: |
|
292 |
"sngl (Node lv l a r) \<Longrightarrow> sngl (Node lv l' a' r)" |
|
293 |
by(cases r) (simp_all) |
|
294 |
||
295 |
||
296 |
declare invarL[simp] invarR[simp] |
|
297 |
||
298 |
lemma pre_cases: |
|
299 |
assumes "pre_adjust (Node lv l x r)" |
|
300 |
obtains |
|
301 |
(tSngl) "invar l \<and> invar r \<and> |
|
302 |
lv = Suc (lvl r) \<and> lvl l = lvl r" | |
|
303 |
(tDouble) "invar l \<and> invar r \<and> |
|
304 |
lv = lvl r \<and> Suc (lvl l) = lvl r \<and> sngl r " | |
|
305 |
(rDown) "invar l \<and> invar r \<and> |
|
306 |
lv = Suc (Suc (lvl r)) \<and> lv = Suc (lvl l)" | |
|
307 |
(lDown_tSngl) "invar l \<and> invar r \<and> |
|
308 |
lv = Suc (lvl r) \<and> lv = Suc (Suc (lvl l))" | |
|
309 |
(lDown_tDouble) "invar l \<and> invar r \<and> |
|
310 |
lv = lvl r \<and> lv = Suc (Suc (lvl l)) \<and> sngl r" |
|
311 |
using assms unfolding pre_adjust.simps |
|
312 |
by auto |
|
313 |
||
314 |
declare invar.simps(2)[simp del] invar_2Nodes[simp add] |
|
315 |
||
316 |
lemma invar_adjust: |
|
317 |
assumes pre: "pre_adjust (Node lv l a r)" |
|
318 |
shows "invar(adjust (Node lv l a r))" |
|
319 |
using pre proof (cases rule: pre_cases) |
|
320 |
case (tDouble) thus ?thesis unfolding adjust_def by (cases r) (auto simp: invar.simps(2)) |
|
321 |
next |
|
322 |
case (rDown) |
|
323 |
from rDown obtain llv ll la lr where l: "l = Node llv ll la lr" by (cases l) auto |
|
324 |
from rDown show ?thesis unfolding adjust_def by (auto simp: l invar.simps(2) split: tree.splits) |
|
325 |
next |
|
326 |
case (lDown_tDouble) |
|
327 |
from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rlv rl ra rr" by (cases r) auto |
|
328 |
from lDown_tDouble and r obtain rrlv rrr rra rrl where |
|
329 |
rr :"rr = Node rrlv rrr rra rrl" by (cases rr) auto |
|
330 |
from lDown_tDouble show ?thesis unfolding adjust_def r rr |
|
63636 | 331 |
apply (cases rl) apply (auto simp add: invar.simps(2) split!: if_split) |
62496 | 332 |
using lDown_tDouble by (auto simp: split_case lvl_0_iff elim:lvl.elims split: tree.split) |
333 |
qed (auto simp: split_case invar.simps(2) adjust_def split: tree.splits) |
|
334 |
||
335 |
lemma lvl_adjust: |
|
336 |
assumes "pre_adjust (Node lv l a r)" |
|
337 |
shows "lv = lvl (adjust(Node lv l a r)) \<or> lv = lvl (adjust(Node lv l a r)) + 1" |
|
338 |
using assms(1) proof(cases rule: pre_cases) |
|
339 |
case lDown_tSngl thus ?thesis |
|
340 |
using lvl_split[of "\<langle>lvl r, l, a, r\<rangle>"] by (auto simp: adjust_def) |
|
341 |
next |
|
342 |
case lDown_tDouble thus ?thesis |
|
343 |
by (auto simp: adjust_def invar.simps(2) split: tree.split) |
|
344 |
qed (auto simp: adjust_def split: tree.splits) |
|
345 |
||
346 |
lemma sngl_adjust: assumes "pre_adjust (Node lv l a r)" |
|
347 |
"sngl \<langle>lv, l, a, r\<rangle>" "lv = lvl (adjust \<langle>lv, l, a, r\<rangle>)" |
|
348 |
shows "sngl (adjust \<langle>lv, l, a, r\<rangle>)" |
|
349 |
using assms proof (cases rule: pre_cases) |
|
350 |
case rDown |
|
351 |
thus ?thesis using assms(2,3) unfolding adjust_def |
|
352 |
by (auto simp add: skew_case) (auto split: tree.split) |
|
353 |
qed (auto simp: adjust_def skew_case split_case split: tree.split) |
|
354 |
||
355 |
definition "post_del t t' == |
|
356 |
invar t' \<and> |
|
357 |
(lvl t' = lvl t \<or> lvl t' + 1 = lvl t) \<and> |
|
358 |
(lvl t' = lvl t \<and> sngl t \<longrightarrow> sngl t')" |
|
359 |
||
360 |
lemma pre_adj_if_postR: |
|
361 |
"invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del r r' \<Longrightarrow> pre_adjust \<langle>lv, l, a, r'\<rangle>" |
|
362 |
by(cases "sngl r") |
|
363 |
(auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) |
|
364 |
||
365 |
lemma pre_adj_if_postL: |
|
366 |
"invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>lv, l', b, r\<rangle>" |
|
367 |
by(cases "sngl r") |
|
368 |
(auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) |
|
369 |
||
370 |
lemma post_del_adjL: |
|
371 |
"\<lbrakk> invar\<langle>lv, l, a, r\<rangle>; pre_adjust \<langle>lv, l', b, r\<rangle> \<rbrakk> |
|
372 |
\<Longrightarrow> post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l', b, r\<rangle>)" |
|
373 |
unfolding post_del_def |
|
374 |
by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2)) |
|
375 |
||
376 |
lemma post_del_adjR: |
|
377 |
assumes "invar\<langle>lv, l, a, r\<rangle>" "pre_adjust \<langle>lv, l, a, r'\<rangle>" "post_del r r'" |
|
378 |
shows "post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l, a, r'\<rangle>)" |
|
379 |
proof(unfold post_del_def, safe del: disjCI) |
|
380 |
let ?t = "\<langle>lv, l, a, r\<rangle>" |
|
381 |
let ?t' = "adjust \<langle>lv, l, a, r'\<rangle>" |
|
382 |
show "invar ?t'" by(rule invar_adjust[OF assms(2)]) |
|
383 |
show "lvl ?t' = lvl ?t \<or> lvl ?t' + 1 = lvl ?t" |
|
384 |
using lvl_adjust[OF assms(2)] by auto |
|
385 |
show "sngl ?t'" if as: "lvl ?t' = lvl ?t" "sngl ?t" |
|
386 |
proof - |
|
387 |
have s: "sngl \<langle>lv, l, a, r'\<rangle>" |
|
388 |
proof(cases r') |
|
389 |
case Leaf thus ?thesis by simp |
|
390 |
next |
|
391 |
case Node thus ?thesis using as(2) assms(1,3) |
|
392 |
by (cases r) (auto simp: post_del_def) |
|
393 |
qed |
|
394 |
show ?thesis using as(1) sngl_adjust[OF assms(2) s] by simp |
|
395 |
qed |
|
396 |
qed |
|
397 |
||
398 |
declare prod.splits[split] |
|
399 |
||
400 |
theorem post_del_max: |
|
401 |
"\<lbrakk> invar t; (t', x) = del_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'" |
|
402 |
proof (induction t arbitrary: t' rule: del_max.induct) |
|
403 |
case (2 lv l a lvr rl ra rr) |
|
404 |
let ?r = "\<langle>lvr, rl, ra, rr\<rangle>" |
|
405 |
let ?t = "\<langle>lv, l, a, ?r\<rangle>" |
|
406 |
from "2.prems"(2) obtain r' where r': "(r', x) = del_max ?r" |
|
407 |
and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto |
|
408 |
from "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp |
|
409 |
note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> post] |
|
410 |
show ?case by (simp add: post_del_adjR[OF "2.prems"(1) preR post]) |
|
411 |
qed (auto simp: post_del_def) |
|
412 |
||
413 |
theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)" |
|
414 |
proof (induction t) |
|
415 |
case (Node lv l a r) |
|
416 |
||
417 |
let ?l' = "delete x l" and ?r' = "delete x r" |
|
418 |
let ?t = "Node lv l a r" let ?t' = "delete x ?t" |
|
419 |
||
420 |
from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) |
|
421 |
||
422 |
note post_l' = Node.IH(1)[OF inv_l] |
|
423 |
note preL = pre_adj_if_postL[OF Node.prems post_l'] |
|
424 |
||
425 |
note post_r' = Node.IH(2)[OF inv_r] |
|
426 |
note preR = pre_adj_if_postR[OF Node.prems post_r'] |
|
427 |
||
428 |
show ?case |
|
429 |
proof (cases rule: linorder_cases[of x a]) |
|
430 |
case less |
|
431 |
thus ?thesis using Node.prems by (simp add: post_del_adjL preL) |
|
432 |
next |
|
433 |
case greater |
|
434 |
thus ?thesis using Node.prems by (simp add: post_del_adjR preR post_r') |
|
435 |
next |
|
436 |
case equal |
|
437 |
show ?thesis |
|
438 |
proof cases |
|
439 |
assume "l = Leaf" thus ?thesis using equal Node.prems |
|
440 |
by(auto simp: post_del_def invar.simps(2)) |
|
441 |
next |
|
442 |
assume "l \<noteq> Leaf" thus ?thesis using equal |
|
443 |
by simp (metis Node.prems inv_l post_del_adjL post_del_max pre_adj_if_postL) |
|
444 |
qed |
|
445 |
qed |
|
446 |
qed (simp add: post_del_def) |
|
447 |
||
448 |
declare invar_2Nodes[simp del] |
|
449 |
||
61793 | 450 |
|
451 |
subsection "Functional Correctness" |
|
452 |
||
62496 | 453 |
|
61793 | 454 |
subsubsection "Proofs for insert" |
455 |
||
456 |
lemma inorder_split: "inorder(split t) = inorder t" |
|
457 |
by(cases t rule: split.cases) (auto) |
|
458 |
||
459 |
lemma inorder_skew: "inorder(skew t) = inorder t" |
|
460 |
by(cases t rule: skew.cases) (auto) |
|
461 |
||
462 |
lemma inorder_insert: |
|
463 |
"sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)" |
|
464 |
by(induction t) (auto simp: ins_list_simps inorder_split inorder_skew) |
|
465 |
||
62496 | 466 |
|
61793 | 467 |
subsubsection "Proofs for delete" |
468 |
||
62496 | 469 |
lemma inorder_adjust: "t \<noteq> Leaf \<Longrightarrow> pre_adjust t \<Longrightarrow> inorder(adjust t) = inorder t" |
62526 | 470 |
by(cases t) |
62496 | 471 |
(auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps |
472 |
split: tree.splits) |
|
473 |
||
61793 | 474 |
lemma del_maxD: |
62496 | 475 |
"\<lbrakk> del_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t" |
61793 | 476 |
by(induction t arbitrary: t' rule: del_max.induct) |
62496 | 477 |
(auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_del_max split: prod.splits) |
61793 | 478 |
|
479 |
lemma inorder_delete: |
|
62496 | 480 |
"invar t \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" |
61793 | 481 |
by(induction t) |
62496 | 482 |
(auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR |
483 |
post_del_max post_delete del_maxD split: prod.splits) |
|
61793 | 484 |
|
62496 | 485 |
interpretation I: Set_by_Ordered |
61793 | 486 |
where empty = Leaf and isin = isin and insert = insert and delete = delete |
62496 | 487 |
and inorder = inorder and inv = invar |
61793 | 488 |
proof (standard, goal_cases) |
489 |
case 1 show ?case by simp |
|
490 |
next |
|
67967 | 491 |
case 2 thus ?case by(simp add: isin_set_inorder) |
61793 | 492 |
next |
493 |
case 3 thus ?case by(simp add: inorder_insert) |
|
494 |
next |
|
495 |
case 4 thus ?case by(simp add: inorder_delete) |
|
62496 | 496 |
next |
497 |
case 5 thus ?case by(simp) |
|
498 |
next |
|
499 |
case 6 thus ?case by(simp add: invar_insert) |
|
500 |
next |
|
501 |
case 7 thus ?case using post_delete by(auto simp: post_del_def) |
|
502 |
qed |
|
61793 | 503 |
|
62390 | 504 |
end |