diff -r acc1749c2be9 -r 7daa65d45462 src/HOL/Data_Structures/Tree23_Map.thy --- a/src/HOL/Data_Structures/Tree23_Map.thy Thu May 16 19:43:21 2019 +0200 +++ b/src/HOL/Data_Structures/Tree23_Map.thy Mon May 20 17:33:13 2019 +0200 @@ -22,37 +22,37 @@ EQ \ Some b2 | GT \ lookup r x))" -fun upd :: "'a::linorder \ 'b \ ('a*'b) tree23 \ ('a*'b) up\<^sub>i" where -"upd x y Leaf = Up\<^sub>i Leaf (x,y) Leaf" | +fun upd :: "'a::linorder \ 'b \ ('a*'b) tree23 \ ('a*'b) upI" where +"upd x y Leaf = OF 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) | + TI l' => TI (Node2 l' ab r) + | OF l1 ab' l2 => TI (Node3 l1 ab' l2 ab r)) | + EQ \ TI (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)))" | + TI r' => TI (Node2 l ab r') + | OF r1 ab' r2 => TI (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) | + TI l' => TI (Node3 l' ab1 m ab2 r) + | OF l1 ab' l2 => OF (Node2 l1 ab' l2) ab1 (Node2 m ab2 r)) | + EQ \ TI (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) | + TI m' => TI (Node3 l ab1 m' ab2 r) + | OF m1 ab' m2 => OF (Node2 l ab1 m1) ab' (Node2 m2 ab2 r)) | + EQ \ TI (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))))" + TI r' => TI (Node3 l ab1 m ab2 r') + | OF r1 ab' r2 => OF (Node2 l ab1 m) ab2 (Node2 r1 ab' r2))))" definition update :: "'a::linorder \ 'b \ ('a*'b) tree23 \ ('a*'b) tree23" where -"update a b t = tree\<^sub>i(upd a b t)" +"update a b t = treeI(upd a b t)" -fun del :: "'a::linorder \ ('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 +fun del :: "'a::linorder \ ('a*'b) tree23 \ ('a*'b) upD" where +"del x Leaf = TD Leaf" | +"del x (Node2 Leaf ab1 Leaf) = (if x=fst ab1 then UF Leaf else TD(Node2 Leaf ab1 Leaf))" | +"del x (Node3 Leaf ab1 Leaf ab2 Leaf) = TD(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 | @@ -67,7 +67,7 @@ GT \ node33 l ab1 m ab2 (del x r)))" definition delete :: "'a::linorder \ ('a*'b) tree23 \ ('a*'b) tree23" where -"delete x t = tree\<^sub>d(del x t)" +"delete x t = treeD(del x t)" subsection \Functional Correctness\ @@ -78,8 +78,8 @@ lemma inorder_upd: - "sorted1(inorder t) \ inorder(tree\<^sub>i(upd x y t)) = upd_list x y (inorder t)" -by(induction t) (auto simp: upd_list_simps split: up\<^sub>i.splits) + "sorted1(inorder t) \ inorder(treeI(upd x y t)) = upd_list x y (inorder t)" +by(induction t) (auto simp: upd_list_simps split: upI.splits) corollary inorder_update: "sorted1(inorder t) \ inorder(update x y t) = upd_list x y (inorder t)" @@ -87,7 +87,7 @@ lemma inorder_del: "\ complete t ; sorted1(inorder t) \ \ - inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)" + inorder(treeD (del x t)) = del_list x (inorder t)" by(induction t rule: del.induct) (auto simp: del_list_simps inorder_nodes split_minD split!: if_split prod.splits) @@ -98,8 +98,8 @@ subsection \Balancedness\ -lemma complete_upd: "complete t \ complete (tree\<^sub>i(upd x y t)) \ height(upd x y t) = height t" -by (induct t) (auto split!: if_split up\<^sub>i.split)(* 16 secs in 2015 *) +lemma complete_upd: "complete t \ complete (treeI(upd x y t)) \ height(upd x y t) = height t" +by (induct t) (auto split!: if_split upI.split)(* 16 secs in 2015 *) corollary complete_update: "complete t \ complete (update x y t)" by (simp add: update_def complete_upd) @@ -109,12 +109,12 @@ by(induction x t rule: del.induct) (auto simp add: heights max_def height_split_min split: prod.split) -lemma complete_tree\<^sub>d_del: "complete t \ complete(tree\<^sub>d(del x t))" +lemma complete_treeD_del: "complete t \ complete(treeD(del x t))" by(induction x t rule: del.induct) (auto simp: completes complete_split_min height_del height_split_min split: prod.split) corollary complete_delete: "complete t \ complete(delete x t)" -by(simp add: delete_def complete_tree\<^sub>d_del) +by(simp add: delete_def complete_treeD_del) subsection \Overall Correctness\