| author | wenzelm | 
| Sat, 05 Jan 2019 17:24:33 +0100 | |
| changeset 69597 | ff784d5a5bfb | 
| parent 69505 | cc2d676d5395 | 
| child 70755 | 3fb16bed5d6c | 
| permissions | -rw-r--r-- | 
| 61793 | 1 | (* | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 2 | Author: Tobias Nipkow, Daniel Stüwe | 
| 61793 | 3 | *) | 
| 4 | ||
| 62130 | 5 | section \<open>AA Tree Implementation of Sets\<close> | 
| 61793 | 6 | |
| 7 | theory AA_Set | |
| 8 | imports | |
| 9 | Isin2 | |
| 10 | Cmp | |
| 11 | begin | |
| 12 | ||
| 13 | type_synonym 'a aa_tree = "('a,nat) tree"
 | |
| 14 | ||
| 68431 | 15 | definition empty :: "'a aa_tree" where | 
| 16 | "empty = Leaf" | |
| 17 | ||
| 61793 | 18 | fun lvl :: "'a aa_tree \<Rightarrow> nat" where | 
| 19 | "lvl Leaf = 0" | | |
| 68413 | 20 | "lvl (Node _ _ lv _) = lv" | 
| 62496 | 21 | |
| 61793 | 22 | fun invar :: "'a aa_tree \<Rightarrow> bool" where | 
| 23 | "invar Leaf = True" | | |
| 68413 | 24 | "invar (Node l a h r) = | 
| 61793 | 25 | (invar l \<and> invar r \<and> | 
| 68413 | 26 | h = lvl l + 1 \<and> (h = lvl r + 1 \<or> (\<exists>lr b rr. r = Node lr b h rr \<and> h = lvl rr + 1)))" | 
| 62496 | 27 | |
| 61793 | 28 | fun skew :: "'a aa_tree \<Rightarrow> 'a aa_tree" where | 
| 68413 | 29 | "skew (Node (Node t1 b lvb t2) a lva t3) = | 
| 30 | (if lva = lvb then Node t1 b lvb (Node t2 a lva t3) else Node (Node t1 b lvb t2) a lva t3)" | | |
| 61793 | 31 | "skew t = t" | 
| 32 | ||
| 33 | fun split :: "'a aa_tree \<Rightarrow> 'a aa_tree" where | |
| 68413 | 34 | "split (Node t1 a lva (Node t2 b lvb (Node t3 c lvc t4))) = | 
| 67369 | 35 | (if lva = lvb \<and> lvb = lvc \<comment> \<open>\<open>lva = lvc\<close> suffices\<close> | 
| 68413 | 36 | then Node (Node t1 a lva t2) b (lva+1) (Node t3 c lva t4) | 
| 37 | else Node t1 a lva (Node t2 b lvb (Node t3 c lvc t4)))" | | |
| 61793 | 38 | "split t = t" | 
| 39 | ||
| 40 | hide_const (open) insert | |
| 41 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 42 | fun insert :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where | 
| 68413 | 43 | "insert x Leaf = Node Leaf x 1 Leaf" | | 
| 44 | "insert x (Node t1 a lv t2) = | |
| 61793 | 45 | (case cmp x a of | 
| 68413 | 46 | LT \<Rightarrow> split (skew (Node (insert x t1) a lv t2)) | | 
| 47 | GT \<Rightarrow> split (skew (Node t1 a lv (insert x t2))) | | |
| 48 | EQ \<Rightarrow> Node t1 x lv t2)" | |
| 61793 | 49 | |
| 50 | fun sngl :: "'a aa_tree \<Rightarrow> bool" where | |
| 51 | "sngl Leaf = False" | | |
| 52 | "sngl (Node _ _ _ Leaf) = True" | | |
| 68413 | 53 | "sngl (Node _ _ lva (Node _ _ lvb _)) = (lva > lvb)" | 
| 61793 | 54 | |
| 55 | definition adjust :: "'a aa_tree \<Rightarrow> 'a aa_tree" where | |
| 56 | "adjust t = | |
| 57 | (case t of | |
| 68413 | 58 | Node l x lv r \<Rightarrow> | 
| 61793 | 59 | (if lvl l >= lv-1 \<and> lvl r >= lv-1 then t else | 
| 68413 | 60 | if lvl r < lv-1 \<and> sngl l then skew (Node l x (lv-1) r) else | 
| 61793 | 61 | if lvl r < lv-1 | 
| 62 | then case l of | |
| 68413 | 63 | Node t1 a lva (Node t2 b lvb t3) | 
| 64 | \<Rightarrow> Node (Node t1 a lva t2) b (lvb+1) (Node t3 x (lv-1) r) | |
| 61793 | 65 | else | 
| 68413 | 66 | if lvl r < lv then split (Node l x (lv-1) r) | 
| 61793 | 67 | else | 
| 68 | case r of | |
| 68413 | 69 | Node t1 b lvb t4 \<Rightarrow> | 
| 61793 | 70 | (case t1 of | 
| 68413 | 71 | Node t2 a lva t3 | 
| 72 | \<Rightarrow> Node (Node l x (lv-1) t2) a (lva+1) | |
| 73 | (split (Node t3 b (if sngl t1 then lva else lva+1) t4)))))" | |
| 62496 | 74 | |
| 69597 | 75 | text\<open>In the paper, the last case of \<^const>\<open>adjust\<close> is expressed with the help of an | 
| 62496 | 76 | incorrect auxiliary function \texttt{nlvl}.
 | 
| 77 | ||
| 69505 | 78 | Function \<open>split_max\<close> below is called \texttt{dellrg} in the paper.
 | 
| 62496 | 79 | The latter is incorrect for two reasons: \texttt{dellrg} is meant to delete the largest
 | 
| 80 | element but recurses on the left instead of the right subtree; the invariant | |
| 67406 | 81 | is not restored.\<close> | 
| 62496 | 82 | |
| 68023 | 83 | fun split_max :: "'a aa_tree \<Rightarrow> 'a aa_tree * 'a" where | 
| 68413 | 84 | "split_max (Node l a lv Leaf) = (l,a)" | | 
| 85 | "split_max (Node l a lv r) = (let (r',b) = split_max r in (adjust(Node l a lv r'), b))" | |
| 61793 | 86 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 87 | fun delete :: "'a::linorder \<Rightarrow> 'a aa_tree \<Rightarrow> 'a aa_tree" where | 
| 61793 | 88 | "delete _ Leaf = Leaf" | | 
| 68413 | 89 | "delete x (Node l a lv r) = | 
| 61793 | 90 | (case cmp x a of | 
| 68413 | 91 | LT \<Rightarrow> adjust (Node (delete x l) a lv r) | | 
| 92 | GT \<Rightarrow> adjust (Node l a lv (delete x r)) | | |
| 61793 | 93 | EQ \<Rightarrow> (if l = Leaf then r | 
| 68413 | 94 | else let (l',b) = split_max l in adjust (Node l' b lv r)))" | 
| 61793 | 95 | |
| 62496 | 96 | fun pre_adjust where | 
| 68413 | 97 | "pre_adjust (Node l a lv r) = (invar l \<and> invar r \<and> | 
| 62496 | 98 | ((lv = lvl l + 1 \<and> (lv = lvl r + 1 \<or> lv = lvl r + 2 \<or> lv = lvl r \<and> sngl r)) \<or> | 
| 99 | (lv = lvl l + 2 \<and> (lv = lvl r + 1 \<or> lv = lvl r \<and> sngl r))))" | |
| 100 | ||
| 101 | declare pre_adjust.simps [simp del] | |
| 102 | ||
| 103 | subsection "Auxiliary Proofs" | |
| 104 | ||
| 105 | lemma split_case: "split t = (case t of | |
| 68413 | 106 | Node t1 x lvx (Node t2 y lvy (Node t3 z lvz t4)) \<Rightarrow> | 
| 62496 | 107 | (if lvx = lvy \<and> lvy = lvz | 
| 68413 | 108 | then Node (Node t1 x lvx t2) y (lvx+1) (Node t3 z lvx t4) | 
| 62496 | 109 | else t) | 
| 110 | | t \<Rightarrow> t)" | |
| 111 | by(auto split: tree.split) | |
| 112 | ||
| 113 | lemma skew_case: "skew t = (case t of | |
| 68413 | 114 | Node (Node t1 y lvy t2) x lvx t3 \<Rightarrow> | 
| 115 | (if lvx = lvy then Node t1 y lvx (Node t2 x lvx t3) else t) | |
| 62496 | 116 | | t \<Rightarrow> t)" | 
| 117 | by(auto split: tree.split) | |
| 118 | ||
| 119 | lemma lvl_0_iff: "invar t \<Longrightarrow> lvl t = 0 \<longleftrightarrow> t = Leaf" | |
| 120 | by(cases t) auto | |
| 121 | ||
| 68413 | 122 | lemma lvl_Suc_iff: "lvl t = Suc n \<longleftrightarrow> (\<exists> l a r. t = Node l a (Suc n) r)" | 
| 62496 | 123 | by(cases t) auto | 
| 124 | ||
| 125 | lemma lvl_skew: "lvl (skew t) = lvl t" | |
| 62526 | 126 | by(cases t rule: skew.cases) auto | 
| 62496 | 127 | |
| 128 | lemma lvl_split: "lvl (split t) = lvl t \<or> lvl (split t) = lvl t + 1 \<and> sngl (split t)" | |
| 62526 | 129 | by(cases t rule: split.cases) auto | 
| 62496 | 130 | |
| 68413 | 131 | lemma invar_2Nodes:"invar (Node l x lv (Node rl rx rlv rr)) = | 
| 132 | (invar l \<and> invar \<langle>rl, rx, rlv, rr\<rangle> \<and> lv = Suc (lvl l) \<and> | |
| 62496 | 133 | (lv = Suc rlv \<or> rlv = lv \<and> lv = Suc (lvl rr)))" | 
| 134 | by simp | |
| 135 | ||
| 136 | lemma invar_NodeLeaf[simp]: | |
| 68413 | 137 | "invar (Node l x lv Leaf) = (invar l \<and> lv = Suc (lvl l) \<and> lv = Suc 0)" | 
| 62496 | 138 | by simp | 
| 139 | ||
| 68413 | 140 | lemma sngl_if_invar: "invar (Node l a n r) \<Longrightarrow> n = lvl r \<Longrightarrow> sngl r" | 
| 62496 | 141 | by(cases r rule: sngl.cases) clarsimp+ | 
| 142 | ||
| 143 | ||
| 144 | subsection "Invariance" | |
| 145 | ||
| 146 | subsubsection "Proofs for insert" | |
| 147 | ||
| 148 | lemma lvl_insert_aux: | |
| 149 | "lvl (insert x t) = lvl t \<or> lvl (insert x t) = lvl t + 1 \<and> sngl (insert x t)" | |
| 150 | apply(induction t) | |
| 151 | apply (auto simp: lvl_skew) | |
| 152 | apply (metis Suc_eq_plus1 lvl.simps(2) lvl_split lvl_skew)+ | |
| 153 | done | |
| 154 | ||
| 155 | lemma lvl_insert: obtains | |
| 156 | (Same) "lvl (insert x t) = lvl t" | | |
| 157 | (Incr) "lvl (insert x t) = lvl t + 1" "sngl (insert x t)" | |
| 158 | using lvl_insert_aux by blast | |
| 159 | ||
| 160 | lemma lvl_insert_sngl: "invar t \<Longrightarrow> sngl t \<Longrightarrow> lvl(insert x t) = lvl t" | |
| 62526 | 161 | proof (induction t rule: insert.induct) | 
| 68413 | 162 | case (2 x t1 a lv t2) | 
| 62496 | 163 | consider (LT) "x < a" | (GT) "x > a" | (EQ) "x = a" | 
| 164 | using less_linear by blast | |
| 165 | thus ?case proof cases | |
| 166 | case LT | |
| 167 | thus ?thesis using 2 by (auto simp add: skew_case split_case split: tree.splits) | |
| 168 | next | |
| 169 | case GT | |
| 170 | thus ?thesis using 2 proof (cases t1) | |
| 171 | case Node | |
| 172 | thus ?thesis using 2 GT | |
| 173 | apply (auto simp add: skew_case split_case split: tree.splits) | |
| 174 | by (metis less_not_refl2 lvl.simps(2) lvl_insert_aux n_not_Suc_n sngl.simps(3))+ | |
| 175 | qed (auto simp add: lvl_0_iff) | |
| 176 | qed simp | |
| 177 | qed simp | |
| 178 | ||
| 179 | lemma skew_invar: "invar t \<Longrightarrow> skew t = t" | |
| 62526 | 180 | by(cases t rule: skew.cases) auto | 
| 62496 | 181 | |
| 182 | lemma split_invar: "invar t \<Longrightarrow> split t = t" | |
| 62526 | 183 | by(cases t rule: split.cases) clarsimp+ | 
| 62496 | 184 | |
| 185 | lemma invar_NodeL: | |
| 68413 | 186 | "\<lbrakk> invar(Node l x n r); invar l'; lvl l' = lvl l \<rbrakk> \<Longrightarrow> invar(Node l' x n r)" | 
| 62496 | 187 | by(auto) | 
| 188 | ||
| 189 | lemma invar_NodeR: | |
| 68413 | 190 | "\<lbrakk> invar(Node l x n r); n = lvl r + 1; invar r'; lvl r' = lvl r \<rbrakk> \<Longrightarrow> invar(Node l x n r')" | 
| 62496 | 191 | by(auto) | 
| 192 | ||
| 193 | lemma invar_NodeR2: | |
| 68413 | 194 | "\<lbrakk> invar(Node l x n r); sngl r'; n = lvl r + 1; invar r'; lvl r' = n \<rbrakk> \<Longrightarrow> invar(Node l x n r')" | 
| 62496 | 195 | by(cases r' rule: sngl.cases) clarsimp+ | 
| 196 | ||
| 197 | ||
| 198 | lemma lvl_insert_incr_iff: "(lvl(insert a t) = lvl t + 1) \<longleftrightarrow> | |
| 68413 | 199 | (\<exists>l x r. insert a t = Node l x (lvl t + 1) r \<and> lvl l = lvl r)" | 
| 62496 | 200 | apply(cases t) | 
| 201 | apply(auto simp add: skew_case split_case split: if_splits) | |
| 202 | apply(auto split: tree.splits if_splits) | |
| 203 | done | |
| 204 | ||
| 205 | lemma invar_insert: "invar t \<Longrightarrow> invar(insert a t)" | |
| 206 | proof(induction t) | |
| 68413 | 207 | case N: (Node l x n r) | 
| 62496 | 208 | hence il: "invar l" and ir: "invar r" by auto | 
| 67040 | 209 | note iil = N.IH(1)[OF il] | 
| 210 | note iir = N.IH(2)[OF ir] | |
| 68413 | 211 | let ?t = "Node l x n r" | 
| 62496 | 212 | have "a < x \<or> a = x \<or> x < a" by auto | 
| 213 | moreover | |
| 67040 | 214 | have ?case if "a < x" | 
| 215 | proof (cases rule: lvl_insert[of a l]) | |
| 216 | case (Same) thus ?thesis | |
| 217 | using \<open>a<x\<close> invar_NodeL[OF N.prems iil Same] | |
| 218 | by (simp add: skew_invar split_invar del: invar.simps) | |
| 219 | next | |
| 220 | case (Incr) | |
| 68413 | 221 | then obtain t1 w t2 where ial[simp]: "insert a l = Node t1 w n t2" | 
| 67040 | 222 | using N.prems by (auto simp: lvl_Suc_iff) | 
| 223 | have l12: "lvl t1 = lvl t2" | |
| 224 | by (metis Incr(1) ial lvl_insert_incr_iff tree.inject) | |
| 68413 | 225 | have "insert a ?t = split(skew(Node (insert a l) x n r))" | 
| 67040 | 226 | by(simp add: \<open>a<x\<close>) | 
| 68413 | 227 | also have "skew(Node (insert a l) x n r) = Node t1 w n (Node t2 x n r)" | 
| 67040 | 228 | by(simp) | 
| 229 | also have "invar(split \<dots>)" | |
| 230 | proof (cases r) | |
| 231 | case Leaf | |
| 232 | hence "l = Leaf" using N.prems by(auto simp: lvl_0_iff) | |
| 233 | thus ?thesis using Leaf ial by simp | |
| 62496 | 234 | next | 
| 68413 | 235 | case [simp]: (Node t3 y m t4) | 
| 67040 | 236 | show ?thesis (*using N(3) iil l12 by(auto)*) | 
| 237 | proof cases | |
| 238 | assume "m = n" thus ?thesis using N(3) iil by(auto) | |
| 62496 | 239 | next | 
| 67040 | 240 | assume "m \<noteq> n" thus ?thesis using N(3) iil l12 by(auto) | 
| 62496 | 241 | qed | 
| 242 | qed | |
| 67040 | 243 | finally show ?thesis . | 
| 244 | qed | |
| 62496 | 245 | moreover | 
| 67040 | 246 | have ?case if "x < a" | 
| 247 | proof - | |
| 62496 | 248 | from \<open>invar ?t\<close> have "n = lvl r \<or> n = lvl r + 1" by auto | 
| 67040 | 249 | thus ?case | 
| 62496 | 250 | proof | 
| 251 | assume 0: "n = lvl r" | |
| 68413 | 252 | have "insert a ?t = split(skew(Node l x n (insert a r)))" | 
| 62496 | 253 | using \<open>a>x\<close> by(auto) | 
| 68413 | 254 | also have "skew(Node l x n (insert a r)) = Node l x n (insert a r)" | 
| 67040 | 255 | using N.prems by(simp add: skew_case split: tree.split) | 
| 62496 | 256 | also have "invar(split \<dots>)" | 
| 257 | proof - | |
| 258 | from lvl_insert_sngl[OF ir sngl_if_invar[OF \<open>invar ?t\<close> 0], of a] | |
| 68413 | 259 | obtain t1 y t2 where iar: "insert a r = Node t1 y n t2" | 
| 67040 | 260 | using N.prems 0 by (auto simp: lvl_Suc_iff) | 
| 261 | from N.prems iar 0 iir | |
| 62496 | 262 | show ?thesis by (auto simp: split_case split: tree.splits) | 
| 263 | qed | |
| 264 | finally show ?thesis . | |
| 265 | next | |
| 266 | assume 1: "n = lvl r + 1" | |
| 267 | hence "sngl ?t" by(cases r) auto | |
| 268 | show ?thesis | |
| 269 | proof (cases rule: lvl_insert[of a r]) | |
| 270 | case (Same) | |
| 67040 | 271 | show ?thesis using \<open>x<a\<close> il ir invar_NodeR[OF N.prems 1 iir Same] | 
| 62496 | 272 | by (auto simp add: skew_invar split_invar) | 
| 273 | next | |
| 274 | case (Incr) | |
| 67406 | 275 | thus ?thesis using invar_NodeR2[OF \<open>invar ?t\<close> Incr(2) 1 iir] 1 \<open>x < a\<close> | 
| 62496 | 276 | by (auto simp add: skew_invar split_invar split: if_splits) | 
| 277 | qed | |
| 278 | qed | |
| 67040 | 279 | qed | 
| 280 | moreover | |
| 281 | have "a = x \<Longrightarrow> ?case" using N.prems by auto | |
| 62496 | 282 | ultimately show ?case by blast | 
| 283 | qed simp | |
| 284 | ||
| 285 | ||
| 286 | subsubsection "Proofs for delete" | |
| 287 | ||
| 68413 | 288 | lemma invarL: "ASSUMPTION(invar \<langle>l, a, lv, r\<rangle>) \<Longrightarrow> invar l" | 
| 62496 | 289 | by(simp add: ASSUMPTION_def) | 
| 290 | ||
| 291 | lemma invarR: "ASSUMPTION(invar \<langle>lv, l, a, r\<rangle>) \<Longrightarrow> invar r" | |
| 292 | by(simp add: ASSUMPTION_def) | |
| 293 | ||
| 294 | lemma sngl_NodeI: | |
| 68413 | 295 | "sngl (Node l a lv r) \<Longrightarrow> sngl (Node l' a' lv r)" | 
| 62496 | 296 | by(cases r) (simp_all) | 
| 297 | ||
| 298 | ||
| 299 | declare invarL[simp] invarR[simp] | |
| 300 | ||
| 301 | lemma pre_cases: | |
| 68413 | 302 | assumes "pre_adjust (Node l x lv r)" | 
| 62496 | 303 | obtains | 
| 304 | (tSngl) "invar l \<and> invar r \<and> | |
| 305 | lv = Suc (lvl r) \<and> lvl l = lvl r" | | |
| 306 | (tDouble) "invar l \<and> invar r \<and> | |
| 307 | lv = lvl r \<and> Suc (lvl l) = lvl r \<and> sngl r " | | |
| 308 | (rDown) "invar l \<and> invar r \<and> | |
| 309 | lv = Suc (Suc (lvl r)) \<and> lv = Suc (lvl l)" | | |
| 310 | (lDown_tSngl) "invar l \<and> invar r \<and> | |
| 311 | lv = Suc (lvl r) \<and> lv = Suc (Suc (lvl l))" | | |
| 312 | (lDown_tDouble) "invar l \<and> invar r \<and> | |
| 313 | lv = lvl r \<and> lv = Suc (Suc (lvl l)) \<and> sngl r" | |
| 314 | using assms unfolding pre_adjust.simps | |
| 315 | by auto | |
| 316 | ||
| 317 | declare invar.simps(2)[simp del] invar_2Nodes[simp add] | |
| 318 | ||
| 319 | lemma invar_adjust: | |
| 68413 | 320 | assumes pre: "pre_adjust (Node l a lv r)" | 
| 321 | shows "invar(adjust (Node l a lv r))" | |
| 62496 | 322 | using pre proof (cases rule: pre_cases) | 
| 323 | case (tDouble) thus ?thesis unfolding adjust_def by (cases r) (auto simp: invar.simps(2)) | |
| 324 | next | |
| 325 | case (rDown) | |
| 68413 | 326 | from rDown obtain llv ll la lr where l: "l = Node ll la llv lr" by (cases l) auto | 
| 62496 | 327 | from rDown show ?thesis unfolding adjust_def by (auto simp: l invar.simps(2) split: tree.splits) | 
| 328 | next | |
| 329 | case (lDown_tDouble) | |
| 68413 | 330 | from lDown_tDouble obtain rlv rr ra rl where r: "r = Node rl ra rlv rr" by (cases r) auto | 
| 62496 | 331 | from lDown_tDouble and r obtain rrlv rrr rra rrl where | 
| 68413 | 332 | rr :"rr = Node rrr rra rrlv rrl" by (cases rr) auto | 
| 62496 | 333 | from lDown_tDouble show ?thesis unfolding adjust_def r rr | 
| 63636 | 334 | apply (cases rl) apply (auto simp add: invar.simps(2) split!: if_split) | 
| 62496 | 335 | using lDown_tDouble by (auto simp: split_case lvl_0_iff elim:lvl.elims split: tree.split) | 
| 336 | qed (auto simp: split_case invar.simps(2) adjust_def split: tree.splits) | |
| 337 | ||
| 338 | lemma lvl_adjust: | |
| 68413 | 339 | assumes "pre_adjust (Node l a lv r)" | 
| 340 | shows "lv = lvl (adjust(Node l a lv r)) \<or> lv = lvl (adjust(Node l a lv r)) + 1" | |
| 62496 | 341 | using assms(1) proof(cases rule: pre_cases) | 
| 342 | case lDown_tSngl thus ?thesis | |
| 68413 | 343 | using lvl_split[of "\<langle>l, a, lvl r, r\<rangle>"] by (auto simp: adjust_def) | 
| 62496 | 344 | next | 
| 345 | case lDown_tDouble thus ?thesis | |
| 346 | by (auto simp: adjust_def invar.simps(2) split: tree.split) | |
| 347 | qed (auto simp: adjust_def split: tree.splits) | |
| 348 | ||
| 68413 | 349 | lemma sngl_adjust: assumes "pre_adjust (Node l a lv r)" | 
| 350 | "sngl \<langle>l, a, lv, r\<rangle>" "lv = lvl (adjust \<langle>l, a, lv, r\<rangle>)" | |
| 351 | shows "sngl (adjust \<langle>l, a, lv, r\<rangle>)" | |
| 62496 | 352 | using assms proof (cases rule: pre_cases) | 
| 353 | case rDown | |
| 354 | thus ?thesis using assms(2,3) unfolding adjust_def | |
| 355 | by (auto simp add: skew_case) (auto split: tree.split) | |
| 356 | qed (auto simp: adjust_def skew_case split_case split: tree.split) | |
| 357 | ||
| 358 | definition "post_del t t' == | |
| 359 | invar t' \<and> | |
| 360 | (lvl t' = lvl t \<or> lvl t' + 1 = lvl t) \<and> | |
| 361 | (lvl t' = lvl t \<and> sngl t \<longrightarrow> sngl t')" | |
| 362 | ||
| 363 | lemma pre_adj_if_postR: | |
| 364 | "invar\<langle>lv, l, a, r\<rangle> \<Longrightarrow> post_del r r' \<Longrightarrow> pre_adjust \<langle>lv, l, a, r'\<rangle>" | |
| 365 | by(cases "sngl r") | |
| 366 | (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) | |
| 367 | ||
| 368 | lemma pre_adj_if_postL: | |
| 68413 | 369 | "invar\<langle>l, a, lv, r\<rangle> \<Longrightarrow> post_del l l' \<Longrightarrow> pre_adjust \<langle>l', b, lv, r\<rangle>" | 
| 62496 | 370 | by(cases "sngl r") | 
| 371 | (auto simp: pre_adjust.simps post_del_def invar.simps(2) elim: sngl.elims) | |
| 372 | ||
| 373 | lemma post_del_adjL: | |
| 68413 | 374 | "\<lbrakk> invar\<langle>l, a, lv, r\<rangle>; pre_adjust \<langle>l', b, lv, r\<rangle> \<rbrakk> | 
| 375 | \<Longrightarrow> post_del \<langle>l, a, lv, r\<rangle> (adjust \<langle>l', b, lv, r\<rangle>)" | |
| 62496 | 376 | unfolding post_del_def | 
| 377 | by (metis invar_adjust lvl_adjust sngl_NodeI sngl_adjust lvl.simps(2)) | |
| 378 | ||
| 379 | lemma post_del_adjR: | |
| 380 | assumes "invar\<langle>lv, l, a, r\<rangle>" "pre_adjust \<langle>lv, l, a, r'\<rangle>" "post_del r r'" | |
| 381 | shows "post_del \<langle>lv, l, a, r\<rangle> (adjust \<langle>lv, l, a, r'\<rangle>)" | |
| 382 | proof(unfold post_del_def, safe del: disjCI) | |
| 383 | let ?t = "\<langle>lv, l, a, r\<rangle>" | |
| 384 | let ?t' = "adjust \<langle>lv, l, a, r'\<rangle>" | |
| 385 | show "invar ?t'" by(rule invar_adjust[OF assms(2)]) | |
| 386 | show "lvl ?t' = lvl ?t \<or> lvl ?t' + 1 = lvl ?t" | |
| 387 | using lvl_adjust[OF assms(2)] by auto | |
| 388 | show "sngl ?t'" if as: "lvl ?t' = lvl ?t" "sngl ?t" | |
| 389 | proof - | |
| 390 | have s: "sngl \<langle>lv, l, a, r'\<rangle>" | |
| 391 | proof(cases r') | |
| 392 | case Leaf thus ?thesis by simp | |
| 393 | next | |
| 394 | case Node thus ?thesis using as(2) assms(1,3) | |
| 395 | by (cases r) (auto simp: post_del_def) | |
| 396 | qed | |
| 397 | show ?thesis using as(1) sngl_adjust[OF assms(2) s] by simp | |
| 398 | qed | |
| 399 | qed | |
| 400 | ||
| 401 | declare prod.splits[split] | |
| 402 | ||
| 68023 | 403 | theorem post_split_max: | 
| 404 | "\<lbrakk> invar t; (t', x) = split_max t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> post_del t t'" | |
| 405 | proof (induction t arbitrary: t' rule: split_max.induct) | |
| 62496 | 406 | case (2 lv l a lvr rl ra rr) | 
| 407 | let ?r = "\<langle>lvr, rl, ra, rr\<rangle>" | |
| 408 | let ?t = "\<langle>lv, l, a, ?r\<rangle>" | |
| 68023 | 409 | from "2.prems"(2) obtain r' where r': "(r', x) = split_max ?r" | 
| 62496 | 410 | and [simp]: "t' = adjust \<langle>lv, l, a, r'\<rangle>" by auto | 
| 411 | from "2.IH"[OF _ r'] \<open>invar ?t\<close> have post: "post_del ?r r'" by simp | |
| 412 | note preR = pre_adj_if_postR[OF \<open>invar ?t\<close> post] | |
| 413 | show ?case by (simp add: post_del_adjR[OF "2.prems"(1) preR post]) | |
| 414 | qed (auto simp: post_del_def) | |
| 415 | ||
| 416 | theorem post_delete: "invar t \<Longrightarrow> post_del t (delete x t)" | |
| 417 | proof (induction t) | |
| 68413 | 418 | case (Node l a lv r) | 
| 62496 | 419 | |
| 420 | let ?l' = "delete x l" and ?r' = "delete x r" | |
| 68413 | 421 | let ?t = "Node l a lv r" let ?t' = "delete x ?t" | 
| 62496 | 422 | |
| 423 | from Node.prems have inv_l: "invar l" and inv_r: "invar r" by (auto) | |
| 424 | ||
| 425 | note post_l' = Node.IH(1)[OF inv_l] | |
| 426 | note preL = pre_adj_if_postL[OF Node.prems post_l'] | |
| 427 | ||
| 428 | note post_r' = Node.IH(2)[OF inv_r] | |
| 429 | note preR = pre_adj_if_postR[OF Node.prems post_r'] | |
| 430 | ||
| 431 | show ?case | |
| 432 | proof (cases rule: linorder_cases[of x a]) | |
| 433 | case less | |
| 434 | thus ?thesis using Node.prems by (simp add: post_del_adjL preL) | |
| 435 | next | |
| 436 | case greater | |
| 437 | thus ?thesis using Node.prems by (simp add: post_del_adjR preR post_r') | |
| 438 | next | |
| 439 | case equal | |
| 440 | show ?thesis | |
| 441 | proof cases | |
| 442 | assume "l = Leaf" thus ?thesis using equal Node.prems | |
| 443 | by(auto simp: post_del_def invar.simps(2)) | |
| 444 | next | |
| 445 | assume "l \<noteq> Leaf" thus ?thesis using equal | |
| 68023 | 446 | by simp (metis Node.prems inv_l post_del_adjL post_split_max pre_adj_if_postL) | 
| 62496 | 447 | qed | 
| 448 | qed | |
| 449 | qed (simp add: post_del_def) | |
| 450 | ||
| 451 | declare invar_2Nodes[simp del] | |
| 452 | ||
| 61793 | 453 | |
| 454 | subsection "Functional Correctness" | |
| 455 | ||
| 62496 | 456 | |
| 61793 | 457 | subsubsection "Proofs for insert" | 
| 458 | ||
| 459 | lemma inorder_split: "inorder(split t) = inorder t" | |
| 460 | by(cases t rule: split.cases) (auto) | |
| 461 | ||
| 462 | lemma inorder_skew: "inorder(skew t) = inorder t" | |
| 463 | by(cases t rule: skew.cases) (auto) | |
| 464 | ||
| 465 | lemma inorder_insert: | |
| 466 | "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)" | |
| 467 | by(induction t) (auto simp: ins_list_simps inorder_split inorder_skew) | |
| 468 | ||
| 62496 | 469 | |
| 61793 | 470 | subsubsection "Proofs for delete" | 
| 471 | ||
| 62496 | 472 | lemma inorder_adjust: "t \<noteq> Leaf \<Longrightarrow> pre_adjust t \<Longrightarrow> inorder(adjust t) = inorder t" | 
| 62526 | 473 | by(cases t) | 
| 62496 | 474 | (auto simp: adjust_def inorder_skew inorder_split invar.simps(2) pre_adjust.simps | 
| 475 | split: tree.splits) | |
| 476 | ||
| 68023 | 477 | lemma split_maxD: | 
| 478 | "\<lbrakk> split_max t = (t',x); t \<noteq> Leaf; invar t \<rbrakk> \<Longrightarrow> inorder t' @ [x] = inorder t" | |
| 479 | by(induction t arbitrary: t' rule: split_max.induct) | |
| 480 | (auto simp: sorted_lems inorder_adjust pre_adj_if_postR post_split_max split: prod.splits) | |
| 61793 | 481 | |
| 482 | lemma inorder_delete: | |
| 62496 | 483 | "invar t \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" | 
| 61793 | 484 | by(induction t) | 
| 62496 | 485 | (auto simp: del_list_simps inorder_adjust pre_adj_if_postL pre_adj_if_postR | 
| 68023 | 486 | post_split_max post_delete split_maxD split: prod.splits) | 
| 61793 | 487 | |
| 68440 | 488 | interpretation S: Set_by_Ordered | 
| 68431 | 489 | where empty = empty and isin = isin and insert = insert and delete = delete | 
| 62496 | 490 | and inorder = inorder and inv = invar | 
| 61793 | 491 | proof (standard, goal_cases) | 
| 68431 | 492 | case 1 show ?case by (simp add: empty_def) | 
| 61793 | 493 | next | 
| 67967 | 494 | case 2 thus ?case by(simp add: isin_set_inorder) | 
| 61793 | 495 | next | 
| 496 | case 3 thus ?case by(simp add: inorder_insert) | |
| 497 | next | |
| 498 | case 4 thus ?case by(simp add: inorder_delete) | |
| 62496 | 499 | next | 
| 68431 | 500 | case 5 thus ?case by(simp add: empty_def) | 
| 62496 | 501 | next | 
| 502 | case 6 thus ?case by(simp add: invar_insert) | |
| 503 | next | |
| 504 | case 7 thus ?case using post_delete by(auto simp: post_del_def) | |
| 505 | qed | |
| 61793 | 506 | |
| 62390 | 507 | end |