| author | wenzelm | 
| Sun, 18 Dec 2016 13:46:57 +0100 | |
| changeset 64597 | 1c252d8b6ca6 | 
| parent 64541 | 3d4331b65861 | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 63829 | 1 | (* Author: Tobias Nipkow *) | 
| 63643 | 2 | |
| 63829 | 3 | section \<open>Creating Balanced Trees\<close> | 
| 63643 | 4 | |
| 63829 | 5 | theory Balance | 
| 63643 | 6 | imports | 
| 64018 | 7 | Complex_Main | 
| 63643 | 8 | "~~/src/HOL/Library/Tree" | 
| 9 | begin | |
| 10 | ||
| 64533 | 11 | (* The following two lemmas should go into theory \<open>Tree\<close>, except that that | 
| 12 | theory would then depend on \<open>Complex_Main\<close>. *) | |
| 13 | ||
| 14 | lemma min_height_balanced: assumes "balanced t" | |
| 15 | shows "min_height t = nat(floor(log 2 (size1 t)))" | |
| 16 | proof cases | |
| 17 | assume *: "complete t" | |
| 18 | hence "size1 t = 2 ^ min_height t" | |
| 19 | by (simp add: complete_iff_height size1_if_complete) | |
| 20 | hence "size1 t = 2 powr min_height t" | |
| 21 | using * by (simp add: powr_realpow) | |
| 22 | hence "min_height t = log 2 (size1 t)" | |
| 23 | by simp | |
| 24 | thus ?thesis | |
| 25 | by linarith | |
| 26 | next | |
| 27 | assume *: "\<not> complete t" | |
| 28 | hence "height t = min_height t + 1" | |
| 64540 | 29 | using assms min_height_le_height[of t] | 
| 64533 | 30 | by(auto simp add: balanced_def complete_iff_height) | 
| 31 | hence "2 ^ min_height t \<le> size1 t \<and> size1 t < 2 ^ (min_height t + 1)" | |
| 32 | by (metis * min_height_size1 size1_height_if_incomplete) | |
| 33 | hence "2 powr min_height t \<le> size1 t \<and> size1 t < 2 powr (min_height t + 1)" | |
| 34 | by(simp only: powr_realpow) | |
| 35 | (metis of_nat_less_iff of_nat_le_iff of_nat_numeral of_nat_power) | |
| 36 | hence "min_height t \<le> log 2 (size1 t) \<and> log 2 (size1 t) < min_height t + 1" | |
| 37 | by(simp add: log_less_iff le_log_iff) | |
| 38 | thus ?thesis by linarith | |
| 39 | qed | |
| 40 | ||
| 41 | lemma height_balanced: assumes "balanced t" | |
| 42 | shows "height t = nat(ceiling(log 2 (size1 t)))" | |
| 43 | proof cases | |
| 44 | assume *: "complete t" | |
| 45 | hence "size1 t = 2 ^ height t" | |
| 46 | by (simp add: size1_if_complete) | |
| 47 | hence "size1 t = 2 powr height t" | |
| 48 | using * by (simp add: powr_realpow) | |
| 49 | hence "height t = log 2 (size1 t)" | |
| 50 | by simp | |
| 51 | thus ?thesis | |
| 52 | by linarith | |
| 53 | next | |
| 54 | assume *: "\<not> complete t" | |
| 55 | hence **: "height t = min_height t + 1" | |
| 64540 | 56 | using assms min_height_le_height[of t] | 
| 64533 | 57 | by(auto simp add: balanced_def complete_iff_height) | 
| 58 | hence 0: "2 ^ min_height t < size1 t \<and> size1 t \<le> 2 ^ (min_height t + 1)" | |
| 59 | by (metis "*" min_height_size1_if_incomplete size1_height) | |
| 60 | hence "2 powr min_height t < size1 t \<and> size1 t \<le> 2 powr (min_height t + 1)" | |
| 61 | by(simp only: powr_realpow) | |
| 62 | (metis of_nat_less_iff of_nat_le_iff of_nat_numeral of_nat_power) | |
| 63 | hence "min_height t < log 2 (size1 t) \<and> log 2 (size1 t) \<le> min_height t + 1" | |
| 64 | by(simp add: log_le_iff less_log_iff) | |
| 65 | thus ?thesis using ** by linarith | |
| 66 | qed | |
| 67 | ||
| 64018 | 68 | (* mv *) | 
| 69 | ||
| 70 | text \<open>The lemmas about \<open>floor\<close> and \<open>ceiling\<close> of \<open>log 2\<close> should be generalized | |
| 64065 | 71 | from 2 to \<open>n\<close> and should be made executable. In the end they should be moved | 
| 72 | to theory \<open>Log_Nat\<close> and \<open>floorlog\<close> should be replaced.\<close> | |
| 64018 | 73 | |
| 64065 | 74 | lemma floor_log_nat_ivl: fixes b n k :: nat | 
| 64018 | 75 | assumes "b \<ge> 2" "b^n \<le> k" "k < b^(n+1)" | 
| 76 | shows "floor (log b (real k)) = int(n)" | |
| 77 | proof - | |
| 78 | have "k \<ge> 1" | |
| 79 | using assms(1,2) one_le_power[of b n] by linarith | |
| 80 | show ?thesis | |
| 81 | proof(rule floor_eq2) | |
| 82 | show "int n \<le> log b k" | |
| 83 | using assms(1,2) \<open>k \<ge> 1\<close> | |
| 84 | by(simp add: powr_realpow le_log_iff of_nat_power[symmetric] del: of_nat_power) | |
| 85 | next | |
| 86 | have "real k < b powr (real(n + 1))" using assms(1,3) | |
| 87 | by (simp only: powr_realpow) (metis of_nat_less_iff of_nat_power) | |
| 88 | thus "log b k < real_of_int (int n) + 1" | |
| 89 | using assms(1) \<open>k \<ge> 1\<close> by(simp add: log_less_iff add_ac) | |
| 90 | qed | |
| 91 | qed | |
| 92 | ||
| 64065 | 93 | lemma ceil_log_nat_ivl: fixes b n k :: nat | 
| 64018 | 94 | assumes "b \<ge> 2" "b^n < k" "k \<le> b^(n+1)" | 
| 95 | shows "ceiling (log b (real k)) = int(n)+1" | |
| 96 | proof(rule ceiling_eq) | |
| 97 | show "int n < log b k" | |
| 98 | using assms(1,2) | |
| 99 | by(simp add: powr_realpow less_log_iff of_nat_power[symmetric] del: of_nat_power) | |
| 100 | next | |
| 101 | have "real k \<le> b powr (real(n + 1))" | |
| 102 | using assms(1,3) | |
| 103 | by (simp only: powr_realpow) (metis of_nat_le_iff of_nat_power) | |
| 104 | thus "log b k \<le> real_of_int (int n) + 1" | |
| 105 | using assms(1,2) by(simp add: log_le_iff add_ac) | |
| 106 | qed | |
| 107 | ||
| 108 | lemma ceil_log2_div2: assumes "n \<ge> 2" | |
| 109 | shows "ceiling(log 2 (real n)) = ceiling(log 2 ((n-1) div 2 + 1)) + 1" | |
| 110 | proof cases | |
| 111 | assume "n=2" | |
| 112 | thus ?thesis by simp | |
| 113 | next | |
| 114 | let ?m = "(n-1) div 2 + 1" | |
| 115 | assume "n\<noteq>2" | |
| 116 | hence "2 \<le> ?m" | |
| 117 | using assms by arith | |
| 118 | then obtain i where i: "2 ^ i < ?m" "?m \<le> 2 ^ (i + 1)" | |
| 119 | using ex_power_ivl2[of 2 ?m] by auto | |
| 120 | have "n \<le> 2*?m" | |
| 121 | by arith | |
| 122 | also have "2*?m \<le> 2 ^ ((i+1)+1)" | |
| 123 | using i(2) by simp | |
| 124 | finally have *: "n \<le> \<dots>" . | |
| 125 | have "2^(i+1) < n" | |
| 126 | using i(1) by (auto simp add: less_Suc_eq_0_disj) | |
| 64065 | 127 | from ceil_log_nat_ivl[OF _ this *] ceil_log_nat_ivl[OF _ i] | 
| 64018 | 128 | show ?thesis by simp | 
| 129 | qed | |
| 130 | ||
| 131 | lemma floor_log2_div2: fixes n :: nat assumes "n \<ge> 2" | |
| 132 | shows "floor(log 2 n) = floor(log 2 (n div 2)) + 1" | |
| 133 | proof cases | |
| 134 | assume "n=2" | |
| 135 | thus ?thesis by simp | |
| 136 | next | |
| 137 | let ?m = "n div 2" | |
| 138 | assume "n\<noteq>2" | |
| 139 | hence "1 \<le> ?m" | |
| 140 | using assms by arith | |
| 141 | then obtain i where i: "2 ^ i \<le> ?m" "?m < 2 ^ (i + 1)" | |
| 142 | using ex_power_ivl1[of 2 ?m] by auto | |
| 143 | have "2^(i+1) \<le> 2*?m" | |
| 144 | using i(1) by simp | |
| 145 | also have "2*?m \<le> n" | |
| 146 | by arith | |
| 147 | finally have *: "2^(i+1) \<le> \<dots>" . | |
| 148 | have "n < 2^(i+1+1)" | |
| 149 | using i(2) by simp | |
| 64065 | 150 | from floor_log_nat_ivl[OF _ * this] floor_log_nat_ivl[OF _ i] | 
| 64018 | 151 | show ?thesis by simp | 
| 152 | qed | |
| 153 | ||
| 64541 | 154 | lemma balanced_Node_if_wbal1: | 
| 155 | assumes "balanced l" "balanced r" "size l = size r + 1" | |
| 156 | shows "balanced \<langle>l, x, r\<rangle>" | |
| 157 | proof - | |
| 158 | from assms(3) have [simp]: "size1 l = size1 r + 1" by(simp add: size1_def) | |
| 159 | have "nat \<lceil>log 2 (1 + size1 r)\<rceil> \<ge> nat \<lceil>log 2 (size1 r)\<rceil>" | |
| 160 | by(rule nat_mono[OF ceiling_mono]) simp | |
| 161 | hence 1: "height(Node l x r) = nat \<lceil>log 2 (1 + size1 r)\<rceil> + 1" | |
| 162 | using height_balanced[OF assms(1)] height_balanced[OF assms(2)] | |
| 163 | by (simp del: nat_ceiling_le_eq add: max_def) | |
| 164 | have "nat \<lfloor>log 2 (1 + size1 r)\<rfloor> \<ge> nat \<lfloor>log 2 (size1 r)\<rfloor>" | |
| 165 | by(rule nat_mono[OF floor_mono]) simp | |
| 166 | hence 2: "min_height(Node l x r) = nat \<lfloor>log 2 (size1 r)\<rfloor> + 1" | |
| 167 | using min_height_balanced[OF assms(1)] min_height_balanced[OF assms(2)] | |
| 168 | by (simp) | |
| 169 | have "size1 r \<ge> 1" by(simp add: size1_def) | |
| 170 | then obtain i where i: "2 ^ i \<le> size1 r" "size1 r < 2 ^ (i + 1)" | |
| 171 | using ex_power_ivl1[of 2 "size1 r"] by auto | |
| 172 | hence i1: "2 ^ i < size1 r + 1" "size1 r + 1 \<le> 2 ^ (i + 1)" by auto | |
| 173 | from 1 2 floor_log_nat_ivl[OF _ i] ceil_log_nat_ivl[OF _ i1] | |
| 174 | show ?thesis by(simp add:balanced_def) | |
| 175 | qed | |
| 176 | ||
| 177 | lemma balanced_sym: "balanced \<langle>l, x, r\<rangle> \<Longrightarrow> balanced \<langle>r, y, l\<rangle>" | |
| 178 | by(auto simp: balanced_def) | |
| 179 | ||
| 180 | lemma balanced_Node_if_wbal2: | |
| 181 | assumes "balanced l" "balanced r" "abs(int(size l) - int(size r)) \<le> 1" | |
| 182 | shows "balanced \<langle>l, x, r\<rangle>" | |
| 183 | proof - | |
| 184 | have "size l = size r \<or> (size l = size r + 1 \<or> size r = size l + 1)" (is "?A \<or> ?B") | |
| 185 | using assms(3) by linarith | |
| 186 | thus ?thesis | |
| 187 | proof | |
| 188 | assume "?A" | |
| 189 | thus ?thesis using assms(1,2) | |
| 190 | apply(simp add: balanced_def min_def max_def) | |
| 191 | by (metis assms(1,2) balanced_optimal le_antisym le_less) | |
| 192 | next | |
| 193 | assume "?B" | |
| 194 | thus ?thesis | |
| 195 | by (meson assms(1,2) balanced_sym balanced_Node_if_wbal1) | |
| 196 | qed | |
| 197 | qed | |
| 198 | ||
| 199 | lemma balanced_if_wbalanced: "wbalanced t \<Longrightarrow> balanced t" | |
| 200 | proof(induction t) | |
| 201 | case Leaf show ?case by (simp add: balanced_def) | |
| 202 | next | |
| 203 | case (Node l x r) | |
| 204 | thus ?case by(simp add: balanced_Node_if_wbal2) | |
| 205 | qed | |
| 206 | ||
| 64018 | 207 | (* end of mv *) | 
| 208 | ||
| 64444 | 209 | fun bal :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree * 'a list" where | 
| 210 | "bal n xs = (if n=0 then (Leaf,xs) else | |
| 63643 | 211 | (let m = n div 2; | 
| 64444 | 212 | (l, ys) = bal m xs; | 
| 213 | (r, zs) = bal (n-1-m) (tl ys) | |
| 63643 | 214 | in (Node l (hd ys) r, zs)))" | 
| 215 | ||
| 216 | declare bal.simps[simp del] | |
| 217 | ||
| 64444 | 218 | definition bal_list :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree" where | 
| 219 | "bal_list n xs = fst (bal n xs)" | |
| 220 | ||
| 63829 | 221 | definition balance_list :: "'a list \<Rightarrow> 'a tree" where | 
| 64444 | 222 | "balance_list xs = bal_list (length xs) xs" | 
| 223 | ||
| 224 | definition bal_tree :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where | |
| 225 | "bal_tree n t = bal_list n (inorder t)" | |
| 63829 | 226 | |
| 227 | definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where | |
| 64444 | 228 | "balance_tree t = bal_tree (size t) t" | 
| 63829 | 229 | |
| 63843 | 230 | lemma bal_simps: | 
| 64444 | 231 | "bal 0 xs = (Leaf, xs)" | 
| 63843 | 232 | "n > 0 \<Longrightarrow> | 
| 64444 | 233 | bal n xs = | 
| 63843 | 234 | (let m = n div 2; | 
| 64444 | 235 | (l, ys) = bal m xs; | 
| 236 | (r, zs) = bal (n-1-m) (tl ys) | |
| 63843 | 237 | in (Node l (hd ys) r, zs))" | 
| 238 | by(simp_all add: bal.simps) | |
| 63643 | 239 | |
| 64444 | 240 | text\<open>Some of the following lemmas take advantage of the fact | 
| 63861 | 241 | that \<open>bal xs n\<close> yields a result even if \<open>n > length xs\<close>.\<close> | 
| 242 | ||
| 64444 | 243 | lemma size_bal: "bal n xs = (t,ys) \<Longrightarrow> size t = n" | 
| 244 | proof(induction n xs arbitrary: t ys rule: bal.induct) | |
| 245 | case (1 n xs) | |
| 63861 | 246 | thus ?case | 
| 247 | by(cases "n=0") | |
| 248 | (auto simp add: bal_simps Let_def split: prod.splits) | |
| 249 | qed | |
| 250 | ||
| 63643 | 251 | lemma bal_inorder: | 
| 64444 | 252 | "\<lbrakk> bal n xs = (t,ys); n \<le> length xs \<rbrakk> | 
| 63755 
182c111190e5
Renamed balanced to complete; added balanced; more about both
 nipkow parents: 
63663diff
changeset | 253 | \<Longrightarrow> inorder t = take n xs \<and> ys = drop n xs" | 
| 64444 | 254 | proof(induction n xs arbitrary: t ys rule: bal.induct) | 
| 255 | case (1 n xs) show ?case | |
| 63643 | 256 | proof cases | 
| 63843 | 257 | assume "n = 0" thus ?thesis using 1 by (simp add: bal_simps) | 
| 63643 | 258 | next | 
| 259 | assume [arith]: "n \<noteq> 0" | |
| 260 | let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" | |
| 261 | from "1.prems" obtain l r xs' where | |
| 64444 | 262 | b1: "bal ?n1 xs = (l,xs')" and | 
| 263 | b2: "bal ?n2 (tl xs') = (r,ys)" and | |
| 63643 | 264 | t: "t = \<langle>l, hd xs', r\<rangle>" | 
| 63843 | 265 | by(auto simp: Let_def bal_simps split: prod.splits) | 
| 63643 | 266 | have IH1: "inorder l = take ?n1 xs \<and> xs' = drop ?n1 xs" | 
| 267 | using b1 "1.prems" by(intro "1.IH"(1)) auto | |
| 268 | have IH2: "inorder r = take ?n2 (tl xs') \<and> ys = drop ?n2 (tl xs')" | |
| 269 | using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto | |
| 270 | have "drop (n div 2) xs \<noteq> []" using "1.prems"(2) by simp | |
| 271 | hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" | |
| 272 | by (metis Suc_eq_plus1 take_Suc) | |
| 273 | hence *: "inorder t = take n xs" using t IH1 IH2 | |
| 274 | using take_add[of ?n1 "?n2+1" xs] by(simp) | |
| 275 | have "n - n div 2 + n div 2 = n" by simp | |
| 276 | hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) | |
| 277 | thus ?thesis using * by blast | |
| 278 | qed | |
| 279 | qed | |
| 280 | ||
| 64444 | 281 | corollary inorder_bal_list[simp]: | 
| 282 | "n \<le> length xs \<Longrightarrow> inorder(bal_list n xs) = take n xs" | |
| 283 | unfolding bal_list_def by (metis bal_inorder eq_fst_iff) | |
| 284 | ||
| 285 | corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs" | |
| 286 | by(simp add: balance_list_def) | |
| 287 | ||
| 288 | corollary inorder_bal_tree: | |
| 289 | "n \<le> size t \<Longrightarrow> inorder(bal_tree n t) = take n (inorder t)" | |
| 290 | by(simp add: bal_tree_def) | |
| 63643 | 291 | |
| 64018 | 292 | corollary inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t" | 
| 64444 | 293 | by(simp add: balance_tree_def inorder_bal_tree) | 
| 294 | ||
| 295 | corollary size_bal_list[simp]: "size(bal_list n xs) = n" | |
| 296 | unfolding bal_list_def by (metis prod.collapse size_bal) | |
| 64018 | 297 | |
| 298 | corollary size_balance_list[simp]: "size(balance_list xs) = length xs" | |
| 64444 | 299 | by (simp add: balance_list_def) | 
| 300 | ||
| 301 | corollary size_bal_tree[simp]: "size(bal_tree n t) = n" | |
| 302 | by(simp add: bal_tree_def) | |
| 64018 | 303 | |
| 304 | corollary size_balance_tree[simp]: "size(balance_tree t) = size t" | |
| 64444 | 305 | by(simp add: balance_tree_def) | 
| 64018 | 306 | |
| 307 | lemma min_height_bal: | |
| 64444 | 308 | "bal n xs = (t,ys) \<Longrightarrow> min_height t = nat(floor(log 2 (n + 1)))" | 
| 309 | proof(induction n xs arbitrary: t ys rule: bal.induct) | |
| 310 | case (1 n xs) show ?case | |
| 63643 | 311 | proof cases | 
| 312 | assume "n = 0" thus ?thesis | |
| 64018 | 313 | using "1.prems" by (simp add: bal_simps) | 
| 63643 | 314 | next | 
| 315 | assume [arith]: "n \<noteq> 0" | |
| 316 | from "1.prems" obtain l r xs' where | |
| 64444 | 317 | b1: "bal (n div 2) xs = (l,xs')" and | 
| 318 | b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and | |
| 63643 | 319 | t: "t = \<langle>l, hd xs', r\<rangle>" | 
| 63843 | 320 | by(auto simp: bal_simps Let_def split: prod.splits) | 
| 64018 | 321 | let ?log1 = "nat (floor(log 2 (n div 2 + 1)))" | 
| 322 | let ?log2 = "nat (floor(log 2 (n - 1 - n div 2 + 1)))" | |
| 63643 | 323 | have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp | 
| 64018 | 324 | have IH2: "min_height r = ?log2" using "1.IH"(2) b1 b2 by simp | 
| 325 | have "(n+1) div 2 \<ge> 1" by arith | |
| 326 | hence 0: "log 2 ((n+1) div 2) \<ge> 0" by simp | |
| 327 | have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith | |
| 328 | hence le: "?log2 \<le> ?log1" | |
| 329 | by(simp add: nat_mono floor_mono) | |
| 330 | have "min_height t = min ?log1 ?log2 + 1" by (simp add: t IH1 IH2) | |
| 331 | also have "\<dots> = ?log2 + 1" using le by (simp add: min_absorb2) | |
| 332 | also have "n - 1 - n div 2 + 1 = (n+1) div 2" by linarith | |
| 333 | also have "nat (floor(log 2 ((n+1) div 2))) + 1 | |
| 334 | = nat (floor(log 2 ((n+1) div 2) + 1))" | |
| 335 | using 0 by linarith | |
| 336 | also have "\<dots> = nat (floor(log 2 (n + 1)))" | |
| 337 | using floor_log2_div2[of "n+1"] by (simp add: log_mult) | |
| 338 | finally show ?thesis . | |
| 339 | qed | |
| 340 | qed | |
| 341 | ||
| 342 | lemma height_bal: | |
| 64444 | 343 | "bal n xs = (t,ys) \<Longrightarrow> height t = nat \<lceil>log 2 (n + 1)\<rceil>" | 
| 344 | proof(induction n xs arbitrary: t ys rule: bal.induct) | |
| 345 | case (1 n xs) show ?case | |
| 64018 | 346 | proof cases | 
| 347 | assume "n = 0" thus ?thesis | |
| 348 | using "1.prems" by (simp add: bal_simps) | |
| 349 | next | |
| 350 | assume [arith]: "n \<noteq> 0" | |
| 351 | from "1.prems" obtain l r xs' where | |
| 64444 | 352 | b1: "bal (n div 2) xs = (l,xs')" and | 
| 353 | b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and | |
| 64018 | 354 | t: "t = \<langle>l, hd xs', r\<rangle>" | 
| 355 | by(auto simp: bal_simps Let_def split: prod.splits) | |
| 356 | let ?log1 = "nat \<lceil>log 2 (n div 2 + 1)\<rceil>" | |
| 357 | let ?log2 = "nat \<lceil>log 2 (n - 1 - n div 2 + 1)\<rceil>" | |
| 358 | have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp | |
| 359 | have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp | |
| 360 | have 0: "log 2 (n div 2 + 1) \<ge> 0" by auto | |
| 361 | have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith | |
| 362 | hence le: "?log2 \<le> ?log1" | |
| 363 | by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq) | |
| 364 | have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) | |
| 365 | also have "\<dots> = ?log1 + 1" using le by (simp add: max_absorb1) | |
| 366 | also have "\<dots> = nat \<lceil>log 2 (n div 2 + 1) + 1\<rceil>" using 0 by linarith | |
| 367 | also have "\<dots> = nat \<lceil>log 2 (n + 1)\<rceil>" | |
| 368 | using ceil_log2_div2[of "n+1"] by (simp) | |
| 63643 | 369 | finally show ?thesis . | 
| 370 | qed | |
| 371 | qed | |
| 372 | ||
| 373 | lemma balanced_bal: | |
| 64444 | 374 | assumes "bal n xs = (t,ys)" shows "balanced t" | 
| 64018 | 375 | unfolding balanced_def | 
| 376 | using height_bal[OF assms] min_height_bal[OF assms] | |
| 377 | by linarith | |
| 63643 | 378 | |
| 64444 | 379 | lemma height_bal_list: | 
| 380 | "n \<le> length xs \<Longrightarrow> height (bal_list n xs) = nat \<lceil>log 2 (n + 1)\<rceil>" | |
| 381 | unfolding bal_list_def by (metis height_bal prod.collapse) | |
| 382 | ||
| 64018 | 383 | lemma height_balance_list: | 
| 384 | "height (balance_list xs) = nat \<lceil>log 2 (length xs + 1)\<rceil>" | |
| 64444 | 385 | by (simp add: balance_list_def height_bal_list) | 
| 386 | ||
| 387 | corollary height_bal_tree: | |
| 388 | "n \<le> length xs \<Longrightarrow> height (bal_tree n t) = nat(ceiling(log 2 (n + 1)))" | |
| 389 | unfolding bal_list_def bal_tree_def | |
| 390 | using height_bal prod.exhaust_sel by blast | |
| 64018 | 391 | |
| 392 | corollary height_balance_tree: | |
| 393 | "height (balance_tree t) = nat(ceiling(log 2 (size t + 1)))" | |
| 64444 | 394 | by (simp add: bal_tree_def balance_tree_def height_bal_list) | 
| 395 | ||
| 396 | corollary balanced_bal_list[simp]: "balanced (bal_list n xs)" | |
| 397 | unfolding bal_list_def by (metis balanced_bal prod.collapse) | |
| 63829 | 398 | |
| 399 | corollary balanced_balance_list[simp]: "balanced (balance_list xs)" | |
| 64444 | 400 | by (simp add: balance_list_def) | 
| 401 | ||
| 402 | corollary balanced_bal_tree[simp]: "balanced (bal_tree n t)" | |
| 403 | by (simp add: bal_tree_def) | |
| 63829 | 404 | |
| 405 | corollary balanced_balance_tree[simp]: "balanced (balance_tree t)" | |
| 406 | by (simp add: balance_tree_def) | |
| 407 | ||
| 64444 | 408 | lemma wbalanced_bal: "bal n xs = (t,ys) \<Longrightarrow> wbalanced t" | 
| 409 | proof(induction n xs arbitrary: t ys rule: bal.induct) | |
| 410 | case (1 n xs) | |
| 63861 | 411 | show ?case | 
| 412 | proof cases | |
| 413 | assume "n = 0" | |
| 414 | thus ?thesis | |
| 415 | using "1.prems" by(simp add: bal_simps) | |
| 416 | next | |
| 417 | assume "n \<noteq> 0" | |
| 418 | with "1.prems" obtain l ys r zs where | |
| 64444 | 419 | rec1: "bal (n div 2) xs = (l, ys)" and | 
| 420 | rec2: "bal (n - 1 - n div 2) (tl ys) = (r, zs)" and | |
| 63861 | 421 | t: "t = \<langle>l, hd ys, r\<rangle>" | 
| 422 | by(auto simp add: bal_simps Let_def split: prod.splits) | |
| 423 | have l: "wbalanced l" using "1.IH"(1)[OF \<open>n\<noteq>0\<close> refl rec1] . | |
| 424 | have "wbalanced r" using "1.IH"(2)[OF \<open>n\<noteq>0\<close> refl rec1[symmetric] refl rec2] . | |
| 425 | with l t size_bal[OF rec1] size_bal[OF rec2] | |
| 426 | show ?thesis by auto | |
| 427 | qed | |
| 428 | qed | |
| 429 | ||
| 64541 | 430 | text\<open>An alternative proof via @{thm balanced_if_wbalanced}:\<close>
 | 
| 431 | lemma "bal n xs = (t,ys) \<Longrightarrow> balanced t" | |
| 432 | by(rule balanced_if_wbalanced[OF wbalanced_bal]) | |
| 433 | ||
| 64444 | 434 | lemma wbalanced_bal_list[simp]: "wbalanced (bal_list n xs)" | 
| 435 | by(simp add: bal_list_def) (metis prod.collapse wbalanced_bal) | |
| 436 | ||
| 437 | lemma wbalanced_balance_list[simp]: "wbalanced (balance_list xs)" | |
| 438 | by(simp add: balance_list_def) | |
| 439 | ||
| 440 | lemma wbalanced_bal_tree[simp]: "wbalanced (bal_tree n t)" | |
| 441 | by(simp add: bal_tree_def) | |
| 442 | ||
| 63861 | 443 | lemma wbalanced_balance_tree: "wbalanced (balance_tree t)" | 
| 64444 | 444 | by (simp add: balance_tree_def) | 
| 63861 | 445 | |
| 63829 | 446 | hide_const (open) bal | 
| 63643 | 447 | |
| 448 | end |