diff -r 947ce60a06e1 -r 00d9682e8dd7 src/HOL/Data_Structures/AVL_Map.thy --- a/src/HOL/Data_Structures/AVL_Map.thy Wed Nov 04 15:07:23 2015 +0100 +++ b/src/HOL/Data_Structures/AVL_Map.thy Thu Nov 05 08:27:14 2015 +0100 @@ -8,36 +8,34 @@ Lookup2 begin -fun update :: "'a::order \ 'b \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where +fun update :: "'a::cmp \ 'b \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where "update x y Leaf = Node 1 Leaf (x,y) Leaf" | -"update x y (Node h l (a,b) r) = - (if x = a then Node h l (x,y) r else - if x < a then node_bal_l (update x y l) (a,b) r - else node_bal_r l (a,b) (update x y r))" +"update x y (Node h l (a,b) r) = (case cmp x a of + EQ \ Node h l (x,y) r | + LT \ balL (update x y l) (a,b) r | + GT \ balR l (a,b) (update x y r))" -fun delete :: "'a::order \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where +fun delete :: "'a::cmp \ ('a*'b) avl_tree \ ('a*'b) avl_tree" where "delete _ Leaf = Leaf" | -"delete x (Node h l (a,b) r) = ( - if x = a then delete_root (Node h l (a,b) r) else - if x < a then node_bal_r (delete x l) (a,b) r - else node_bal_l l (a,b) (delete x r))" +"delete x (Node h l (a,b) r) = (case cmp x a of + EQ \ delete_root (Node h l (a,b) r) | + LT \ balR (delete x l) (a,b) r | + GT \ balL l (a,b) (delete x r))" subsection {* Functional Correctness Proofs *} theorem inorder_update: "sorted1(inorder t) \ inorder(update x y t) = upd_list x y (inorder t)" -by (induct t) - (auto simp: upd_list_simps inorder_node_bal_l inorder_node_bal_r) +by (induct t) (auto simp: upd_list_simps inorder_balL inorder_balR) theorem inorder_delete: "sorted1(inorder t) \ inorder (delete x t) = del_list x (inorder t)" by(induction t) - (auto simp: del_list_simps inorder_node_bal_l inorder_node_bal_r + (auto simp: del_list_simps inorder_balL inorder_balR inorder_delete_root inorder_delete_maxD split: prod.splits) - interpretation Map_by_Ordered where empty = Leaf and lookup = lookup and update = update and delete = delete and inorder = inorder and wf = "\_. True"