# HG changeset patch # User haftmann # Date 1267601313 -3600 # Node ID 14d8d72f8b1fd1836471e3ec81d981ef60612738 # Parent 5d88ffdb290c0d5ae3b9498950ba51c99c7683ce more explicit naming scheme diff -r 5d88ffdb290c -r 14d8d72f8b1f src/HOL/Library/RBT.thy --- a/src/HOL/Library/RBT.thy Tue Mar 02 15:39:15 2010 +0100 +++ b/src/HOL/Library/RBT.thy Wed Mar 03 08:28:33 2010 +0100 @@ -11,135 +11,151 @@ begin datatype color = R | B -datatype ('a,'b)"rbt" = Empty | Tr color "('a,'b)rbt" 'a 'b "('a,'b)rbt" +datatype ('a, 'b) rbt = Empty | Branch color "('a, 'b) rbt" 'a 'b "('a, 'b) rbt" + +lemma rbt_cases: + obtains (Empty) "t = Empty" + | (Red) l k v r where "t = Branch R l k v r" + | (Black) l k v r where "t = Branch B l k v r" +proof (cases t) + case Empty with that show thesis by blast +next + case (Branch c) with that show thesis by (cases c) blast+ +qed + +text {* Content of a tree *} + +primrec entries +where + "entries Empty = []" +| "entries (Branch _ l k v r) = entries l @ (k,v) # entries r" text {* Search tree properties *} -primrec - pin_tree :: "'a \ 'b \ ('a,'b) rbt \ bool" +primrec entry_in_tree :: "'a \ 'b \ ('a, 'b) rbt \ bool" where - "pin_tree k v Empty = False" -| "pin_tree k v (Tr c l x y r) = (k = x \ v = y \ pin_tree k v l \ pin_tree k v r)" + "entry_in_tree k v Empty = False" +| "entry_in_tree k v (Branch c l x y r) \ k = x \ v = y \ entry_in_tree k v l \ entry_in_tree k v r" -primrec - keys :: "('k,'v) rbt \ 'k set" +primrec keys :: "('k, 'v) rbt \ 'k set" where "keys Empty = {}" -| "keys (Tr _ l k _ r) = { k } \ keys l \ keys r" +| "keys (Branch _ l k _ r) = { k } \ keys l \ keys r" -lemma pint_keys: "pin_tree k v t \ k \ keys t" by (induct t) auto +lemma entry_in_tree_keys: + "entry_in_tree k v t \ k \ keys t" + by (induct t) auto -primrec tlt :: "'a\order \ ('a,'b) rbt \ bool" +definition tree_less :: "'a\order \ ('a, 'b) rbt \ bool" where - "tlt k Empty = True" -| "tlt k (Tr c lt kt v rt) = (kt < k \ tlt k lt \ tlt k rt)" + tree_less_prop: "tree_less k t \ (\x\keys t. x < k)" + +abbreviation tree_less_symbol (infix "|\" 50) +where "t |\ x \ tree_less x t" -abbreviation tllt (infix "|\" 50) -where "t |\ x == tlt x t" +definition tree_greater :: "'a\order \ ('a, 'b) rbt \ bool" (infix "\|" 50) +where + tree_greater_prop: "tree_greater k t = (\x\keys t. k < x)" -primrec tgt :: "'a\order \ ('a,'b) rbt \ bool" (infix "\|" 50) -where - "tgt k Empty = True" -| "tgt k (Tr c lt kt v rt) = (k < kt \ tgt k lt \ tgt k rt)" +lemma tree_less_simps [simp]: + "tree_less k Empty = True" + "tree_less k (Branch c lt kt v rt) \ kt < k \ tree_less k lt \ tree_less k rt" + by (auto simp add: tree_less_prop) -lemma tlt_prop: "(t |\ k) = (\x\keys t. x < k)" by (induct t) auto -lemma tgt_prop: "(k \| t) = (\x\keys t. k < x)" by (induct t) auto -lemmas tlgt_props = tlt_prop tgt_prop +lemma tree_greater_simps [simp]: + "tree_greater k Empty = True" + "tree_greater k (Branch c lt kt v rt) \ k < kt \ tree_greater k lt \ tree_greater k rt" + by (auto simp add: tree_greater_prop) -lemmas tgt_nit = tgt_prop pint_keys -lemmas tlt_nit = tlt_prop pint_keys +lemmas tree_ord_props = tree_less_prop tree_greater_prop -lemma tlt_trans: "\ t |\ x; x < y \ \ t |\ y" - and tgt_trans: "\ x < y; y \| t\ \ x \| t" -by (auto simp: tlgt_props) - +lemmas tree_greater_nit = tree_greater_prop entry_in_tree_keys +lemmas tree_less_nit = tree_less_prop entry_in_tree_keys -primrec st :: "('a::linorder, 'b) rbt \ bool" -where - "st Empty = True" -| "st (Tr c l k v r) = (l |\ k \ k \| r \ st l \ st r)" +lemma tree_less_trans: "t |\ x \ x < y \ t |\ y" + and tree_greater_trans: "x < y \ y \| t \ x \| t" +by (auto simp: tree_ord_props) -primrec map_of :: "('a\linorder, 'b) rbt \ 'a \ 'b" +primrec sorted :: "('a::linorder, 'b) rbt \ bool" where - "map_of Empty k = None" -| "map_of (Tr _ l x y r) k = (if k < x then map_of l k else if x < k then map_of r k else Some y)" + "sorted Empty = True" +| "sorted (Branch c l k v r) = (l |\ k \ k \| r \ sorted l \ sorted r)" -lemma map_of_tlt[simp]: "t |\ k \ map_of t k = None" +primrec lookup :: "('a\linorder, 'b) rbt \ 'a \ 'b" +where + "lookup Empty k = None" +| "lookup (Branch _ l x y r) k = (if k < x then lookup l k else if x < k then lookup r k else Some y)" + +lemma lookup_tree_less[simp]: "t |\ k \ lookup t k = None" by (induct t) auto -lemma map_of_tgt[simp]: "k \| t \ map_of t k = None" +lemma lookup_tree_greater[simp]: "k \| t \ lookup t k = None" by (induct t) auto -lemma mapof_keys: "st t \ dom (map_of t) = keys t" -by (induct t) (auto simp: dom_def tgt_prop tlt_prop) +lemma lookup_keys: "sorted t \ dom (lookup t) = keys t" +by (induct t) (auto simp: dom_def tree_greater_prop tree_less_prop) -lemma mapof_pit: "st t \ (map_of t k = Some v) = pin_tree k v t" -by (induct t) (auto simp: tlt_prop tgt_prop pint_keys) +lemma lookup_pit: "sorted t \ (lookup t k = Some v) = entry_in_tree k v t" +by (induct t) (auto simp: tree_less_prop tree_greater_prop entry_in_tree_keys) -lemma map_of_Empty: "map_of Empty = empty" +lemma lookup_Empty: "lookup Empty = empty" by (rule ext) simp (* a kind of extensionality *) -lemma mapof_from_pit: - assumes st: "st t1" "st t2" - and eq: "\v. pin_tree (k\'a\linorder) v t1 = pin_tree k v t2" - shows "map_of t1 k = map_of t2 k" -proof (cases "map_of t1 k") +lemma lookup_from_pit: + assumes sorted: "sorted t1" "sorted t2" + and eq: "\v. entry_in_tree (k\'a\linorder) v t1 = entry_in_tree k v t2" + shows "lookup t1 k = lookup t2 k" +proof (cases "lookup t1 k") case None - then have "\v. \ pin_tree k v t1" - by (simp add: mapof_pit[symmetric] st) + then have "\v. \ entry_in_tree k v t1" + by (simp add: lookup_pit[symmetric] sorted) with None show ?thesis - by (cases "map_of t2 k") (auto simp: mapof_pit st eq) + by (cases "lookup t2 k") (auto simp: lookup_pit sorted eq) next case (Some a) then show ?thesis - apply (cases "map_of t2 k") - apply (auto simp: mapof_pit st eq) - by (auto simp add: mapof_pit[symmetric] st Some) + apply (cases "lookup t2 k") + apply (auto simp: lookup_pit sorted eq) + by (auto simp add: lookup_pit[symmetric] sorted Some) qed subsection {* Red-black properties *} -primrec treec :: "('a,'b) rbt \ color" +primrec color_of :: "('a, 'b) rbt \ color" where - "treec Empty = B" -| "treec (Tr c _ _ _ _) = c" + "color_of Empty = B" +| "color_of (Branch c _ _ _ _) = c" -primrec inv1 :: "('a,'b) rbt \ bool" +primrec bheight :: "('a,'b) rbt \ nat" +where + "bheight Empty = 0" +| "bheight (Branch c lt k v rt) = (if c = B then Suc (bheight lt) else bheight lt)" + +primrec inv1 :: "('a, 'b) rbt \ bool" where "inv1 Empty = True" -| "inv1 (Tr c lt k v rt) = (inv1 lt \ inv1 rt \ (c = B \ treec lt = B \ treec rt = B))" +| "inv1 (Branch c lt k v rt) \ inv1 lt \ inv1 rt \ (c = B \ color_of lt = B \ color_of rt = B)" -(* Weaker version *) -primrec inv1l :: "('a,'b) rbt \ bool" +primrec inv1l :: "('a, 'b) rbt \ bool" -- {* Weaker version *} where "inv1l Empty = True" -| "inv1l (Tr c l k v r) = (inv1 l \ inv1 r)" +| "inv1l (Branch c l k v r) = (inv1 l \ inv1 r)" lemma [simp]: "inv1 t \ inv1l t" by (cases t) simp+ -primrec bh :: "('a,'b) rbt \ nat" -where - "bh Empty = 0" -| "bh (Tr c lt k v rt) = (if c = B then Suc (bh lt) else bh lt)" - -primrec inv2 :: "('a,'b) rbt \ bool" +primrec inv2 :: "('a, 'b) rbt \ bool" where "inv2 Empty = True" -| "inv2 (Tr c lt k v rt) = (inv2 lt \ inv2 rt \ bh lt = bh rt)" +| "inv2 (Branch c lt k v rt) = (inv2 lt \ inv2 rt \ bheight lt = bheight rt)" -definition - "isrbt t = (inv1 t \ inv2 t \ treec t = B \ st t)" - -lemma isrbt_st[simp]: "isrbt t \ st t" by (simp add: isrbt_def) +definition is_rbt :: "('a\linorder, 'b) rbt \ bool" where + "is_rbt t \ inv1 t \ inv2 t \ color_of t = B \ sorted t" -lemma rbt_cases: - obtains (Empty) "t = Empty" - | (Red) l k v r where "t = Tr R l k v r" - | (Black) l k v r where "t = Tr B l k v r" -by (cases t, simp) (case_tac "color", auto) +lemma is_rbt_sorted [simp]: + "is_rbt t \ sorted t" by (simp add: is_rbt_def) -theorem Empty_isrbt[simp]: "isrbt Empty" -unfolding isrbt_def by simp +theorem Empty_is_rbt [simp]: + "is_rbt Empty" by (simp add: is_rbt_def) subsection {* Insertion *} @@ -147,80 +163,80 @@ fun (* slow, due to massive case splitting *) balance :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "balance (Tr R a w x b) s t (Tr R c y z d) = Tr R (Tr B a w x b) s t (Tr B c y z d)" | - "balance (Tr R (Tr R a w x b) s t c) y z d = Tr R (Tr B a w x b) s t (Tr B c y z d)" | - "balance (Tr R a w x (Tr R b s t c)) y z d = Tr R (Tr B a w x b) s t (Tr B c y z d)" | - "balance a w x (Tr R b s t (Tr R c y z d)) = Tr R (Tr B a w x b) s t (Tr B c y z d)" | - "balance a w x (Tr R (Tr R b s t c) y z d) = Tr R (Tr B a w x b) s t (Tr B c y z d)" | - "balance a s t b = Tr B a s t b" + "balance (Branch R a w x b) s t (Branch R c y z d) = Branch R (Branch B a w x b) s t (Branch B c y z d)" | + "balance (Branch R (Branch R a w x b) s t c) y z d = Branch R (Branch B a w x b) s t (Branch B c y z d)" | + "balance (Branch R a w x (Branch R b s t c)) y z d = Branch R (Branch B a w x b) s t (Branch B c y z d)" | + "balance a w x (Branch R b s t (Branch R c y z d)) = Branch R (Branch B a w x b) s t (Branch B c y z d)" | + "balance a w x (Branch R (Branch R b s t c) y z d) = Branch R (Branch B a w x b) s t (Branch B c y z d)" | + "balance a s t b = Branch B a s t b" lemma balance_inv1: "\inv1l l; inv1l r\ \ inv1 (balance l k v r)" by (induct l k v r rule: balance.induct) auto -lemma balance_bh: "bh l = bh r \ bh (balance l k v r) = Suc (bh l)" +lemma balance_bheight: "bheight l = bheight r \ bheight (balance l k v r) = Suc (bheight l)" by (induct l k v r rule: balance.induct) auto lemma balance_inv2: - assumes "inv2 l" "inv2 r" "bh l = bh r" + assumes "inv2 l" "inv2 r" "bheight l = bheight r" shows "inv2 (balance l k v r)" using assms by (induct l k v r rule: balance.induct) auto -lemma balance_tgt[simp]: "(v \| balance a k x b) = (v \| a \ v \| b \ v < k)" +lemma balance_tree_greater[simp]: "(v \| balance a k x b) = (v \| a \ v \| b \ v < k)" by (induct a k x b rule: balance.induct) auto -lemma balance_tlt[simp]: "(balance a k x b |\ v) = (a |\ v \ b |\ v \ k < v)" +lemma balance_tree_less[simp]: "(balance a k x b |\ v) = (a |\ v \ b |\ v \ k < v)" by (induct a k x b rule: balance.induct) auto -lemma balance_st: +lemma balance_sorted: fixes k :: "'a::linorder" - assumes "st l" "st r" "l |\ k" "k \| r" - shows "st (balance l k v r)" + assumes "sorted l" "sorted r" "l |\ k" "k \| r" + shows "sorted (balance l k v r)" using assms proof (induct l k v r rule: balance.induct) case ("2_2" a x w b y t c z s va vb vd vc) - hence "y < z \ z \| Tr B va vb vd vc" - by (auto simp add: tlgt_props) - hence "tgt y (Tr B va vb vd vc)" by (blast dest: tgt_trans) + hence "y < z \ z \| Branch B va vb vd vc" + by (auto simp add: tree_ord_props) + hence "tree_greater y (Branch B va vb vd vc)" by (blast dest: tree_greater_trans) with "2_2" show ?case by simp next case ("3_2" va vb vd vc x w b y s c z) - from "3_2" have "x < y \ tlt x (Tr B va vb vd vc)" - by (simp add: tlt.simps tgt.simps) - hence "tlt y (Tr B va vb vd vc)" by (blast dest: tlt_trans) + from "3_2" have "x < y \ tree_less x (Branch B va vb vd vc)" + by simp + hence "tree_less y (Branch B va vb vd vc)" by (blast dest: tree_less_trans) with "3_2" show ?case by simp next case ("3_3" x w b y s c z t va vb vd vc) - from "3_3" have "y < z \ tgt z (Tr B va vb vd vc)" by simp - hence "tgt y (Tr B va vb vd vc)" by (blast dest: tgt_trans) + from "3_3" have "y < z \ tree_greater z (Branch B va vb vd vc)" by simp + hence "tree_greater y (Branch B va vb vd vc)" by (blast dest: tree_greater_trans) with "3_3" show ?case by simp next case ("3_4" vd ve vg vf x w b y s c z t va vb vii vc) - hence "x < y \ tlt x (Tr B vd ve vg vf)" by simp - hence 1: "tlt y (Tr B vd ve vg vf)" by (blast dest: tlt_trans) - from "3_4" have "y < z \ tgt z (Tr B va vb vii vc)" by simp - hence "tgt y (Tr B va vb vii vc)" by (blast dest: tgt_trans) + hence "x < y \ tree_less x (Branch B vd ve vg vf)" by simp + hence 1: "tree_less y (Branch B vd ve vg vf)" by (blast dest: tree_less_trans) + from "3_4" have "y < z \ tree_greater z (Branch B va vb vii vc)" by simp + hence "tree_greater y (Branch B va vb vii vc)" by (blast dest: tree_greater_trans) with 1 "3_4" show ?case by simp next case ("4_2" va vb vd vc x w b y s c z t dd) - hence "x < y \ tlt x (Tr B va vb vd vc)" by simp - hence "tlt y (Tr B va vb vd vc)" by (blast dest: tlt_trans) + hence "x < y \ tree_less x (Branch B va vb vd vc)" by simp + hence "tree_less y (Branch B va vb vd vc)" by (blast dest: tree_less_trans) with "4_2" show ?case by simp next case ("5_2" x w b y s c z t va vb vd vc) - hence "y < z \ tgt z (Tr B va vb vd vc)" by simp - hence "tgt y (Tr B va vb vd vc)" by (blast dest: tgt_trans) + hence "y < z \ tree_greater z (Branch B va vb vd vc)" by simp + hence "tree_greater y (Branch B va vb vd vc)" by (blast dest: tree_greater_trans) with "5_2" show ?case by simp next case ("5_3" va vb vd vc x w b y s c z t) - hence "x < y \ tlt x (Tr B va vb vd vc)" by simp - hence "tlt y (Tr B va vb vd vc)" by (blast dest: tlt_trans) + hence "x < y \ tree_less x (Branch B va vb vd vc)" by simp + hence "tree_less y (Branch B va vb vd vc)" by (blast dest: tree_less_trans) with "5_3" show ?case by simp next case ("5_4" va vb vg vc x w b y s c z t vd ve vii vf) - hence "x < y \ tlt x (Tr B va vb vg vc)" by simp - hence 1: "tlt y (Tr B va vb vg vc)" by (blast dest: tlt_trans) - from "5_4" have "y < z \ tgt z (Tr B vd ve vii vf)" by simp - hence "tgt y (Tr B vd ve vii vf)" by (blast dest: tgt_trans) + hence "x < y \ tree_less x (Branch B va vb vg vc)" by simp + hence 1: "tree_less y (Branch B va vb vg vc)" by (blast dest: tree_less_trans) + from "5_4" have "y < z \ tree_greater z (Branch B vd ve vii vf)" by simp + hence "tree_greater y (Branch B vd ve vii vf)" by (blast dest: tree_greater_trans) with 1 "5_4" show ?case by simp qed simp+ @@ -229,62 +245,62 @@ by (induct l k v r rule: balance.induct) auto lemma balance_pit: - "pin_tree k x (balance l v y r) = (pin_tree k x l \ k = v \ x = y \ pin_tree k x r)" + "entry_in_tree k x (balance l v y r) = (entry_in_tree k x l \ k = v \ x = y \ entry_in_tree k x r)" by (induct l v y r rule: balance.induct) auto -lemma map_of_balance[simp]: +lemma lookup_balance[simp]: fixes k :: "'a::linorder" -assumes "st l" "st r" "l |\ k" "k \| r" -shows "map_of (balance l k v r) x = map_of (Tr B l k v r) x" -by (rule mapof_from_pit) (auto simp:assms balance_pit balance_st) +assumes "sorted l" "sorted r" "l |\ k" "k \| r" +shows "lookup (balance l k v r) x = lookup (Branch B l k v r) x" +by (rule lookup_from_pit) (auto simp:assms balance_pit balance_sorted) primrec paint :: "color \ ('a,'b) rbt \ ('a,'b) rbt" where "paint c Empty = Empty" -| "paint c (Tr _ l k v r) = Tr c l k v r" +| "paint c (Branch _ l k v r) = Branch c l k v r" lemma paint_inv1l[simp]: "inv1l t \ inv1l (paint c t)" by (cases t) auto lemma paint_inv1[simp]: "inv1l t \ inv1 (paint B t)" by (cases t) auto lemma paint_inv2[simp]: "inv2 t \ inv2 (paint c t)" by (cases t) auto -lemma paint_treec[simp]: "treec (paint B t) = B" by (cases t) auto -lemma paint_st[simp]: "st t \ st (paint c t)" by (cases t) auto -lemma paint_pit[simp]: "pin_tree k x (paint c t) = pin_tree k x t" by (cases t) auto -lemma paint_mapof[simp]: "map_of (paint c t) = map_of t" by (rule ext) (cases t, auto) -lemma paint_tgt[simp]: "(v \| paint c t) = (v \| t)" by (cases t) auto -lemma paint_tlt[simp]: "(paint c t |\ v) = (t |\ v)" by (cases t) auto +lemma paint_color_of[simp]: "color_of (paint B t) = B" by (cases t) auto +lemma paint_sorted[simp]: "sorted t \ sorted (paint c t)" by (cases t) auto +lemma paint_pit[simp]: "entry_in_tree k x (paint c t) = entry_in_tree k x t" by (cases t) auto +lemma paint_lookup[simp]: "lookup (paint c t) = lookup t" by (rule ext) (cases t, auto) +lemma paint_tree_greater[simp]: "(v \| paint c t) = (v \| t)" by (cases t) auto +lemma paint_tree_less[simp]: "(paint c t |\ v) = (t |\ v)" by (cases t) auto fun ins :: "('a\linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "ins f k v Empty = Tr R Empty k v Empty" | - "ins f k v (Tr B l x y r) = (if k < x then balance (ins f k v l) x y r + "ins f k v Empty = Branch R Empty k v Empty" | + "ins f k v (Branch B l x y r) = (if k < x then balance (ins f k v l) x y r else if k > x then balance l x y (ins f k v r) - else Tr B l x (f k y v) r)" | - "ins f k v (Tr R l x y r) = (if k < x then Tr R (ins f k v l) x y r - else if k > x then Tr R l x y (ins f k v r) - else Tr R l x (f k y v) r)" + else Branch B l x (f k y v) r)" | + "ins f k v (Branch R l x y r) = (if k < x then Branch R (ins f k v l) x y r + else if k > x then Branch R l x y (ins f k v r) + else Branch R l x (f k y v) r)" lemma ins_inv1_inv2: assumes "inv1 t" "inv2 t" - shows "inv2 (ins f k x t)" "bh (ins f k x t) = bh t" - "treec t = B \ inv1 (ins f k x t)" "inv1l (ins f k x t)" + shows "inv2 (ins f k x t)" "bheight (ins f k x t) = bheight t" + "color_of t = B \ inv1 (ins f k x t)" "inv1l (ins f k x t)" using assms - by (induct f k x t rule: ins.induct) (auto simp: balance_inv1 balance_inv2 balance_bh) + by (induct f k x t rule: ins.induct) (auto simp: balance_inv1 balance_inv2 balance_bheight) -lemma ins_tgt[simp]: "(v \| ins f k x t) = (v \| t \ k > v)" +lemma ins_tree_greater[simp]: "(v \| ins f k x t) = (v \| t \ k > v)" by (induct f k x t rule: ins.induct) auto -lemma ins_tlt[simp]: "(ins f k x t |\ v) = (t |\ v \ k < v)" +lemma ins_tree_less[simp]: "(ins f k x t |\ v) = (t |\ v \ k < v)" by (induct f k x t rule: ins.induct) auto -lemma ins_st[simp]: "st t \ st (ins f k x t)" - by (induct f k x t rule: ins.induct) (auto simp: balance_st) +lemma ins_sorted[simp]: "sorted t \ sorted (ins f k x t)" + by (induct f k x t rule: ins.induct) (auto simp: balance_sorted) lemma keys_ins: "keys (ins f k v t) = { k } \ keys t" by (induct f k v t rule: ins.induct) auto -lemma map_of_ins: +lemma lookup_ins: fixes k :: "'a::linorder" - assumes "st t" - shows "map_of (ins f k v t) x = ((map_of t)(k |-> case map_of t k of None \ v + assumes "sorted t" + shows "lookup (ins f k v t) x = ((lookup t)(k |-> case lookup t k of None \ v | Some w \ f k w v)) x" using assms by (induct f k v t rule: ins.induct) auto @@ -293,98 +309,97 @@ where "insertwithkey f k v t = paint B (ins f k v t)" -lemma insertwk_st: "st t \ st (insertwithkey f k x t)" +lemma insertwk_sorted: "sorted t \ sorted (insertwithkey f k x t)" by (auto simp: insertwithkey_def) -theorem insertwk_isrbt: - assumes inv: "isrbt t" - shows "isrbt (insertwithkey f k x t)" +theorem insertwk_is_rbt: + assumes inv: "is_rbt t" + shows "is_rbt (insertwithkey f k x t)" using assms -unfolding insertwithkey_def isrbt_def +unfolding insertwithkey_def is_rbt_def by (auto simp: ins_inv1_inv2) -lemma map_of_insertwk: - assumes "st t" - shows "map_of (insertwithkey f k v t) x = ((map_of t)(k |-> case map_of t k of None \ v +lemma lookup_insertwk: + assumes "sorted t" + shows "lookup (insertwithkey f k v t) x = ((lookup t)(k |-> case lookup t k of None \ v | Some w \ f k w v)) x" unfolding insertwithkey_def using assms -by (simp add:map_of_ins) +by (simp add:lookup_ins) definition insertw_def: "insertwith f = insertwithkey (\_. f)" -lemma insertw_st: "st t \ st (insertwith f k v t)" by (simp add: insertwk_st insertw_def) -theorem insertw_isrbt: "isrbt t \ isrbt (insertwith f k v t)" by (simp add: insertwk_isrbt insertw_def) +lemma insertw_sorted: "sorted t \ sorted (insertwith f k v t)" by (simp add: insertwk_sorted insertw_def) +theorem insertw_is_rbt: "is_rbt t \ is_rbt (insertwith f k v t)" by (simp add: insertwk_is_rbt insertw_def) -lemma map_of_insertw: - assumes "isrbt t" - shows "map_of (insertwith f k v t) = (map_of t)(k \ (if k:dom (map_of t) then f (the (map_of t k)) v else v))" +lemma lookup_insertw: + assumes "is_rbt t" + shows "lookup (insertwith f k v t) = (lookup t)(k \ (if k:dom (lookup t) then f (the (lookup t k)) v else v))" using assms unfolding insertw_def -by (rule_tac ext) (cases "map_of t k", auto simp:map_of_insertwk dom_def) - +by (rule_tac ext) (cases "lookup t k", auto simp:lookup_insertwk dom_def) -definition - "insrt k v t = insertwithkey (\_ _ nv. nv) k v t" +definition insert :: "'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" where + "insert k v t = insertwithkey (\_ _ nv. nv) k v t" -lemma insrt_st: "st t \ st (insrt k v t)" by (simp add: insertwk_st insrt_def) -theorem insrt_isrbt: "isrbt t \ isrbt (insrt k v t)" by (simp add: insertwk_isrbt insrt_def) +lemma insert_sorted: "sorted t \ sorted (insert k v t)" by (simp add: insertwk_sorted insert_def) +theorem insert_is_rbt: "is_rbt t \ is_rbt (insert k v t)" by (simp add: insertwk_is_rbt insert_def) -lemma map_of_insert: - assumes "isrbt t" - shows "map_of (insrt k v t) = (map_of t)(k\v)" -unfolding insrt_def +lemma lookup_insert: + assumes "is_rbt t" + shows "lookup (insert k v t) = (lookup t)(k\v)" +unfolding insert_def using assms -by (rule_tac ext) (simp add: map_of_insertwk split:option.split) +by (rule_tac ext) (simp add: lookup_insertwk split:option.split) subsection {* Deletion *} -lemma bh_paintR'[simp]: "treec t = B \ bh (paint R t) = bh t - 1" +lemma bheight_paintR'[simp]: "color_of t = B \ bheight (paint R t) = bheight t - 1" by (cases t rule: rbt_cases) auto fun balleft :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "balleft (Tr R a k x b) s y c = Tr R (Tr B a k x b) s y c" | - "balleft bl k x (Tr B a s y b) = balance bl k x (Tr R a s y b)" | - "balleft bl k x (Tr R (Tr B a s y b) t z c) = Tr R (Tr B bl k x a) s y (balance b t z (paint R c))" | + "balleft (Branch R a k x b) s y c = Branch R (Branch B a k x b) s y c" | + "balleft bl k x (Branch B a s y b) = balance bl k x (Branch R a s y b)" | + "balleft bl k x (Branch R (Branch B a s y b) t z c) = Branch R (Branch B bl k x a) s y (balance b t z (paint R c))" | "balleft t k x s = Empty" lemma balleft_inv2_with_inv1: - assumes "inv2 lt" "inv2 rt" "bh lt + 1 = bh rt" "inv1 rt" - shows "bh (balleft lt k v rt) = bh lt + 1" + assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "inv1 rt" + shows "bheight (balleft lt k v rt) = bheight lt + 1" and "inv2 (balleft lt k v rt)" using assms -by (induct lt k v rt rule: balleft.induct) (auto simp: balance_inv2 balance_bh) +by (induct lt k v rt rule: balleft.induct) (auto simp: balance_inv2 balance_bheight) lemma balleft_inv2_app: - assumes "inv2 lt" "inv2 rt" "bh lt + 1 = bh rt" "treec rt = B" + assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "color_of rt = B" shows "inv2 (balleft lt k v rt)" - "bh (balleft lt k v rt) = bh rt" + "bheight (balleft lt k v rt) = bheight rt" using assms -by (induct lt k v rt rule: balleft.induct) (auto simp add: balance_inv2 balance_bh)+ +by (induct lt k v rt rule: balleft.induct) (auto simp add: balance_inv2 balance_bheight)+ -lemma balleft_inv1: "\inv1l a; inv1 b; treec b = B\ \ inv1 (balleft a k x b)" +lemma balleft_inv1: "\inv1l a; inv1 b; color_of b = B\ \ inv1 (balleft a k x b)" by (induct a k x b rule: balleft.induct) (simp add: balance_inv1)+ lemma balleft_inv1l: "\ inv1l lt; inv1 rt \ \ inv1l (balleft lt k x rt)" by (induct lt k x rt rule: balleft.induct) (auto simp: balance_inv1) -lemma balleft_st: "\ st l; st r; tlt k l; tgt k r \ \ st (balleft l k v r)" +lemma balleft_sorted: "\ sorted l; sorted r; tree_less k l; tree_greater k r \ \ sorted (balleft l k v r)" apply (induct l k v r rule: balleft.induct) -apply (auto simp: balance_st) -apply (unfold tgt_prop tlt_prop) +apply (auto simp: balance_sorted) +apply (unfold tree_greater_prop tree_less_prop) by force+ -lemma balleft_tgt: +lemma balleft_tree_greater: fixes k :: "'a::order" assumes "k \| a" "k \| b" "k < x" shows "k \| balleft a x t b" using assms by (induct a x t b rule: balleft.induct) auto -lemma balleft_tlt: +lemma balleft_tree_less: fixes k :: "'a::order" assumes "a |\ k" "b |\ k" "x < k" shows "balleft a x t b |\ k" @@ -392,52 +407,52 @@ by (induct a x t b rule: balleft.induct) auto lemma balleft_pit: - assumes "inv1l l" "inv1 r" "bh l + 1 = bh r" - shows "pin_tree k v (balleft l a b r) = (pin_tree k v l \ k = a \ v = b \ pin_tree k v r)" + assumes "inv1l l" "inv1 r" "bheight l + 1 = bheight r" + shows "entry_in_tree k v (balleft l a b r) = (entry_in_tree k v l \ k = a \ v = b \ entry_in_tree k v r)" using assms by (induct l k v r rule: balleft.induct) (auto simp: balance_pit) fun balright :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "balright a k x (Tr R b s y c) = Tr R a k x (Tr B b s y c)" | - "balright (Tr B a k x b) s y bl = balance (Tr R a k x b) s y bl" | - "balright (Tr R a k x (Tr B b s y c)) t z bl = Tr R (balance (paint R a) k x b) s y (Tr B c t z bl)" | + "balright a k x (Branch R b s y c) = Branch R a k x (Branch B b s y c)" | + "balright (Branch B a k x b) s y bl = balance (Branch R a k x b) s y bl" | + "balright (Branch R a k x (Branch B b s y c)) t z bl = Branch R (balance (paint R a) k x b) s y (Branch B c t z bl)" | "balright t k x s = Empty" lemma balright_inv2_with_inv1: - assumes "inv2 lt" "inv2 rt" "bh lt = bh rt + 1" "inv1 lt" - shows "inv2 (balright lt k v rt) \ bh (balright lt k v rt) = bh lt" + assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt + 1" "inv1 lt" + shows "inv2 (balright lt k v rt) \ bheight (balright lt k v rt) = bheight lt" using assms -by (induct lt k v rt rule: balright.induct) (auto simp: balance_inv2 balance_bh) +by (induct lt k v rt rule: balright.induct) (auto simp: balance_inv2 balance_bheight) -lemma balright_inv1: "\inv1 a; inv1l b; treec a = B\ \ inv1 (balright a k x b)" +lemma balright_inv1: "\inv1 a; inv1l b; color_of a = B\ \ inv1 (balright a k x b)" by (induct a k x b rule: balright.induct) (simp add: balance_inv1)+ lemma balright_inv1l: "\ inv1 lt; inv1l rt \ \inv1l (balright lt k x rt)" by (induct lt k x rt rule: balright.induct) (auto simp: balance_inv1) -lemma balright_st: "\ st l; st r; tlt k l; tgt k r \ \ st (balright l k v r)" +lemma balright_sorted: "\ sorted l; sorted r; tree_less k l; tree_greater k r \ \ sorted (balright l k v r)" apply (induct l k v r rule: balright.induct) -apply (auto simp:balance_st) -apply (unfold tlt_prop tgt_prop) +apply (auto simp:balance_sorted) +apply (unfold tree_less_prop tree_greater_prop) by force+ -lemma balright_tgt: +lemma balright_tree_greater: fixes k :: "'a::order" assumes "k \| a" "k \| b" "k < x" shows "k \| balright a x t b" using assms by (induct a x t b rule: balright.induct) auto -lemma balright_tlt: +lemma balright_tree_less: fixes k :: "'a::order" assumes "a |\ k" "b |\ k" "x < k" shows "balright a x t b |\ k" using assms by (induct a x t b rule: balright.induct) auto lemma balright_pit: - assumes "inv1 l" "inv1l r" "bh l = bh r + 1" "inv2 l" "inv2 r" - shows "pin_tree x y (balright l k v r) = (pin_tree x y l \ x = k \ y = v \ pin_tree x y r)" + assumes "inv1 l" "inv1l r" "bheight l = bheight r + 1" "inv2 l" "inv2 r" + shows "entry_in_tree x y (balright l k v r) = (entry_in_tree x y l \ x = k \ y = v \ entry_in_tree x y r)" using assms by (induct l k v r rule: balright.induct) (auto simp: balance_pit) @@ -448,50 +463,50 @@ where "app Empty x = x" | "app x Empty = x" -| "app (Tr R a k x b) (Tr R c s y d) = (case (app b c) of - Tr R b2 t z c2 \ (Tr R (Tr R a k x b2) t z (Tr R c2 s y d)) | - bc \ Tr R a k x (Tr R bc s y d))" -| "app (Tr B a k x b) (Tr B c s y d) = (case (app b c) of - Tr R b2 t z c2 \ Tr R (Tr B a k x b2) t z (Tr B c2 s y d) | - bc \ balleft a k x (Tr B bc s y d))" -| "app a (Tr R b k x c) = Tr R (app a b) k x c" -| "app (Tr R a k x b) c = Tr R a k x (app b c)" +| "app (Branch R a k x b) (Branch R c s y d) = (case (app b c) of + Branch R b2 t z c2 \ (Branch R (Branch R a k x b2) t z (Branch R c2 s y d)) | + bc \ Branch R a k x (Branch R bc s y d))" +| "app (Branch B a k x b) (Branch B c s y d) = (case (app b c) of + Branch R b2 t z c2 \ Branch R (Branch B a k x b2) t z (Branch B c2 s y d) | + bc \ balleft a k x (Branch B bc s y d))" +| "app a (Branch R b k x c) = Branch R (app a b) k x c" +| "app (Branch R a k x b) c = Branch R a k x (app b c)" lemma app_inv2: - assumes "inv2 lt" "inv2 rt" "bh lt = bh rt" - shows "bh (app lt rt) = bh lt" "inv2 (app lt rt)" + assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt" + shows "bheight (app lt rt) = bheight lt" "inv2 (app lt rt)" using assms by (induct lt rt rule: app.induct) (auto simp: balleft_inv2_app split: rbt.splits color.splits) lemma app_inv1: assumes "inv1 lt" "inv1 rt" - shows "treec lt = B \ treec rt = B \ inv1 (app lt rt)" + shows "color_of lt = B \ color_of rt = B \ inv1 (app lt rt)" "inv1l (app lt rt)" using assms by (induct lt rt rule: app.induct) (auto simp: balleft_inv1 split: rbt.splits color.splits) -lemma app_tgt[simp]: +lemma app_tree_greater[simp]: fixes k :: "'a::linorder" assumes "k \| l" "k \| r" shows "k \| app l r" using assms by (induct l r rule: app.induct) - (auto simp: balleft_tgt split:rbt.splits color.splits) + (auto simp: balleft_tree_greater split:rbt.splits color.splits) -lemma app_tlt[simp]: +lemma app_tree_less[simp]: fixes k :: "'a::linorder" assumes "l |\ k" "r |\ k" shows "app l r |\ k" using assms by (induct l r rule: app.induct) - (auto simp: balleft_tlt split:rbt.splits color.splits) + (auto simp: balleft_tree_less split:rbt.splits color.splits) -lemma app_st: +lemma app_sorted: fixes k :: "'a::linorder" - assumes "st l" "st r" "l |\ k" "k \| r" - shows "st (app l r)" + assumes "sorted l" "sorted r" "l |\ k" "k \| r" + shows "sorted (app l r)" using assms proof (induct l r rule: app.induct) case (3 a x v b c y w d) hence ineqs: "a |\ x" "x \| b" "b |\ k" "k \| c" "c |\ y" "y \| d" @@ -500,55 +515,55 @@ show ?case apply (cases "app b c" rule: rbt_cases) apply auto - by (metis app_tgt app_tlt ineqs ineqs tlt.simps(2) tgt.simps(2) tgt_trans tlt_trans)+ + by (metis app_tree_greater app_tree_less ineqs ineqs tree_less_simps(2) tree_greater_simps(2) tree_greater_trans tree_less_trans)+ next case (4 a x v b c y w d) - hence "x < k \ tgt k c" by simp - hence "tgt x c" by (blast dest: tgt_trans) - with 4 have 2: "tgt x (app b c)" by (simp add: app_tgt) - from 4 have "k < y \ tlt k b" by simp - hence "tlt y b" by (blast dest: tlt_trans) - with 4 have 3: "tlt y (app b c)" by (simp add: app_tlt) + hence "x < k \ tree_greater k c" by simp + hence "tree_greater x c" by (blast dest: tree_greater_trans) + with 4 have 2: "tree_greater x (app b c)" by (simp add: app_tree_greater) + from 4 have "k < y \ tree_less k b" by simp + hence "tree_less y b" by (blast dest: tree_less_trans) + with 4 have 3: "tree_less y (app b c)" by (simp add: app_tree_less) show ?case proof (cases "app b c" rule: rbt_cases) case Empty - from 4 have "x < y \ tgt y d" by auto - hence "tgt x d" by (blast dest: tgt_trans) - with 4 Empty have "st a" and "st (Tr B Empty y w d)" and "tlt x a" and "tgt x (Tr B Empty y w d)" by auto - with Empty show ?thesis by (simp add: balleft_st) + from 4 have "x < y \ tree_greater y d" by auto + hence "tree_greater x d" by (blast dest: tree_greater_trans) + with 4 Empty have "sorted a" and "sorted (Branch B Empty y w d)" and "tree_less x a" and "tree_greater x (Branch B Empty y w d)" by auto + with Empty show ?thesis by (simp add: balleft_sorted) next case (Red lta va ka rta) - with 2 4 have "x < va \ tlt x a" by simp - hence 5: "tlt va a" by (blast dest: tlt_trans) - from Red 3 4 have "va < y \ tgt y d" by simp - hence "tgt va d" by (blast dest: tgt_trans) + with 2 4 have "x < va \ tree_less x a" by simp + hence 5: "tree_less va a" by (blast dest: tree_less_trans) + from Red 3 4 have "va < y \ tree_greater y d" by simp + hence "tree_greater va d" by (blast dest: tree_greater_trans) with Red 2 3 4 5 show ?thesis by simp next case (Black lta va ka rta) - from 4 have "x < y \ tgt y d" by auto - hence "tgt x d" by (blast dest: tgt_trans) - with Black 2 3 4 have "st a" and "st (Tr B (app b c) y w d)" and "tlt x a" and "tgt x (Tr B (app b c) y w d)" by auto - with Black show ?thesis by (simp add: balleft_st) + from 4 have "x < y \ tree_greater y d" by auto + hence "tree_greater x d" by (blast dest: tree_greater_trans) + with Black 2 3 4 have "sorted a" and "sorted (Branch B (app b c) y w d)" and "tree_less x a" and "tree_greater x (Branch B (app b c) y w d)" by auto + with Black show ?thesis by (simp add: balleft_sorted) qed next case (5 va vb vd vc b x w c) - hence "k < x \ tlt k (Tr B va vb vd vc)" by simp - hence "tlt x (Tr B va vb vd vc)" by (blast dest: tlt_trans) - with 5 show ?case by (simp add: app_tlt) + hence "k < x \ tree_less k (Branch B va vb vd vc)" by simp + hence "tree_less x (Branch B va vb vd vc)" by (blast dest: tree_less_trans) + with 5 show ?case by (simp add: app_tree_less) next case (6 a x v b va vb vd vc) - hence "x < k \ tgt k (Tr B va vb vd vc)" by simp - hence "tgt x (Tr B va vb vd vc)" by (blast dest: tgt_trans) - with 6 show ?case by (simp add: app_tgt) + hence "x < k \ tree_greater k (Branch B va vb vd vc)" by simp + hence "tree_greater x (Branch B va vb vd vc)" by (blast dest: tree_greater_trans) + with 6 show ?case by (simp add: app_tree_greater) qed simp+ lemma app_pit: - assumes "inv2 l" "inv2 r" "bh l = bh r" "inv1 l" "inv1 r" - shows "pin_tree k v (app l r) = (pin_tree k v l \ pin_tree k v r)" + assumes "inv2 l" "inv2 r" "bheight l = bheight r" "inv1 l" "inv1 r" + shows "entry_in_tree k v (app l r) = (entry_in_tree k v l \ entry_in_tree k v r)" using assms proof (induct l r rule: app.induct) case (4 _ _ _ b c) - hence a: "bh (app b c) = bh b" by (simp add: app_inv2) + hence a: "bheight (app b c) = bheight b" by (simp add: app_inv2) from 4 have b: "inv1l (app b c)" by (simp add: app_inv1) show ?case @@ -570,21 +585,21 @@ del :: "('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt" where "del x Empty = Empty" | - "del x (Tr c a y s b) = (if x < y then delformLeft x a y s b else (if x > y then delformRight x a y s b else app a b))" | - "delformLeft x (Tr B lt z v rt) y s b = balleft (del x (Tr B lt z v rt)) y s b" | - "delformLeft x a y s b = Tr R (del x a) y s b" | - "delformRight x a y s (Tr B lt z v rt) = balright a y s (del x (Tr B lt z v rt))" | - "delformRight x a y s b = Tr R a y s (del x b)" + "del x (Branch c a y s b) = (if x < y then delformLeft x a y s b else (if x > y then delformRight x a y s b else app a b))" | + "delformLeft x (Branch B lt z v rt) y s b = balleft (del x (Branch B lt z v rt)) y s b" | + "delformLeft x a y s b = Branch R (del x a) y s b" | + "delformRight x a y s (Branch B lt z v rt) = balright a y s (del x (Branch B lt z v rt))" | + "delformRight x a y s b = Branch R a y s (del x b)" lemma assumes "inv2 lt" "inv1 lt" shows - "\inv2 rt; bh lt = bh rt; inv1 rt\ \ - inv2 (delformLeft x lt k v rt) \ bh (delformLeft x lt k v rt) = bh lt \ (treec lt = B \ treec rt = B \ inv1 (delformLeft x lt k v rt) \ (treec lt \ B \ treec rt \ B) \ inv1l (delformLeft x lt k v rt))" - and "\inv2 rt; bh lt = bh rt; inv1 rt\ \ - inv2 (delformRight x lt k v rt) \ bh (delformRight x lt k v rt) = bh lt \ (treec lt = B \ treec rt = B \ inv1 (delformRight x lt k v rt) \ (treec lt \ B \ treec rt \ B) \ inv1l (delformRight x lt k v rt))" - and del_inv1_inv2: "inv2 (del x lt) \ (treec lt = R \ bh (del x lt) = bh lt \ inv1 (del x lt) - \ treec lt = B \ bh (del x lt) = bh lt - 1 \ inv1l (del x lt))" + "\inv2 rt; bheight lt = bheight rt; inv1 rt\ \ + inv2 (delformLeft x lt k v rt) \ bheight (delformLeft x lt k v rt) = bheight lt \ (color_of lt = B \ color_of rt = B \ inv1 (delformLeft x lt k v rt) \ (color_of lt \ B \ color_of rt \ B) \ inv1l (delformLeft x lt k v rt))" + and "\inv2 rt; bheight lt = bheight rt; inv1 rt\ \ + inv2 (delformRight x lt k v rt) \ bheight (delformRight x lt k v rt) = bheight lt \ (color_of lt = B \ color_of rt = B \ inv1 (delformRight x lt k v rt) \ (color_of lt \ B \ color_of rt \ B) \ inv1l (delformRight x lt k v rt))" + and del_inv1_inv2: "inv2 (del x lt) \ (color_of lt = R \ bheight (del x lt) = bheight lt \ inv1 (del x lt) + \ color_of lt = B \ bheight (del x lt) = bheight lt - 1 \ inv1l (del x lt))" using assms proof (induct x lt k v rt and x lt k v rt and x lt rule: delformLeft_delformRight_del.induct) case (2 y c _ y') @@ -601,55 +616,55 @@ qed next case (3 y lt z v rta y' ss bb) - thus ?case by (cases "treec (Tr B lt z v rta) = B \ treec bb = B") (simp add: balleft_inv2_with_inv1 balleft_inv1 balleft_inv1l)+ + thus ?case by (cases "color_of (Branch B lt z v rta) = B \ color_of bb = B") (simp add: balleft_inv2_with_inv1 balleft_inv1 balleft_inv1l)+ next case (5 y a y' ss lt z v rta) - thus ?case by (cases "treec a = B \ treec (Tr B lt z v rta) = B") (simp add: balright_inv2_with_inv1 balright_inv1 balright_inv1l)+ + thus ?case by (cases "color_of a = B \ color_of (Branch B lt z v rta) = B") (simp add: balright_inv2_with_inv1 balright_inv1 balright_inv1l)+ next - case ("6_1" y a y' ss) thus ?case by (cases "treec a = B \ treec Empty = B") simp+ + case ("6_1" y a y' ss) thus ?case by (cases "color_of a = B \ color_of Empty = B") simp+ qed auto lemma - delformLeft_tlt: "\tlt v lt; tlt v rt; k < v\ \ tlt v (delformLeft x lt k y rt)" - and delformRight_tlt: "\tlt v lt; tlt v rt; k < v\ \ tlt v (delformRight x lt k y rt)" - and del_tlt: "tlt v lt \ tlt v (del x lt)" + delformLeft_tree_less: "\tree_less v lt; tree_less v rt; k < v\ \ tree_less v (delformLeft x lt k y rt)" + and delformRight_tree_less: "\tree_less v lt; tree_less v rt; k < v\ \ tree_less v (delformRight x lt k y rt)" + and del_tree_less: "tree_less v lt \ tree_less v (del x lt)" by (induct x lt k y rt and x lt k y rt and x lt rule: delformLeft_delformRight_del.induct) - (auto simp: balleft_tlt balright_tlt) + (auto simp: balleft_tree_less balright_tree_less) -lemma delformLeft_tgt: "\tgt v lt; tgt v rt; k > v\ \ tgt v (delformLeft x lt k y rt)" - and delformRight_tgt: "\tgt v lt; tgt v rt; k > v\ \ tgt v (delformRight x lt k y rt)" - and del_tgt: "tgt v lt \ tgt v (del x lt)" +lemma delformLeft_tree_greater: "\tree_greater v lt; tree_greater v rt; k > v\ \ tree_greater v (delformLeft x lt k y rt)" + and delformRight_tree_greater: "\tree_greater v lt; tree_greater v rt; k > v\ \ tree_greater v (delformRight x lt k y rt)" + and del_tree_greater: "tree_greater v lt \ tree_greater v (del x lt)" by (induct x lt k y rt and x lt k y rt and x lt rule: delformLeft_delformRight_del.induct) - (auto simp: balleft_tgt balright_tgt) + (auto simp: balleft_tree_greater balright_tree_greater) -lemma "\st lt; st rt; tlt k lt; tgt k rt\ \ st (delformLeft x lt k y rt)" - and "\st lt; st rt; tlt k lt; tgt k rt\ \ st (delformRight x lt k y rt)" - and del_st: "st lt \ st (del x lt)" +lemma "\sorted lt; sorted rt; tree_less k lt; tree_greater k rt\ \ sorted (delformLeft x lt k y rt)" + and "\sorted lt; sorted rt; tree_less k lt; tree_greater k rt\ \ sorted (delformRight x lt k y rt)" + and del_sorted: "sorted lt \ sorted (del x lt)" proof (induct x lt k y rt and x lt k y rt and x lt rule: delformLeft_delformRight_del.induct) case (3 x lta zz v rta yy ss bb) - from 3 have "tlt yy (Tr B lta zz v rta)" by simp - hence "tlt yy (del x (Tr B lta zz v rta))" by (rule del_tlt) - with 3 show ?case by (simp add: balleft_st) + from 3 have "tree_less yy (Branch B lta zz v rta)" by simp + hence "tree_less yy (del x (Branch B lta zz v rta))" by (rule del_tree_less) + with 3 show ?case by (simp add: balleft_sorted) next case ("4_2" x vaa vbb vdd vc yy ss bb) - hence "tlt yy (Tr R vaa vbb vdd vc)" by simp - hence "tlt yy (del x (Tr R vaa vbb vdd vc))" by (rule del_tlt) + hence "tree_less yy (Branch R vaa vbb vdd vc)" by simp + hence "tree_less yy (del x (Branch R vaa vbb vdd vc))" by (rule del_tree_less) with "4_2" show ?case by simp next case (5 x aa yy ss lta zz v rta) - hence "tgt yy (Tr B lta zz v rta)" by simp - hence "tgt yy (del x (Tr B lta zz v rta))" by (rule del_tgt) - with 5 show ?case by (simp add: balright_st) + hence "tree_greater yy (Branch B lta zz v rta)" by simp + hence "tree_greater yy (del x (Branch B lta zz v rta))" by (rule del_tree_greater) + with 5 show ?case by (simp add: balright_sorted) next case ("6_2" x aa yy ss vaa vbb vdd vc) - hence "tgt yy (Tr R vaa vbb vdd vc)" by simp - hence "tgt yy (del x (Tr R vaa vbb vdd vc))" by (rule del_tgt) + hence "tree_greater yy (Branch R vaa vbb vdd vc)" by simp + hence "tree_greater yy (del x (Branch R vaa vbb vdd vc))" by (rule del_tree_greater) with "6_2" show ?case by simp -qed (auto simp: app_st) +qed (auto simp: app_sorted) -lemma "\st lt; st rt; tlt kt lt; tgt kt rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bh lt = bh rt; x < kt\ \ pin_tree k v (delformLeft x lt kt y rt) = (False \ (x \ k \ pin_tree k v (Tr c lt kt y rt)))" - and "\st lt; st rt; tlt kt lt; tgt kt rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bh lt = bh rt; x > kt\ \ pin_tree k v (delformRight x lt kt y rt) = (False \ (x \ k \ pin_tree k v (Tr c lt kt y rt)))" - and del_pit: "\st t; inv1 t; inv2 t\ \ pin_tree k v (del x t) = (False \ (x \ k \ pin_tree k v t))" +lemma "\sorted lt; sorted rt; tree_less kt lt; tree_greater kt rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bheight lt = bheight rt; x < kt\ \ entry_in_tree k v (delformLeft x lt kt y rt) = (False \ (x \ k \ entry_in_tree k v (Branch c lt kt y rt)))" + and "\sorted lt; sorted rt; tree_less kt lt; tree_greater kt rt; inv1 lt; inv1 rt; inv2 lt; inv2 rt; bheight lt = bheight rt; x > kt\ \ entry_in_tree k v (delformRight x lt kt y rt) = (False \ (x \ k \ entry_in_tree k v (Branch c lt kt y rt)))" + and del_pit: "\sorted t; inv1 t; inv2 t\ \ entry_in_tree k v (del x t) = (False \ (x \ k \ entry_in_tree k v t))" proof (induct x lt kt y rt and x lt kt y rt and x t rule: delformLeft_delformRight_del.induct) case (2 xx c aa yy ss bb) have "xx = yy \ xx < yy \ xx > yy" by auto @@ -657,68 +672,68 @@ assume "xx = yy" with 2 show ?thesis proof (cases "xx = k") case True - from 2 `xx = yy` `xx = k` have "st (Tr c aa yy ss bb) \ k = yy" by simp - hence "\ pin_tree k v aa" "\ pin_tree k v bb" by (auto simp: tlt_nit tgt_prop) + from 2 `xx = yy` `xx = k` have "sorted (Branch c aa yy ss bb) \ k = yy" by simp + hence "\ entry_in_tree k v aa" "\ entry_in_tree k v bb" by (auto simp: tree_less_nit tree_greater_prop) with `xx = yy` 2 `xx = k` show ?thesis by (simp add: app_pit) qed (simp add: app_pit) qed simp+ next case (3 xx lta zz vv rta yy ss bb) - def mt[simp]: mt == "Tr B lta zz vv rta" + def mt[simp]: mt == "Branch B lta zz vv rta" from 3 have "inv2 mt \ inv1 mt" by simp - hence "inv2 (del xx mt) \ (treec mt = R \ bh (del xx mt) = bh mt \ inv1 (del xx mt) \ treec mt = B \ bh (del xx mt) = bh mt - 1 \ inv1l (del xx mt))" by (blast dest: del_inv1_inv2) - with 3 have 4: "pin_tree k v (delformLeft xx mt yy ss bb) = (False \ xx \ k \ pin_tree k v mt \ (k = yy \ v = ss) \ pin_tree k v bb)" by (simp add: balleft_pit) + hence "inv2 (del xx mt) \ (color_of mt = R \ bheight (del xx mt) = bheight mt \ inv1 (del xx mt) \ color_of mt = B \ bheight (del xx mt) = bheight mt - 1 \ inv1l (del xx mt))" by (blast dest: del_inv1_inv2) + with 3 have 4: "entry_in_tree k v (delformLeft xx mt yy ss bb) = (False \ xx \ k \ entry_in_tree k v mt \ (k = yy \ v = ss) \ entry_in_tree k v bb)" by (simp add: balleft_pit) thus ?case proof (cases "xx = k") case True - from 3 True have "tgt yy bb \ yy > k" by simp - hence "tgt k bb" by (blast dest: tgt_trans) - with 3 4 True show ?thesis by (auto simp: tgt_nit) + from 3 True have "tree_greater yy bb \ yy > k" by simp + hence "tree_greater k bb" by (blast dest: tree_greater_trans) + with 3 4 True show ?thesis by (auto simp: tree_greater_nit) qed auto next case ("4_1" xx yy ss bb) show ?case proof (cases "xx = k") case True - with "4_1" have "tgt yy bb \ k < yy" by simp - hence "tgt k bb" by (blast dest: tgt_trans) + with "4_1" have "tree_greater yy bb \ k < yy" by simp + hence "tree_greater k bb" by (blast dest: tree_greater_trans) with "4_1" `xx = k` - have "pin_tree k v (Tr R Empty yy ss bb) = pin_tree k v Empty" by (auto simp: tgt_nit) + have "entry_in_tree k v (Branch R Empty yy ss bb) = entry_in_tree k v Empty" by (auto simp: tree_greater_nit) thus ?thesis by auto qed simp+ next case ("4_2" xx vaa vbb vdd vc yy ss bb) thus ?case proof (cases "xx = k") case True - with "4_2" have "k < yy \ tgt yy bb" by simp - hence "tgt k bb" by (blast dest: tgt_trans) - with True "4_2" show ?thesis by (auto simp: tgt_nit) + with "4_2" have "k < yy \ tree_greater yy bb" by simp + hence "tree_greater k bb" by (blast dest: tree_greater_trans) + with True "4_2" show ?thesis by (auto simp: tree_greater_nit) qed simp next case (5 xx aa yy ss lta zz vv rta) - def mt[simp]: mt == "Tr B lta zz vv rta" + def mt[simp]: mt == "Branch B lta zz vv rta" from 5 have "inv2 mt \ inv1 mt" by simp - hence "inv2 (del xx mt) \ (treec mt = R \ bh (del xx mt) = bh mt \ inv1 (del xx mt) \ treec mt = B \ bh (del xx mt) = bh mt - 1 \ inv1l (del xx mt))" by (blast dest: del_inv1_inv2) - with 5 have 3: "pin_tree k v (delformRight xx aa yy ss mt) = (pin_tree k v aa \ (k = yy \ v = ss) \ False \ xx \ k \ pin_tree k v mt)" by (simp add: balright_pit) + hence "inv2 (del xx mt) \ (color_of mt = R \ bheight (del xx mt) = bheight mt \ inv1 (del xx mt) \ color_of mt = B \ bheight (del xx mt) = bheight mt - 1 \ inv1l (del xx mt))" by (blast dest: del_inv1_inv2) + with 5 have 3: "entry_in_tree k v (delformRight xx aa yy ss mt) = (entry_in_tree k v aa \ (k = yy \ v = ss) \ False \ xx \ k \ entry_in_tree k v mt)" by (simp add: balright_pit) thus ?case proof (cases "xx = k") case True - from 5 True have "tlt yy aa \ yy < k" by simp - hence "tlt k aa" by (blast dest: tlt_trans) - with 3 5 True show ?thesis by (auto simp: tlt_nit) + from 5 True have "tree_less yy aa \ yy < k" by simp + hence "tree_less k aa" by (blast dest: tree_less_trans) + with 3 5 True show ?thesis by (auto simp: tree_less_nit) qed auto next case ("6_1" xx aa yy ss) show ?case proof (cases "xx = k") case True - with "6_1" have "tlt yy aa \ k > yy" by simp - hence "tlt k aa" by (blast dest: tlt_trans) - with "6_1" `xx = k` show ?thesis by (auto simp: tlt_nit) + with "6_1" have "tree_less yy aa \ k > yy" by simp + hence "tree_less k aa" by (blast dest: tree_less_trans) + with "6_1" `xx = k` show ?thesis by (auto simp: tree_less_nit) qed simp next case ("6_2" xx aa yy ss vaa vbb vdd vc) thus ?case proof (cases "xx = k") case True - with "6_2" have "k > yy \ tlt yy aa" by simp - hence "tlt k aa" by (blast dest: tlt_trans) - with True "6_2" show ?thesis by (auto simp: tlt_nit) + with "6_2" have "k > yy \ tree_less yy aa" by simp + hence "tree_less k aa" by (blast dest: tree_less_trans) + with True "6_2" show ?thesis by (auto simp: tree_less_nit) qed simp qed simp @@ -726,36 +741,36 @@ definition delete where delete_def: "delete k t = paint B (del k t)" -theorem delete_isrbt[simp]: assumes "isrbt t" shows "isrbt (delete k t)" +theorem delete_is_rbt[simp]: assumes "is_rbt t" shows "is_rbt (delete k t)" proof - - from assms have "inv2 t" and "inv1 t" unfolding isrbt_def by auto - hence "inv2 (del k t) \ (treec t = R \ bh (del k t) = bh t \ inv1 (del k t) \ treec t = B \ bh (del k t) = bh t - 1 \ inv1l (del k t))" by (rule del_inv1_inv2) - hence "inv2 (del k t) \ inv1l (del k t)" by (cases "treec t") auto + from assms have "inv2 t" and "inv1 t" unfolding is_rbt_def by auto + hence "inv2 (del k t) \ (color_of t = R \ bheight (del k t) = bheight t \ inv1 (del k t) \ color_of t = B \ bheight (del k t) = bheight t - 1 \ inv1l (del k t))" by (rule del_inv1_inv2) + hence "inv2 (del k t) \ inv1l (del k t)" by (cases "color_of t") auto with assms show ?thesis - unfolding isrbt_def delete_def - by (auto intro: paint_st del_st) + unfolding is_rbt_def delete_def + by (auto intro: paint_sorted del_sorted) qed lemma delete_pit: - assumes "isrbt t" - shows "pin_tree k v (delete x t) = (x \ k \ pin_tree k v t)" - using assms unfolding isrbt_def delete_def + assumes "is_rbt t" + shows "entry_in_tree k v (delete x t) = (x \ k \ entry_in_tree k v t)" + using assms unfolding is_rbt_def delete_def by (auto simp: del_pit) -lemma map_of_delete: - assumes isrbt: "isrbt t" - shows "map_of (delete k t) = (map_of t)|`(-{k})" +lemma lookup_delete: + assumes is_rbt: "is_rbt t" + shows "lookup (delete k t) = (lookup t)|`(-{k})" proof fix x - show "map_of (delete k t) x = (map_of t |` (-{k})) x" + show "lookup (delete k t) x = (lookup t |` (-{k})) x" proof (cases "x = k") assume "x = k" - with isrbt show ?thesis - by (cases "map_of (delete k t) k") (auto simp: mapof_pit delete_pit) + with is_rbt show ?thesis + by (cases "lookup (delete k t) k") (auto simp: lookup_pit delete_pit) next assume "x \ k" thus ?thesis - by auto (metis isrbt delete_isrbt delete_pit isrbt_st mapof_from_pit) + by auto (metis is_rbt delete_is_rbt delete_pit is_rbt_sorted lookup_from_pit) qed qed @@ -765,43 +780,43 @@ unionwithkey :: "('a\linorder \ 'b \ 'b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt" where "unionwithkey f t Empty = t" -| "unionwithkey f t (Tr c lt k v rt) = unionwithkey f (unionwithkey f (insertwithkey f k v t) lt) rt" +| "unionwithkey f t (Branch c lt k v rt) = unionwithkey f (unionwithkey f (insertwithkey f k v t) lt) rt" -lemma unionwk_st: "st lt \ st (unionwithkey f lt rt)" - by (induct rt arbitrary: lt) (auto simp: insertwk_st) -theorem unionwk_isrbt[simp]: "isrbt lt \ isrbt (unionwithkey f lt rt)" - by (induct rt arbitrary: lt) (simp add: insertwk_isrbt)+ +lemma unionwk_sorted: "sorted lt \ sorted (unionwithkey f lt rt)" + by (induct rt arbitrary: lt) (auto simp: insertwk_sorted) +theorem unionwk_is_rbt[simp]: "is_rbt lt \ is_rbt (unionwithkey f lt rt)" + by (induct rt arbitrary: lt) (simp add: insertwk_is_rbt)+ definition unionwith where "unionwith f = unionwithkey (\_. f)" -theorem unionw_isrbt: "isrbt lt \ isrbt (unionwith f lt rt)" unfolding unionwith_def by simp +theorem unionw_is_rbt: "is_rbt lt \ is_rbt (unionwith f lt rt)" unfolding unionwith_def by simp definition union where "union = unionwithkey (%_ _ rv. rv)" -theorem union_isrbt: "isrbt lt \ isrbt (union lt rt)" unfolding union_def by simp +theorem union_is_rbt: "is_rbt lt \ is_rbt (union lt rt)" unfolding union_def by simp -lemma union_Tr[simp]: - "union t (Tr c lt k v rt) = union (union (insrt k v t) lt) rt" - unfolding union_def insrt_def +lemma union_Branch[simp]: + "union t (Branch c lt k v rt) = union (union (insert k v t) lt) rt" + unfolding union_def insert_def by simp -lemma map_of_union: - assumes "isrbt s" "st t" - shows "map_of (union s t) = map_of s ++ map_of t" +lemma lookup_union: + assumes "is_rbt s" "sorted t" + shows "lookup (union s t) = lookup s ++ lookup t" using assms proof (induct t arbitrary: s) case Empty thus ?case by (auto simp: union_def) next - case (Tr c l k v r s) - hence strl: "st r" "st l" "l |\ k" "k \| r" by auto + case (Branch c l k v r s) + hence sortedrl: "sorted r" "sorted l" "l |\ k" "k \| r" by auto - have meq: "map_of s(k \ v) ++ map_of l ++ map_of r = - map_of s ++ - (\a. if a < k then map_of l a - else if k < a then map_of r a else Some v)" (is "?m1 = ?m2") + have meq: "lookup s(k \ v) ++ lookup l ++ lookup r = + lookup s ++ + (\a. if a < k then lookup l a + else if k < a then lookup r a else Some v)" (is "?m1 = ?m2") proof (rule ext) fix a @@ -809,7 +824,7 @@ thus "?m1 a = ?m2 a" proof (elim disjE) assume "k < a" - with `l |\ k` have "l |\ a" by (rule tlt_trans) + with `l |\ k` have "l |\ a" by (rule tree_less_trans) with `k < a` show ?thesis by (auto simp: map_add_def split: option.splits) next @@ -818,20 +833,20 @@ show ?thesis by (auto simp: map_add_def) next assume "a < k" - from this `k \| r` have "a \| r" by (rule tgt_trans) + from this `k \| r` have "a \| r" by (rule tree_greater_trans) with `a < k` show ?thesis by (auto simp: map_add_def split: option.splits) qed qed - from Tr + from Branch have IHs: - "map_of (union (union (insrt k v s) l) r) = map_of (union (insrt k v s) l) ++ map_of r" - "map_of (union (insrt k v s) l) = map_of (insrt k v s) ++ map_of l" - by (auto intro: union_isrbt insrt_isrbt) + "lookup (union (union (insert k v s) l) r) = lookup (union (insert k v s) l) ++ lookup r" + "lookup (union (insert k v s) l) = lookup (insert k v s) ++ lookup l" + by (auto intro: union_is_rbt insert_is_rbt) with meq show ?case - by (auto simp: map_of_insert[OF Tr(3)]) + by (auto simp: lookup_insert[OF Branch(3)]) qed subsection {* Adjust *} @@ -840,33 +855,33 @@ adjustwithkey :: "('a \ 'b \ 'b) \ ('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt" where "adjustwithkey f k Empty = Empty" -| "adjustwithkey f k (Tr c lt x v rt) = (if k < x then (Tr c (adjustwithkey f k lt) x v rt) else if k > x then (Tr c lt x v (adjustwithkey f k rt)) else (Tr c lt x (f x v) rt))" +| "adjustwithkey f k (Branch c lt x v rt) = (if k < x then (Branch c (adjustwithkey f k lt) x v rt) else if k > x then (Branch c lt x v (adjustwithkey f k rt)) else (Branch c lt x (f x v) rt))" -lemma adjustwk_treec: "treec (adjustwithkey f k t) = treec t" by (induct t) simp+ -lemma adjustwk_inv1: "inv1 (adjustwithkey f k t) = inv1 t" by (induct t) (simp add: adjustwk_treec)+ -lemma adjustwk_inv2: "inv2 (adjustwithkey f k t) = inv2 t" "bh (adjustwithkey f k t) = bh t" by (induct t) simp+ -lemma adjustwk_tgt: "tgt k (adjustwithkey f kk t) = tgt k t" by (induct t) simp+ -lemma adjustwk_tlt: "tlt k (adjustwithkey f kk t) = tlt k t" by (induct t) simp+ -lemma adjustwk_st: "st (adjustwithkey f k t) = st t" by (induct t) (simp add: adjustwk_tlt adjustwk_tgt)+ +lemma adjustwk_color_of: "color_of (adjustwithkey f k t) = color_of t" by (induct t) simp+ +lemma adjustwk_inv1: "inv1 (adjustwithkey f k t) = inv1 t" by (induct t) (simp add: adjustwk_color_of)+ +lemma adjustwk_inv2: "inv2 (adjustwithkey f k t) = inv2 t" "bheight (adjustwithkey f k t) = bheight t" by (induct t) simp+ +lemma adjustwk_tree_greater: "tree_greater k (adjustwithkey f kk t) = tree_greater k t" by (induct t) simp+ +lemma adjustwk_tree_less: "tree_less k (adjustwithkey f kk t) = tree_less k t" by (induct t) simp+ +lemma adjustwk_sorted: "sorted (adjustwithkey f k t) = sorted t" by (induct t) (simp add: adjustwk_tree_less adjustwk_tree_greater)+ -theorem adjustwk_isrbt[simp]: "isrbt (adjustwithkey f k t) = isrbt t" -unfolding isrbt_def by (simp add: adjustwk_inv2 adjustwk_treec adjustwk_st adjustwk_inv1 ) +theorem adjustwk_is_rbt[simp]: "is_rbt (adjustwithkey f k t) = is_rbt t" +unfolding is_rbt_def by (simp add: adjustwk_inv2 adjustwk_color_of adjustwk_sorted adjustwk_inv1 ) theorem adjustwithkey_map[simp]: - "map_of (adjustwithkey f k t) x = - (if x = k then case map_of t x of None \ None | Some y \ Some (f k y) - else map_of t x)" + "lookup (adjustwithkey f k t) x = + (if x = k then case lookup t x of None \ None | Some y \ Some (f k y) + else lookup t x)" by (induct t arbitrary: x) (auto split:option.splits) definition adjust where "adjust f = adjustwithkey (\_. f)" -theorem adjust_isrbt[simp]: "isrbt (adjust f k t) = isrbt t" unfolding adjust_def by simp +theorem adjust_is_rbt[simp]: "is_rbt (adjust f k t) = is_rbt t" unfolding adjust_def by simp theorem adjust_map[simp]: - "map_of (adjust f k t) x = - (if x = k then case map_of t x of None \ None | Some y \ Some (f y) - else map_of t x)" + "lookup (adjust f k t) x = + (if x = k then case lookup t x of None \ None | Some y \ Some (f y) + else lookup t x)" unfolding adjust_def by simp subsection {* Map *} @@ -875,27 +890,27 @@ mapwithkey :: "('a::linorder \ 'b \ 'c) \ ('a,'b) rbt \ ('a,'c) rbt" where "mapwithkey f Empty = Empty" -| "mapwithkey f (Tr c lt k v rt) = Tr c (mapwithkey f lt) k (f k v) (mapwithkey f rt)" +| "mapwithkey f (Branch c lt k v rt) = Branch c (mapwithkey f lt) k (f k v) (mapwithkey f rt)" theorem mapwk_keys[simp]: "keys (mapwithkey f t) = keys t" by (induct t) auto -lemma mapwk_tgt: "tgt k (mapwithkey f t) = tgt k t" by (induct t) simp+ -lemma mapwk_tlt: "tlt k (mapwithkey f t) = tlt k t" by (induct t) simp+ -lemma mapwk_st: "st (mapwithkey f t) = st t" by (induct t) (simp add: mapwk_tlt mapwk_tgt)+ -lemma mapwk_treec: "treec (mapwithkey f t) = treec t" by (induct t) simp+ -lemma mapwk_inv1: "inv1 (mapwithkey f t) = inv1 t" by (induct t) (simp add: mapwk_treec)+ -lemma mapwk_inv2: "inv2 (mapwithkey f t) = inv2 t" "bh (mapwithkey f t) = bh t" by (induct t) simp+ -theorem mapwk_isrbt[simp]: "isrbt (mapwithkey f t) = isrbt t" -unfolding isrbt_def by (simp add: mapwk_inv1 mapwk_inv2 mapwk_st mapwk_treec) +lemma mapwk_tree_greater: "tree_greater k (mapwithkey f t) = tree_greater k t" by (induct t) simp+ +lemma mapwk_tree_less: "tree_less k (mapwithkey f t) = tree_less k t" by (induct t) simp+ +lemma mapwk_sorted: "sorted (mapwithkey f t) = sorted t" by (induct t) (simp add: mapwk_tree_less mapwk_tree_greater)+ +lemma mapwk_color_of: "color_of (mapwithkey f t) = color_of t" by (induct t) simp+ +lemma mapwk_inv1: "inv1 (mapwithkey f t) = inv1 t" by (induct t) (simp add: mapwk_color_of)+ +lemma mapwk_inv2: "inv2 (mapwithkey f t) = inv2 t" "bheight (mapwithkey f t) = bheight t" by (induct t) simp+ +theorem mapwk_is_rbt[simp]: "is_rbt (mapwithkey f t) = is_rbt t" +unfolding is_rbt_def by (simp add: mapwk_inv1 mapwk_inv2 mapwk_sorted mapwk_color_of) -theorem map_of_mapwk[simp]: "map_of (mapwithkey f t) x = Option.map (f x) (map_of t x)" +theorem lookup_mapwk[simp]: "lookup (mapwithkey f t) x = Option.map (f x) (lookup t x)" by (induct t) auto definition map where map_def: "map f == mapwithkey (\_. f)" theorem map_keys[simp]: "keys (map f t) = keys t" unfolding map_def by simp -theorem map_isrbt[simp]: "isrbt (map f t) = isrbt t" unfolding map_def by simp -theorem map_of_map[simp]: "map_of (map f t) = Option.map f o map_of t" +theorem map_is_rbt[simp]: "is_rbt (map f t) = is_rbt t" unfolding map_def by simp +theorem lookup_map[simp]: "lookup (map f t) = Option.map f o lookup t" by (rule ext) (simp add:map_def) subsection {* Fold *} @@ -906,62 +921,57 @@ foldwithkey :: "('a::linorder \ 'b \ 'c \ 'c) \ ('a,'b) rbt \ 'c \ 'c" where "foldwithkey f Empty v = v" -| "foldwithkey f (Tr c lt k x rt) v = foldwithkey f rt (f k x (foldwithkey f lt v))" +| "foldwithkey f (Branch c lt k x rt) v = foldwithkey f rt (f k x (foldwithkey f lt v))" -primrec alist_of -where - "alist_of Empty = []" -| "alist_of (Tr _ l k v r) = alist_of l @ (k,v) # alist_of r" - -lemma map_of_alist_of_aux: "st (Tr c t1 k v t2) \ RBT.map_of (Tr c t1 k v t2) = RBT.map_of t2 ++ [k\v] ++ RBT.map_of t1" +lemma lookup_entries_aux: "sorted (Branch c t1 k v t2) \ RBT.lookup (Branch c t1 k v t2) = RBT.lookup t2 ++ [k\v] ++ RBT.lookup t1" proof (rule ext) fix x - assume ST: "st (Tr c t1 k v t2)" - let ?thesis = "RBT.map_of (Tr c t1 k v t2) x = (RBT.map_of t2 ++ [k \ v] ++ RBT.map_of t1) x" + assume SORTED: "sorted (Branch c t1 k v t2)" + let ?thesis = "RBT.lookup (Branch c t1 k v t2) x = (RBT.lookup t2 ++ [k \ v] ++ RBT.lookup t1) x" - have DOM_T1: "!!k'. k'\dom (RBT.map_of t1) \ k>k'" + have DOM_T1: "!!k'. k'\dom (RBT.lookup t1) \ k>k'" proof - fix k' - from ST have "t1 |\ k" by simp - with tlt_prop have "\k'\keys t1. k>k'" by auto - moreover assume "k'\dom (RBT.map_of t1)" - ultimately show "k>k'" using RBT.mapof_keys ST by auto + from SORTED have "t1 |\ k" by simp + with tree_less_prop have "\k'\keys t1. k>k'" by auto + moreover assume "k'\dom (RBT.lookup t1)" + ultimately show "k>k'" using RBT.lookup_keys SORTED by auto qed - have DOM_T2: "!!k'. k'\dom (RBT.map_of t2) \ kdom (RBT.lookup t2) \ k| t2" by simp - with tgt_prop have "\k'\keys t2. kdom (RBT.map_of t2)" - ultimately show "k| t2" by simp + with tree_greater_prop have "\k'\keys t2. kdom (RBT.lookup t2)" + ultimately show "kdom [k\v]" by simp - moreover have "x\dom (RBT.map_of t2)" proof - assume "x\dom (RBT.map_of t2)" + moreover have "x\dom (RBT.lookup t2)" proof + assume "x\dom (RBT.lookup t2)" with DOM_T2 have "k v] x" by simp - moreover have "x\dom (RBT.map_of t1)" proof - assume "x\dom (RBT.map_of t1)" + hence "RBT.lookup (Branch c t1 k v t2) x = [k \ v] x" by simp + moreover have "x\dom (RBT.lookup t1)" proof + assume "x\dom (RBT.lookup t1)" with DOM_T1 have "k>x" by blast thus False by simp qed ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps) } moreover { assume C: "x>k" - hence "RBT.map_of (Tr c t1 k v t2) x = RBT.map_of t2 x" by (simp add: less_not_sym[of k x]) + hence "RBT.lookup (Branch c t1 k v t2) x = RBT.lookup t2 x" by (simp add: less_not_sym[of k x]) moreover from C have "x\dom [k\v]" by simp - moreover have "x\dom (RBT.map_of t1)" proof - assume "x\dom (RBT.map_of t1)" + moreover have "x\dom (RBT.lookup t1)" proof + assume "x\dom (RBT.lookup t1)" with DOM_T1 have "k>x" by simp with C show False by simp qed @@ -969,35 +979,38 @@ } ultimately show ?thesis using less_linear by blast qed -lemma map_of_alist_of: - shows "st t \ Map.map_of (alist_of t) = map_of t" +lemma map_of_entries: + shows "sorted t \ map_of (entries t) = lookup t" proof (induct t) - case Empty thus ?case by (simp add: RBT.map_of_Empty) + case Empty thus ?case by (simp add: RBT.lookup_Empty) next - case (Tr c t1 k v t2) - hence "Map.map_of (alist_of (Tr c t1 k v t2)) = RBT.map_of t2 ++ [k \ v] ++ RBT.map_of t1" by simp - also note map_of_alist_of_aux[OF Tr.prems,symmetric] + case (Branch c t1 k v t2) + hence "map_of (entries (Branch c t1 k v t2)) = RBT.lookup t2 ++ [k \ v] ++ RBT.lookup t1" by simp + also note lookup_entries_aux [OF Branch.prems,symmetric] finally show ?case . qed -lemma fold_alist_fold: - "foldwithkey f t x = foldl (\x (k,v). f k v x) x (alist_of t)" +lemma fold_entries_fold: + "foldwithkey f t x = foldl (\x (k,v). f k v x) x (entries t)" by (induct t arbitrary: x) auto -lemma alist_pit[simp]: "(k, v) \ set (alist_of t) = pin_tree k v t" +lemma entries_pit[simp]: "(k, v) \ set (entries t) = entry_in_tree k v t" by (induct t) auto -lemma sorted_alist: - "st t \ sorted (List.map fst (alist_of t))" +lemma sorted_entries: + "sorted t \ List.sorted (List.map fst (entries t))" by (induct t) - (force simp: sorted_append sorted_Cons tlgt_props - dest!:pint_keys)+ + (force simp: sorted_append sorted_Cons tree_ord_props + dest!: entry_in_tree_keys)+ -lemma distinct_alist: - "st t \ distinct (List.map fst (alist_of t))" +lemma distinct_entries: + "sorted t \ distinct (List.map fst (entries t))" by (induct t) - (force simp: sorted_append sorted_Cons tlgt_props - dest!:pint_keys)+ + (force simp: sorted_append sorted_Cons tree_ord_props + dest!: entry_in_tree_keys)+ + +hide (open) const Empty insert delete entries lookup map fold union adjust sorted + (*>*) text {* @@ -1010,20 +1023,20 @@ text {* The type @{typ "('k, 'v) rbt"} denotes red-black trees with keys of type @{typ "'k"} and values of type @{typ "'v"}. To function - properly, the key type must belong to the @{text "linorder"} class. + properly, the key type musorted belong to the @{text "linorder"} class. A value @{term t} of this type is a valid red-black tree if it - satisfies the invariant @{text "isrbt t"}. + satisfies the invariant @{text "is_rbt t"}. This theory provides lemmas to prove that the invariant is satisfied throughout the computation. - The interpretation function @{const "map_of"} returns the partial + The interpretation function @{const "RBT.lookup"} returns the partial map represented by a red-black tree: - @{term_type[display] "map_of"} + @{term_type[display] "RBT.lookup"} This function should be used for reasoning about the semantics of the RBT operations. Furthermore, it implements the lookup functionality for - the data structure: It is executable and the lookup is performed in + the data sortedructure: It is executable and the lookup is performed in $O(\log n)$. *} @@ -1032,19 +1045,19 @@ text {* Currently, the following operations are supported: - @{term_type[display] "Empty"} + @{term_type[display] "RBT.Empty"} Returns the empty tree. $O(1)$ - @{term_type[display] "insrt"} + @{term_type[display] "RBT.insert"} Updates the map at a given position. $O(\log n)$ - @{term_type[display] "delete"} + @{term_type[display] "RBT.delete"} Deletes a map entry at a given position. $O(\log n)$ - @{term_type[display] "union"} + @{term_type[display] "RBT.union"} Forms the union of two trees, preferring entries from the first one. - @{term_type[display] "map"} + @{term_type[display] "RBT.map"} Maps a function over the values of a map. $O(n)$ *} @@ -1053,47 +1066,47 @@ text {* \noindent - @{thm Empty_isrbt}\hfill(@{text "Empty_isrbt"}) + @{thm Empty_is_rbt}\hfill(@{text "Empty_is_rbt"}) \noindent - @{thm insrt_isrbt}\hfill(@{text "insrt_isrbt"}) + @{thm insert_is_rbt}\hfill(@{text "insert_is_rbt"}) \noindent - @{thm delete_isrbt}\hfill(@{text "delete_isrbt"}) + @{thm delete_is_rbt}\hfill(@{text "delete_is_rbt"}) \noindent - @{thm union_isrbt}\hfill(@{text "union_isrbt"}) + @{thm union_is_rbt}\hfill(@{text "union_is_rbt"}) \noindent - @{thm map_isrbt}\hfill(@{text "map_isrbt"}) + @{thm map_is_rbt}\hfill(@{text "map_is_rbt"}) *} subsection {* Map Semantics *} text {* \noindent - \underline{@{text "map_of_Empty"}} - @{thm[display] map_of_Empty} + \underline{@{text "lookup_Empty"}} + @{thm[display] lookup_Empty} \vspace{1ex} \noindent - \underline{@{text "map_of_insert"}} - @{thm[display] map_of_insert} + \underline{@{text "lookup_insert"}} + @{thm[display] lookup_insert} \vspace{1ex} \noindent - \underline{@{text "map_of_delete"}} - @{thm[display] map_of_delete} + \underline{@{text "lookup_delete"}} + @{thm[display] lookup_delete} \vspace{1ex} \noindent - \underline{@{text "map_of_union"}} - @{thm[display] map_of_union} + \underline{@{text "lookup_union"}} + @{thm[display] lookup_union} \vspace{1ex} \noindent - \underline{@{text "map_of_map"}} - @{thm[display] map_of_map} + \underline{@{text "lookup_map"}} + @{thm[display] lookup_map} \vspace{1ex} *}