| author | nipkow |
| Tue, 29 Nov 2016 10:53:52 +0100 | |
| changeset 64533 | 172f3a047f4a |
| parent 64444 | daae191c9344 |
| child 64540 | f1f4ba6d02c9 |
| 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" |
|
29 |
using assms min_hight_le_height[of t] |
|
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" |
|
56 |
using assms min_hight_le_height[of t] |
|
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 |
||
154 |
(* end of mv *) |
|
155 |
||
| 64444 | 156 |
fun bal :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree * 'a list" where |
157 |
"bal n xs = (if n=0 then (Leaf,xs) else |
|
| 63643 | 158 |
(let m = n div 2; |
| 64444 | 159 |
(l, ys) = bal m xs; |
160 |
(r, zs) = bal (n-1-m) (tl ys) |
|
| 63643 | 161 |
in (Node l (hd ys) r, zs)))" |
162 |
||
163 |
declare bal.simps[simp del] |
|
164 |
||
| 64444 | 165 |
definition bal_list :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree" where |
166 |
"bal_list n xs = fst (bal n xs)" |
|
167 |
||
| 63829 | 168 |
definition balance_list :: "'a list \<Rightarrow> 'a tree" where |
| 64444 | 169 |
"balance_list xs = bal_list (length xs) xs" |
170 |
||
171 |
definition bal_tree :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
|
172 |
"bal_tree n t = bal_list n (inorder t)" |
|
| 63829 | 173 |
|
174 |
definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where |
|
| 64444 | 175 |
"balance_tree t = bal_tree (size t) t" |
| 63829 | 176 |
|
| 63843 | 177 |
lemma bal_simps: |
| 64444 | 178 |
"bal 0 xs = (Leaf, xs)" |
| 63843 | 179 |
"n > 0 \<Longrightarrow> |
| 64444 | 180 |
bal n xs = |
| 63843 | 181 |
(let m = n div 2; |
| 64444 | 182 |
(l, ys) = bal m xs; |
183 |
(r, zs) = bal (n-1-m) (tl ys) |
|
| 63843 | 184 |
in (Node l (hd ys) r, zs))" |
185 |
by(simp_all add: bal.simps) |
|
| 63643 | 186 |
|
| 64444 | 187 |
text\<open>Some of the following lemmas take advantage of the fact |
| 63861 | 188 |
that \<open>bal xs n\<close> yields a result even if \<open>n > length xs\<close>.\<close> |
189 |
||
| 64444 | 190 |
lemma size_bal: "bal n xs = (t,ys) \<Longrightarrow> size t = n" |
191 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
192 |
case (1 n xs) |
|
| 63861 | 193 |
thus ?case |
194 |
by(cases "n=0") |
|
195 |
(auto simp add: bal_simps Let_def split: prod.splits) |
|
196 |
qed |
|
197 |
||
| 63643 | 198 |
lemma bal_inorder: |
| 64444 | 199 |
"\<lbrakk> bal n xs = (t,ys); n \<le> length xs \<rbrakk> |
|
63755
182c111190e5
Renamed balanced to complete; added balanced; more about both
nipkow
parents:
63663
diff
changeset
|
200 |
\<Longrightarrow> inorder t = take n xs \<and> ys = drop n xs" |
| 64444 | 201 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
202 |
case (1 n xs) show ?case |
|
| 63643 | 203 |
proof cases |
| 63843 | 204 |
assume "n = 0" thus ?thesis using 1 by (simp add: bal_simps) |
| 63643 | 205 |
next |
206 |
assume [arith]: "n \<noteq> 0" |
|
207 |
let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" |
|
208 |
from "1.prems" obtain l r xs' where |
|
| 64444 | 209 |
b1: "bal ?n1 xs = (l,xs')" and |
210 |
b2: "bal ?n2 (tl xs') = (r,ys)" and |
|
| 63643 | 211 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
| 63843 | 212 |
by(auto simp: Let_def bal_simps split: prod.splits) |
| 63643 | 213 |
have IH1: "inorder l = take ?n1 xs \<and> xs' = drop ?n1 xs" |
214 |
using b1 "1.prems" by(intro "1.IH"(1)) auto |
|
215 |
have IH2: "inorder r = take ?n2 (tl xs') \<and> ys = drop ?n2 (tl xs')" |
|
216 |
using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto |
|
217 |
have "drop (n div 2) xs \<noteq> []" using "1.prems"(2) by simp |
|
218 |
hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" |
|
219 |
by (metis Suc_eq_plus1 take_Suc) |
|
220 |
hence *: "inorder t = take n xs" using t IH1 IH2 |
|
221 |
using take_add[of ?n1 "?n2+1" xs] by(simp) |
|
222 |
have "n - n div 2 + n div 2 = n" by simp |
|
223 |
hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) |
|
224 |
thus ?thesis using * by blast |
|
225 |
qed |
|
226 |
qed |
|
227 |
||
| 64444 | 228 |
corollary inorder_bal_list[simp]: |
229 |
"n \<le> length xs \<Longrightarrow> inorder(bal_list n xs) = take n xs" |
|
230 |
unfolding bal_list_def by (metis bal_inorder eq_fst_iff) |
|
231 |
||
232 |
corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs" |
|
233 |
by(simp add: balance_list_def) |
|
234 |
||
235 |
corollary inorder_bal_tree: |
|
236 |
"n \<le> size t \<Longrightarrow> inorder(bal_tree n t) = take n (inorder t)" |
|
237 |
by(simp add: bal_tree_def) |
|
| 63643 | 238 |
|
| 64018 | 239 |
corollary inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t" |
| 64444 | 240 |
by(simp add: balance_tree_def inorder_bal_tree) |
241 |
||
242 |
corollary size_bal_list[simp]: "size(bal_list n xs) = n" |
|
243 |
unfolding bal_list_def by (metis prod.collapse size_bal) |
|
| 64018 | 244 |
|
245 |
corollary size_balance_list[simp]: "size(balance_list xs) = length xs" |
|
| 64444 | 246 |
by (simp add: balance_list_def) |
247 |
||
248 |
corollary size_bal_tree[simp]: "size(bal_tree n t) = n" |
|
249 |
by(simp add: bal_tree_def) |
|
| 64018 | 250 |
|
251 |
corollary size_balance_tree[simp]: "size(balance_tree t) = size t" |
|
| 64444 | 252 |
by(simp add: balance_tree_def) |
| 64018 | 253 |
|
254 |
lemma min_height_bal: |
|
| 64444 | 255 |
"bal n xs = (t,ys) \<Longrightarrow> min_height t = nat(floor(log 2 (n + 1)))" |
256 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
257 |
case (1 n xs) show ?case |
|
| 63643 | 258 |
proof cases |
259 |
assume "n = 0" thus ?thesis |
|
| 64018 | 260 |
using "1.prems" by (simp add: bal_simps) |
| 63643 | 261 |
next |
262 |
assume [arith]: "n \<noteq> 0" |
|
263 |
from "1.prems" obtain l r xs' where |
|
| 64444 | 264 |
b1: "bal (n div 2) xs = (l,xs')" and |
265 |
b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and |
|
| 63643 | 266 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
| 63843 | 267 |
by(auto simp: bal_simps Let_def split: prod.splits) |
| 64018 | 268 |
let ?log1 = "nat (floor(log 2 (n div 2 + 1)))" |
269 |
let ?log2 = "nat (floor(log 2 (n - 1 - n div 2 + 1)))" |
|
| 63643 | 270 |
have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp |
| 64018 | 271 |
have IH2: "min_height r = ?log2" using "1.IH"(2) b1 b2 by simp |
272 |
have "(n+1) div 2 \<ge> 1" by arith |
|
273 |
hence 0: "log 2 ((n+1) div 2) \<ge> 0" by simp |
|
274 |
have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith |
|
275 |
hence le: "?log2 \<le> ?log1" |
|
276 |
by(simp add: nat_mono floor_mono) |
|
277 |
have "min_height t = min ?log1 ?log2 + 1" by (simp add: t IH1 IH2) |
|
278 |
also have "\<dots> = ?log2 + 1" using le by (simp add: min_absorb2) |
|
279 |
also have "n - 1 - n div 2 + 1 = (n+1) div 2" by linarith |
|
280 |
also have "nat (floor(log 2 ((n+1) div 2))) + 1 |
|
281 |
= nat (floor(log 2 ((n+1) div 2) + 1))" |
|
282 |
using 0 by linarith |
|
283 |
also have "\<dots> = nat (floor(log 2 (n + 1)))" |
|
284 |
using floor_log2_div2[of "n+1"] by (simp add: log_mult) |
|
285 |
finally show ?thesis . |
|
286 |
qed |
|
287 |
qed |
|
288 |
||
289 |
lemma height_bal: |
|
| 64444 | 290 |
"bal n xs = (t,ys) \<Longrightarrow> height t = nat \<lceil>log 2 (n + 1)\<rceil>" |
291 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
292 |
case (1 n xs) show ?case |
|
| 64018 | 293 |
proof cases |
294 |
assume "n = 0" thus ?thesis |
|
295 |
using "1.prems" by (simp add: bal_simps) |
|
296 |
next |
|
297 |
assume [arith]: "n \<noteq> 0" |
|
298 |
from "1.prems" obtain l r xs' where |
|
| 64444 | 299 |
b1: "bal (n div 2) xs = (l,xs')" and |
300 |
b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and |
|
| 64018 | 301 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
302 |
by(auto simp: bal_simps Let_def split: prod.splits) |
|
303 |
let ?log1 = "nat \<lceil>log 2 (n div 2 + 1)\<rceil>" |
|
304 |
let ?log2 = "nat \<lceil>log 2 (n - 1 - n div 2 + 1)\<rceil>" |
|
305 |
have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp |
|
306 |
have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp |
|
307 |
have 0: "log 2 (n div 2 + 1) \<ge> 0" by auto |
|
308 |
have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith |
|
309 |
hence le: "?log2 \<le> ?log1" |
|
310 |
by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq) |
|
311 |
have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) |
|
312 |
also have "\<dots> = ?log1 + 1" using le by (simp add: max_absorb1) |
|
313 |
also have "\<dots> = nat \<lceil>log 2 (n div 2 + 1) + 1\<rceil>" using 0 by linarith |
|
314 |
also have "\<dots> = nat \<lceil>log 2 (n + 1)\<rceil>" |
|
315 |
using ceil_log2_div2[of "n+1"] by (simp) |
|
| 63643 | 316 |
finally show ?thesis . |
317 |
qed |
|
318 |
qed |
|
319 |
||
320 |
lemma balanced_bal: |
|
| 64444 | 321 |
assumes "bal n xs = (t,ys)" shows "balanced t" |
| 64018 | 322 |
unfolding balanced_def |
323 |
using height_bal[OF assms] min_height_bal[OF assms] |
|
324 |
by linarith |
|
| 63643 | 325 |
|
| 64444 | 326 |
lemma height_bal_list: |
327 |
"n \<le> length xs \<Longrightarrow> height (bal_list n xs) = nat \<lceil>log 2 (n + 1)\<rceil>" |
|
328 |
unfolding bal_list_def by (metis height_bal prod.collapse) |
|
329 |
||
| 64018 | 330 |
lemma height_balance_list: |
331 |
"height (balance_list xs) = nat \<lceil>log 2 (length xs + 1)\<rceil>" |
|
| 64444 | 332 |
by (simp add: balance_list_def height_bal_list) |
333 |
||
334 |
corollary height_bal_tree: |
|
335 |
"n \<le> length xs \<Longrightarrow> height (bal_tree n t) = nat(ceiling(log 2 (n + 1)))" |
|
336 |
unfolding bal_list_def bal_tree_def |
|
337 |
using height_bal prod.exhaust_sel by blast |
|
| 64018 | 338 |
|
339 |
corollary height_balance_tree: |
|
340 |
"height (balance_tree t) = nat(ceiling(log 2 (size t + 1)))" |
|
| 64444 | 341 |
by (simp add: bal_tree_def balance_tree_def height_bal_list) |
342 |
||
343 |
corollary balanced_bal_list[simp]: "balanced (bal_list n xs)" |
|
344 |
unfolding bal_list_def by (metis balanced_bal prod.collapse) |
|
| 63829 | 345 |
|
346 |
corollary balanced_balance_list[simp]: "balanced (balance_list xs)" |
|
| 64444 | 347 |
by (simp add: balance_list_def) |
348 |
||
349 |
corollary balanced_bal_tree[simp]: "balanced (bal_tree n t)" |
|
350 |
by (simp add: bal_tree_def) |
|
| 63829 | 351 |
|
352 |
corollary balanced_balance_tree[simp]: "balanced (balance_tree t)" |
|
353 |
by (simp add: balance_tree_def) |
|
354 |
||
| 64444 | 355 |
lemma wbalanced_bal: "bal n xs = (t,ys) \<Longrightarrow> wbalanced t" |
356 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
357 |
case (1 n xs) |
|
| 63861 | 358 |
show ?case |
359 |
proof cases |
|
360 |
assume "n = 0" |
|
361 |
thus ?thesis |
|
362 |
using "1.prems" by(simp add: bal_simps) |
|
363 |
next |
|
364 |
assume "n \<noteq> 0" |
|
365 |
with "1.prems" obtain l ys r zs where |
|
| 64444 | 366 |
rec1: "bal (n div 2) xs = (l, ys)" and |
367 |
rec2: "bal (n - 1 - n div 2) (tl ys) = (r, zs)" and |
|
| 63861 | 368 |
t: "t = \<langle>l, hd ys, r\<rangle>" |
369 |
by(auto simp add: bal_simps Let_def split: prod.splits) |
|
370 |
have l: "wbalanced l" using "1.IH"(1)[OF \<open>n\<noteq>0\<close> refl rec1] . |
|
371 |
have "wbalanced r" using "1.IH"(2)[OF \<open>n\<noteq>0\<close> refl rec1[symmetric] refl rec2] . |
|
372 |
with l t size_bal[OF rec1] size_bal[OF rec2] |
|
373 |
show ?thesis by auto |
|
374 |
qed |
|
375 |
qed |
|
376 |
||
| 64444 | 377 |
lemma wbalanced_bal_list[simp]: "wbalanced (bal_list n xs)" |
378 |
by(simp add: bal_list_def) (metis prod.collapse wbalanced_bal) |
|
379 |
||
380 |
lemma wbalanced_balance_list[simp]: "wbalanced (balance_list xs)" |
|
381 |
by(simp add: balance_list_def) |
|
382 |
||
383 |
lemma wbalanced_bal_tree[simp]: "wbalanced (bal_tree n t)" |
|
384 |
by(simp add: bal_tree_def) |
|
385 |
||
| 63861 | 386 |
lemma wbalanced_balance_tree: "wbalanced (balance_tree t)" |
| 64444 | 387 |
by (simp add: balance_tree_def) |
| 63861 | 388 |
|
| 63829 | 389 |
hide_const (open) bal |
| 63643 | 390 |
|
391 |
end |