| author | wenzelm |
| Mon, 13 Apr 2020 17:40:44 +0200 | |
| changeset 71749 | 77232ff6b8f6 |
| parent 71722 | 1cffe8f4d7b3 |
| child 71795 | 8ad9b2d3dd81 |
| permissions | -rw-r--r-- |
| 61232 | 1 |
(* |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
2 |
Author: Tobias Nipkow, Daniel Stüwe |
| 71635 | 3 |
Based on the AFP entry AVL. |
| 61232 | 4 |
*) |
5 |
||
6 |
section "AVL Tree Implementation of Sets" |
|
7 |
||
8 |
theory AVL_Set |
|
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
9 |
imports |
| 67964 | 10 |
Cmp |
11 |
Isin2 |
|
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63411
diff
changeset
|
12 |
"HOL-Number_Theory.Fib" |
| 61232 | 13 |
begin |
14 |
||
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
15 |
type_synonym 'a avl_tree = "('a*nat) tree"
|
| 61232 | 16 |
|
| 68431 | 17 |
definition empty :: "'a avl_tree" where |
18 |
"empty = Leaf" |
|
19 |
||
| 67406 | 20 |
text \<open>Invariant:\<close> |
| 61232 | 21 |
|
22 |
fun avl :: "'a avl_tree \<Rightarrow> bool" where |
|
23 |
"avl Leaf = True" | |
|
| 71466 | 24 |
"avl (Node l (a,n) r) = |
| 61232 | 25 |
((height l = height r \<or> height l = height r + 1 \<or> height r = height l + 1) \<and> |
| 71466 | 26 |
n = max (height l) (height r) + 1 \<and> avl l \<and> avl r)" |
| 61232 | 27 |
|
28 |
fun ht :: "'a avl_tree \<Rightarrow> nat" where |
|
29 |
"ht Leaf = 0" | |
|
| 71466 | 30 |
"ht (Node l (a,n) r) = n" |
| 61232 | 31 |
|
32 |
definition node :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where |
|
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
33 |
"node l a r = Node l (a, max (ht l) (ht r) + 1) r" |
| 61232 | 34 |
|
| 61581 | 35 |
definition balL :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where |
| 71635 | 36 |
"balL AB b C = |
37 |
(if ht AB = ht C + 2 then |
|
38 |
case AB of |
|
39 |
Node A (a, _) B \<Rightarrow> |
|
40 |
if ht A \<ge> ht B then node A a (node B b C) |
|
41 |
else |
|
42 |
case B of |
|
43 |
Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C) |
|
44 |
else node AB b C)" |
|
| 61232 | 45 |
|
| 61581 | 46 |
definition balR :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where |
| 71635 | 47 |
"balR A a BC = |
48 |
(if ht BC = ht A + 2 then |
|
49 |
case BC of |
|
50 |
Node B (b, _) C \<Rightarrow> |
|
51 |
if ht B \<le> ht C then node (node A a B) b C |
|
52 |
else |
|
53 |
case B of |
|
54 |
Node B\<^sub>1 (ab, _) B\<^sub>2 \<Rightarrow> node (node A a B\<^sub>1) ab (node B\<^sub>2 b C) |
|
55 |
else node A a BC)" |
|
| 61232 | 56 |
|
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
57 |
fun insert :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
58 |
"insert x Leaf = Node Leaf (x, 1) Leaf" | |
| 71466 | 59 |
"insert x (Node l (a, n) r) = (case cmp x a of |
60 |
EQ \<Rightarrow> Node l (a, n) r | |
|
| 61581 | 61 |
LT \<Rightarrow> balL (insert x l) a r | |
62 |
GT \<Rightarrow> balR l a (insert x r))" |
|
| 61232 | 63 |
|
| 68023 | 64 |
fun split_max :: "'a avl_tree \<Rightarrow> 'a avl_tree * 'a" where |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
65 |
"split_max (Node l (a, _) r) = |
| 68023 | 66 |
(if r = Leaf then (l,a) else let (r',a') = split_max r in (balL l a r', a'))" |
| 61232 | 67 |
|
| 68023 | 68 |
lemmas split_max_induct = split_max.induct[case_names Node Leaf] |
| 61232 | 69 |
|
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
70 |
fun delete :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where |
| 61232 | 71 |
"delete _ Leaf = Leaf" | |
| 71466 | 72 |
"delete x (Node l (a, n) r) = |
| 61678 | 73 |
(case cmp x a of |
| 71636 | 74 |
EQ \<Rightarrow> if l = Leaf then r |
75 |
else let (l', a') = split_max l in balR l' a' r | |
|
| 61678 | 76 |
LT \<Rightarrow> balR (delete x l) a r | |
77 |
GT \<Rightarrow> balL l a (delete x r))" |
|
| 61232 | 78 |
|
79 |
||
| 67406 | 80 |
subsection \<open>Functional Correctness Proofs\<close> |
| 61232 | 81 |
|
| 67406 | 82 |
text\<open>Very different from the AFP/AVL proofs\<close> |
| 61232 | 83 |
|
84 |
||
85 |
subsubsection "Proofs for insert" |
|
86 |
||
| 61581 | 87 |
lemma inorder_balL: |
88 |
"inorder (balL l a r) = inorder l @ a # inorder r" |
|
89 |
by (auto simp: node_def balL_def split:tree.splits) |
|
| 61232 | 90 |
|
| 61581 | 91 |
lemma inorder_balR: |
92 |
"inorder (balR l a r) = inorder l @ a # inorder r" |
|
93 |
by (auto simp: node_def balR_def split:tree.splits) |
|
| 61232 | 94 |
|
95 |
theorem inorder_insert: |
|
96 |
"sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)" |
|
97 |
by (induct t) |
|
| 61581 | 98 |
(auto simp: ins_list_simps inorder_balL inorder_balR) |
| 61232 | 99 |
|
100 |
||
101 |
subsubsection "Proofs for delete" |
|
102 |
||
| 68023 | 103 |
lemma inorder_split_maxD: |
104 |
"\<lbrakk> split_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow> |
|
| 61232 | 105 |
inorder t' @ [a] = inorder t" |
| 68023 | 106 |
by(induction t arbitrary: t' rule: split_max.induct) |
| 61647 | 107 |
(auto simp: inorder_balL split: if_splits prod.splits tree.split) |
| 61232 | 108 |
|
109 |
theorem inorder_delete: |
|
110 |
"sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)" |
|
111 |
by(induction t) |
|
| 71636 | 112 |
(auto simp: del_list_simps inorder_balL inorder_balR inorder_split_maxD split: prod.splits) |
| 61232 | 113 |
|
114 |
||
| 67406 | 115 |
subsection \<open>AVL invariants\<close> |
| 61232 | 116 |
|
| 67406 | 117 |
text\<open>Essentially the AFP/AVL proofs\<close> |
| 61232 | 118 |
|
119 |
||
| 67406 | 120 |
subsubsection \<open>Insertion maintains AVL balance\<close> |
| 61232 | 121 |
|
122 |
declare Let_def [simp] |
|
123 |
||
| 70571 | 124 |
lemma ht_height[simp]: "avl t \<Longrightarrow> ht t = height t" |
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
125 |
by (cases t rule: tree2_cases) simp_all |
| 61232 | 126 |
|
| 71635 | 127 |
text \<open>First, a fast but relatively manual proof with many lemmas:\<close> |
128 |
||
| 61581 | 129 |
lemma height_balL: |
| 71722 | 130 |
"\<lbrakk> avl l; avl r; height l = height r + 2 \<rbrakk> \<Longrightarrow> |
| 61581 | 131 |
height (balL l a r) = height r + 2 \<or> |
132 |
height (balL l a r) = height r + 3" |
|
133 |
by (cases l) (auto simp:node_def balL_def split:tree.split) |
|
| 61232 | 134 |
|
| 61581 | 135 |
lemma height_balR: |
| 71722 | 136 |
"\<lbrakk> avl l; avl r; height r = height l + 2 \<rbrakk> \<Longrightarrow> |
| 61581 | 137 |
height (balR l a r) = height l + 2 \<or> |
138 |
height (balR l a r) = height l + 3" |
|
139 |
by (cases r) (auto simp add:node_def balR_def split:tree.split) |
|
| 61232 | 140 |
|
| 70571 | 141 |
lemma height_node[simp]: "height(node l a r) = max (height l) (height r) + 1" |
| 61232 | 142 |
by (simp add: node_def) |
143 |
||
| 61581 | 144 |
lemma height_balL2: |
| 61232 | 145 |
"\<lbrakk> avl l; avl r; height l \<noteq> height r + 2 \<rbrakk> \<Longrightarrow> |
| 71722 | 146 |
height (balL l a r) = 1 + max (height l) (height r)" |
| 61581 | 147 |
by (cases l, cases r) (simp_all add: balL_def) |
| 61232 | 148 |
|
| 61581 | 149 |
lemma height_balR2: |
| 61232 | 150 |
"\<lbrakk> avl l; avl r; height r \<noteq> height l + 2 \<rbrakk> \<Longrightarrow> |
| 71722 | 151 |
height (balR l a r) = 1 + max (height l) (height r)" |
| 61581 | 152 |
by (cases l, cases r) (simp_all add: balR_def) |
| 61232 | 153 |
|
| 61581 | 154 |
lemma avl_balL: |
| 71722 | 155 |
"\<lbrakk> avl l; avl r; height r - 1 \<le> height l \<and> height l \<le> height r + 2 \<rbrakk> \<Longrightarrow> avl(balL l a r)" |
156 |
by(cases l, cases r) |
|
157 |
(auto simp: balL_def node_def split!: if_split tree.split) |
|
| 61232 | 158 |
|
| 61581 | 159 |
lemma avl_balR: |
| 71722 | 160 |
"\<lbrakk> avl l; avl r; height l - 1 \<le> height r \<and> height r \<le> height l + 2 \<rbrakk> \<Longrightarrow> avl(balR l a r)" |
161 |
by(cases r, cases l) |
|
162 |
(auto simp: balR_def node_def split!: if_split tree.split) |
|
| 61232 | 163 |
|
| 71635 | 164 |
text\<open>Insertion maintains the AVL property. Requires simultaneous proof.\<close> |
| 61232 | 165 |
|
| 68422 | 166 |
theorem avl_insert: |
| 61232 | 167 |
assumes "avl t" |
168 |
shows "avl(insert x t)" |
|
169 |
"(height (insert x t) = height t \<or> height (insert x t) = height t + 1)" |
|
170 |
using assms |
|
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
171 |
proof (induction t rule: tree2_induct) |
| 71466 | 172 |
case (Node l a _ r) |
| 61232 | 173 |
case 1 |
| 68422 | 174 |
show ?case |
| 61232 | 175 |
proof(cases "x = a") |
| 68422 | 176 |
case True with Node 1 show ?thesis by simp |
| 61232 | 177 |
next |
178 |
case False |
|
| 68422 | 179 |
show ?thesis |
| 61232 | 180 |
proof(cases "x<a") |
| 68422 | 181 |
case True with Node 1 show ?thesis by (auto simp add:avl_balL) |
| 61232 | 182 |
next |
| 68422 | 183 |
case False with Node 1 \<open>x\<noteq>a\<close> show ?thesis by (auto simp add:avl_balR) |
| 61232 | 184 |
qed |
185 |
qed |
|
186 |
case 2 |
|
| 68422 | 187 |
show ?case |
| 61232 | 188 |
proof(cases "x = a") |
| 68422 | 189 |
case True with Node 1 show ?thesis by simp |
| 61232 | 190 |
next |
191 |
case False |
|
| 68422 | 192 |
show ?thesis |
193 |
proof(cases "x<a") |
|
| 61232 | 194 |
case True |
| 68422 | 195 |
show ?thesis |
| 61232 | 196 |
proof(cases "height (insert x l) = height r + 2") |
| 67406 | 197 |
case False with Node 2 \<open>x < a\<close> show ?thesis by (auto simp: height_balL2) |
| 61232 | 198 |
next |
199 |
case True |
|
| 61581 | 200 |
hence "(height (balL (insert x l) a r) = height r + 2) \<or> |
201 |
(height (balL (insert x l) a r) = height r + 3)" (is "?A \<or> ?B") |
|
202 |
using Node 2 by (intro height_balL) simp_all |
|
| 61232 | 203 |
thus ?thesis |
204 |
proof |
|
| 68422 | 205 |
assume ?A with 2 \<open>x < a\<close> show ?thesis by (auto) |
| 61232 | 206 |
next |
| 68422 | 207 |
assume ?B with True 1 Node(2) \<open>x < a\<close> show ?thesis by (simp) arith |
| 61232 | 208 |
qed |
209 |
qed |
|
210 |
next |
|
211 |
case False |
|
| 68422 | 212 |
show ?thesis |
| 61232 | 213 |
proof(cases "height (insert x r) = height l + 2") |
| 68422 | 214 |
case False with Node 2 \<open>\<not>x < a\<close> show ?thesis by (auto simp: height_balR2) |
| 61232 | 215 |
next |
216 |
case True |
|
| 61581 | 217 |
hence "(height (balR l a (insert x r)) = height l + 2) \<or> |
218 |
(height (balR l a (insert x r)) = height l + 3)" (is "?A \<or> ?B") |
|
219 |
using Node 2 by (intro height_balR) simp_all |
|
| 61232 | 220 |
thus ?thesis |
221 |
proof |
|
| 68422 | 222 |
assume ?A with 2 \<open>\<not>x < a\<close> show ?thesis by (auto) |
| 61232 | 223 |
next |
| 68422 | 224 |
assume ?B with True 1 Node(4) \<open>\<not>x < a\<close> show ?thesis by (simp) arith |
| 61232 | 225 |
qed |
226 |
qed |
|
227 |
qed |
|
228 |
qed |
|
229 |
qed simp_all |
|
230 |
||
| 71635 | 231 |
text \<open>Now an automatic proof without lemmas:\<close> |
232 |
||
233 |
theorem avl_insert_auto: "avl t \<Longrightarrow> |
|
234 |
avl(insert x t) \<and> height (insert x t) \<in> {height t, height t + 1}"
|
|
235 |
apply (induction t rule: tree2_induct) |
|
| 71722 | 236 |
(* if you want to save a few secs: apply (auto split!: if_split) *) |
237 |
apply (auto simp: balL_def balR_def node_def max_absorb2 split!: if_split tree.split) |
|
| 71635 | 238 |
done |
239 |
||
| 61232 | 240 |
|
| 67406 | 241 |
subsubsection \<open>Deletion maintains AVL balance\<close> |
| 61232 | 242 |
|
| 68023 | 243 |
lemma avl_split_max: |
| 71722 | 244 |
"\<lbrakk> avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> |
245 |
avl (fst (split_max t)) \<and> |
|
246 |
height t \<in> {height(fst (split_max t)), height(fst (split_max t)) + 1}"
|
|
247 |
by(induct t rule: split_max_induct) |
|
248 |
(auto simp: balL_def node_def max_absorb2 split!: prod.split if_split tree.split) |
|
| 61232 | 249 |
|
| 67406 | 250 |
text\<open>Deletion maintains the AVL property:\<close> |
| 61232 | 251 |
|
| 68422 | 252 |
theorem avl_delete: |
| 61232 | 253 |
assumes "avl t" |
254 |
shows "avl(delete x t)" and "height t = (height (delete x t)) \<or> height t = height (delete x t) + 1" |
|
255 |
using assms |
|
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
256 |
proof (induct t rule: tree2_induct) |
| 71466 | 257 |
case (Node l a n r) |
| 61232 | 258 |
case 1 |
| 71722 | 259 |
thus ?case |
260 |
using Node avl_split_max[of l] by (auto simp: avl_balL avl_balR split: prod.split) |
|
| 61232 | 261 |
case 2 |
| 68422 | 262 |
show ?case |
| 71466 | 263 |
proof(cases "x = a") |
| 71636 | 264 |
case True then show ?thesis using 1 avl_split_max[of l] |
265 |
by(auto simp: balR_def max_absorb2 split!: if_splits prod.split tree.split) |
|
| 61232 | 266 |
next |
267 |
case False |
|
| 68422 | 268 |
show ?thesis |
| 71466 | 269 |
proof(cases "x<a") |
| 61232 | 270 |
case True |
271 |
show ?thesis |
|
272 |
proof(cases "height r = height (delete x l) + 2") |
|
| 71466 | 273 |
case False with Node 1 \<open>x < a\<close> show ?thesis by(auto simp: balR_def) |
| 61232 | 274 |
next |
275 |
case True |
|
| 71466 | 276 |
hence "(height (balR (delete x l) a r) = height (delete x l) + 2) \<or> |
277 |
height (balR (delete x l) a r) = height (delete x l) + 3" (is "?A \<or> ?B") |
|
| 61581 | 278 |
using Node 2 by (intro height_balR) auto |
| 61232 | 279 |
thus ?thesis |
280 |
proof |
|
| 71466 | 281 |
assume ?A with \<open>x < a\<close> Node 2 show ?thesis by(auto simp: balR_def) |
| 61232 | 282 |
next |
| 71466 | 283 |
assume ?B with \<open>x < a\<close> Node 2 show ?thesis by(auto simp: balR_def) |
| 61232 | 284 |
qed |
285 |
qed |
|
286 |
next |
|
287 |
case False |
|
288 |
show ?thesis |
|
289 |
proof(cases "height l = height (delete x r) + 2") |
|
| 71466 | 290 |
case False with Node 1 \<open>\<not>x < a\<close> \<open>x \<noteq> a\<close> show ?thesis by(auto simp: balL_def) |
| 61232 | 291 |
next |
292 |
case True |
|
| 71466 | 293 |
hence "(height (balL l a (delete x r)) = height (delete x r) + 2) \<or> |
294 |
height (balL l a (delete x r)) = height (delete x r) + 3" (is "?A \<or> ?B") |
|
| 61581 | 295 |
using Node 2 by (intro height_balL) auto |
| 61232 | 296 |
thus ?thesis |
297 |
proof |
|
| 71466 | 298 |
assume ?A with \<open>\<not>x < a\<close> \<open>x \<noteq> a\<close> Node 2 show ?thesis by(auto simp: balL_def) |
| 61232 | 299 |
next |
| 71466 | 300 |
assume ?B with \<open>\<not>x < a\<close> \<open>x \<noteq> a\<close> Node 2 show ?thesis by(auto simp: balL_def) |
| 61232 | 301 |
qed |
302 |
qed |
|
303 |
qed |
|
304 |
qed |
|
305 |
qed simp_all |
|
306 |
||
| 71722 | 307 |
text \<open>A more automatic proof. |
308 |
Complete automation as for insertion seems hard due to resource requirements.\<close> |
|
309 |
||
310 |
theorem avl_delete_auto: |
|
311 |
assumes "avl t" |
|
312 |
shows "avl(delete x t)" and "height t \<in> {height (delete x t), height (delete x t) + 1}"
|
|
313 |
using assms |
|
314 |
proof (induct t rule: tree2_induct) |
|
315 |
case (Node l a n r) |
|
316 |
case 1 |
|
317 |
show ?case |
|
318 |
proof(cases "x = a") |
|
319 |
case True thus ?thesis |
|
320 |
using 1 avl_split_max[of l] by (auto simp: avl_balR split: prod.split) |
|
321 |
next |
|
322 |
case False thus ?thesis |
|
323 |
using Node 1 by (auto simp: avl_balL avl_balR) |
|
324 |
qed |
|
325 |
case 2 |
|
326 |
show ?case |
|
327 |
proof(cases "x = a") |
|
328 |
case True thus ?thesis |
|
329 |
using 1 avl_split_max[of l] |
|
330 |
by(auto simp: balR_def max_absorb2 split!: if_splits prod.split tree.split) |
|
331 |
next |
|
332 |
case False thus ?thesis |
|
333 |
using height_balL[of l "delete x r" a] height_balR[of "delete x l" r a] Node 1 |
|
334 |
by(auto simp: balL_def balR_def split!: if_split) |
|
335 |
qed |
|
336 |
qed simp_all |
|
337 |
||
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
338 |
|
| 68422 | 339 |
subsection "Overall correctness" |
340 |
||
| 68440 | 341 |
interpretation S: Set_by_Ordered |
| 68431 | 342 |
where empty = empty and isin = isin and insert = insert and delete = delete |
| 68422 | 343 |
and inorder = inorder and inv = avl |
344 |
proof (standard, goal_cases) |
|
| 68431 | 345 |
case 1 show ?case by (simp add: empty_def) |
| 68422 | 346 |
next |
347 |
case 2 thus ?case by(simp add: isin_set_inorder) |
|
348 |
next |
|
349 |
case 3 thus ?case by(simp add: inorder_insert) |
|
350 |
next |
|
351 |
case 4 thus ?case by(simp add: inorder_delete) |
|
352 |
next |
|
| 68431 | 353 |
case 5 thus ?case by (simp add: empty_def) |
| 68422 | 354 |
next |
355 |
case 6 thus ?case by (simp add: avl_insert(1)) |
|
356 |
next |
|
357 |
case 7 thus ?case by (simp add: avl_delete(1)) |
|
358 |
qed |
|
359 |
||
360 |
||
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
361 |
subsection \<open>Height-Size Relation\<close> |
|
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
362 |
|
| 71487 | 363 |
text \<open>Based on theorems by Daniel St\"uwe, Manuel Eberl and Peter Lammich, much simplified.\<close> |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
364 |
|
| 71466 | 365 |
text \<open>Any AVL tree of height \<open>n\<close> has at least \<open>fib (n+2)\<close> leaves:\<close> |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
366 |
|
| 71466 | 367 |
lemma avl_fib_bound: "avl t \<Longrightarrow> height t = n \<Longrightarrow> fib (n+2) \<le> size1 t" |
368 |
proof (induction n arbitrary: t rule: fib.induct) |
|
| 71487 | 369 |
case 1 thus ?case by (simp) |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
370 |
next |
| 71487 | 371 |
case 2 thus ?case by (cases t) (auto) |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
372 |
next |
| 68313 | 373 |
case (3 h) |
374 |
from "3.prems" obtain l a r where |
|
|
70755
3fb16bed5d6c
replaced new type ('a,'b) tree by old type ('a*'b) tree.
nipkow
parents:
70585
diff
changeset
|
375 |
[simp]: "t = Node l (a,Suc(Suc h)) r" "avl l" "avl r" |
| 68313 | 376 |
and C: " |
377 |
height r = Suc h \<and> height l = Suc h |
|
378 |
\<or> height r = Suc h \<and> height l = h |
|
379 |
\<or> height r = h \<and> height l = Suc h" (is "?C1 \<or> ?C2 \<or> ?C3") |
|
380 |
by (cases t) (simp, fastforce) |
|
381 |
{
|
|
382 |
assume ?C1 |
|
383 |
with "3.IH"(1) |
|
384 |
have "fib (h + 3) \<le> size1 l" "fib (h + 3) \<le> size1 r" |
|
385 |
by (simp_all add: eval_nat_numeral) |
|
386 |
hence ?case by (auto simp: eval_nat_numeral) |
|
387 |
} moreover {
|
|
388 |
assume ?C2 |
|
389 |
hence ?case using "3.IH"(1)[of r] "3.IH"(2)[of l] by auto |
|
390 |
} moreover {
|
|
391 |
assume ?C3 |
|
392 |
hence ?case using "3.IH"(1)[of l] "3.IH"(2)[of r] by auto |
|
393 |
} ultimately show ?case using C by blast |
|
394 |
qed |
|
395 |
||
| 69597 | 396 |
text \<open>An exponential lower bound for \<^const>\<open>fib\<close>:\<close> |
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
397 |
|
| 68313 | 398 |
lemma fib_lowerbound: |
399 |
defines "\<phi> \<equiv> (1 + sqrt 5) / 2" |
|
| 71486 | 400 |
shows "real (fib(n+2)) \<ge> \<phi> ^ n" |
401 |
proof (induction n rule: fib.induct) |
|
402 |
case 1 |
|
403 |
then show ?case by simp |
|
404 |
next |
|
405 |
case 2 |
|
406 |
then show ?case by (simp add: \<phi>_def real_le_lsqrt) |
|
407 |
next |
|
408 |
case (3 n) term ?case |
|
409 |
have "\<phi> ^ Suc (Suc n) = \<phi> ^ 2 * \<phi> ^ n" |
|
410 |
by (simp add: field_simps power2_eq_square) |
|
411 |
also have "\<dots> = (\<phi> + 1) * \<phi> ^ n" |
|
412 |
by (simp_all add: \<phi>_def power2_eq_square field_simps) |
|
413 |
also have "\<dots> = \<phi> ^ Suc n + \<phi> ^ n" |
|
| 68313 | 414 |
by (simp add: field_simps) |
| 71486 | 415 |
also have "\<dots> \<le> real (fib (Suc n + 2)) + real (fib (n + 2))" |
416 |
by (intro add_mono "3.IH") |
|
417 |
finally show ?case by simp |
|
| 68313 | 418 |
qed |
419 |
||
420 |
text \<open>The size of an AVL tree is (at least) exponential in its height:\<close> |
|
421 |
||
| 68342 | 422 |
lemma avl_size_lowerbound: |
| 68313 | 423 |
defines "\<phi> \<equiv> (1 + sqrt 5) / 2" |
424 |
assumes "avl t" |
|
| 68342 | 425 |
shows "\<phi> ^ (height t) \<le> size1 t" |
| 68313 | 426 |
proof - |
| 71486 | 427 |
have "\<phi> ^ height t \<le> fib (height t + 2)" |
428 |
unfolding \<phi>_def by(rule fib_lowerbound) |
|
| 68342 | 429 |
also have "\<dots> \<le> size1 t" |
| 68313 | 430 |
using avl_fib_bound[of t "height t"] assms by simp |
431 |
finally show ?thesis . |
|
432 |
qed |
|
|
63411
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents:
62526
diff
changeset
|
433 |
|
| 69597 | 434 |
text \<open>The height of an AVL tree is most \<^term>\<open>(1/log 2 \<phi>)\<close> \<open>\<approx> 1.44\<close> times worse |
435 |
than \<^term>\<open>log 2 (size1 t)\<close>:\<close> |
|
| 68342 | 436 |
|
437 |
lemma avl_height_upperbound: |
|
438 |
defines "\<phi> \<equiv> (1 + sqrt 5) / 2" |
|
439 |
assumes "avl t" |
|
440 |
shows "height t \<le> (1/log 2 \<phi>) * log 2 (size1 t)" |
|
441 |
proof - |
|
442 |
have "\<phi> > 0" "\<phi> > 1" by(auto simp: \<phi>_def pos_add_strict) |
|
443 |
hence "height t = log \<phi> (\<phi> ^ height t)" by(simp add: log_nat_power) |
|
444 |
also have "\<dots> \<le> log \<phi> (size1 t)" |
|
| 70350 | 445 |
using avl_size_lowerbound[OF assms(2), folded \<phi>_def] \<open>1 < \<phi>\<close> |
446 |
by (simp add: le_log_of_power) |
|
| 68342 | 447 |
also have "\<dots> = (1/log 2 \<phi>) * log 2 (size1 t)" |
448 |
by(simp add: log_base_change[of 2 \<phi>]) |
|
449 |
finally show ?thesis . |
|
450 |
qed |
|
451 |
||
| 61232 | 452 |
end |