author | nipkow |
Sat, 26 Aug 2017 16:47:25 +0200 | |
changeset 66515 | 85c505c98332 |
parent 66510 | ca7a369301f6 |
child 66516 | 97c2d3846e10 |
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 |
66510 | 7 |
"HOL-Library.Tree_Real" |
63643 | 8 |
begin |
9 |
||
66515 | 10 |
(* FIXME rm floor_eq_iff / rename unique \<rightarrow> eq *) |
64018 | 11 |
|
64444 | 12 |
fun bal :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree * 'a list" where |
13 |
"bal n xs = (if n=0 then (Leaf,xs) else |
|
63643 | 14 |
(let m = n div 2; |
64444 | 15 |
(l, ys) = bal m xs; |
16 |
(r, zs) = bal (n-1-m) (tl ys) |
|
63643 | 17 |
in (Node l (hd ys) r, zs)))" |
18 |
||
19 |
declare bal.simps[simp del] |
|
20 |
||
64444 | 21 |
definition bal_list :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a tree" where |
22 |
"bal_list n xs = fst (bal n xs)" |
|
23 |
||
63829 | 24 |
definition balance_list :: "'a list \<Rightarrow> 'a tree" where |
64444 | 25 |
"balance_list xs = bal_list (length xs) xs" |
26 |
||
27 |
definition bal_tree :: "nat \<Rightarrow> 'a tree \<Rightarrow> 'a tree" where |
|
28 |
"bal_tree n t = bal_list n (inorder t)" |
|
63829 | 29 |
|
30 |
definition balance_tree :: "'a tree \<Rightarrow> 'a tree" where |
|
64444 | 31 |
"balance_tree t = bal_tree (size t) t" |
63829 | 32 |
|
63843 | 33 |
lemma bal_simps: |
64444 | 34 |
"bal 0 xs = (Leaf, xs)" |
63843 | 35 |
"n > 0 \<Longrightarrow> |
64444 | 36 |
bal n xs = |
63843 | 37 |
(let m = n div 2; |
64444 | 38 |
(l, ys) = bal m xs; |
39 |
(r, zs) = bal (n-1-m) (tl ys) |
|
63843 | 40 |
in (Node l (hd ys) r, zs))" |
41 |
by(simp_all add: bal.simps) |
|
63643 | 42 |
|
64444 | 43 |
text\<open>Some of the following lemmas take advantage of the fact |
63861 | 44 |
that \<open>bal xs n\<close> yields a result even if \<open>n > length xs\<close>.\<close> |
45 |
||
64444 | 46 |
lemma size_bal: "bal n xs = (t,ys) \<Longrightarrow> size t = n" |
47 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
48 |
case (1 n xs) |
|
63861 | 49 |
thus ?case |
50 |
by(cases "n=0") |
|
51 |
(auto simp add: bal_simps Let_def split: prod.splits) |
|
52 |
qed |
|
53 |
||
63643 | 54 |
lemma bal_inorder: |
64444 | 55 |
"\<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
|
56 |
\<Longrightarrow> inorder t = take n xs \<and> ys = drop n xs" |
64444 | 57 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
58 |
case (1 n xs) show ?case |
|
63643 | 59 |
proof cases |
63843 | 60 |
assume "n = 0" thus ?thesis using 1 by (simp add: bal_simps) |
63643 | 61 |
next |
62 |
assume [arith]: "n \<noteq> 0" |
|
63 |
let ?n1 = "n div 2" let ?n2 = "n - 1 - ?n1" |
|
64 |
from "1.prems" obtain l r xs' where |
|
64444 | 65 |
b1: "bal ?n1 xs = (l,xs')" and |
66 |
b2: "bal ?n2 (tl xs') = (r,ys)" and |
|
63643 | 67 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
63843 | 68 |
by(auto simp: Let_def bal_simps split: prod.splits) |
63643 | 69 |
have IH1: "inorder l = take ?n1 xs \<and> xs' = drop ?n1 xs" |
70 |
using b1 "1.prems" by(intro "1.IH"(1)) auto |
|
71 |
have IH2: "inorder r = take ?n2 (tl xs') \<and> ys = drop ?n2 (tl xs')" |
|
72 |
using b1 b2 IH1 "1.prems" by(intro "1.IH"(2)) auto |
|
73 |
have "drop (n div 2) xs \<noteq> []" using "1.prems"(2) by simp |
|
74 |
hence "hd (drop ?n1 xs) # take ?n2 (tl (drop ?n1 xs)) = take (?n2 + 1) (drop ?n1 xs)" |
|
75 |
by (metis Suc_eq_plus1 take_Suc) |
|
76 |
hence *: "inorder t = take n xs" using t IH1 IH2 |
|
77 |
using take_add[of ?n1 "?n2+1" xs] by(simp) |
|
78 |
have "n - n div 2 + n div 2 = n" by simp |
|
79 |
hence "ys = drop n xs" using IH1 IH2 by (simp add: drop_Suc[symmetric]) |
|
80 |
thus ?thesis using * by blast |
|
81 |
qed |
|
82 |
qed |
|
83 |
||
64444 | 84 |
corollary inorder_bal_list[simp]: |
85 |
"n \<le> length xs \<Longrightarrow> inorder(bal_list n xs) = take n xs" |
|
86 |
unfolding bal_list_def by (metis bal_inorder eq_fst_iff) |
|
87 |
||
88 |
corollary inorder_balance_list[simp]: "inorder(balance_list xs) = xs" |
|
89 |
by(simp add: balance_list_def) |
|
90 |
||
91 |
corollary inorder_bal_tree: |
|
92 |
"n \<le> size t \<Longrightarrow> inorder(bal_tree n t) = take n (inorder t)" |
|
93 |
by(simp add: bal_tree_def) |
|
63643 | 94 |
|
64018 | 95 |
corollary inorder_balance_tree[simp]: "inorder(balance_tree t) = inorder t" |
64444 | 96 |
by(simp add: balance_tree_def inorder_bal_tree) |
97 |
||
98 |
corollary size_bal_list[simp]: "size(bal_list n xs) = n" |
|
99 |
unfolding bal_list_def by (metis prod.collapse size_bal) |
|
64018 | 100 |
|
101 |
corollary size_balance_list[simp]: "size(balance_list xs) = length xs" |
|
64444 | 102 |
by (simp add: balance_list_def) |
103 |
||
104 |
corollary size_bal_tree[simp]: "size(bal_tree n t) = n" |
|
105 |
by(simp add: bal_tree_def) |
|
64018 | 106 |
|
107 |
corollary size_balance_tree[simp]: "size(balance_tree t) = size t" |
|
64444 | 108 |
by(simp add: balance_tree_def) |
64018 | 109 |
|
110 |
lemma min_height_bal: |
|
66515 | 111 |
"bal n xs = (t,ys) \<Longrightarrow> min_height t = nat(\<lfloor>log 2 (n + 1)\<rfloor>)" |
64444 | 112 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
113 |
case (1 n xs) show ?case |
|
63643 | 114 |
proof cases |
115 |
assume "n = 0" thus ?thesis |
|
64018 | 116 |
using "1.prems" by (simp add: bal_simps) |
63643 | 117 |
next |
118 |
assume [arith]: "n \<noteq> 0" |
|
119 |
from "1.prems" obtain l r xs' where |
|
64444 | 120 |
b1: "bal (n div 2) xs = (l,xs')" and |
121 |
b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and |
|
63643 | 122 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
63843 | 123 |
by(auto simp: bal_simps Let_def split: prod.splits) |
64018 | 124 |
let ?log1 = "nat (floor(log 2 (n div 2 + 1)))" |
125 |
let ?log2 = "nat (floor(log 2 (n - 1 - n div 2 + 1)))" |
|
63643 | 126 |
have IH1: "min_height l = ?log1" using "1.IH"(1) b1 by simp |
64018 | 127 |
have IH2: "min_height r = ?log2" using "1.IH"(2) b1 b2 by simp |
128 |
have "(n+1) div 2 \<ge> 1" by arith |
|
129 |
hence 0: "log 2 ((n+1) div 2) \<ge> 0" by simp |
|
130 |
have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith |
|
131 |
hence le: "?log2 \<le> ?log1" |
|
132 |
by(simp add: nat_mono floor_mono) |
|
133 |
have "min_height t = min ?log1 ?log2 + 1" by (simp add: t IH1 IH2) |
|
134 |
also have "\<dots> = ?log2 + 1" using le by (simp add: min_absorb2) |
|
135 |
also have "n - 1 - n div 2 + 1 = (n+1) div 2" by linarith |
|
136 |
also have "nat (floor(log 2 ((n+1) div 2))) + 1 |
|
137 |
= nat (floor(log 2 ((n+1) div 2) + 1))" |
|
138 |
using 0 by linarith |
|
139 |
also have "\<dots> = nat (floor(log 2 (n + 1)))" |
|
140 |
using floor_log2_div2[of "n+1"] by (simp add: log_mult) |
|
141 |
finally show ?thesis . |
|
142 |
qed |
|
143 |
qed |
|
144 |
||
145 |
lemma height_bal: |
|
64444 | 146 |
"bal n xs = (t,ys) \<Longrightarrow> height t = nat \<lceil>log 2 (n + 1)\<rceil>" |
147 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
148 |
case (1 n xs) show ?case |
|
64018 | 149 |
proof cases |
150 |
assume "n = 0" thus ?thesis |
|
151 |
using "1.prems" by (simp add: bal_simps) |
|
152 |
next |
|
153 |
assume [arith]: "n \<noteq> 0" |
|
154 |
from "1.prems" obtain l r xs' where |
|
64444 | 155 |
b1: "bal (n div 2) xs = (l,xs')" and |
156 |
b2: "bal (n - 1 - n div 2) (tl xs') = (r,ys)" and |
|
64018 | 157 |
t: "t = \<langle>l, hd xs', r\<rangle>" |
158 |
by(auto simp: bal_simps Let_def split: prod.splits) |
|
159 |
let ?log1 = "nat \<lceil>log 2 (n div 2 + 1)\<rceil>" |
|
160 |
let ?log2 = "nat \<lceil>log 2 (n - 1 - n div 2 + 1)\<rceil>" |
|
161 |
have IH1: "height l = ?log1" using "1.IH"(1) b1 by simp |
|
162 |
have IH2: "height r = ?log2" using "1.IH"(2) b1 b2 by simp |
|
163 |
have 0: "log 2 (n div 2 + 1) \<ge> 0" by auto |
|
164 |
have "n - 1 - n div 2 + 1 \<le> n div 2 + 1" by arith |
|
165 |
hence le: "?log2 \<le> ?log1" |
|
166 |
by(simp add: nat_mono ceiling_mono del: nat_ceiling_le_eq) |
|
167 |
have "height t = max ?log1 ?log2 + 1" by (simp add: t IH1 IH2) |
|
168 |
also have "\<dots> = ?log1 + 1" using le by (simp add: max_absorb1) |
|
169 |
also have "\<dots> = nat \<lceil>log 2 (n div 2 + 1) + 1\<rceil>" using 0 by linarith |
|
170 |
also have "\<dots> = nat \<lceil>log 2 (n + 1)\<rceil>" |
|
66515 | 171 |
using ceiling_log2_div2[of "n+1"] by (simp) |
63643 | 172 |
finally show ?thesis . |
173 |
qed |
|
174 |
qed |
|
175 |
||
176 |
lemma balanced_bal: |
|
64444 | 177 |
assumes "bal n xs = (t,ys)" shows "balanced t" |
64018 | 178 |
unfolding balanced_def |
179 |
using height_bal[OF assms] min_height_bal[OF assms] |
|
180 |
by linarith |
|
63643 | 181 |
|
64444 | 182 |
lemma height_bal_list: |
183 |
"n \<le> length xs \<Longrightarrow> height (bal_list n xs) = nat \<lceil>log 2 (n + 1)\<rceil>" |
|
184 |
unfolding bal_list_def by (metis height_bal prod.collapse) |
|
185 |
||
64018 | 186 |
lemma height_balance_list: |
187 |
"height (balance_list xs) = nat \<lceil>log 2 (length xs + 1)\<rceil>" |
|
64444 | 188 |
by (simp add: balance_list_def height_bal_list) |
189 |
||
190 |
corollary height_bal_tree: |
|
66515 | 191 |
"n \<le> length xs \<Longrightarrow> height (bal_tree n t) = nat\<lceil>log 2 (n + 1)\<rceil>" |
64444 | 192 |
unfolding bal_list_def bal_tree_def |
193 |
using height_bal prod.exhaust_sel by blast |
|
64018 | 194 |
|
195 |
corollary height_balance_tree: |
|
66515 | 196 |
"height (balance_tree t) = nat\<lceil>log 2 (size t + 1)\<rceil>" |
64444 | 197 |
by (simp add: bal_tree_def balance_tree_def height_bal_list) |
198 |
||
199 |
corollary balanced_bal_list[simp]: "balanced (bal_list n xs)" |
|
200 |
unfolding bal_list_def by (metis balanced_bal prod.collapse) |
|
63829 | 201 |
|
202 |
corollary balanced_balance_list[simp]: "balanced (balance_list xs)" |
|
64444 | 203 |
by (simp add: balance_list_def) |
204 |
||
205 |
corollary balanced_bal_tree[simp]: "balanced (bal_tree n t)" |
|
206 |
by (simp add: bal_tree_def) |
|
63829 | 207 |
|
208 |
corollary balanced_balance_tree[simp]: "balanced (balance_tree t)" |
|
209 |
by (simp add: balance_tree_def) |
|
210 |
||
64444 | 211 |
lemma wbalanced_bal: "bal n xs = (t,ys) \<Longrightarrow> wbalanced t" |
212 |
proof(induction n xs arbitrary: t ys rule: bal.induct) |
|
213 |
case (1 n xs) |
|
63861 | 214 |
show ?case |
215 |
proof cases |
|
216 |
assume "n = 0" |
|
217 |
thus ?thesis |
|
218 |
using "1.prems" by(simp add: bal_simps) |
|
219 |
next |
|
220 |
assume "n \<noteq> 0" |
|
221 |
with "1.prems" obtain l ys r zs where |
|
64444 | 222 |
rec1: "bal (n div 2) xs = (l, ys)" and |
223 |
rec2: "bal (n - 1 - n div 2) (tl ys) = (r, zs)" and |
|
63861 | 224 |
t: "t = \<langle>l, hd ys, r\<rangle>" |
225 |
by(auto simp add: bal_simps Let_def split: prod.splits) |
|
226 |
have l: "wbalanced l" using "1.IH"(1)[OF \<open>n\<noteq>0\<close> refl rec1] . |
|
227 |
have "wbalanced r" using "1.IH"(2)[OF \<open>n\<noteq>0\<close> refl rec1[symmetric] refl rec2] . |
|
228 |
with l t size_bal[OF rec1] size_bal[OF rec2] |
|
229 |
show ?thesis by auto |
|
230 |
qed |
|
231 |
qed |
|
232 |
||
64541 | 233 |
text\<open>An alternative proof via @{thm balanced_if_wbalanced}:\<close> |
234 |
lemma "bal n xs = (t,ys) \<Longrightarrow> balanced t" |
|
235 |
by(rule balanced_if_wbalanced[OF wbalanced_bal]) |
|
236 |
||
64444 | 237 |
lemma wbalanced_bal_list[simp]: "wbalanced (bal_list n xs)" |
238 |
by(simp add: bal_list_def) (metis prod.collapse wbalanced_bal) |
|
239 |
||
240 |
lemma wbalanced_balance_list[simp]: "wbalanced (balance_list xs)" |
|
241 |
by(simp add: balance_list_def) |
|
242 |
||
243 |
lemma wbalanced_bal_tree[simp]: "wbalanced (bal_tree n t)" |
|
244 |
by(simp add: bal_tree_def) |
|
245 |
||
63861 | 246 |
lemma wbalanced_balance_tree: "wbalanced (balance_tree t)" |
64444 | 247 |
by (simp add: balance_tree_def) |
63861 | 248 |
|
63829 | 249 |
hide_const (open) bal |
63643 | 250 |
|
251 |
end |