src/HOL/Data_Structures/RBT_Map.thy
author wenzelm
Mon, 23 Nov 2015 16:57:01 +0100
changeset 61737 b91b1ebfc8a0
parent 61686 e6784939d645
child 61749 7f530d7e552d
permissions -rw-r--r--
bundle main sources read-only, to avoid accidental editing of imported theories etc.;

(* Author: Tobias Nipkow *)

section \<open>Red-Black Tree Implementation of Maps\<close>

theory RBT_Map
imports
  RBT_Set
  Lookup2
begin

fun update :: "'a::cmp \<Rightarrow> 'b \<Rightarrow> ('a*'b) rbt \<Rightarrow> ('a*'b) rbt" where
"update x y Leaf = R Leaf (x,y) Leaf" |
"update x y (B l (a,b) r) = (case cmp x a of
  LT \<Rightarrow> bal (update x y l) (a,b) r |
  GT \<Rightarrow> bal l (a,b) (update x y r) |
  EQ \<Rightarrow> B l (x,y) r)" |
"update x y (R l (a,b) r) = (case cmp x a of
  LT \<Rightarrow> R (update x y l) (a,b) r |
  GT \<Rightarrow> R l (a,b) (update x y r) |
  EQ \<Rightarrow> R l (x,y) r)"

fun delete :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
and deleteL :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
and deleteR :: "'a::cmp \<Rightarrow> ('a*'b)rbt \<Rightarrow> 'a*'b \<Rightarrow> ('a*'b)rbt \<Rightarrow> ('a*'b)rbt"
where
"delete x Leaf = Leaf" |
"delete x (Node c t1 (a,b) t2) = (case cmp x a of
  LT \<Rightarrow> deleteL x t1 (a,b) t2 |
  GT \<Rightarrow> deleteR x t1 (a,b) t2 |
  EQ \<Rightarrow> combine t1 t2)" |
"deleteL x (B t1 a t2) b t3 = balL (delete x (B t1 a t2)) b t3" |
"deleteL x t1 a t2 = R (delete x t1) a t2" |
"deleteR x t1 a (B t2 b t3) = balR t1 a (delete x (B t2 b t3))" | 
"deleteR x t1 a t2 = R t1 a (delete x t2)"


subsection "Functional Correctness Proofs"

lemma inorder_update:
  "sorted1(inorder t) \<Longrightarrow> inorder(update x y t) = upd_list x y (inorder t)"
by(induction x y t rule: update.induct)
  (auto simp: upd_list_simps inorder_bal)


lemma inorder_delete:
 "sorted1(inorder t1) \<Longrightarrow>  inorder(delete x t1) = del_list x (inorder t1)" and
 "sorted1(inorder t1) \<Longrightarrow>  inorder(deleteL x t1 a t2) =
    del_list x (inorder t1) @ a # inorder t2" and
 "sorted1(inorder t2) \<Longrightarrow>  inorder(deleteR x t1 a t2) =
    inorder t1 @ a # del_list x (inorder t2)"
by(induction x t1 and x t1 a t2 and x t1 a t2 rule: delete_deleteL_deleteR.induct)
  (auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)

interpretation Map_by_Ordered
where empty = Leaf and lookup = lookup and update = update and delete = delete
and inorder = inorder and inv = "\<lambda>_. True"
proof (standard, goal_cases)
  case 1 show ?case by simp
next
  case 2 thus ?case by(simp add: lookup_eq)
next
  case 3 thus ?case by(simp add: inorder_update)
next
  case 4 thus ?case by(simp add: inorder_delete)
qed auto

end