# HG changeset patch # User haftmann # Date 1267633305 -3600 # Node ID e2bc7f8d8d51159a74949ebc68a535a85772aa5e # Parent ede0b67432f334d6a665fd20934eaa4fbb996641 restructured RBT theory diff -r ede0b67432f3 -r e2bc7f8d8d51 src/HOL/Library/RBT.thy --- a/src/HOL/Library/RBT.thy Wed Mar 03 10:06:12 2010 +0100 +++ b/src/HOL/Library/RBT.thy Wed Mar 03 17:21:45 2010 +0100 @@ -10,6 +10,8 @@ imports Main AssocList begin +subsection {* Datatype of RB trees *} + datatype color = R | B datatype ('a, 'b) rbt = Empty | Branch color "('a, 'b) rbt" 'a 'b "('a, 'b) rbt" @@ -23,39 +25,48 @@ case (Branch c) with that show thesis by (cases c) blast+ qed -text {* Content of a tree *} +subsection {* Tree properties *} -primrec entries +subsubsection {* Content of a tree *} + +primrec entries :: "('a, 'b) rbt \ ('a \ 'b) list" where "entries Empty = []" | "entries (Branch _ l k v r) = entries l @ (k,v) # entries r" -text {* Search tree properties *} - -primrec entry_in_tree :: "'a \ 'b \ ('a, 'b) rbt \ bool" +abbreviation (input) entry_in_tree :: "'a \ 'b \ ('a, 'b) rbt \ bool" where - "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" + "entry_in_tree k v t \ (k, v) \ set (entries t)" + +definition keys :: "('a, 'b) rbt \ 'a list" where + "keys t = map fst (entries t)" -primrec keys :: "('k, 'v) rbt \ 'k set" -where - "keys Empty = {}" -| "keys (Branch _ l k _ r) = { k } \ keys l \ keys r" +lemma keys_simps [simp, code]: + "keys Empty = []" + "keys (Branch c l k v r) = keys l @ k # keys r" + by (simp_all add: keys_def) lemma entry_in_tree_keys: - "entry_in_tree k v t \ k \ keys t" - by (induct t) auto + assumes "(k, v) \ set (entries t)" + shows "k \ set (keys t)" +proof - + from assms have "fst (k, v) \ fst ` set (entries t)" by (rule imageI) + then show ?thesis by (simp add: keys_def) +qed + + +subsubsection {* Search tree properties *} definition tree_less :: "'a\order \ ('a, 'b) rbt \ bool" where - tree_less_prop: "tree_less k t \ (\x\keys t. x < k)" + tree_less_prop: "tree_less k t \ (\x\set (keys t). x < k)" abbreviation tree_less_symbol (infix "|\" 50) where "t |\ x \ tree_less 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)" + tree_greater_prop: "tree_greater k t = (\x\set (keys t). k < x)" lemma tree_less_simps [simp]: "tree_less k Empty = True" @@ -72,55 +83,172 @@ lemmas tree_greater_nit = tree_greater_prop entry_in_tree_keys lemmas tree_less_nit = tree_less_prop entry_in_tree_keys -lemma tree_less_trans: "t |\ x \ x < y \ t |\ y" +lemma tree_less_eq_trans: "l |\ u \ u \ v \ l |\ v" + and tree_less_trans: "t |\ x \ x < y \ t |\ y" + and tree_greater_eq_trans: "u \ v \ v \| r \ u \| r" and tree_greater_trans: "x < y \ y \| t \ x \| t" -by (auto simp: tree_ord_props) + by (auto simp: tree_ord_props) primrec sorted :: "('a::linorder, 'b) rbt \ bool" where "sorted Empty = True" | "sorted (Branch c l k v r) = (l |\ k \ k \| r \ sorted l \ sorted r)" +lemma sorted_entries: + "sorted t \ List.sorted (List.map fst (entries t))" +by (induct t) + (force simp: sorted_append sorted_Cons tree_ord_props + dest!: entry_in_tree_keys)+ + +lemma distinct_entries: + "sorted t \ distinct (List.map fst (entries t))" +by (induct t) + (force simp: sorted_append sorted_Cons tree_ord_props + dest!: entry_in_tree_keys)+ + + +subsubsection {* Tree lookup *} + 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_keys: "sorted t \ dom (lookup t) = set (keys t)" + by (induct t) (auto simp: dom_def tree_greater_prop tree_less_prop) + +lemma dom_lookup_Branch: + "sorted (Branch c t1 k v t2) \ + dom (lookup (Branch c t1 k v t2)) + = Set.insert k (dom (lookup t1) \ dom (lookup t2))" +proof - + assume "sorted (Branch c t1 k v t2)" + moreover from this have "sorted t1" "sorted t2" by simp_all + ultimately show ?thesis by (simp add: lookup_keys) +qed + +lemma finite_dom_lookup [simp, intro!]: "finite (dom (lookup t))" +proof (induct t) + case Empty then show ?case by simp +next + case (Branch color t1 a b t2) + let ?A = "Set.insert a (dom (lookup t1) \ dom (lookup t2))" + have "dom (lookup (Branch color t1 a b t2)) \ ?A" by (auto split: split_if_asm) + moreover from Branch have "finite (insert a (dom (lookup t1) \ dom (lookup t2)))" by simp + ultimately show ?case by (rule finite_subset) +qed + lemma lookup_tree_less[simp]: "t |\ k \ lookup t k = None" by (induct t) auto lemma lookup_tree_greater[simp]: "k \| t \ lookup t k = None" by (induct t) auto -lemma lookup_keys: "sorted t \ dom (lookup t) = keys t" -by (induct t) (auto simp: dom_def tree_greater_prop tree_less_prop) - -lemma lookup_pit: "sorted t \ (lookup t k = Some v) = entry_in_tree k v t" +lemma lookup_in_tree: "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 lookup_Empty: "lookup Empty = empty" by (rule ext) simp +lemma lookup_map_of_entries: + "sorted t \ lookup t = map_of (entries t)" +proof (induct t) + case Empty thus ?case by (simp add: lookup_Empty) +next + case (Branch c t1 k v t2) + have "lookup (Branch c t1 k v t2) = lookup t2 ++ [k\v] ++ lookup t1" + proof (rule ext) + fix x + from Branch have SORTED: "sorted (Branch c t1 k v t2)" by simp + let ?thesis = "lookup (Branch c t1 k v t2) x = (lookup t2 ++ [k \ v] ++ lookup t1) x" + + have DOM_T1: "!!k'. k'\dom (lookup t1) \ k>k'" + proof - + fix k' + from SORTED have "t1 |\ k" by simp + with tree_less_prop have "\k'\set (keys t1). k>k'" by auto + moreover assume "k'\dom (lookup t1)" + ultimately show "k>k'" using lookup_keys SORTED by auto + qed + + have DOM_T2: "!!k'. k'\dom (lookup t2) \ k| t2" by simp + with tree_greater_prop have "\k'\set (keys t2). kdom (lookup t2)" + ultimately show "kdom [k\v]" by simp + moreover have "x\dom (lookup t2)" proof + assume "x\dom (lookup t2)" + with DOM_T2 have "k v] x" by simp + moreover have "x\dom (lookup t1)" proof + assume "x\dom (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 "lookup (Branch c t1 k v t2) x = 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 (lookup t1)" proof + assume "x\dom (lookup t1)" + with DOM_T1 have "k>x" by simp + with C show False by simp + qed + ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps) + } ultimately show ?thesis using less_linear by blast + qed + also from Branch have "lookup t2 ++ [k \ v] ++ lookup t1 = map_of (entries (Branch c t1 k v t2))" by simp + finally show ?case . +qed + +(*lemma map_of_inject: + assumes distinct: "distinct (map fst xs)" "distinct (map fst ys)" + shows "map_of xs = map_of ys \ set xs = set ys" + +lemma entries_eqI: + assumes sorted: "sorted t1" "sorted t2" + assumes lookup: "lookup t1 = lookup t2" + shows entries: "entries t1 = entries t2" +proof - + from sorted lookup have "map_of (entries t1) = map_of (entries t2)" + by (simp_all add: lookup_map_of_entries) +qed*) + (* a kind of extensionality *) -lemma lookup_from_pit: +lemma lookup_from_in_tree: 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. \ entry_in_tree k v t1" - by (simp add: lookup_pit[symmetric] sorted) + by (simp add: lookup_in_tree[symmetric] sorted) with None show ?thesis - by (cases "lookup t2 k") (auto simp: lookup_pit sorted eq) + by (cases "lookup t2 k") (auto simp: lookup_in_tree sorted eq) next case (Some a) then show ?thesis apply (cases "lookup t2 k") - apply (auto simp: lookup_pit sorted eq) - by (auto simp add: lookup_pit[symmetric] sorted Some) + apply (auto simp: lookup_in_tree sorted eq) + by (auto simp add: lookup_in_tree[symmetric] sorted Some) qed -subsection {* Red-black properties *} + +subsubsection {* Red-black properties *} primrec color_of :: "('a, 'b) rbt \ color" where @@ -240,19 +368,23 @@ with 1 "5_4" show ?case by simp qed simp+ -lemma keys_balance[simp]: - "keys (balance l k v r) = { k } \ keys l \ keys r" -by (induct l k v r rule: balance.induct) auto +lemma entries_balance [simp]: + "entries (balance l k v r) = entries l @ (k, v) # entries r" + by (induct l k v r rule: balance.induct) auto -lemma balance_pit: - "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 keys_balance [simp]: + "keys (balance l k v r) = keys l @ k # keys r" + by (simp add: keys_def) + +lemma balance_in_tree: + "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 (auto simp add: keys_def) lemma lookup_balance[simp]: fixes k :: "'a::linorder" 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) +by (rule lookup_from_in_tree) (auto simp:assms balance_in_tree balance_sorted) primrec paint :: "color \ ('a,'b) rbt \ ('a,'b) rbt" where @@ -264,7 +396,7 @@ lemma paint_inv2[simp]: "inv2 t \ inv2 (paint c t)" 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_in_tree[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 @@ -294,8 +426,8 @@ 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 keys_ins: "set (keys (ins f k v t)) = { k } \ set (keys t)" + by (induct f k v t rule: ins.induct) auto lemma lookup_ins: fixes k :: "'a::linorder" @@ -305,45 +437,45 @@ using assms by (induct f k v t rule: ins.induct) auto definition - insertwithkey :: "('a\linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" + insert_with_key :: "('a\linorder \ 'b \ 'b \ 'b) \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "insertwithkey f k v t = paint B (ins f k v t)" + "insert_with_key f k v t = paint B (ins f k v t)" -lemma insertwk_sorted: "sorted t \ sorted (insertwithkey f k x t)" - by (auto simp: insertwithkey_def) +lemma insertwk_sorted: "sorted t \ sorted (insert_with_key f k x t)" + by (auto simp: insert_with_key_def) theorem insertwk_is_rbt: assumes inv: "is_rbt t" - shows "is_rbt (insertwithkey f k x t)" + shows "is_rbt (insert_with_key f k x t)" using assms -unfolding insertwithkey_def is_rbt_def +unfolding insert_with_key_def is_rbt_def by (auto simp: ins_inv1_inv2) lemma lookup_insertwk: assumes "sorted t" - shows "lookup (insertwithkey f k v t) x = ((lookup t)(k |-> case lookup t k of None \ v + shows "lookup (insert_with_key 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 +unfolding insert_with_key_def using assms by (simp add:lookup_ins) definition - insertw_def: "insertwith f = insertwithkey (\_. f)" + insertw_def: "insert_with f = insert_with_key (\_. f)" -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 insertw_sorted: "sorted t \ sorted (insert_with f k v t)" by (simp add: insertwk_sorted insertw_def) +theorem insertw_is_rbt: "is_rbt t \ is_rbt (insert_with f k v t)" by (simp add: insertwk_is_rbt insertw_def) 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))" + shows "lookup (insert_with 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 "lookup t k", auto simp:lookup_insertwk dom_def) definition insert :: "'a\linorder \ 'b \ ('a, 'b) rbt \ ('a, 'b) rbt" where - "insert k v t = insertwithkey (\_ _ nv. nv) k v t" + "insert = insert_with_key (\_ _ nv. nv)" 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) +theorem insert_is_rbt [simp]: "is_rbt t \ is_rbt (insert k v t)" by (simp add: insertwk_is_rbt insert_def) lemma lookup_insert: assumes "is_rbt t" @@ -359,178 +491,174 @@ by (cases t rule: rbt_cases) auto fun - balleft :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" + balance_left :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "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" + "balance_left (Branch R a k x b) s y c = Branch R (Branch B a k x b) s y c" | + "balance_left bl k x (Branch B a s y b) = balance bl k x (Branch R a s y b)" | + "balance_left 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))" | + "balance_left t k x s = Empty" -lemma balleft_inv2_with_inv1: +lemma balance_left_inv2_with_inv1: 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)" + shows "bheight (balance_left lt k v rt) = bheight lt + 1" + and "inv2 (balance_left lt k v rt)" using assms -by (induct lt k v rt rule: balleft.induct) (auto simp: balance_inv2 balance_bheight) +by (induct lt k v rt rule: balance_left.induct) (auto simp: balance_inv2 balance_bheight) -lemma balleft_inv2_app: +lemma balance_left_inv2_app: assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "color_of rt = B" - shows "inv2 (balleft lt k v rt)" - "bheight (balleft lt k v rt) = bheight rt" + shows "inv2 (balance_left lt k v rt)" + "bheight (balance_left lt k v rt) = bheight rt" using assms -by (induct lt k v rt rule: balleft.induct) (auto simp add: balance_inv2 balance_bheight)+ +by (induct lt k v rt rule: balance_left.induct) (auto simp add: balance_inv2 balance_bheight)+ -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 balance_left_inv1: "\inv1l a; inv1 b; color_of b = B\ \ inv1 (balance_left a k x b)" + by (induct a k x b rule: balance_left.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 balance_left_inv1l: "\ inv1l lt; inv1 rt \ \ inv1l (balance_left lt k x rt)" +by (induct lt k x rt rule: balance_left.induct) (auto simp: balance_inv1) -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) +lemma balance_left_sorted: "\ sorted l; sorted r; tree_less k l; tree_greater k r \ \ sorted (balance_left l k v r)" +apply (induct l k v r rule: balance_left.induct) apply (auto simp: balance_sorted) apply (unfold tree_greater_prop tree_less_prop) by force+ -lemma balleft_tree_greater: +lemma balance_left_tree_greater: fixes k :: "'a::order" assumes "k \| a" "k \| b" "k < x" - shows "k \| balleft a x t b" + shows "k \| balance_left a x t b" using assms -by (induct a x t b rule: balleft.induct) auto +by (induct a x t b rule: balance_left.induct) auto -lemma balleft_tree_less: +lemma balance_left_tree_less: fixes k :: "'a::order" assumes "a |\ k" "b |\ k" "x < k" - shows "balleft a x t b |\ k" + shows "balance_left a x t b |\ k" using assms -by (induct a x t b rule: balleft.induct) auto +by (induct a x t b rule: balance_left.induct) auto -lemma balleft_pit: +lemma balance_left_in_tree: 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)" + shows "entry_in_tree k v (balance_left 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) +by (induct l k v r rule: balance_left.induct) (auto simp: balance_in_tree) fun - balright :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" + balance_right :: "('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" where - "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" + "balance_right a k x (Branch R b s y c) = Branch R a k x (Branch B b s y c)" | + "balance_right (Branch B a k x b) s y bl = balance (Branch R a k x b) s y bl" | + "balance_right (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)" | + "balance_right t k x s = Empty" -lemma balright_inv2_with_inv1: +lemma balance_right_inv2_with_inv1: 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" + shows "inv2 (balance_right lt k v rt) \ bheight (balance_right lt k v rt) = bheight lt" using assms -by (induct lt k v rt rule: balright.induct) (auto simp: balance_inv2 balance_bheight) +by (induct lt k v rt rule: balance_right.induct) (auto simp: balance_inv2 balance_bheight) -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 balance_right_inv1: "\inv1 a; inv1l b; color_of a = B\ \ inv1 (balance_right a k x b)" +by (induct a k x b rule: balance_right.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 balance_right_inv1l: "\ inv1 lt; inv1l rt \ \inv1l (balance_right lt k x rt)" +by (induct lt k x rt rule: balance_right.induct) (auto simp: balance_inv1) -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) +lemma balance_right_sorted: "\ sorted l; sorted r; tree_less k l; tree_greater k r \ \ sorted (balance_right l k v r)" +apply (induct l k v r rule: balance_right.induct) apply (auto simp:balance_sorted) apply (unfold tree_less_prop tree_greater_prop) by force+ -lemma balright_tree_greater: +lemma balance_right_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 + shows "k \| balance_right a x t b" +using assms by (induct a x t b rule: balance_right.induct) auto -lemma balright_tree_less: +lemma balance_right_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 + shows "balance_right a x t b |\ k" +using assms by (induct a x t b rule: balance_right.induct) auto -lemma balright_pit: +lemma balance_right_in_tree: 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) - - -text {* app *} + shows "entry_in_tree x y (balance_right 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: balance_right.induct) (auto simp: balance_in_tree) fun - app :: "('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt" + combine :: "('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt" where - "app Empty x = x" -| "app x Empty = x" -| "app (Branch R a k x b) (Branch R c s y d) = (case (app b c) of + "combine Empty x = x" +| "combine x Empty = x" +| "combine (Branch R a k x b) (Branch R c s y d) = (case (combine 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 +| "combine (Branch B a k x b) (Branch B c s y d) = (case (combine 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)" + bc \ balance_left a k x (Branch B bc s y d))" +| "combine a (Branch R b k x c) = Branch R (combine a b) k x c" +| "combine (Branch R a k x b) c = Branch R a k x (combine b c)" -lemma app_inv2: +lemma combine_inv2: assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt" - shows "bheight (app lt rt) = bheight lt" "inv2 (app lt rt)" + shows "bheight (combine lt rt) = bheight lt" "inv2 (combine lt rt)" using assms -by (induct lt rt rule: app.induct) - (auto simp: balleft_inv2_app split: rbt.splits color.splits) +by (induct lt rt rule: combine.induct) + (auto simp: balance_left_inv2_app split: rbt.splits color.splits) -lemma app_inv1: +lemma combine_inv1: assumes "inv1 lt" "inv1 rt" - shows "color_of lt = B \ color_of rt = B \ inv1 (app lt rt)" - "inv1l (app lt rt)" + shows "color_of lt = B \ color_of rt = B \ inv1 (combine lt rt)" + "inv1l (combine lt rt)" using assms -by (induct lt rt rule: app.induct) - (auto simp: balleft_inv1 split: rbt.splits color.splits) +by (induct lt rt rule: combine.induct) + (auto simp: balance_left_inv1 split: rbt.splits color.splits) -lemma app_tree_greater[simp]: +lemma combine_tree_greater[simp]: fixes k :: "'a::linorder" assumes "k \| l" "k \| r" - shows "k \| app l r" + shows "k \| combine l r" using assms -by (induct l r rule: app.induct) - (auto simp: balleft_tree_greater split:rbt.splits color.splits) +by (induct l r rule: combine.induct) + (auto simp: balance_left_tree_greater split:rbt.splits color.splits) -lemma app_tree_less[simp]: +lemma combine_tree_less[simp]: fixes k :: "'a::linorder" assumes "l |\ k" "r |\ k" - shows "app l r |\ k" + shows "combine l r |\ k" using assms -by (induct l r rule: app.induct) - (auto simp: balleft_tree_less split:rbt.splits color.splits) +by (induct l r rule: combine.induct) + (auto simp: balance_left_tree_less split:rbt.splits color.splits) -lemma app_sorted: +lemma combine_sorted: fixes k :: "'a::linorder" assumes "sorted l" "sorted r" "l |\ k" "k \| r" - shows "sorted (app l r)" -using assms proof (induct l r rule: app.induct) + shows "sorted (combine l r)" +using assms proof (induct l r rule: combine.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" by auto with 3 show ?case - apply (cases "app b c" rule: rbt_cases) - apply auto - by (metis app_tree_greater app_tree_less ineqs ineqs tree_less_simps(2) tree_greater_simps(2) tree_greater_trans tree_less_trans)+ + by (cases "combine b c" rule: rbt_cases) + (auto, (metis combine_tree_greater combine_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 \ 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) + with 4 have 2: "tree_greater x (combine b c)" by (simp add: combine_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) + with 4 have 3: "tree_less y (combine b c)" by (simp add: combine_tree_less) show ?case - proof (cases "app b c" rule: rbt_cases) + proof (cases "combine b c" rule: rbt_cases) case Empty 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) + with Empty show ?thesis by (simp add: balance_left_sorted) next case (Red lta va ka rta) with 2 4 have "x < va \ tree_less x a" by simp @@ -542,71 +670,71 @@ case (Black lta va ka rta) 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) + with Black 2 3 4 have "sorted a" and "sorted (Branch B (combine b c) y w d)" and "tree_less x a" and "tree_greater x (Branch B (combine b c) y w d)" by auto + with Black show ?thesis by (simp add: balance_left_sorted) qed next case (5 va vb vd vc b x w c) 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) + with 5 show ?case by (simp add: combine_tree_less) next case (6 a x v b va vb vd vc) 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) + with 6 show ?case by (simp add: combine_tree_greater) qed simp+ -lemma app_pit: +lemma combine_in_tree: 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)" + shows "entry_in_tree k v (combine l r) = (entry_in_tree k v l \ entry_in_tree k v r)" using assms -proof (induct l r rule: app.induct) +proof (induct l r rule: combine.induct) case (4 _ _ _ b c) - 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) + hence a: "bheight (combine b c) = bheight b" by (simp add: combine_inv2) + from 4 have b: "inv1l (combine b c)" by (simp add: combine_inv1) show ?case - proof (cases "app b c" rule: rbt_cases) + proof (cases "combine b c" rule: rbt_cases) case Empty - with 4 a show ?thesis by (auto simp: balleft_pit) + with 4 a show ?thesis by (auto simp: balance_left_in_tree) next case (Red lta ka va rta) with 4 show ?thesis by auto next case (Black lta ka va rta) - with a b 4 show ?thesis by (auto simp: balleft_pit) + with a b 4 show ?thesis by (auto simp: balance_left_in_tree) qed qed (auto split: rbt.splits color.splits) fun - delformLeft :: "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" and - delformRight :: "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" and + del_from_left :: "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" and + del_from_right :: "('a\linorder) \ ('a,'b) rbt \ 'a \ 'b \ ('a,'b) rbt \ ('a,'b) rbt" and del :: "('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt" where "del x Empty = Empty" | - "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)" + "del x (Branch c a y s b) = (if x < y then del_from_left x a y s b else (if x > y then del_from_right x a y s b else combine a b))" | + "del_from_left x (Branch B lt z v rt) y s b = balance_left (del x (Branch B lt z v rt)) y s b" | + "del_from_left x a y s b = Branch R (del x a) y s b" | + "del_from_right x a y s (Branch B lt z v rt) = balance_right a y s (del x (Branch B lt z v rt))" | + "del_from_right x a y s b = Branch R a y s (del x b)" lemma assumes "inv2 lt" "inv1 lt" shows "\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))" + inv2 (del_from_left x lt k v rt) \ bheight (del_from_left x lt k v rt) = bheight lt \ (color_of lt = B \ color_of rt = B \ inv1 (del_from_left x lt k v rt) \ (color_of lt \ B \ color_of rt \ B) \ inv1l (del_from_left 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))" + inv2 (del_from_right x lt k v rt) \ bheight (del_from_right x lt k v rt) = bheight lt \ (color_of lt = B \ color_of rt = B \ inv1 (del_from_right x lt k v rt) \ (color_of lt \ B \ color_of rt \ B) \ inv1l (del_from_right 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) +proof (induct x lt k v rt and x lt k v rt and x lt rule: del_from_left_del_from_right_del.induct) case (2 y c _ y') have "y = y' \ y < y' \ y > y'" by auto thus ?case proof (elim disjE) assume "y = y'" - with 2 show ?thesis by (cases c) (simp add: app_inv2 app_inv1)+ + with 2 show ?thesis by (cases c) (simp add: combine_inv2 combine_inv1)+ next assume "y < y'" with 2 show ?thesis by (cases c) auto @@ -616,35 +744,35 @@ qed next case (3 y lt z v rta y' ss bb) - 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)+ + thus ?case by (cases "color_of (Branch B lt z v rta) = B \ color_of bb = B") (simp add: balance_left_inv2_with_inv1 balance_left_inv1 balance_left_inv1l)+ next case (5 y a y' ss lt z v rta) - 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)+ + thus ?case by (cases "color_of a = B \ color_of (Branch B lt z v rta) = B") (simp add: balance_right_inv2_with_inv1 balance_right_inv1 balance_right_inv1l)+ next case ("6_1" y a y' ss) thus ?case by (cases "color_of a = B \ color_of Empty = B") simp+ qed auto lemma - 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)" + del_from_left_tree_less: "\tree_less v lt; tree_less v rt; k < v\ \ tree_less v (del_from_left x lt k y rt)" + and del_from_right_tree_less: "\tree_less v lt; tree_less v rt; k < v\ \ tree_less v (del_from_right 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_tree_less balright_tree_less) +by (induct x lt k y rt and x lt k y rt and x lt rule: del_from_left_del_from_right_del.induct) + (auto simp: balance_left_tree_less balance_right_tree_less) -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)" +lemma del_from_left_tree_greater: "\tree_greater v lt; tree_greater v rt; k > v\ \ tree_greater v (del_from_left x lt k y rt)" + and del_from_right_tree_greater: "\tree_greater v lt; tree_greater v rt; k > v\ \ tree_greater v (del_from_right 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_tree_greater balright_tree_greater) +by (induct x lt k y rt and x lt k y rt and x lt rule: del_from_left_del_from_right_del.induct) + (auto simp: balance_left_tree_greater balance_right_tree_greater) -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)" +lemma "\sorted lt; sorted rt; tree_less k lt; tree_greater k rt\ \ sorted (del_from_left x lt k y rt)" + and "\sorted lt; sorted rt; tree_less k lt; tree_greater k rt\ \ sorted (del_from_right 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) +proof (induct x lt k y rt and x lt k y rt and x lt rule: del_from_left_del_from_right_del.induct) case (3 x lta zz v rta yy ss bb) 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) + with 3 show ?case by (simp add: balance_left_sorted) next case ("4_2" x vaa vbb vdd vc yy ss bb) hence "tree_less yy (Branch R vaa vbb vdd vc)" by simp @@ -654,18 +782,18 @@ case (5 x aa yy ss lta zz v rta) 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) + with 5 show ?case by (simp add: balance_right_sorted) next case ("6_2" x aa yy ss vaa vbb vdd vc) 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_sorted) +qed (auto simp: combine_sorted) -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) +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 (del_from_left 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 (del_from_right x lt kt y rt) = (False \ (x \ k \ entry_in_tree k v (Branch c lt kt y rt)))" + and del_in_tree: "\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: del_from_left_del_from_right_del.induct) case (2 xx c aa yy ss bb) have "xx = yy \ xx < yy \ xx > yy" by auto from this 2 show ?case proof (elim disjE) @@ -674,15 +802,15 @@ case True 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) + with `xx = yy` 2 `xx = k` show ?thesis by (simp add: combine_in_tree) + qed (simp add: combine_in_tree) qed simp+ next case (3 xx lta zz vv rta yy ss bb) def mt[simp]: mt == "Branch B lta zz vv rta" from 3 have "inv2 mt \ inv1 mt" by simp 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) + with 3 have 4: "entry_in_tree k v (del_from_left 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: balance_left_in_tree) thus ?case proof (cases "xx = k") case True from 3 True have "tree_greater yy bb \ yy > k" by simp @@ -706,13 +834,13 @@ 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 + qed auto next case (5 xx aa yy ss 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) \ (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) + with 5 have 3: "entry_in_tree k v (del_from_right 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: balance_right_in_tree) thus ?case proof (cases "xx = k") case True from 5 True have "tree_less yy aa \ yy < k" by simp @@ -734,14 +862,14 @@ 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 auto qed simp definition delete where delete_def: "delete k t = paint B (del k t)" -theorem delete_is_rbt[simp]: assumes "is_rbt t" shows "is_rbt (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 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) @@ -751,11 +879,11 @@ by (auto intro: paint_sorted del_sorted) qed -lemma delete_pit: +lemma delete_in_tree: 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) + by (auto simp: del_in_tree) lemma lookup_delete: assumes is_rbt: "is_rbt t" @@ -766,35 +894,36 @@ proof (cases "x = k") assume "x = k" with is_rbt show ?thesis - by (cases "lookup (delete k t) k") (auto simp: lookup_pit delete_pit) + by (cases "lookup (delete k t) k") (auto simp: lookup_in_tree delete_in_tree) next assume "x \ k" thus ?thesis - by auto (metis is_rbt delete_is_rbt delete_pit is_rbt_sorted lookup_from_pit) + by auto (metis is_rbt delete_is_rbt delete_in_tree is_rbt_sorted lookup_from_in_tree) qed qed + subsection {* Union *} primrec - unionwithkey :: "('a\linorder \ 'b \ 'b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt" + union_with_key :: "('a\linorder \ 'b \ 'b \ 'b) \ ('a,'b) rbt \ ('a,'b) rbt \ ('a,'b) rbt" where - "unionwithkey f t Empty = t" -| "unionwithkey f t (Branch c lt k v rt) = unionwithkey f (unionwithkey f (insertwithkey f k v t) lt) rt" + "union_with_key f t Empty = t" +| "union_with_key f t (Branch c lt k v rt) = union_with_key f (union_with_key f (insert_with_key f k v t) lt) rt" -lemma unionwk_sorted: "sorted lt \ sorted (unionwithkey f lt rt)" +lemma unionwk_sorted: "sorted lt \ sorted (union_with_key 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)" +theorem unionwk_is_rbt[simp]: "is_rbt lt \ is_rbt (union_with_key f lt rt)" by (induct rt arbitrary: lt) (simp add: insertwk_is_rbt)+ definition - unionwith where - "unionwith f = unionwithkey (\_. f)" + union_with where + "union_with f = union_with_key (\_. f)" -theorem unionw_is_rbt: "is_rbt lt \ is_rbt (unionwith f lt rt)" unfolding unionwith_def by simp +theorem unionw_is_rbt: "is_rbt lt \ is_rbt (union_with f lt rt)" unfolding union_with_def by simp definition union where - "union = unionwithkey (%_ _ rv. rv)" + "union = union_with_key (%_ _ rv. rv)" theorem union_is_rbt: "is_rbt lt \ is_rbt (union lt rt)" unfolding union_def by simp @@ -811,7 +940,7 @@ case Empty thus ?case by (auto simp: union_def) next case (Branch c l k v r s) - hence sortedrl: "sorted r" "sorted l" "l |\ k" "k \| r" by auto + then have "sorted r" "sorted l" "l |\ k" "k \| r" by auto have meq: "lookup s(k \ v) ++ lookup l ++ lookup r = lookup s ++ @@ -839,178 +968,79 @@ qed qed - from Branch - have IHs: + from Branch have is_rbt: "is_rbt (RBT.union (RBT.insert k v s) l)" + by (auto intro: union_is_rbt insert_is_rbt) + with Branch have IHs: "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) + by auto with meq show ?case by (auto simp: lookup_insert[OF Branch(3)]) + qed -subsection {* Adjust *} + +subsection {* Modifying existing entries *} primrec - adjustwithkey :: "('a \ 'b \ 'b) \ ('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt" + map_entry :: "('a \ 'b \ 'b) \ ('a\linorder) \ ('a,'b) rbt \ ('a,'b) rbt" where - "adjustwithkey f k Empty = Empty" -| "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))" + "map_entry f k Empty = Empty" +| "map_entry f k (Branch c lt x v rt) = (if k < x then (Branch c (map_entry f k lt) x v rt) else if k > x then (Branch c lt x v (map_entry f k rt)) else (Branch c lt x (f x v) rt))" -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)+ +lemma map_entrywk_color_of: "color_of (map_entry f k t) = color_of t" by (induct t) simp+ +lemma map_entrywk_inv1: "inv1 (map_entry f k t) = inv1 t" by (induct t) (simp add: map_entrywk_color_of)+ +lemma map_entrywk_inv2: "inv2 (map_entry f k t) = inv2 t" "bheight (map_entry f k t) = bheight t" by (induct t) simp+ +lemma map_entrywk_tree_greater: "tree_greater k (map_entry f kk t) = tree_greater k t" by (induct t) simp+ +lemma map_entrywk_tree_less: "tree_less k (map_entry f kk t) = tree_less k t" by (induct t) simp+ +lemma map_entrywk_sorted: "sorted (map_entry f k t) = sorted t" by (induct t) (simp add: map_entrywk_tree_less map_entrywk_tree_greater)+ -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 map_entrywk_is_rbt [simp]: "is_rbt (map_entry f k t) = is_rbt t" +unfolding is_rbt_def by (simp add: map_entrywk_inv2 map_entrywk_color_of map_entrywk_sorted map_entrywk_inv1 ) -theorem adjustwithkey_map[simp]: - "lookup (adjustwithkey f k t) x = +theorem map_entry_map [simp]: + "lookup (map_entry 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_is_rbt[simp]: "is_rbt (adjust f k t) = is_rbt t" unfolding adjust_def by simp - -theorem adjust_map[simp]: - "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 *} - -primrec - mapwithkey :: "('a::linorder \ 'b \ 'c) \ ('a,'b) rbt \ ('a,'c) rbt" -where - "mapwithkey f Empty = Empty" -| "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_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 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_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 *} - -text {* The following is still incomplete... *} +subsection {* Mapping all entries *} primrec - foldwithkey :: "('a::linorder \ 'b \ 'c \ 'c) \ ('a,'b) rbt \ 'c \ 'c" + map :: "('a::linorder \ 'b \ 'c) \ ('a,'b) rbt \ ('a,'c) rbt" where - "foldwithkey f Empty v = v" -| "foldwithkey f (Branch c lt k x rt) v = foldwithkey f rt (f k x (foldwithkey f lt v))" - -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 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.lookup t1) \ k>k'" - proof - - fix k' - 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.lookup t2) \ 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.lookup t2)" proof - assume "x\dom (RBT.lookup t2)" - with DOM_T2 have "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.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.lookup t1)" proof - assume "x\dom (RBT.lookup t1)" - with DOM_T1 have "k>x" by simp - with C show False by simp - qed - ultimately have ?thesis by (simp add: map_add_upd_left map_add_dom_app_simps) - } ultimately show ?thesis using less_linear by blast -qed +lemma map_entries [simp]: "entries (map f t) = List.map (\(k, v). (k, f k v)) (entries t)" + by (induct t) auto +lemma map_keys [simp]: "keys (map f t) = keys t" by (simp add: keys_def split_def) +lemma map_tree_greater: "tree_greater k (map f t) = tree_greater k t" by (induct t) simp+ +lemma map_tree_less: "tree_less k (map f t) = tree_less k t" by (induct t) simp+ +lemma map_sorted: "sorted (map f t) = sorted t" by (induct t) (simp add: map_tree_less map_tree_greater)+ +lemma map_color_of: "color_of (map f t) = color_of t" by (induct t) simp+ +lemma map_inv1: "inv1 (map f t) = inv1 t" by (induct t) (simp add: map_color_of)+ +lemma map_inv2: "inv2 (map f t) = inv2 t" "bheight (map f t) = bheight t" by (induct t) simp+ +theorem map_is_rbt [simp]: "is_rbt (map f t) = is_rbt t" +unfolding is_rbt_def by (simp add: map_inv1 map_inv2 map_sorted map_color_of) -lemma map_of_entries: - shows "sorted t \ map_of (entries t) = lookup t" -proof (induct t) - case Empty thus ?case by (simp add: RBT.lookup_Empty) -next - 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_entries_fold: - "foldwithkey f t x = foldl (\x (k,v). f k v x) x (entries t)" -by (induct t arbitrary: x) auto - -lemma entries_pit[simp]: "(k, v) \ set (entries t) = entry_in_tree k v t" +theorem lookup_map [simp]: "lookup (map f t) x = Option.map (f x) (lookup t x)" by (induct t) auto -lemma sorted_entries: - "sorted t \ List.sorted (List.map fst (entries t))" -by (induct t) - (force simp: sorted_append sorted_Cons tree_ord_props - dest!: entry_in_tree_keys)+ + +subsection {* Folding over entries *} + +definition fold :: "('a \ 'b \ 'c \ 'c) \ ('a, 'b) rbt \ 'c \ 'c" where + "fold f t s = foldl (\s (k, v). f k v s) s (entries t)" -lemma distinct_entries: - "sorted t \ distinct (List.map fst (entries t))" -by (induct t) - (force simp: sorted_append sorted_Cons tree_ord_props - dest!: entry_in_tree_keys)+ +lemma fold_simps [simp, code]: + "fold f Empty = id" + "fold f (Branch c lt k v rt) = fold f rt \ f k v \ fold f lt" + by (simp_all add: fold_def expand_fun_eq) -hide (open) const Empty insert delete entries lookup map fold union adjust sorted - +hide (open) const Empty insert delete entries lookup map_entry map fold union sorted (*>*) text {* @@ -1018,6 +1048,7 @@ used as an efficient representation of finite maps. *} + subsection {* Data type and invariant *} text {* @@ -1040,6 +1071,7 @@ $O(\log n)$. *} + subsection {* Operations *} text {* @@ -1081,6 +1113,7 @@ @{thm map_is_rbt}\hfill(@{text "map_is_rbt"}) *} + subsection {* Map Semantics *} text {*