section ‹Red-Black Tree Implementation of Sets›
theory RBT_Set
imports
RBT
Cmp
Isin2
begin
fun ins :: "'a::cmp ⇒ 'a rbt ⇒ 'a rbt" where
"ins x Leaf = R Leaf x Leaf" |
"ins x (B l a r) =
(case cmp x a of
LT ⇒ bal (ins x l) a r |
GT ⇒ bal l a (ins x r) |
EQ ⇒ B l a r)" |
"ins x (R l a r) =
(case cmp x a of
LT ⇒ R (ins x l) a r |
GT ⇒ R l a (ins x r) |
EQ ⇒ R l a r)"
definition insert :: "'a::cmp ⇒ 'a rbt ⇒ 'a rbt" where
"insert x t = paint Black (ins x t)"
fun del :: "'a::cmp ⇒ 'a rbt ⇒ 'a rbt"
and delL :: "'a::cmp ⇒ 'a rbt ⇒ 'a ⇒ 'a rbt ⇒ 'a rbt"
and delR :: "'a::cmp ⇒ 'a rbt ⇒ 'a ⇒ 'a rbt ⇒ 'a rbt"
where
"del x Leaf = Leaf" |
"del x (Node _ l a r) =
(case cmp x a of
LT ⇒ delL x l a r |
GT ⇒ delR x l a r |
EQ ⇒ combine l r)" |
"delL x (B t1 a t2) b t3 = balL (del x (B t1 a t2)) b t3" |
"delL x l a r = R (del x l) a r" |
"delR x t1 a (B t2 b t3) = balR t1 a (del x (B t2 b t3))" |
"delR x l a r = R l a (del x r)"
definition delete :: "'a::cmp ⇒ 'a rbt ⇒ 'a rbt" where
"delete x t = paint Black (del x t)"
subsection "Functional Correctness Proofs"
lemma inorder_paint: "inorder(paint c t) = inorder t"
by(induction t) (auto)
lemma inorder_bal:
"inorder(bal l a r) = inorder l @ a # inorder r"
by(induction l a r rule: bal.induct) (auto)
lemma inorder_ins:
"sorted(inorder t) ⟹ inorder(ins x t) = ins_list x (inorder t)"
by(induction x t rule: ins.induct) (auto simp: ins_list_simps inorder_bal)
lemma inorder_insert:
"sorted(inorder t) ⟹ inorder(insert x t) = ins_list x (inorder t)"
by (simp add: insert_def inorder_ins inorder_paint)
lemma inorder_balL:
"inorder(balL l a r) = inorder l @ a # inorder r"
by(induction l a r rule: balL.induct)(auto simp: inorder_bal inorder_paint)
lemma inorder_balR:
"inorder(balR l a r) = inorder l @ a # inorder r"
by(induction l a r rule: balR.induct) (auto simp: inorder_bal inorder_paint)
lemma inorder_combine:
"inorder(combine l r) = inorder l @ inorder r"
by(induction l r rule: combine.induct)
(auto simp: inorder_balL inorder_balR split: tree.split color.split)
lemma inorder_del:
"sorted(inorder t) ⟹ inorder(del x t) = del_list x (inorder t)"
"sorted(inorder l) ⟹ inorder(delL x l a r) =
del_list x (inorder l) @ a # inorder r"
"sorted(inorder r) ⟹ inorder(delR x l a r) =
inorder l @ a # del_list x (inorder r)"
by(induction x t and x l a r and x l a r rule: del_delL_delR.induct)
(auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)
lemma inorder_delete:
"sorted(inorder t) ⟹ inorder(delete x t) = del_list x (inorder t)"
by (auto simp: delete_def inorder_del inorder_paint)
interpretation Set_by_Ordered
where empty = Leaf and isin = isin and insert = insert and delete = delete
and inorder = inorder and inv = "λ_. True"
proof (standard, goal_cases)
case 1 show ?case by simp
next
case 2 thus ?case by(simp add: isin_set)
next
case 3 thus ?case by(simp add: inorder_insert)
next
case 4 thus ?case by(simp add: inorder_delete)
qed auto
subsection ‹Structural invariants›
fun color :: "'a rbt ⇒ color" where
"color Leaf = Black" |
"color (Node c _ _ _) = c"
fun bheight :: "'a rbt ⇒ nat" where
"bheight Leaf = 0" |
"bheight (Node c l x r) = (if c = Black then Suc(bheight l) else bheight l)"
fun inv1 :: "'a rbt ⇒ bool" where
"inv1 Leaf = True" |
"inv1 (Node c l a r) =
(inv1 l ∧ inv1 r ∧ (c = Black ∨ color l = Black ∧ color r = Black))"
fun inv1_root :: "'a rbt ⇒ bool" ― ‹Weaker version› where
"inv1_root Leaf = True" |
"inv1_root (Node c l a r) = (inv1 l ∧ inv1 r)"
fun inv2 :: "'a rbt ⇒ bool" where
"inv2 Leaf = True" |
"inv2 (Node c l x r) = (inv2 l ∧ inv2 r ∧ bheight l = bheight r)"
lemma inv1_rootI[simp]: "inv1 t ⟹ inv1_root t"
by (cases t) simp+
definition rbt :: "'a rbt ⇒ bool" where
"rbt t = (inv1 t ∧ inv2 t ∧ color t = Black)"
lemma color_paint_Black: "color (paint Black t) = Black"
by (cases t) auto
theorem rbt_Leaf: "rbt Leaf"
by (simp add: rbt_def)
lemma paint_inv1_root: "inv1_root t ⟹ inv1_root (paint c t)"
by (cases t) auto
lemma inv1_paint_Black: "inv1_root t ⟹ inv1 (paint Black t)"
by (cases t) auto
lemma inv2_paint: "inv2 t ⟹ inv2 (paint c t)"
by (cases t) auto
lemma inv1_bal: "⟦inv1_root l; inv1_root r⟧ ⟹ inv1 (bal l a r)"
by (induct l a r rule: bal.induct) auto
lemma bheight_bal:
"bheight l = bheight r ⟹ bheight (bal l a r) = Suc (bheight l)"
by (induct l a r rule: bal.induct) auto
lemma inv2_bal:
"⟦ inv2 l; inv2 r; bheight l = bheight r ⟧ ⟹ inv2 (bal l a r)"
by (induct l a r rule: bal.induct) auto
subsubsection ‹Insertion›
lemma inv1_ins: assumes "inv1 t"
shows "color t = Black ⟹ inv1 (ins x t)" "inv1_root (ins x t)"
using assms
by (induct x t rule: ins.induct) (auto simp: inv1_bal)
lemma inv2_ins: assumes "inv2 t"
shows "inv2 (ins x t)" "bheight (ins x t) = bheight t"
using assms
by (induct x t rule: ins.induct) (auto simp: inv2_bal bheight_bal)
theorem rbt_ins: "rbt t ⟹ rbt (insert x t)"
by (simp add: inv1_ins inv2_ins color_paint_Black inv1_paint_Black inv2_paint
rbt_def insert_def)
end