| author | paulson | 
| Tue, 18 May 2021 20:25:19 +0100 | |
| changeset 73708 | 2e3a60ce5a9f | 
| parent 72566 | 831f17da1aab | 
| permissions | -rw-r--r-- | 
| 63829 | 1  | 
(* Author: Tobias Nipkow *)  | 
| 63643 | 2  | 
|
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
3  | 
section \<open>Creating Almost Complete Trees\<close>  | 
| 63643 | 4  | 
|
| 63829 | 5  | 
theory Balance  | 
| 63643 | 6  | 
imports  | 
| 66510 | 7  | 
"HOL-Library.Tree_Real"  | 
| 63643 | 8  | 
begin  | 
9  | 
||
| 64444 | 10  | 
fun bal :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree * 'a list" where  | 
11  | 
"bal n xs = (if n=0 then (Leaf,xs) else  | 
|
| 63643 | 12  | 
(let m = n div 2;  | 
| 64444 | 13  | 
(l, ys) = bal m xs;  | 
14  | 
(r, zs) = bal (n-1-m) (tl ys)  | 
|
| 63643 | 15  | 
in (Node l (hd ys) r, zs)))"  | 
16  | 
||
17  | 
declare bal.simps[simp del]  | 
|
| 71846 | 18  | 
declare Let_def[simp]  | 
| 63643 | 19  | 
|
| 64444 | 20  | 
definition bal_list :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree" where  | 
21  | 
"bal_list n xs = fst (bal n xs)"  | 
|
22  | 
||
| 63829 | 23  | 
definition balance_list :: "'a list \<Rightarrow> 'a tree" where  | 
| 64444 | 24  | 
"balance_list xs = bal_list (length xs) xs"  | 
25  | 
||
26  | 
definition bal_tree :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where  | 
|
27  | 
"bal_tree n t = bal_list n (inorder t)"  | 
|
| 63829 | 28  | 
|
29  | 
definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where  | 
|
| 64444 | 30  | 
"balance_tree t = bal_tree (size t) t"  | 
| 63829 | 31  | 
|
| 63843 | 32  | 
lemma bal_simps:  | 
| 64444 | 33  | 
"bal 0 xs = (Leaf, xs)"  | 
| 63843 | 34  | 
"n > 0 \<Longrightarrow>  | 
| 64444 | 35  | 
bal n xs =  | 
| 63843 | 36  | 
(let m = n div 2;  | 
| 64444 | 37  | 
(l, ys) = bal m xs;  | 
38  | 
(r, zs) = bal (n-1-m) (tl ys)  | 
|
| 63843 | 39  | 
in (Node l (hd ys) r, zs))"  | 
40  | 
by(simp_all add: bal.simps)  | 
|
| 63643 | 41  | 
|
42  | 
lemma bal_inorder:  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
43  | 
"\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk>  | 
| 70751 | 44  | 
\<Longrightarrow> xs = inorder t @ zs \<and> size t = n"  | 
| 71858 | 45  | 
proof(induction n arbitrary: xs t zs rule: less_induct)  | 
46  | 
case (less n) show ?case  | 
|
| 63643 | 47  | 
proof cases  | 
| 71858 | 48  | 
assume "n = 0" thus ?thesis using less.prems by (simp add: bal_simps)  | 
| 63643 | 49  | 
next  | 
50  | 
assume [arith]: "n \<noteq> 0"  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
51  | 
let ?m = "n div 2" let ?m' = "n - 1 - ?m"  | 
| 71858 | 52  | 
from less.prems(2) obtain l r ys where  | 
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
53  | 
b1: "bal ?m xs = (l,ys)" and  | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
54  | 
b2: "bal ?m' (tl ys) = (r,zs)" and  | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
55  | 
t: "t = \<langle>l, hd ys, r\<rangle>"  | 
| 71846 | 56  | 
by(auto simp: bal_simps split: prod.splits)  | 
| 70751 | 57  | 
have IH1: "xs = inorder l @ ys \<and> size l = ?m"  | 
| 71858 | 58  | 
using b1 less.prems(1) by(intro less.IH) auto  | 
| 70751 | 59  | 
have IH2: "tl ys = inorder r @ zs \<and> size r = ?m'"  | 
| 71858 | 60  | 
using b2 IH1 less.prems(1) by(intro less.IH) auto  | 
61  | 
show ?thesis using t IH1 IH2 less.prems(1) hd_Cons_tl[of ys] by fastforce  | 
|
| 63643 | 62  | 
qed  | 
63  | 
qed  | 
|
64  | 
||
| 64444 | 65  | 
corollary inorder_bal_list[simp]:  | 
66  | 
"n \<le> length xs \<Longrightarrow> inorder(bal_list n xs) = take n xs"  | 
|
| 70751 | 67  | 
unfolding bal_list_def  | 
68  | 
by (metis (mono_tags) prod.collapse[of "bal n xs"] append_eq_conv_conj bal_inorder length_inorder)  | 
|
| 64444 | 69  | 
|
70  | 
corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs"  | 
|
71  | 
by(simp add: balance_list_def)  | 
|
72  | 
||
73  | 
corollary inorder_bal_tree:  | 
|
74  | 
"n \<le> size t \<Longrightarrow> inorder(bal_tree n t) = take n (inorder t)"  | 
|
75  | 
by(simp add: bal_tree_def)  | 
|
| 63643 | 76  | 
|
| 64018 | 77  | 
corollary inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t"  | 
| 64444 | 78  | 
by(simp add: balance_tree_def inorder_bal_tree)  | 
79  | 
||
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
80  | 
|
| 70751 | 81  | 
text\<open>The length/size lemmas below do not require the precondition @{prop"n \<le> length xs"}
 | 
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
82  | 
(or  @{prop"n \<le> size t"}) that they come with. They could take advantage of the fact
 | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
83  | 
that @{term "bal xs n"} yields a result even if @{prop "n > length xs"}.
 | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
84  | 
In that case the result will contain one or more occurrences of @{term "hd []"}.
 | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
85  | 
However, this is counter-intuitive and does not reflect the execution  | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
86  | 
in an eager functional language.\<close>  | 
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
87  | 
|
| 70751 | 88  | 
lemma bal_length: "\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> length zs = length xs - n"  | 
89  | 
using bal_inorder by fastforce  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
90  | 
|
| 
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
91  | 
corollary size_bal_list[simp]: "n \<le> length xs \<Longrightarrow> size(bal_list n xs) = n"  | 
| 70751 | 92  | 
unfolding bal_list_def using bal_inorder prod.exhaust_sel by blast  | 
| 64018 | 93  | 
|
94  | 
corollary size_balance_list[simp]: "size(balance_list xs) = length xs"  | 
|
| 64444 | 95  | 
by (simp add: balance_list_def)  | 
96  | 
||
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
97  | 
corollary size_bal_tree[simp]: "n \<le> size t \<Longrightarrow> size(bal_tree n t) = n"  | 
| 64444 | 98  | 
by(simp add: bal_tree_def)  | 
| 64018 | 99  | 
|
100  | 
corollary size_balance_tree[simp]: "size(balance_tree t) = size t"  | 
|
| 64444 | 101  | 
by(simp add: balance_tree_def)  | 
| 64018 | 102  | 
|
103  | 
lemma min_height_bal:  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
104  | 
"\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> min_height t = nat(\<lfloor>log 2 (n + 1)\<rfloor>)"  | 
| 71858 | 105  | 
proof(induction n arbitrary: xs t zs rule: less_induct)  | 
106  | 
case (less n)  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
107  | 
show ?case  | 
| 63643 | 108  | 
proof cases  | 
| 71858 | 109  | 
assume "n = 0" thus ?thesis using less.prems(2) by (simp add: bal_simps)  | 
| 63643 | 110  | 
next  | 
111  | 
assume [arith]: "n \<noteq> 0"  | 
|
| 70751 | 112  | 
let ?m = "n div 2" let ?m' = "n - 1 - ?m"  | 
| 71858 | 113  | 
from less.prems obtain l r ys where  | 
| 70751 | 114  | 
b1: "bal ?m xs = (l,ys)" and  | 
115  | 
b2: "bal ?m' (tl ys) = (r,zs)" and  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
116  | 
t: "t = \<langle>l, hd ys, r\<rangle>"  | 
| 71846 | 117  | 
by(auto simp: bal_simps split: prod.splits)  | 
| 70751 | 118  | 
let ?hl = "nat (floor(log 2 (?m + 1)))"  | 
119  | 
let ?hr = "nat (floor(log 2 (?m' + 1)))"  | 
|
| 71858 | 120  | 
have IH1: "min_height l = ?hl" using less.IH[OF _ _ b1] less.prems(1) by simp  | 
| 70751 | 121  | 
have IH2: "min_height r = ?hr"  | 
| 71858 | 122  | 
using less.prems(1) bal_length[OF _ b1] b2 by(intro less.IH) auto  | 
| 64018 | 123  | 
have "(n+1) div 2 \<ge> 1" by arith  | 
124  | 
hence 0: "log 2 ((n+1) div 2) \<ge> 0" by simp  | 
|
| 70751 | 125  | 
have "?m' \<le> ?m" by arith  | 
126  | 
hence le: "?hr \<le> ?hl" by(simp add: nat_mono floor_mono)  | 
|
127  | 
have "min_height t = min ?hl ?hr + 1" by (simp add: t IH1 IH2)  | 
|
128  | 
also have "\<dots> = ?hr + 1" using le by (simp add: min_absorb2)  | 
|
129  | 
also have "?m' + 1 = (n+1) div 2" by linarith  | 
|
| 64018 | 130  | 
also have "nat (floor(log 2 ((n+1) div 2))) + 1  | 
131  | 
= nat (floor(log 2 ((n+1) div 2) + 1))"  | 
|
132  | 
using 0 by linarith  | 
|
133  | 
also have "\<dots> = nat (floor(log 2 (n + 1)))"  | 
|
134  | 
using floor_log2_div2[of "n+1"] by (simp add: log_mult)  | 
|
135  | 
finally show ?thesis .  | 
|
136  | 
qed  | 
|
137  | 
qed  | 
|
138  | 
||
139  | 
lemma height_bal:  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
140  | 
"\<lbrakk> n \<le> length xs; bal n xs = (t,zs) \<rbrakk> \<Longrightarrow> height t = nat \<lceil>log 2 (n + 1)\<rceil>"  | 
| 71858 | 141  | 
proof(induction n arbitrary: xs t zs rule: less_induct)  | 
142  | 
case (less n) show ?case  | 
|
| 64018 | 143  | 
proof cases  | 
144  | 
assume "n = 0" thus ?thesis  | 
|
| 71858 | 145  | 
using less.prems by (simp add: bal_simps)  | 
| 64018 | 146  | 
next  | 
147  | 
assume [arith]: "n \<noteq> 0"  | 
|
| 70751 | 148  | 
let ?m = "n div 2" let ?m' = "n - 1 - ?m"  | 
| 71858 | 149  | 
from less.prems obtain l r ys where  | 
| 70751 | 150  | 
b1: "bal ?m xs = (l,ys)" and  | 
151  | 
b2: "bal ?m' (tl ys) = (r,zs)" and  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
152  | 
t: "t = \<langle>l, hd ys, r\<rangle>"  | 
| 71846 | 153  | 
by(auto simp: bal_simps split: prod.splits)  | 
| 70751 | 154  | 
let ?hl = "nat \<lceil>log 2 (?m + 1)\<rceil>"  | 
155  | 
let ?hr = "nat \<lceil>log 2 (?m' + 1)\<rceil>"  | 
|
| 71858 | 156  | 
have IH1: "height l = ?hl" using less.IH[OF _ _ b1] less.prems(1) by simp  | 
| 70751 | 157  | 
have IH2: "height r = ?hr"  | 
| 71858 | 158  | 
using b2 bal_length[OF _ b1] less.prems(1) by(intro less.IH) auto  | 
| 70751 | 159  | 
have 0: "log 2 (?m + 1) \<ge> 0" by simp  | 
160  | 
have "?m' \<le> ?m" by arith  | 
|
161  | 
hence le: "?hr \<le> ?hl"  | 
|
| 64018 | 162  | 
by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq)  | 
| 70751 | 163  | 
have "height t = max ?hl ?hr + 1" by (simp add: t IH1 IH2)  | 
164  | 
also have "\<dots> = ?hl + 1" using le by (simp add: max_absorb1)  | 
|
165  | 
also have "\<dots> = nat \<lceil>log 2 (?m + 1) + 1\<rceil>" using 0 by linarith  | 
|
| 64018 | 166  | 
also have "\<dots> = nat \<lceil>log 2 (n + 1)\<rceil>"  | 
| 66515 | 167  | 
using ceiling_log2_div2[of "n+1"] by (simp)  | 
| 63643 | 168  | 
finally show ?thesis .  | 
169  | 
qed  | 
|
170  | 
qed  | 
|
171  | 
||
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
172  | 
lemma acomplete_bal:  | 
| 
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
173  | 
assumes "n \<le> length xs" "bal n xs = (t,ys)" shows "acomplete t"  | 
| 
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
174  | 
unfolding acomplete_def  | 
| 64018 | 175  | 
using height_bal[OF assms] min_height_bal[OF assms]  | 
176  | 
by linarith  | 
|
| 63643 | 177  | 
|
| 64444 | 178  | 
lemma height_bal_list:  | 
179  | 
"n \<le> length xs \<Longrightarrow> height (bal_list n xs) = nat \<lceil>log 2 (n + 1)\<rceil>"  | 
|
180  | 
unfolding bal_list_def by (metis height_bal prod.collapse)  | 
|
181  | 
||
| 64018 | 182  | 
lemma height_balance_list:  | 
183  | 
"height (balance_list xs) = nat \<lceil>log 2 (length xs + 1)\<rceil>"  | 
|
| 64444 | 184  | 
by (simp add: balance_list_def height_bal_list)  | 
185  | 
||
186  | 
corollary height_bal_tree:  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
187  | 
"n \<le> size t \<Longrightarrow> height (bal_tree n t) = nat\<lceil>log 2 (n + 1)\<rceil>"  | 
| 64444 | 188  | 
unfolding bal_list_def bal_tree_def  | 
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
189  | 
by (metis bal_list_def height_bal_list length_inorder)  | 
| 64018 | 190  | 
|
191  | 
corollary height_balance_tree:  | 
|
| 66515 | 192  | 
"height (balance_tree t) = nat\<lceil>log 2 (size t + 1)\<rceil>"  | 
| 64444 | 193  | 
by (simp add: bal_tree_def balance_tree_def height_bal_list)  | 
194  | 
||
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
195  | 
corollary acomplete_bal_list[simp]: "n \<le> length xs \<Longrightarrow> acomplete (bal_list n xs)"  | 
| 
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
196  | 
unfolding bal_list_def by (metis acomplete_bal prod.collapse)  | 
| 63829 | 197  | 
|
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
198  | 
corollary acomplete_balance_list[simp]: "acomplete (balance_list xs)"  | 
| 64444 | 199  | 
by (simp add: balance_list_def)  | 
200  | 
||
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
201  | 
corollary acomplete_bal_tree[simp]: "n \<le> size t \<Longrightarrow> acomplete (bal_tree n t)"  | 
| 64444 | 202  | 
by (simp add: bal_tree_def)  | 
| 63829 | 203  | 
|
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
204  | 
corollary acomplete_balance_tree[simp]: "acomplete (balance_tree t)"  | 
| 63829 | 205  | 
by (simp add: balance_tree_def)  | 
206  | 
||
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
207  | 
lemma wbalanced_bal: "\<lbrakk> n \<le> length xs; bal n xs = (t,ys) \<rbrakk> \<Longrightarrow> wbalanced t"  | 
| 71858 | 208  | 
proof(induction n arbitrary: xs t ys rule: less_induct)  | 
209  | 
case (less n)  | 
|
| 63861 | 210  | 
show ?case  | 
211  | 
proof cases  | 
|
212  | 
assume "n = 0"  | 
|
| 71858 | 213  | 
thus ?thesis using less.prems(2) by(simp add: bal_simps)  | 
| 63861 | 214  | 
next  | 
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
215  | 
assume [arith]: "n \<noteq> 0"  | 
| 71858 | 216  | 
with less.prems obtain l ys r zs where  | 
| 64444 | 217  | 
rec1: "bal (n div 2) xs = (l, ys)" and  | 
218  | 
rec2: "bal (n - 1 - n div 2) (tl ys) = (r, zs)" and  | 
|
| 63861 | 219  | 
t: "t = \<langle>l, hd ys, r\<rangle>"  | 
| 71846 | 220  | 
by(auto simp add: bal_simps split: prod.splits)  | 
| 71858 | 221  | 
have l: "wbalanced l" using less.IH[OF _ _ rec1] less.prems(1) by linarith  | 
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
222  | 
have "wbalanced r"  | 
| 71858 | 223  | 
using rec1 rec2 bal_length[OF _ rec1] less.prems(1) by(intro less.IH) auto  | 
224  | 
with l t bal_length[OF _ rec1] less.prems(1) bal_inorder[OF _ rec1] bal_inorder[OF _ rec2]  | 
|
| 63861 | 225  | 
show ?thesis by auto  | 
226  | 
qed  | 
|
227  | 
qed  | 
|
228  | 
||
| 
72566
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
229  | 
text\<open>An alternative proof via @{thm acomplete_if_wbalanced}:\<close>
 | 
| 
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
230  | 
lemma "\<lbrakk> n \<le> length xs; bal n xs = (t,ys) \<rbrakk> \<Longrightarrow> acomplete t"  | 
| 
 
831f17da1aab
renamed "balanced" -> "acomplete" because balanced has other meanings in the literature
 
nipkow 
parents: 
71858 
diff
changeset
 | 
231  | 
by(rule acomplete_if_wbalanced[OF wbalanced_bal])  | 
| 64541 | 232  | 
|
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
233  | 
lemma wbalanced_bal_list[simp]: "n \<le> length xs \<Longrightarrow> wbalanced (bal_list n xs)"  | 
| 64444 | 234  | 
by(simp add: bal_list_def) (metis prod.collapse wbalanced_bal)  | 
235  | 
||
236  | 
lemma wbalanced_balance_list[simp]: "wbalanced (balance_list xs)"  | 
|
237  | 
by(simp add: balance_list_def)  | 
|
238  | 
||
| 
70747
 
548420d389ea
Enforced precodition "n <= length xs" to avoid relying on "hd []".
 
nipkow 
parents: 
66516 
diff
changeset
 | 
239  | 
lemma wbalanced_bal_tree[simp]: "n \<le> size t \<Longrightarrow> wbalanced (bal_tree n t)"  | 
| 64444 | 240  | 
by(simp add: bal_tree_def)  | 
241  | 
||
| 63861 | 242  | 
lemma wbalanced_balance_tree: "wbalanced (balance_tree t)"  | 
| 64444 | 243  | 
by (simp add: balance_tree_def)  | 
| 63861 | 244  | 
|
| 63829 | 245  | 
hide_const (open) bal  | 
| 63643 | 246  | 
|
247  | 
end  |