diff -r 7ffc9c4f1f74 -r 44c9198f210c src/HOL/Data_Structures/Tree23_Map.thy --- a/src/HOL/Data_Structures/Tree23_Map.thy Wed Nov 11 16:42:30 2015 +0100 +++ b/src/HOL/Data_Structures/Tree23_Map.thy Wed Nov 11 18:32:26 2015 +0100 @@ -1,136 +1,136 @@ -(* Author: Tobias Nipkow *) - -section \A 2-3 Tree Implementation of Maps\ - -theory Tree23_Map -imports - Tree23_Set - Map_by_Ordered -begin - -fun lookup :: "('a::cmp * 'b) tree23 \ 'a \ 'b option" where -"lookup Leaf x = None" | -"lookup (Node2 l (a,b) r) x = (case cmp x a of - LT \ lookup l x | - GT \ lookup r x | - EQ \ Some b)" | -"lookup (Node3 l (a1,b1) m (a2,b2) r) x = (case cmp x a1 of - LT \ lookup l x | - EQ \ Some b1 | - GT \ (case cmp x a2 of - LT \ lookup m x | - EQ \ Some b2 | - GT \ lookup r x))" - -fun upd :: "'a::cmp \ 'b \ ('a*'b) tree23 \ ('a*'b) up\<^sub>i" where -"upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" | -"upd x y (Node2 l ab r) = (case cmp x (fst ab) of - LT \ (case upd x y l of - T\<^sub>i l' => T\<^sub>i (Node2 l' ab r) - | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) | - EQ \ T\<^sub>i (Node2 l (x,y) r) | - GT \ (case upd x y r of - T\<^sub>i r' => T\<^sub>i (Node2 l ab r') - | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" | -"upd x y (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of - LT \ (case upd x y l of - T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r) - | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) | - EQ \ T\<^sub>i (Node3 l (x,y) m ab2 r) | - GT \ (case cmp x (fst ab2) of - LT \ (case upd x y m of - T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r) - | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) | - EQ \ T\<^sub>i (Node3 l ab1 m (x,y) r) | - GT \ (case upd x y r of - T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r') - | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))" - -definition update :: "'a::cmp \ 'b \ ('a*'b) tree23 \ ('a*'b) tree23" where -"update a b t = tree\<^sub>i(upd a b t)" - -fun del :: "'a::cmp \ ('a*'b) tree23 \ ('a*'b) up\<^sub>d" where -"del x Leaf = T\<^sub>d Leaf" | -"del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" | -"del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf - else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" | -"del x (Node2 l ab1 r) = (case cmp x (fst ab1) of - LT \ node21 (del x l) ab1 r | - GT \ node22 l ab1 (del x r) | - EQ \ let (ab1',t) = del_min r in node22 l ab1' t)" | -"del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of - LT \ node31 (del x l) ab1 m ab2 r | - EQ \ let (ab1',m') = del_min m in node32 l ab1' m' ab2 r | - GT \ (case cmp x (fst ab2) of - LT \ node32 l ab1 (del x m) ab2 r | - EQ \ let (ab2',r') = del_min r in node33 l ab1 m ab2' r' | - GT \ node33 l ab1 m ab2 (del x r)))" - -definition delete :: "'a::cmp \ ('a*'b) tree23 \ ('a*'b) tree23" where -"delete x t = tree\<^sub>d(del x t)" - - -subsection \Functional Correctness\ - -lemma lookup: "sorted1(inorder t) \ lookup t x = map_of (inorder t) x" -by (induction t) (auto simp: map_of_simps split: option.split) - - -lemma inorder_upd: - "sorted1(inorder t) \ inorder(tree\<^sub>i(upd a b t)) = upd_list a b (inorder t)" -by(induction t) (auto simp: upd_list_simps split: up\<^sub>i.splits) - -corollary inorder_update: - "sorted1(inorder t) \ inorder(update a b t) = upd_list a b (inorder t)" -by(simp add: update_def inorder_upd) - - -lemma inorder_del: "\ bal t ; sorted1(inorder t) \ \ - inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)" -by(induction t rule: del.induct) - (auto simp: del_list_simps inorder_nodes del_minD split: prod.splits) - -corollary inorder_delete: "\ bal t ; sorted1(inorder t) \ \ - inorder(delete x t) = del_list x (inorder t)" -by(simp add: delete_def inorder_del) - - -subsection \Balancedness\ - -lemma bal_upd: "bal t \ bal (tree\<^sub>i(upd a b t)) \ height(upd a b t) = height t" -by (induct t) (auto split: up\<^sub>i.split)(* 16 secs in 2015 *) - -corollary bal_update: "bal t \ bal (update a b t)" -by (simp add: update_def bal_upd) - - -lemma height_del: "bal t \ height(del x t) = height t" -by(induction x t rule: del.induct) - (auto simp add: heights max_def height_del_min split: prod.split) - -lemma bal_tree\<^sub>d_del: "bal t \ bal(tree\<^sub>d(del x t))" -by(induction x t rule: del.induct) - (auto simp: bals bal_del_min height_del height_del_min split: prod.split) - -corollary bal_delete: "bal t \ bal(delete x t)" -by(simp add: delete_def bal_tree\<^sub>d_del) - - -subsection \Overall Correctness\ - -interpretation T23_Map: Map_by_Ordered -where empty = Leaf and lookup = lookup and update = update and delete = delete -and inorder = inorder and wf = bal -proof (standard, goal_cases) - case 2 thus ?case by(simp add: lookup) -next - case 3 thus ?case by(simp add: inorder_update) -next - case 4 thus ?case by(simp add: inorder_delete) -next - case 6 thus ?case by(simp add: bal_update) -next - case 7 thus ?case by(simp add: bal_delete) -qed simp+ - -end +(* Author: Tobias Nipkow *) + +section \A 2-3 Tree Implementation of Maps\ + +theory Tree23_Map +imports + Tree23_Set + Map_by_Ordered +begin + +fun lookup :: "('a::cmp * 'b) tree23 \ 'a \ 'b option" where +"lookup Leaf x = None" | +"lookup (Node2 l (a,b) r) x = (case cmp x a of + LT \ lookup l x | + GT \ lookup r x | + EQ \ Some b)" | +"lookup (Node3 l (a1,b1) m (a2,b2) r) x = (case cmp x a1 of + LT \ lookup l x | + EQ \ Some b1 | + GT \ (case cmp x a2 of + LT \ lookup m x | + EQ \ Some b2 | + GT \ lookup r x))" + +fun upd :: "'a::cmp \ 'b \ ('a*'b) tree23 \ ('a*'b) up\<^sub>i" where +"upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" | +"upd x y (Node2 l ab r) = (case cmp x (fst ab) of + LT \ (case upd x y l of + T\<^sub>i l' => T\<^sub>i (Node2 l' ab r) + | Up\<^sub>i l1 ab' l2 => T\<^sub>i (Node3 l1 ab' l2 ab r)) | + EQ \ T\<^sub>i (Node2 l (x,y) r) | + GT \ (case upd x y r of + T\<^sub>i r' => T\<^sub>i (Node2 l ab r') + | Up\<^sub>i r1 ab' r2 => T\<^sub>i (Node3 l ab r1 ab' r2)))" | +"upd x y (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of + LT \ (case upd x y l of + T\<^sub>i l' => T\<^sub>i (Node3 l' ab1 m ab2 r) + | Up\<^sub>i l1 ab' l2 => Up\<^sub>i (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) | + EQ \ T\<^sub>i (Node3 l (x,y) m ab2 r) | + GT \ (case cmp x (fst ab2) of + LT \ (case upd x y m of + T\<^sub>i m' => T\<^sub>i (Node3 l ab1 m' ab2 r) + | Up\<^sub>i m1 ab' m2 => Up\<^sub>i (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) | + EQ \ T\<^sub>i (Node3 l ab1 m (x,y) r) | + GT \ (case upd x y r of + T\<^sub>i r' => T\<^sub>i (Node3 l ab1 m ab2 r') + | Up\<^sub>i r1 ab' r2 => Up\<^sub>i (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))" + +definition update :: "'a::cmp \ 'b \ ('a*'b) tree23 \ ('a*'b) tree23" where +"update a b t = tree\<^sub>i(upd a b t)" + +fun del :: "'a::cmp \ ('a*'b) tree23 \ ('a*'b) up\<^sub>d" where +"del x Leaf = T\<^sub>d Leaf" | +"del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf ab1 Leaf))" | +"del x (Node3 Leaf ab1 Leaf ab2 Leaf) = T\<^sub>d(if x=fst ab1 then Node2 Leaf ab2 Leaf + else if x=fst ab2 then Node2 Leaf ab1 Leaf else Node3 Leaf ab1 Leaf ab2 Leaf)" | +"del x (Node2 l ab1 r) = (case cmp x (fst ab1) of + LT \ node21 (del x l) ab1 r | + GT \ node22 l ab1 (del x r) | + EQ \ let (ab1',t) = del_min r in node22 l ab1' t)" | +"del x (Node3 l ab1 m ab2 r) = (case cmp x (fst ab1) of + LT \ node31 (del x l) ab1 m ab2 r | + EQ \ let (ab1',m') = del_min m in node32 l ab1' m' ab2 r | + GT \ (case cmp x (fst ab2) of + LT \ node32 l ab1 (del x m) ab2 r | + EQ \ let (ab2',r') = del_min r in node33 l ab1 m ab2' r' | + GT \ node33 l ab1 m ab2 (del x r)))" + +definition delete :: "'a::cmp \ ('a*'b) tree23 \ ('a*'b) tree23" where +"delete x t = tree\<^sub>d(del x t)" + + +subsection \Functional Correctness\ + +lemma lookup: "sorted1(inorder t) \ lookup t x = map_of (inorder t) x" +by (induction t) (auto simp: map_of_simps split: option.split) + + +lemma inorder_upd: + "sorted1(inorder t) \ inorder(tree\<^sub>i(upd a b t)) = upd_list a b (inorder t)" +by(induction t) (auto simp: upd_list_simps split: up\<^sub>i.splits) + +corollary inorder_update: + "sorted1(inorder t) \ inorder(update a b t) = upd_list a b (inorder t)" +by(simp add: update_def inorder_upd) + + +lemma inorder_del: "\ bal t ; sorted1(inorder t) \ \ + inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)" +by(induction t rule: del.induct) + (auto simp: del_list_simps inorder_nodes del_minD split: prod.splits) + +corollary inorder_delete: "\ bal t ; sorted1(inorder t) \ \ + inorder(delete x t) = del_list x (inorder t)" +by(simp add: delete_def inorder_del) + + +subsection \Balancedness\ + +lemma bal_upd: "bal t \ bal (tree\<^sub>i(upd a b t)) \ height(upd a b t) = height t" +by (induct t) (auto split: up\<^sub>i.split)(* 16 secs in 2015 *) + +corollary bal_update: "bal t \ bal (update a b t)" +by (simp add: update_def bal_upd) + + +lemma height_del: "bal t \ height(del x t) = height t" +by(induction x t rule: del.induct) + (auto simp add: heights max_def height_del_min split: prod.split) + +lemma bal_tree\<^sub>d_del: "bal t \ bal(tree\<^sub>d(del x t))" +by(induction x t rule: del.induct) + (auto simp: bals bal_del_min height_del height_del_min split: prod.split) + +corollary bal_delete: "bal t \ bal(delete x t)" +by(simp add: delete_def bal_tree\<^sub>d_del) + + +subsection \Overall Correctness\ + +interpretation T23_Map: Map_by_Ordered +where empty = Leaf and lookup = lookup and update = update and delete = delete +and inorder = inorder and wf = bal +proof (standard, goal_cases) + case 2 thus ?case by(simp add: lookup) +next + case 3 thus ?case by(simp add: inorder_update) +next + case 4 thus ?case by(simp add: inorder_delete) +next + case 6 thus ?case by(simp add: bal_update) +next + case 7 thus ?case by(simp add: bal_delete) +qed simp+ + +end