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