61224
|
1 |
(* Author: Tobias Nipkow *)
|
|
2 |
|
|
3 |
section \<open>Red-Black Tree Implementation of Maps\<close>
|
|
4 |
|
|
5 |
theory RBT_Map
|
|
6 |
imports
|
|
7 |
RBT_Set
|
61231
|
8 |
Lookup2
|
61224
|
9 |
begin
|
|
10 |
|
61581
|
11 |
fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
|
61224
|
12 |
"update x y Leaf = R Leaf (x,y) Leaf" |
|
61581
|
13 |
"update x y (B l (a,b) r) = (case cmp x a of
|
|
14 |
LT \<Rightarrow> bal (update x y l) (a,b) r |
|
|
15 |
GT \<Rightarrow> bal l (a,b) (update x y r) |
|
|
16 |
EQ \<Rightarrow> B l (x,y) r)" |
|
|
17 |
"update x y (R l (a,b) r) = (case cmp x a of
|
|
18 |
LT \<Rightarrow> R (update x y l) (a,b) r |
|
|
19 |
GT \<Rightarrow> R l (a,b) (update x y r) |
|
|
20 |
EQ \<Rightarrow> R l (x,y) r)"
|
61224
|
21 |
|
61581
|
22 |
fun delete :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
|
|
23 |
and deleteL :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
|
|
24 |
and deleteR :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
|
61224
|
25 |
where
|
|
26 |
"delete x Leaf = Leaf" |
|
61581
|
27 |
"delete x (Node c t1 (a,b) t2) = (case cmp x a of
|
|
28 |
LT \<Rightarrow> deleteL x t1 (a,b) t2 |
|
|
29 |
GT \<Rightarrow> deleteR x t1 (a,b) t2 |
|
|
30 |
EQ \<Rightarrow> combine t1 t2)" |
|
61224
|
31 |
"deleteL x (B t1 a t2) b t3 = balL (delete x (B t1 a t2)) b t3" |
|
|
32 |
"deleteL x t1 a t2 = R (delete x t1) a t2" |
|
|
33 |
"deleteR x t1 a (B t2 b t3) = balR t1 a (delete x (B t2 b t3))" |
|
|
34 |
"deleteR x t1 a t2 = R t1 a (delete x t2)"
|
|
35 |
|
|
36 |
|
|
37 |
subsection "Functional Correctness Proofs"
|
|
38 |
|
|
39 |
lemma inorder_update:
|
|
40 |
"sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
|
|
41 |
by(induction x y t rule: update.induct)
|
|
42 |
(auto simp: upd_list_simps inorder_bal)
|
|
43 |
|
|
44 |
|
|
45 |
lemma inorder_delete:
|
|
46 |
"sorted1(inorder t1) \<Longrightarrow> inorder(delete x t1) = del_list x (inorder t1)" and
|
|
47 |
"sorted1(inorder t1) \<Longrightarrow> inorder(deleteL x t1 a t2) =
|
|
48 |
del_list x (inorder t1) @ a # inorder t2" and
|
|
49 |
"sorted1(inorder t2) \<Longrightarrow> inorder(deleteR x t1 a t2) =
|
|
50 |
inorder t1 @ a # del_list x (inorder t2)"
|
|
51 |
by(induction x t1 and x t1 a t2 and x t1 a t2 rule: delete_deleteL_deleteR.induct)
|
61231
|
52 |
(auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)
|
61224
|
53 |
|
|
54 |
interpretation Map_by_Ordered
|
|
55 |
where empty = Leaf and lookup = lookup and update = update and delete = delete
|
61686
|
56 |
and inorder = inorder and inv = "\<lambda>_. True"
|
61224
|
57 |
proof (standard, goal_cases)
|
|
58 |
case 1 show ?case by simp
|
|
59 |
next
|
|
60 |
case 2 thus ?case by(simp add: lookup_eq)
|
|
61 |
next
|
|
62 |
case 3 thus ?case by(simp add: inorder_update)
|
|
63 |
next
|
|
64 |
case 4 thus ?case by(simp add: inorder_delete)
|
61686
|
65 |
qed auto
|
61224
|
66 |
|
|
67 |
end
|