# HG changeset patch # User nipkow # Date 1456909291 -3600 # Node ID f187aaf602c4f9cc29a010d5b71b1fade062f1f4 # Parent 83db706d7771404c66b607bab2b3958b85e8861a added invariant proofs to AA trees diff -r 83db706d7771 -r f187aaf602c4 src/HOL/Data_Structures/AA_Map.thy --- a/src/HOL/Data_Structures/AA_Map.thy Tue Mar 01 22:49:33 2016 +0100 +++ b/src/HOL/Data_Structures/AA_Map.thy Wed Mar 02 10:01:31 2016 +0100 @@ -1,6 +1,6 @@ (* Author: Tobias Nipkow *) -section "AA Implementation of Maps" +section "AA Tree Implementation of Maps" theory AA_Map imports @@ -26,21 +26,190 @@ else let (l',ab') = del_max l in adjust (Node lv l' ab' r)))" +subsection "Invariance" + +subsubsection "Proofs for insert" + +lemma lvl_update_aux: + "lvl (update x y t) = lvl t \ lvl (update x y t) = lvl t + 1 \ sngl (update x y t)" +apply(induction t) +apply (auto simp: lvl_skew) +apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+ +done + +lemma lvl_update: obtains + (Same) "lvl (update x y t) = lvl t" | + (Incr) "lvl (update x y t) = lvl t + 1" "sngl (update x y t)" +using lvl_update_aux by fastforce + +declare invar.simps(2)[simp] + +lemma lvl_update_sngl: "invar t \ sngl t \ lvl(update x y t) = lvl t" +proof (induction t rule: update.induct) + case (2 x y lv t1 a b t2) + consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" + using less_linear by blast + thus ?case proof cases + case LT + thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits) + next + case GT + thus ?thesis using 2 proof (cases t1) + case Node + thus ?thesis using 2 GT + apply (auto simp add: skew_case split_case split: tree.splits) + by (metis less_not_refl2 lvl.simps(2) lvl_update_aux n_not_Suc_n sngl.simps(3))+ + qed (auto simp add: lvl_0_iff) + qed simp +qed simp + +lemma lvl_update_incr_iff: "(lvl(update a b t) = lvl t + 1) \ + (EX l x r. update a b t = Node (lvl t + 1) l x r \ lvl l = lvl r)" +apply(cases t) +apply(auto simp add: skew_case split_case split: if_splits) +apply(auto split: tree.splits if_splits) +done + +lemma invar_update: "invar t \ invar(update a b t)" +proof(induction t) + case (Node n l xy r) + hence il: "invar l" and ir: "invar r" by auto + obtain x y where [simp]: "xy = (x,y)" by fastforce + note N = Node + let ?t = "Node n l xy r" + have "a < x \ a = x \ x < a" by auto + moreover + { assume "a < x" + note iil = Node.IH(1)[OF il] + have ?case + proof (cases rule: lvl_update[of a b l]) + case (Same) thus ?thesis + using \a invar_NodeL[OF Node.prems iil Same] + by (simp add: skew_invar split_invar del: invar.simps) + next + case (Incr) + then obtain t1 w t2 where ial[simp]: "update a b l = Node n t1 w t2" + using Node.prems by (auto simp: lvl_Suc_iff) + have l12: "lvl t1 = lvl t2" + by (metis Incr(1) ial lvl_update_incr_iff tree.inject) + have "update a b ?t = split(skew(Node n (update a b l) xy r))" + by(simp add: \a) + also have "skew(Node n (update a b l) xy r) = Node n t1 w (Node n t2 xy r)" + by(simp) + also have "invar(split \)" + proof (cases r) + case Leaf + hence "l = Leaf" using Node.prems by(auto simp: lvl_0_iff) + thus ?thesis using Leaf ial by simp + next + case [simp]: (Node m t3 y t4) + show ?thesis (*using N(3) iil l12 by(auto)*) + proof cases + assume "m = n" thus ?thesis using N(3) iil by(auto) + next + assume "m \ n" thus ?thesis using N(3) iil l12 by(auto) + qed + qed + finally show ?thesis . + qed + } + moreover + { assume "x < a" + note iir = Node.IH(2)[OF ir] + from \invar ?t\ have "n = lvl r \ n = lvl r + 1" by auto + hence ?case + proof + assume 0: "n = lvl r" + have "update a b ?t = split(skew(Node n l xy (update a b r)))" + using \a>x\ by(auto) + also have "skew(Node n l xy (update a b r)) = Node n l xy (update a b r)" + using Node.prems by(simp add: skew_case split: tree.split) + also have "invar(split \)" + proof - + from lvl_update_sngl[OF ir sngl_if_invar[OF \invar ?t\ 0], of a b] + obtain t1 p t2 where iar: "update a b r = Node n t1 p t2" + using Node.prems 0 by (auto simp: lvl_Suc_iff) + from Node.prems iar 0 iir + show ?thesis by (auto simp: split_case split: tree.splits) + qed + finally show ?thesis . + next + assume 1: "n = lvl r + 1" + hence "sngl ?t" by(cases r) auto + show ?thesis + proof (cases rule: lvl_update[of a b r]) + case (Same) + show ?thesis using \x il ir invar_NodeR[OF Node.prems 1 iir Same] + by (auto simp add: skew_invar split_invar) + next + case (Incr) + thus ?thesis using invar_NodeR2[OF `invar ?t` Incr(2) 1 iir] 1 \x < a\ + by (auto simp add: skew_invar split_invar split: if_splits) + qed + qed + } + moreover { assume "a = x" hence ?case using Node.prems by auto } + ultimately show ?case by blast +qed simp + +subsubsection "Proofs for delete" + +declare invar.simps(2)[simp del] + +theorem post_delete: "invar t \ post_del t (delete x t)" +proof (induction t) + case (Node lv l ab r) + + obtain a b where [simp]: "ab = (a,b)" by fastforce + + let ?l' = "delete x l" and ?r' = "delete x r" + let ?t = "Node lv l ab r" let ?t' = "delete x ?t" + + from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) + + note post_l' = Node.IH(1)[OF inv_l] + note preL = pre_adj_if_postL[OF Node.prems post_l'] + + note post_r' = Node.IH(2)[OF inv_r] + note preR = pre_adj_if_postR[OF Node.prems post_r'] + + show ?case + proof (cases rule: linorder_cases[of x a]) + case less + thus ?thesis using Node.prems by (simp add: post_del_adjL preL) + next + case greater + thus ?thesis using Node.prems preR by (simp add: post_del_adjR post_r') + next + case equal + show ?thesis + proof cases + assume "l = Leaf" thus ?thesis using equal Node.prems + by(auto simp: post_del_def invar.simps(2)) + next + assume "l \ Leaf" thus ?thesis using equal Node.prems + by simp (metis inv_l post_del_adjL post_del_max pre_adj_if_postL) + qed + qed +qed (simp add: post_del_def) + + 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_split inorder_skew) - theorem inorder_delete: - "sorted1(inorder t) \ inorder (delete x t) = del_list x (inorder t)" + "\invar t; sorted1(inorder t)\ \ + inorder (delete x t) = del_list x (inorder t)" by(induction t) - (auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits) + (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR + post_del_max post_delete del_maxD split: prod.splits) -interpretation Map_by_Ordered +interpretation I: Map_by_Ordered where empty = Leaf and lookup = lookup and update = update and delete = delete -and inorder = inorder and inv = "\_. True" +and inorder = inorder and inv = invar proof (standard, goal_cases) case 1 show ?case by simp next @@ -49,6 +218,12 @@ case 3 thus ?case by(simp add: inorder_update) next case 4 thus ?case by(simp add: inorder_delete) -qed auto +next + case 5 thus ?case by(simp) +next + case 6 thus ?case by(simp add: invar_update) +next + case 7 thus ?case using post_delete by(auto simp: post_del_def) +qed end diff -r 83db706d7771 -r f187aaf602c4 src/HOL/Data_Structures/AA_Set.thy --- a/src/HOL/Data_Structures/AA_Set.thy Tue Mar 01 22:49:33 2016 +0100 +++ b/src/HOL/Data_Structures/AA_Set.thy Wed Mar 02 10:01:31 2016 +0100 @@ -1,7 +1,5 @@ (* -Author: Tobias Nipkow - -Added trivial cases to function `adjust' to obviate invariants. +Author: Tobias Nipkow and Daniel Stüwe *) section \AA Tree Implementation of Sets\ @@ -17,13 +15,13 @@ fun lvl :: "'a aa_tree \ nat" where "lvl Leaf = 0" | "lvl (Node lv _ _ _) = lv" -(* + fun invar :: "'a aa_tree \ bool" where "invar Leaf = True" | "invar (Node h l a r) = (invar l \ invar r \ h = lvl l + 1 \ (h = lvl r + 1 \ (\lr b rr. r = Node h lr b rr \ h = lvl rr + 1)))" -*) + fun skew :: "'a aa_tree \ 'a aa_tree" where "skew (Node lva (Node lvb t1 b t2) a t3) = (if lva = lvb then Node lva t1 b (Node lva t2 a t3) else Node lva (Node lvb t1 b t2) a t3)" | @@ -46,11 +44,6 @@ GT \ split (skew (Node lv t1 a (insert x t2))) | EQ \ Node lv t1 x t2)" -(* wrong in paper! *) -fun del_max :: "'a aa_tree \ 'a aa_tree * 'a" where -"del_max (Node lv l a Leaf) = (l,a)" | -"del_max (Node lv l a r) = (let (r',b) = del_max r in (Node lv l a r', b))" - fun sngl :: "'a aa_tree \ bool" where "sngl Leaf = False" | "sngl (Node _ _ _ Leaf) = True" | @@ -65,19 +58,28 @@ if lvl r < lv-1 then case l of Node lva t1 a (Node lvb t2 b t3) - \ Node (lvb+1) (Node lva t1 a t2) b (Node (lv-1) t3 x r) | - _ \ t (* unreachable *) + \ Node (lvb+1) (Node lva t1 a t2) b (Node (lv-1) t3 x r) else if lvl r < lv then split (Node (lv-1) l x r) else case r of - Leaf \ Leaf (* unreachable *) | Node lvb t1 b t4 \ (case t1 of Node lva t2 a t3 \ Node (lva+1) (Node (lv-1) l x t2) a - (split (Node (if sngl t1 then lva-1 else lva) t3 b t4)) - | _ \ t (* unreachable *))))" + (split (Node (if sngl t1 then lva else lva+1) t3 b t4)))))" + +text{* In the paper, the last case of @{const adjust} is expressed with the help of an +incorrect auxiliary function \texttt{nlvl}. + +Function @{text del_max} below is called \texttt{dellrg} in the paper. +The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest +element but recurses on the left instead of the right subtree; the invariant +is not restored.*} + +fun del_max :: "'a aa_tree \ 'a aa_tree * 'a" where +"del_max (Node lv l a Leaf) = (l,a)" | +"del_max (Node lv l a r) = (let (r',b) = del_max r in (adjust(Node lv l a r'), b))" fun delete :: "'a::cmp \ 'a aa_tree \ 'a aa_tree" where "delete _ Leaf = Leaf" | @@ -88,9 +90,368 @@ EQ \ (if l = Leaf then r else let (l',b) = del_max l in adjust (Node lv l' b r)))" +fun pre_adjust where +"pre_adjust (Node lv l a r) = (invar l \ invar r \ + ((lv = lvl l + 1 \ (lv = lvl r + 1 \ lv = lvl r + 2 \ lv = lvl r \ sngl r)) \ + (lv = lvl l + 2 \ (lv = lvl r + 1 \ lv = lvl r \ sngl r))))" + +declare pre_adjust.simps [simp del] + +subsection "Auxiliary Proofs" + +lemma split_case: "split t = (case t of + Node lvx a x (Node lvy b y (Node lvz c z d)) \ + (if lvx = lvy \ lvy = lvz + then Node (lvx+1) (Node lvx a x b) y (Node lvx c z d) + else t) + | t \ t)" +by(auto split: tree.split) + +lemma skew_case: "skew t = (case t of + Node lvx (Node lvy a y b) x c \ + (if lvx = lvy then Node lvx a y (Node lvx b x c) else t) + | t \ t)" +by(auto split: tree.split) + +lemma lvl_0_iff: "invar t \ lvl t = 0 \ t = Leaf" +by(cases t) auto + +lemma lvl_Suc_iff: "lvl t = Suc n \ (\ l a r. t = Node (Suc n) l a r)" +by(cases t) auto + +lemma lvl_skew: "lvl (skew t) = lvl t" +by(induction t rule: skew.induct) auto + +lemma lvl_split: "lvl (split t) = lvl t \ lvl (split t) = lvl t + 1 \ sngl (split t)" +by(induction t rule: split.induct) auto + +lemma invar_2Nodes:"invar (Node lv l x (Node rlv rl rx rr)) = + (invar l \ invar \rlv, rl, rx, rr\ \ lv = Suc (lvl l) \ + (lv = Suc rlv \ rlv = lv \ lv = Suc (lvl rr)))" +by simp + +lemma invar_NodeLeaf[simp]: + "invar (Node lv l x Leaf) = (invar l \ lv = Suc (lvl l) \ lv = Suc 0)" +by simp + +lemma sngl_if_invar: "invar (Node n l a r) \ n = lvl r \ sngl r" +by(cases r rule: sngl.cases) clarsimp+ + + +subsection "Invariance" + +subsubsection "Proofs for insert" + +lemma lvl_insert_aux: + "lvl (insert x t) = lvl t \ lvl (insert x t) = lvl t + 1 \ sngl (insert x t)" +apply(induction t) +apply (auto simp: lvl_skew) +apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+ +done + +lemma lvl_insert: obtains + (Same) "lvl (insert x t) = lvl t" | + (Incr) "lvl (insert x t) = lvl t + 1" "sngl (insert x t)" +using lvl_insert_aux by blast + +lemma lvl_insert_sngl: "invar t \ sngl t \ lvl(insert x t) = lvl t" +proof (induction t rule: "insert.induct" ) + case (2 x lv t1 a t2) + consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" + using less_linear by blast + thus ?case proof cases + case LT + thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits) + next + case GT + thus ?thesis using 2 proof (cases t1) + case Node + thus ?thesis using 2 GT + apply (auto simp add: skew_case split_case split: tree.splits) + by (metis less_not_refl2 lvl.simps(2) lvl_insert_aux n_not_Suc_n sngl.simps(3))+ + qed (auto simp add: lvl_0_iff) + qed simp +qed simp + +lemma skew_invar: "invar t \ skew t = t" +by(induction t rule: skew.induct) auto + +lemma split_invar: "invar t \ split t = t" +by(induction t rule: split.induct) clarsimp+ + +lemma invar_NodeL: + "\ invar(Node n l x r); invar l'; lvl l' = lvl l \ \ invar(Node n l' x r)" +by(auto) + +lemma invar_NodeR: + "\ invar(Node n l x r); n = lvl r + 1; invar r'; lvl r' = lvl r \ \ invar(Node n l x r')" +by(auto) + +lemma invar_NodeR2: + "\ invar(Node n l x r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \ \ invar(Node n l x r')" +by(cases r' rule: sngl.cases) clarsimp+ + + +lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \ + (EX l x r. insert a t = Node (lvl t + 1) l x r \ lvl l = lvl r)" +apply(cases t) +apply(auto simp add: skew_case split_case split: if_splits) +apply(auto split: tree.splits if_splits) +done + +lemma invar_insert: "invar t \ invar(insert a t)" +proof(induction t) + case (Node n l x r) + hence il: "invar l" and ir: "invar r" by auto + note N = Node + let ?t = "Node n l x r" + have "a < x \ a = x \ x < a" by auto + moreover + { assume "a < x" + note iil = Node.IH(1)[OF il] + have ?case + proof (cases rule: lvl_insert[of a l]) + case (Same) thus ?thesis + using \a invar_NodeL[OF Node.prems iil Same] + by (simp add: skew_invar split_invar del: invar.simps) + next + case (Incr) + then obtain t1 w t2 where ial[simp]: "insert a l = Node n t1 w t2" + using Node.prems by (auto simp: lvl_Suc_iff) + have l12: "lvl t1 = lvl t2" + by (metis Incr(1) ial lvl_insert_incr_iff tree.inject) + have "insert a ?t = split(skew(Node n (insert a l) x r))" + by(simp add: \a) + also have "skew(Node n (insert a l) x r) = Node n t1 w (Node n t2 x r)" + by(simp) + also have "invar(split \)" + proof (cases r) + case Leaf + hence "l = Leaf" using Node.prems by(auto simp: lvl_0_iff) + thus ?thesis using Leaf ial by simp + next + case [simp]: (Node m t3 y t4) + show ?thesis (*using N(3) iil l12 by(auto)*) + proof cases + assume "m = n" thus ?thesis using N(3) iil by(auto) + next + assume "m \ n" thus ?thesis using N(3) iil l12 by(auto) + qed + qed + finally show ?thesis . + qed + } + moreover + { assume "x < a" + note iir = Node.IH(2)[OF ir] + from \invar ?t\ have "n = lvl r \ n = lvl r + 1" by auto + hence ?case + proof + assume 0: "n = lvl r" + have "insert a ?t = split(skew(Node n l x (insert a r)))" + using \a>x\ by(auto) + also have "skew(Node n l x (insert a r)) = Node n l x (insert a r)" + using Node.prems by(simp add: skew_case split: tree.split) + also have "invar(split \)" + proof - + from lvl_insert_sngl[OF ir sngl_if_invar[OF \invar ?t\ 0], of a] + obtain t1 y t2 where iar: "insert a r = Node n t1 y t2" + using Node.prems 0 by (auto simp: lvl_Suc_iff) + from Node.prems iar 0 iir + show ?thesis by (auto simp: split_case split: tree.splits) + qed + finally show ?thesis . + next + assume 1: "n = lvl r + 1" + hence "sngl ?t" by(cases r) auto + show ?thesis + proof (cases rule: lvl_insert[of a r]) + case (Same) + show ?thesis using \x il ir invar_NodeR[OF Node.prems 1 iir Same] + by (auto simp add: skew_invar split_invar) + next + case (Incr) + thus ?thesis using invar_NodeR2[OF `invar ?t` Incr(2) 1 iir] 1 \x < a\ + by (auto simp add: skew_invar split_invar split: if_splits) + qed + qed + } + moreover { assume "a = x" hence ?case using Node.prems by auto } + ultimately show ?case by blast +qed simp + + +subsubsection "Proofs for delete" + +lemma invarL: "ASSUMPTION(invar \lv, l, a, r\) \ invar l" +by(simp add: ASSUMPTION_def) + +lemma invarR: "ASSUMPTION(invar \lv, l, a, r\) \ invar r" +by(simp add: ASSUMPTION_def) + +lemma sngl_NodeI: + "sngl (Node lv l a r) \ sngl (Node lv l' a' r)" +by(cases r) (simp_all) + + +declare invarL[simp] invarR[simp] + +lemma pre_cases: +assumes "pre_adjust (Node lv l x r)" +obtains + (tSngl) "invar l \ invar r \ + lv = Suc (lvl r) \ lvl l = lvl r" | + (tDouble) "invar l \ invar r \ + lv = lvl r \ Suc (lvl l) = lvl r \ sngl r " | + (rDown) "invar l \ invar r \ + lv = Suc (Suc (lvl r)) \ lv = Suc (lvl l)" | + (lDown_tSngl) "invar l \ invar r \ + lv = Suc (lvl r) \ lv = Suc (Suc (lvl l))" | + (lDown_tDouble) "invar l \ invar r \ + lv = lvl r \ lv = Suc (Suc (lvl l)) \ sngl r" +using assms unfolding pre_adjust.simps +by auto + +declare invar.simps(2)[simp del] invar_2Nodes[simp add] + +lemma invar_adjust: + assumes pre: "pre_adjust (Node lv l a r)" + shows "invar(adjust (Node lv l a r))" +using pre proof (cases rule: pre_cases) + case (tDouble) thus ?thesis unfolding adjust_def by (cases r) (auto simp: invar.simps(2)) +next + case (rDown) + from rDown obtain llv ll la lr where l: "l = Node llv ll la lr" by (cases l) auto + from rDown show ?thesis unfolding adjust_def by (auto simp: l invar.simps(2) split: tree.splits) +next + case (lDown_tDouble) + from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rlv rl ra rr" by (cases r) auto + from lDown_tDouble and r obtain rrlv rrr rra rrl where + rr :"rr = Node rrlv rrr rra rrl" by (cases rr) auto + from lDown_tDouble show ?thesis unfolding adjust_def r rr + apply (cases rl) apply (auto simp add: invar.simps(2)) + using lDown_tDouble by (auto simp: split_case lvl_0_iff elim:lvl.elims split: tree.split) +qed (auto simp: split_case invar.simps(2) adjust_def split: tree.splits) + +lemma lvl_adjust: + assumes "pre_adjust (Node lv l a r)" + shows "lv = lvl (adjust(Node lv l a r)) \ lv = lvl (adjust(Node lv l a r)) + 1" +using assms(1) proof(cases rule: pre_cases) + case lDown_tSngl thus ?thesis + using lvl_split[of "\lvl r, l, a, r\"] by (auto simp: adjust_def) +next + case lDown_tDouble thus ?thesis + by (auto simp: adjust_def invar.simps(2) split: tree.split) +qed (auto simp: adjust_def split: tree.splits) + +lemma sngl_adjust: assumes "pre_adjust (Node lv l a r)" + "sngl \lv, l, a, r\" "lv = lvl (adjust \lv, l, a, r\)" + shows "sngl (adjust \lv, l, a, r\)" +using assms proof (cases rule: pre_cases) + case rDown + thus ?thesis using assms(2,3) unfolding adjust_def + by (auto simp add: skew_case) (auto split: tree.split) +qed (auto simp: adjust_def skew_case split_case split: tree.split) + +definition "post_del t t' == + invar t' \ + (lvl t' = lvl t \ lvl t' + 1 = lvl t) \ + (lvl t' = lvl t \ sngl t \ sngl t')" + +lemma pre_adj_if_postR: + "invar\lv, l, a, r\ \ post_del r r' \ pre_adjust \lv, l, a, r'\" +by(cases "sngl r") + (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) + +lemma pre_adj_if_postL: + "invar\lv, l, a, r\ \ post_del l l' \ pre_adjust \lv, l', b, r\" +by(cases "sngl r") + (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) + +lemma post_del_adjL: + "\ invar\lv, l, a, r\; pre_adjust \lv, l', b, r\ \ + \ post_del \lv, l, a, r\ (adjust \lv, l', b, r\)" +unfolding post_del_def +by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2)) + +lemma post_del_adjR: +assumes "invar\lv, l, a, r\" "pre_adjust \lv, l, a, r'\" "post_del r r'" +shows "post_del \lv, l, a, r\ (adjust \lv, l, a, r'\)" +proof(unfold post_del_def, safe del: disjCI) + let ?t = "\lv, l, a, r\" + let ?t' = "adjust \lv, l, a, r'\" + show "invar ?t'" by(rule invar_adjust[OF assms(2)]) + show "lvl ?t' = lvl ?t \ lvl ?t' + 1 = lvl ?t" + using lvl_adjust[OF assms(2)] by auto + show "sngl ?t'" if as: "lvl ?t' = lvl ?t" "sngl ?t" + proof - + have s: "sngl \lv, l, a, r'\" + proof(cases r') + case Leaf thus ?thesis by simp + next + case Node thus ?thesis using as(2) assms(1,3) + by (cases r) (auto simp: post_del_def) + qed + show ?thesis using as(1) sngl_adjust[OF assms(2) s] by simp + qed +qed + +declare prod.splits[split] + +theorem post_del_max: + "\ invar t; (t', x) = del_max t; t \ Leaf \ \ post_del t t'" +proof (induction t arbitrary: t' rule: del_max.induct) + case (2 lv l a lvr rl ra rr) + let ?r = "\lvr, rl, ra, rr\" + let ?t = "\lv, l, a, ?r\" + from "2.prems"(2) obtain r' where r': "(r', x) = del_max ?r" + and [simp]: "t' = adjust \lv, l, a, r'\" by auto + from "2.IH"[OF _ r'] \invar ?t\ have post: "post_del ?r r'" by simp + note preR = pre_adj_if_postR[OF \invar ?t\ post] + show ?case by (simp add: post_del_adjR[OF "2.prems"(1) preR post]) +qed (auto simp: post_del_def) + +theorem post_delete: "invar t \ post_del t (delete x t)" +proof (induction t) + case (Node lv l a r) + + let ?l' = "delete x l" and ?r' = "delete x r" + let ?t = "Node lv l a r" let ?t' = "delete x ?t" + + from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) + + note post_l' = Node.IH(1)[OF inv_l] + note preL = pre_adj_if_postL[OF Node.prems post_l'] + + note post_r' = Node.IH(2)[OF inv_r] + note preR = pre_adj_if_postR[OF Node.prems post_r'] + + show ?case + proof (cases rule: linorder_cases[of x a]) + case less + thus ?thesis using Node.prems by (simp add: post_del_adjL preL) + next + case greater + thus ?thesis using Node.prems by (simp add: post_del_adjR preR post_r') + next + case equal + show ?thesis + proof cases + assume "l = Leaf" thus ?thesis using equal Node.prems + by(auto simp: post_del_def invar.simps(2)) + next + assume "l \ Leaf" thus ?thesis using equal + by simp (metis Node.prems inv_l post_del_adjL post_del_max pre_adj_if_postL) + qed + qed +qed (simp add: post_del_def) + +declare invar_2Nodes[simp del] + subsection "Functional Correctness" + subsubsection "Proofs for insert" lemma inorder_split: "inorder(split t) = inorder t" @@ -103,28 +464,28 @@ "sorted(inorder t) \ inorder(insert x t) = ins_list x (inorder t)" by(induction t) (auto simp: ins_list_simps inorder_split inorder_skew) + subsubsection "Proofs for delete" +lemma inorder_adjust: "t \ Leaf \ pre_adjust t \ inorder(adjust t) = inorder t" +by(induction t) + (auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps + split: tree.splits) + lemma del_maxD: - "\ del_max t = (t',x); t \ Leaf \ \ inorder t' @ [x] = inorder t" + "\ del_max t = (t',x); t \ Leaf; invar t \ \ inorder t' @ [x] = inorder t" by(induction t arbitrary: t' rule: del_max.induct) - (auto simp: sorted_lems split: prod.splits) - -lemma inorder_adjust: "t \ Leaf \ inorder(adjust t) = inorder t" -by(induction t) - (auto simp: adjust_def inorder_skew inorder_split split: tree.splits) + (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_del_max split: prod.splits) lemma inorder_delete: - "sorted(inorder t) \ inorder(delete x t) = del_list x (inorder t)" + "invar t \ sorted(inorder t) \ inorder(delete x t) = del_list x (inorder t)" by(induction t) - (auto simp: del_list_simps inorder_adjust del_maxD split: prod.splits) - + (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR + post_del_max post_delete del_maxD split: prod.splits) -subsection "Overall correctness" - -interpretation Set_by_Ordered +interpretation I: Set_by_Ordered where empty = Leaf and isin = isin and insert = insert and delete = delete -and inorder = inorder and inv = "\_. True" +and inorder = inorder and inv = invar proof (standard, goal_cases) case 1 show ?case by simp next @@ -133,6 +494,12 @@ case 3 thus ?case by(simp add: inorder_insert) next case 4 thus ?case by(simp add: inorder_delete) -qed auto +next + case 5 thus ?case by(simp) +next + case 6 thus ?case by(simp add: invar_insert) +next + case 7 thus ?case using post_delete by(auto simp: post_del_def) +qed end diff -r 83db706d7771 -r f187aaf602c4 src/HOL/Data_Structures/document/root.bib --- a/src/HOL/Data_Structures/document/root.bib Tue Mar 01 22:49:33 2016 +0100 +++ b/src/HOL/Data_Structures/document/root.bib Wed Mar 02 10:01:31 2016 +0100 @@ -1,3 +1,8 @@ +@inproceedings{Andersson-WADS93,author={Arne Andersson}, +title={Balanced search trees made simple},pages={60--71},year=1993, +booktitle={Algorithms and Data Structures (WADS '93)}, +series={LNCS},volume={709},publisher={Springer}} + @article{Hinze-bro12,author={Ralf Hinze}, title={Purely Functional 1-2 Brother Trees}, journal={J. Functional Programming}, @@ -14,6 +19,10 @@ @misc{Kahrs-html,author={Stefan Kahrs},title={Red Black Trees}, note={\url{http://www.cs.ukc.ac.uk/people/staff/smk/redblack/rb.html}}} +@unpublished{Nipkow16,author={Tobias Nipkow}, +title={Automatic Functional Correctness Proofs for Functional Search Trees}, +year=2016,month=feb,note={\url{http://www.in.tum.de/~nipkow/pubs/trees.html}}} + @book{Okasaki,author={Chris Okasaki},title="Purely Functional Data Structures", publisher="Cambridge University Press",year=1998} @@ -25,6 +34,11 @@ title={1-2 Brother Trees or {AVL} Trees Revisited},journal={Comput. J.}, volume=23,number=3,pages={248--255},year=1980} +@inproceedings{Ragde14,author={Prabhakar Ragde}, +title={Simple Balanced Binary Search Trees},pages={78--87},year=2014, +booktitle={Trends in Functional Programming in Education}, +series={EPTCS},volume=170,editor={Caldwell and H\"olzenspies and Achten}} + @article{Reade-SCP92,author={Chris Reade}, title={Balanced Trees with Removals: An Exercise in Rewriting and Proof}, journal={Sci. Comput. Program.},volume=18,number=2,pages={181--204},year=1992} diff -r 83db706d7771 -r f187aaf602c4 src/HOL/Data_Structures/document/root.tex --- a/src/HOL/Data_Structures/document/root.tex Tue Mar 01 22:49:33 2016 +0100 +++ b/src/HOL/Data_Structures/document/root.tex Wed Mar 02 10:01:31 2016 +0100 @@ -27,6 +27,8 @@ A collection of verified functional data structures. The emphasis is on conciseness of algorithms and succinctness of proofs, more in the style of a textbook than a library of efficient algorithms. + +For more details see \cite{Nipkow16}. \end{abstract} \setcounter{tocdepth}{1} @@ -52,6 +54,11 @@ They were invented by Ottmann and Six~\cite{OttmannS76,OttmannW-CJ80}. The functional version is due to Hinze~\cite{Hinze-bro12}. +\paragraph{AA trees} +They were invented by Arne Anderson \cite{Andersson-WADS93}. +Our formalisation follows Ragde~\cite{Ragde14} but fixes a number +of mistakes. + \paragraph{Splay trees} They were invented by Sleator and Tarjan \cite{SleatorT-JACM85}. Our formalisation follows Schoenmakers \cite{Schoenmakers-IPL93}.