| author | wenzelm | 
| Mon, 11 Sep 2023 19:31:09 +0200 | |
| changeset 78660 | 0d2ea608d223 | 
| parent 72851 | f0fa51227a23 | 
| child 80576 | 37482793a949 | 
| permissions | -rw-r--r-- | 
| 64951 
140addd19343
removed contribution by Daniel Stuewe, too detailed.
 nipkow parents: 
64950diff
changeset | 1 | (* Author: Tobias Nipkow *) | 
| 61224 | 2 | |
| 3 | section \<open>Red-Black Tree Implementation of Sets\<close> | |
| 4 | ||
| 5 | theory RBT_Set | |
| 6 | imports | |
| 64950 | 7 | Complex_Main | 
| 61224 | 8 | RBT | 
| 61581 | 9 | Cmp | 
| 61224 | 10 | Isin2 | 
| 11 | begin | |
| 12 | ||
| 68431 | 13 | definition empty :: "'a rbt" where | 
| 14 | "empty = Leaf" | |
| 15 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 16 | fun ins :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where | 
| 61749 | 17 | "ins x Leaf = R Leaf x Leaf" | | 
| 18 | "ins x (B l a r) = | |
| 61678 | 19 | (case cmp x a of | 
| 64960 | 20 | LT \<Rightarrow> baliL (ins x l) a r | | 
| 21 | GT \<Rightarrow> baliR l a (ins x r) | | |
| 61678 | 22 | EQ \<Rightarrow> B l a r)" | | 
| 61749 | 23 | "ins x (R l a r) = | 
| 61678 | 24 | (case cmp x a of | 
| 61749 | 25 | LT \<Rightarrow> R (ins x l) a r | | 
| 26 | GT \<Rightarrow> R l a (ins x r) | | |
| 61678 | 27 | EQ \<Rightarrow> R l a r)" | 
| 61224 | 28 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 29 | definition insert :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where | 
| 61749 | 30 | "insert x t = paint Black (ins x t)" | 
| 31 | ||
| 66087 | 32 | fun color :: "'a rbt \<Rightarrow> color" where | 
| 33 | "color Leaf = Black" | | |
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 34 | "color (Node _ (_, c) _) = c" | 
| 66087 | 35 | |
| 36 | fun del :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where | |
| 61749 | 37 | "del x Leaf = Leaf" | | 
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 38 | "del x (Node l (a, _) r) = | 
| 61678 | 39 | (case cmp x a of | 
| 66087 | 40 | LT \<Rightarrow> if l \<noteq> Leaf \<and> color l = Black | 
| 41 | then baldL (del x l) a r else R (del x l) a r | | |
| 42 | GT \<Rightarrow> if r \<noteq> Leaf\<and> color r = Black | |
| 43 | then baldR l a (del x r) else R l a (del x r) | | |
| 71830 | 44 | EQ \<Rightarrow> join l r)" | 
| 61749 | 45 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 46 | definition delete :: "'a::linorder \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where | 
| 61749 | 47 | "delete x t = paint Black (del x t)" | 
| 61224 | 48 | |
| 49 | ||
| 50 | subsection "Functional Correctness Proofs" | |
| 51 | ||
| 61749 | 52 | lemma inorder_paint: "inorder(paint c t) = inorder t" | 
| 62526 | 53 | by(cases t) (auto) | 
| 61749 | 54 | |
| 64960 | 55 | lemma inorder_baliL: | 
| 56 | "inorder(baliL l a r) = inorder l @ a # inorder r" | |
| 57 | by(cases "(l,a,r)" rule: baliL.cases) (auto) | |
| 58 | ||
| 59 | lemma inorder_baliR: | |
| 60 | "inorder(baliR l a r) = inorder l @ a # inorder r" | |
| 61 | by(cases "(l,a,r)" rule: baliR.cases) (auto) | |
| 61224 | 62 | |
| 61749 | 63 | lemma inorder_ins: | 
| 64 | "sorted(inorder t) \<Longrightarrow> inorder(ins x t) = ins_list x (inorder t)" | |
| 64960 | 65 | by(induction x t rule: ins.induct) | 
| 66 | (auto simp: ins_list_simps inorder_baliL inorder_baliR) | |
| 61749 | 67 | |
| 61224 | 68 | lemma inorder_insert: | 
| 61749 | 69 | "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)" | 
| 70 | by (simp add: insert_def inorder_ins inorder_paint) | |
| 61224 | 71 | |
| 64960 | 72 | lemma inorder_baldL: | 
| 73 | "inorder(baldL l a r) = inorder l @ a # inorder r" | |
| 74 | by(cases "(l,a,r)" rule: baldL.cases) | |
| 75 | (auto simp: inorder_baliL inorder_baliR inorder_paint) | |
| 61224 | 76 | |
| 64960 | 77 | lemma inorder_baldR: | 
| 78 | "inorder(baldR l a r) = inorder l @ a # inorder r" | |
| 79 | by(cases "(l,a,r)" rule: baldR.cases) | |
| 80 | (auto simp: inorder_baliL inorder_baliR inorder_paint) | |
| 61224 | 81 | |
| 71830 | 82 | lemma inorder_join: | 
| 83 | "inorder(join l r) = inorder l @ inorder r" | |
| 84 | by(induction l r rule: join.induct) | |
| 64960 | 85 | (auto simp: inorder_baldL inorder_baldR split: tree.split color.split) | 
| 61224 | 86 | |
| 61749 | 87 | lemma inorder_del: | 
| 88 | "sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)" | |
| 66087 | 89 | by(induction x t rule: del.induct) | 
| 71830 | 90 | (auto simp: del_list_simps inorder_join inorder_baldL inorder_baldR) | 
| 61224 | 91 | |
| 61749 | 92 | lemma inorder_delete: | 
| 93 | "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)" | |
| 94 | by (auto simp: delete_def inorder_del inorder_paint) | |
| 95 | ||
| 61581 | 96 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 97 | subsection \<open>Structural invariants\<close> | 
| 61224 | 98 | |
| 71350 | 99 | lemma neq_Black[simp]: "(c \<noteq> Black) = (c = Red)" | 
| 100 | by (cases c) auto | |
| 101 | ||
| 64952 | 102 | text\<open>The proofs are due to Markus Reiter and Alexander Krauss.\<close> | 
| 61754 | 103 | |
| 104 | fun bheight :: "'a rbt \<Rightarrow> nat" where | |
| 105 | "bheight Leaf = 0" | | |
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 106 | "bheight (Node l (x, c) r) = (if c = Black then bheight l + 1 else bheight l)" | 
| 61754 | 107 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 108 | fun invc :: "'a rbt \<Rightarrow> bool" where | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 109 | "invc Leaf = True" | | 
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 110 | "invc (Node l (a,c) r) = | 
| 72586 | 111 | ((c = Red \<longrightarrow> color l = Black \<and> color r = Black) \<and> invc l \<and> invc r)" | 
| 61754 | 112 | |
| 70708 | 113 | text \<open>Weaker version:\<close> | 
| 114 | abbreviation invc2 :: "'a rbt \<Rightarrow> bool" where | |
| 115 | "invc2 t \<equiv> invc(paint Black t)" | |
| 61754 | 116 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 117 | fun invh :: "'a rbt \<Rightarrow> bool" where | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 118 | "invh Leaf = True" | | 
| 72586 | 119 | "invh (Node l (x, c) r) = (bheight l = bheight r \<and> invh l \<and> invh r)" | 
| 61754 | 120 | |
| 64953 | 121 | lemma invc2I: "invc t \<Longrightarrow> invc2 t" | 
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 122 | by (cases t rule: tree2_cases) simp+ | 
| 61754 | 123 | |
| 124 | definition rbt :: "'a rbt \<Rightarrow> bool" where | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 125 | "rbt t = (invc t \<and> invh t \<and> color t = Black)" | 
| 61754 | 126 | |
| 127 | lemma color_paint_Black: "color (paint Black t) = Black" | |
| 128 | by (cases t) auto | |
| 129 | ||
| 70708 | 130 | lemma paint2: "paint c2 (paint c1 t) = paint c2 t" | 
| 61754 | 131 | by (cases t) auto | 
| 132 | ||
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 133 | lemma invh_paint: "invh t \<Longrightarrow> invh (paint c t)" | 
| 61754 | 134 | by (cases t) auto | 
| 135 | ||
| 64960 | 136 | lemma invc_baliL: | 
| 137 | "\<lbrakk>invc2 l; invc r\<rbrakk> \<Longrightarrow> invc (baliL l a r)" | |
| 138 | by (induct l a r rule: baliL.induct) auto | |
| 139 | ||
| 140 | lemma invc_baliR: | |
| 141 | "\<lbrakk>invc l; invc2 r\<rbrakk> \<Longrightarrow> invc (baliR l a r)" | |
| 142 | by (induct l a r rule: baliR.induct) auto | |
| 143 | ||
| 144 | lemma bheight_baliL: | |
| 145 | "bheight l = bheight r \<Longrightarrow> bheight (baliL l a r) = Suc (bheight l)" | |
| 146 | by (induct l a r rule: baliL.induct) auto | |
| 61754 | 147 | |
| 64960 | 148 | lemma bheight_baliR: | 
| 149 | "bheight l = bheight r \<Longrightarrow> bheight (baliR l a r) = Suc (bheight l)" | |
| 150 | by (induct l a r rule: baliR.induct) auto | |
| 61754 | 151 | |
| 64960 | 152 | lemma invh_baliL: | 
| 153 | "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk> \<Longrightarrow> invh (baliL l a r)" | |
| 154 | by (induct l a r rule: baliL.induct) auto | |
| 155 | ||
| 156 | lemma invh_baliR: | |
| 157 | "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk> \<Longrightarrow> invh (baliR l a r)" | |
| 158 | by (induct l a r rule: baliR.induct) auto | |
| 61754 | 159 | |
| 70708 | 160 | text \<open>All in one:\<close> | 
| 161 | ||
| 162 | lemma inv_baliR: "\<lbrakk> invh l; invh r; invc l; invc2 r; bheight l = bheight r \<rbrakk> | |
| 163 | \<Longrightarrow> invc (baliR l a r) \<and> invh (baliR l a r) \<and> bheight (baliR l a r) = Suc (bheight l)" | |
| 164 | by (induct l a r rule: baliR.induct) auto | |
| 165 | ||
| 166 | lemma inv_baliL: "\<lbrakk> invh l; invh r; invc2 l; invc r; bheight l = bheight r \<rbrakk> | |
| 167 | \<Longrightarrow> invc (baliL l a r) \<and> invh (baliL l a r) \<and> bheight (baliL l a r) = Suc (bheight l)" | |
| 168 | by (induct l a r rule: baliL.induct) auto | |
| 61754 | 169 | |
| 170 | subsubsection \<open>Insertion\<close> | |
| 171 | ||
| 70708 | 172 | lemma invc_ins: "invc t \<longrightarrow> invc2 (ins x t) \<and> (color t = Black \<longrightarrow> invc (ins x t))" | 
| 64960 | 173 | by (induct x t rule: ins.induct) (auto simp: invc_baliL invc_baliR invc2I) | 
| 61754 | 174 | |
| 70708 | 175 | lemma invh_ins: "invh t \<Longrightarrow> invh (ins x t) \<and> bheight (ins x t) = bheight t" | 
| 64960 | 176 | by(induct x t rule: ins.induct) | 
| 177 | (auto simp: invh_baliL invh_baliR bheight_baliL bheight_baliR) | |
| 61754 | 178 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 179 | theorem rbt_insert: "rbt t \<Longrightarrow> rbt (insert x t)" | 
| 70708 | 180 | by (simp add: invc_ins invh_ins color_paint_Black invh_paint rbt_def insert_def) | 
| 181 | ||
| 71350 | 182 | text \<open>All in one:\<close> | 
| 70708 | 183 | |
| 184 | lemma inv_ins: "\<lbrakk> invc t; invh t \<rbrakk> \<Longrightarrow> | |
| 185 | invc2 (ins x t) \<and> (color t = Black \<longrightarrow> invc (ins x t)) \<and> | |
| 186 | invh(ins x t) \<and> bheight (ins x t) = bheight t" | |
| 187 | by (induct x t rule: ins.induct) (auto simp: inv_baliL inv_baliR invc2I) | |
| 188 | ||
| 189 | theorem rbt_insert2: "rbt t \<Longrightarrow> rbt (insert x t)" | |
| 190 | by (simp add: inv_ins color_paint_Black invh_paint rbt_def insert_def) | |
| 61754 | 191 | |
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 192 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 193 | subsubsection \<open>Deletion\<close> | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 194 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 195 | lemma bheight_paint_Red: | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 196 | "color t = Black \<Longrightarrow> bheight (paint Red t) = bheight t - 1" | 
| 61754 | 197 | by (cases t) auto | 
| 198 | ||
| 66087 | 199 | lemma invh_baldL_invc: | 
| 200 | "\<lbrakk> invh l; invh r; bheight l + 1 = bheight r; invc r \<rbrakk> | |
| 71348 | 201 | \<Longrightarrow> invh (baldL l a r) \<and> bheight (baldL l a r) = bheight r" | 
| 64960 | 202 | by (induct l a r rule: baldL.induct) | 
| 203 | (auto simp: invh_baliR invh_paint bheight_baliR bheight_paint_Red) | |
| 61754 | 204 | |
| 66087 | 205 | lemma invh_baldL_Black: | 
| 206 | "\<lbrakk> invh l; invh r; bheight l + 1 = bheight r; color r = Black \<rbrakk> | |
| 207 | \<Longrightarrow> invh (baldL l a r) \<and> bheight (baldL l a r) = bheight r" | |
| 64960 | 208 | by (induct l a r rule: baldL.induct) (auto simp add: invh_baliR bheight_baliR) | 
| 61754 | 209 | |
| 66087 | 210 | lemma invc_baldL: "\<lbrakk>invc2 l; invc r; color r = Black\<rbrakk> \<Longrightarrow> invc (baldL l a r)" | 
| 64960 | 211 | by (induct l a r rule: baldL.induct) (simp_all add: invc_baliR) | 
| 61754 | 212 | |
| 66087 | 213 | lemma invc2_baldL: "\<lbrakk> invc2 l; invc r \<rbrakk> \<Longrightarrow> invc2 (baldL l a r)" | 
| 70708 | 214 | by (induct l a r rule: baldL.induct) (auto simp: invc_baliR paint2 invc2I) | 
| 61754 | 215 | |
| 66087 | 216 | lemma invh_baldR_invc: | 
| 217 | "\<lbrakk> invh l; invh r; bheight l = bheight r + 1; invc l \<rbrakk> | |
| 218 | \<Longrightarrow> invh (baldR l a r) \<and> bheight (baldR l a r) = bheight l" | |
| 64960 | 219 | by(induct l a r rule: baldR.induct) | 
| 220 | (auto simp: invh_baliL bheight_baliL invh_paint bheight_paint_Red) | |
| 61754 | 221 | |
| 70571 | 222 | lemma invc_baldR: "\<lbrakk>invc l; invc2 r; color l = Black\<rbrakk> \<Longrightarrow> invc (baldR l a r)" | 
| 223 | by (induct l a r rule: baldR.induct) (simp_all add: invc_baliL) | |
| 61754 | 224 | |
| 70571 | 225 | lemma invc2_baldR: "\<lbrakk> invc l; invc2 r \<rbrakk> \<Longrightarrow>invc2 (baldR l a r)" | 
| 70708 | 226 | by (induct l a r rule: baldR.induct) (auto simp: invc_baliL paint2 invc2I) | 
| 61754 | 227 | |
| 71830 | 228 | lemma invh_join: | 
| 66087 | 229 | "\<lbrakk> invh l; invh r; bheight l = bheight r \<rbrakk> | 
| 71830 | 230 | \<Longrightarrow> invh (join l r) \<and> bheight (join l r) = bheight l" | 
| 231 | by (induct l r rule: join.induct) | |
| 66087 | 232 | (auto simp: invh_baldL_Black split: tree.splits color.splits) | 
| 61754 | 233 | |
| 71830 | 234 | lemma invc_join: | 
| 70708 | 235 | "\<lbrakk> invc l; invc r \<rbrakk> \<Longrightarrow> | 
| 71830 | 236 | (color l = Black \<and> color r = Black \<longrightarrow> invc (join l r)) \<and> invc2 (join l r)" | 
| 237 | by (induct l r rule: join.induct) | |
| 66087 | 238 | (auto simp: invc_baldL invc2I split: tree.splits color.splits) | 
| 61754 | 239 | |
| 71350 | 240 | text \<open>All in one:\<close> | 
| 241 | ||
| 242 | lemma inv_baldL: | |
| 243 | "\<lbrakk> invh l; invh r; bheight l + 1 = bheight r; invc2 l; invc r \<rbrakk> | |
| 244 | \<Longrightarrow> invh (baldL l a r) \<and> bheight (baldL l a r) = bheight r | |
| 245 | \<and> invc2 (baldL l a r) \<and> (color r = Black \<longrightarrow> invc (baldL l a r))" | |
| 246 | by (induct l a r rule: baldL.induct) | |
| 247 | (auto simp: inv_baliR invh_paint bheight_baliR bheight_paint_Red paint2 invc2I) | |
| 248 | ||
| 249 | lemma inv_baldR: | |
| 250 | "\<lbrakk> invh l; invh r; bheight l = bheight r + 1; invc l; invc2 r \<rbrakk> | |
| 251 | \<Longrightarrow> invh (baldR l a r) \<and> bheight (baldR l a r) = bheight l | |
| 252 | \<and> invc2 (baldR l a r) \<and> (color l = Black \<longrightarrow> invc (baldR l a r))" | |
| 253 | by (induct l a r rule: baldR.induct) | |
| 254 | (auto simp: inv_baliL invh_paint bheight_baliL bheight_paint_Red paint2 invc2I) | |
| 255 | ||
| 71830 | 256 | lemma inv_join: | 
| 71350 | 257 | "\<lbrakk> invh l; invh r; bheight l = bheight r; invc l; invc r \<rbrakk> | 
| 71830 | 258 | \<Longrightarrow> invh (join l r) \<and> bheight (join l r) = bheight l | 
| 259 | \<and> invc2 (join l r) \<and> (color l = Black \<and> color r = Black \<longrightarrow> invc (join l r))" | |
| 260 | by (induct l r rule: join.induct) | |
| 71350 | 261 | (auto simp: invh_baldL_Black inv_baldL invc2I split: tree.splits color.splits) | 
| 262 | ||
| 70755 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 263 | lemma neq_LeafD: "t \<noteq> Leaf \<Longrightarrow> \<exists>l x c r. t = Node l (x,c) r" | 
| 
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
 nipkow parents: 
70708diff
changeset | 264 | by(cases t rule: tree2_cases) auto | 
| 66087 | 265 | |
| 71350 | 266 | lemma inv_del: "\<lbrakk> invh t; invc t \<rbrakk> \<Longrightarrow> | 
| 267 | invh (del x t) \<and> | |
| 72851 | 268 | (color t = Red \<longrightarrow> bheight (del x t) = bheight t \<and> invc (del x t)) \<and> | 
| 269 | (color t = Black \<longrightarrow> bheight (del x t) = bheight t - 1 \<and> invc2 (del x t))" | |
| 71350 | 270 | by(induct x t rule: del.induct) | 
| 71830 | 271 | (auto simp: inv_baldL inv_baldR inv_join dest!: neq_LeafD) | 
| 61754 | 272 | |
| 70571 | 273 | theorem rbt_delete: "rbt t \<Longrightarrow> rbt (delete x t)" | 
| 72851 | 274 | by (metis delete_def rbt_def color_paint_Black inv_del invh_paint) | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 275 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 276 | text \<open>Overall correctness:\<close> | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 277 | |
| 68440 | 278 | interpretation S: Set_by_Ordered | 
| 68431 | 279 | where empty = empty and isin = isin and insert = insert and delete = delete | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 280 | and inorder = inorder and inv = rbt | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 281 | proof (standard, goal_cases) | 
| 68431 | 282 | case 1 show ?case by (simp add: empty_def) | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 283 | next | 
| 67967 | 284 | case 2 thus ?case by(simp add: isin_set_inorder) | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 285 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 286 | case 3 thus ?case by(simp add: inorder_insert) | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 287 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 288 | case 4 thus ?case by(simp add: inorder_delete) | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 289 | next | 
| 68431 | 290 | case 5 thus ?case by (simp add: rbt_def empty_def) | 
| 63411 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 291 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 292 | case 6 thus ?case by (simp add: rbt_insert) | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 293 | next | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 294 | case 7 thus ?case by (simp add: rbt_delete) | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 295 | qed | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 296 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 297 | |
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 298 | subsection \<open>Height-Size Relation\<close> | 
| 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 nipkow parents: 
62526diff
changeset | 299 | |
| 67963 | 300 | lemma rbt_height_bheight_if: "invc t \<Longrightarrow> invh t \<Longrightarrow> | 
| 71346 | 301 | height t \<le> 2 * bheight t + (if color t = Black then 0 else 1)" | 
| 64950 | 302 | by(induction t) (auto split: if_split_asm) | 
| 303 | ||
| 304 | lemma rbt_height_bheight: "rbt t \<Longrightarrow> height t / 2 \<le> bheight t " | |
| 305 | by(auto simp: rbt_def dest: rbt_height_bheight_if) | |
| 306 | ||
| 67963 | 307 | lemma bheight_size_bound: "invc t \<Longrightarrow> invh t \<Longrightarrow> 2 ^ (bheight t) \<le> size1 t" | 
| 64950 | 308 | by (induction t) auto | 
| 309 | ||
| 310 | lemma rbt_height_le: assumes "rbt t" shows "height t \<le> 2 * log 2 (size1 t)" | |
| 311 | proof - | |
| 312 | have "2 powr (height t / 2) \<le> 2 powr bheight t" | |
| 313 | using rbt_height_bheight[OF assms] by (simp) | |
| 314 | also have "\<dots> \<le> size1 t" using assms | |
| 315 | by (simp add: powr_realpow bheight_size_bound rbt_def) | |
| 316 | finally have "2 powr (height t / 2) \<le> size1 t" . | |
| 317 | hence "height t / 2 \<le> log 2 (size1 t)" | |
| 68998 | 318 | by (simp add: le_log_iff size1_size del: divide_le_eq_numeral1(1)) | 
| 64950 | 319 | thus ?thesis by simp | 
| 320 | qed | |
| 321 | ||
| 61224 | 322 | end |