| author | nipkow | 
| Wed, 17 Oct 2018 21:00:53 +0200 | |
| changeset 69145 | 806be481aa57 | 
| parent 68440 | 6826718f732d | 
| child 69597 | ff784d5a5bfb | 
| 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  | 
| 
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
3  | 
Largely derived from 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  | 
||
15  | 
type_synonym 'a avl_tree = "('a,nat) tree"
 | 
|
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" |  | 
|
| 68413 | 24  | 
"avl (Node l a h r) =  | 
| 61232 | 25  | 
((height l = height r \<or> height l = height r + 1 \<or> height r = height l + 1) \<and>  | 
26  | 
h = max (height l) (height r) + 1 \<and> avl l \<and> avl r)"  | 
|
27  | 
||
28  | 
fun ht :: "'a avl_tree \<Rightarrow> nat" where  | 
|
29  | 
"ht Leaf = 0" |  | 
|
| 68413 | 30  | 
"ht (Node l a h r) = h"  | 
| 61232 | 31  | 
|
32  | 
definition node :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where  | 
|
| 68413 | 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  | 
| 61678 | 36  | 
"balL l a r =  | 
37  | 
(if ht l = ht r + 2 then  | 
|
38  | 
case l of  | 
|
| 68413 | 39  | 
Node bl b _ br \<Rightarrow>  | 
| 61678 | 40  | 
if ht bl < ht br then  | 
41  | 
case br of  | 
|
| 68413 | 42  | 
Node cl c _ cr \<Rightarrow> node (node bl b cl) c (node cr a r)  | 
| 61678 | 43  | 
else node bl b (node br a r)  | 
44  | 
else node l a r)"  | 
|
| 61232 | 45  | 
|
| 61581 | 46  | 
definition balR :: "'a avl_tree \<Rightarrow> 'a \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where  | 
| 61678 | 47  | 
"balR l a r =  | 
48  | 
(if ht r = ht l + 2 then  | 
|
49  | 
case r of  | 
|
| 68413 | 50  | 
Node bl b _ br \<Rightarrow>  | 
| 61678 | 51  | 
if ht bl > ht br then  | 
52  | 
case bl of  | 
|
| 68413 | 53  | 
Node cl c _ cr \<Rightarrow> node (node l a cl) c (node cr b br)  | 
| 61678 | 54  | 
else node (node l a bl) b br  | 
| 61232 | 55  | 
else node l a r)"  | 
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  | 
| 68413 | 58  | 
"insert x Leaf = Node Leaf x 1 Leaf" |  | 
59  | 
"insert x (Node l a h r) = (case cmp x a of  | 
|
60  | 
EQ \<Rightarrow> Node l a h 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  | 
| 68413 | 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  | 
|
| 61647 | 70  | 
fun del_root :: "'a avl_tree \<Rightarrow> 'a avl_tree" where  | 
| 68413 | 71  | 
"del_root (Node Leaf a h r) = r" |  | 
72  | 
"del_root (Node l a h Leaf) = l" |  | 
|
73  | 
"del_root (Node l a h r) = (let (l', a') = split_max l in balR l' a' r)"  | 
|
| 61232 | 74  | 
|
| 61647 | 75  | 
lemmas del_root_cases = del_root.cases[case_names Leaf_t Node_Leaf Node_Node]  | 
| 61232 | 76  | 
|
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
77  | 
fun delete :: "'a::linorder \<Rightarrow> 'a avl_tree \<Rightarrow> 'a avl_tree" where  | 
| 61232 | 78  | 
"delete _ Leaf = Leaf" |  | 
| 68413 | 79  | 
"delete x (Node l a h r) =  | 
| 61678 | 80  | 
(case cmp x a of  | 
| 68413 | 81  | 
EQ \<Rightarrow> del_root (Node l a h r) |  | 
| 61678 | 82  | 
LT \<Rightarrow> balR (delete x l) a r |  | 
83  | 
GT \<Rightarrow> balL l a (delete x r))"  | 
|
| 61232 | 84  | 
|
85  | 
||
| 67406 | 86  | 
subsection \<open>Functional Correctness Proofs\<close>  | 
| 61232 | 87  | 
|
| 67406 | 88  | 
text\<open>Very different from the AFP/AVL proofs\<close>  | 
| 61232 | 89  | 
|
90  | 
||
91  | 
subsubsection "Proofs for insert"  | 
|
92  | 
||
| 61581 | 93  | 
lemma inorder_balL:  | 
94  | 
"inorder (balL l a r) = inorder l @ a # inorder r"  | 
|
95  | 
by (auto simp: node_def balL_def split:tree.splits)  | 
|
| 61232 | 96  | 
|
| 61581 | 97  | 
lemma inorder_balR:  | 
98  | 
"inorder (balR l a r) = inorder l @ a # inorder r"  | 
|
99  | 
by (auto simp: node_def balR_def split:tree.splits)  | 
|
| 61232 | 100  | 
|
101  | 
theorem inorder_insert:  | 
|
102  | 
"sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"  | 
|
103  | 
by (induct t)  | 
|
| 61581 | 104  | 
(auto simp: ins_list_simps inorder_balL inorder_balR)  | 
| 61232 | 105  | 
|
106  | 
||
107  | 
subsubsection "Proofs for delete"  | 
|
108  | 
||
| 68023 | 109  | 
lemma inorder_split_maxD:  | 
110  | 
"\<lbrakk> split_max t = (t',a); t \<noteq> Leaf \<rbrakk> \<Longrightarrow>  | 
|
| 61232 | 111  | 
inorder t' @ [a] = inorder t"  | 
| 68023 | 112  | 
by(induction t arbitrary: t' rule: split_max.induct)  | 
| 61647 | 113  | 
(auto simp: inorder_balL split: if_splits prod.splits tree.split)  | 
| 61232 | 114  | 
|
| 61647 | 115  | 
lemma inorder_del_root:  | 
| 68413 | 116  | 
"inorder (del_root (Node l a h r)) = inorder l @ inorder r"  | 
117  | 
by(cases "Node l a h r" rule: del_root.cases)  | 
|
| 68023 | 118  | 
(auto simp: inorder_balL inorder_balR inorder_split_maxD split: if_splits prod.splits)  | 
| 61232 | 119  | 
|
120  | 
theorem inorder_delete:  | 
|
121  | 
"sorted(inorder t) \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"  | 
|
122  | 
by(induction t)  | 
|
| 61581 | 123  | 
(auto simp: del_list_simps inorder_balL inorder_balR  | 
| 68023 | 124  | 
inorder_del_root inorder_split_maxD split: prod.splits)  | 
| 61232 | 125  | 
|
126  | 
||
| 67406 | 127  | 
subsection \<open>AVL invariants\<close>  | 
| 61232 | 128  | 
|
| 67406 | 129  | 
text\<open>Essentially the AFP/AVL proofs\<close>  | 
| 61232 | 130  | 
|
131  | 
||
| 67406 | 132  | 
subsubsection \<open>Insertion maintains AVL balance\<close>  | 
| 61232 | 133  | 
|
134  | 
declare Let_def [simp]  | 
|
135  | 
||
136  | 
lemma [simp]: "avl t \<Longrightarrow> ht t = height t"  | 
|
137  | 
by (induct t) simp_all  | 
|
138  | 
||
| 61581 | 139  | 
lemma height_balL:  | 
| 61232 | 140  | 
"\<lbrakk> height l = height r + 2; avl l; avl r \<rbrakk> \<Longrightarrow>  | 
| 61581 | 141  | 
height (balL l a r) = height r + 2 \<or>  | 
142  | 
height (balL l a r) = height r + 3"  | 
|
143  | 
by (cases l) (auto simp:node_def balL_def split:tree.split)  | 
|
| 61232 | 144  | 
|
| 61581 | 145  | 
lemma height_balR:  | 
| 61232 | 146  | 
"\<lbrakk> height r = height l + 2; avl l; avl r \<rbrakk> \<Longrightarrow>  | 
| 61581 | 147  | 
height (balR l a r) = height l + 2 \<or>  | 
148  | 
height (balR l a r) = height l + 3"  | 
|
149  | 
by (cases r) (auto simp add:node_def balR_def split:tree.split)  | 
|
| 61232 | 150  | 
|
151  | 
lemma [simp]: "height(node l a r) = max (height l) (height r) + 1"  | 
|
152  | 
by (simp add: node_def)  | 
|
153  | 
||
154  | 
lemma avl_node:  | 
|
155  | 
"\<lbrakk> avl l; avl r;  | 
|
156  | 
height l = height r \<or> height l = height r + 1 \<or> height r = height l + 1  | 
|
157  | 
\<rbrakk> \<Longrightarrow> avl(node l a r)"  | 
|
158  | 
by (auto simp add:max_def node_def)  | 
|
159  | 
||
| 61581 | 160  | 
lemma height_balL2:  | 
| 61232 | 161  | 
"\<lbrakk> avl l; avl r; height l \<noteq> height r + 2 \<rbrakk> \<Longrightarrow>  | 
| 61581 | 162  | 
height (balL l a r) = (1 + max (height l) (height r))"  | 
163  | 
by (cases l, cases r) (simp_all add: balL_def)  | 
|
| 61232 | 164  | 
|
| 61581 | 165  | 
lemma height_balR2:  | 
| 61232 | 166  | 
"\<lbrakk> avl l; avl r; height r \<noteq> height l + 2 \<rbrakk> \<Longrightarrow>  | 
| 61581 | 167  | 
height (balR l a r) = (1 + max (height l) (height r))"  | 
168  | 
by (cases l, cases r) (simp_all add: balR_def)  | 
|
| 61232 | 169  | 
|
| 61581 | 170  | 
lemma avl_balL:  | 
| 61232 | 171  | 
assumes "avl l" "avl r" and "height l = height r \<or> height l = height r + 1  | 
172  | 
\<or> height r = height l + 1 \<or> height l = height r + 2"  | 
|
| 61581 | 173  | 
shows "avl(balL l a r)"  | 
| 61232 | 174  | 
proof(cases l)  | 
175  | 
case Leaf  | 
|
| 61581 | 176  | 
with assms show ?thesis by (simp add: node_def balL_def)  | 
| 61232 | 177  | 
next  | 
| 68413 | 178  | 
case Node  | 
| 61232 | 179  | 
with assms show ?thesis  | 
180  | 
proof(cases "height l = height r + 2")  | 
|
181  | 
case True  | 
|
182  | 
from True Node assms show ?thesis  | 
|
| 61581 | 183  | 
by (auto simp: balL_def intro!: avl_node split: tree.split) arith+  | 
| 61232 | 184  | 
next  | 
185  | 
case False  | 
|
| 61581 | 186  | 
with assms show ?thesis by (simp add: avl_node balL_def)  | 
| 61232 | 187  | 
qed  | 
188  | 
qed  | 
|
189  | 
||
| 61581 | 190  | 
lemma avl_balR:  | 
| 61232 | 191  | 
assumes "avl l" and "avl r" and "height l = height r \<or> height l = height r + 1  | 
192  | 
\<or> height r = height l + 1 \<or> height r = height l + 2"  | 
|
| 61581 | 193  | 
shows "avl(balR l a r)"  | 
| 61232 | 194  | 
proof(cases r)  | 
195  | 
case Leaf  | 
|
| 61581 | 196  | 
with assms show ?thesis by (simp add: node_def balR_def)  | 
| 61232 | 197  | 
next  | 
| 68413 | 198  | 
case Node  | 
| 61232 | 199  | 
with assms show ?thesis  | 
200  | 
proof(cases "height r = height l + 2")  | 
|
201  | 
case True  | 
|
202  | 
from True Node assms show ?thesis  | 
|
| 61581 | 203  | 
by (auto simp: balR_def intro!: avl_node split: tree.split) arith+  | 
| 61232 | 204  | 
next  | 
205  | 
case False  | 
|
| 61581 | 206  | 
with assms show ?thesis by (simp add: balR_def avl_node)  | 
| 61232 | 207  | 
qed  | 
208  | 
qed  | 
|
209  | 
||
210  | 
(* It appears that these two properties need to be proved simultaneously: *)  | 
|
211  | 
||
| 67406 | 212  | 
text\<open>Insertion maintains the AVL property:\<close>  | 
| 61232 | 213  | 
|
| 68422 | 214  | 
theorem avl_insert:  | 
| 61232 | 215  | 
assumes "avl t"  | 
216  | 
shows "avl(insert x t)"  | 
|
217  | 
"(height (insert x t) = height t \<or> height (insert x t) = height t + 1)"  | 
|
218  | 
using assms  | 
|
219  | 
proof (induction t)  | 
|
| 68413 | 220  | 
case (Node l a h r)  | 
| 61232 | 221  | 
case 1  | 
| 68422 | 222  | 
show ?case  | 
| 61232 | 223  | 
proof(cases "x = a")  | 
| 68422 | 224  | 
case True with Node 1 show ?thesis by simp  | 
| 61232 | 225  | 
next  | 
226  | 
case False  | 
|
| 68422 | 227  | 
show ?thesis  | 
| 61232 | 228  | 
proof(cases "x<a")  | 
| 68422 | 229  | 
case True with Node 1 show ?thesis by (auto simp add:avl_balL)  | 
| 61232 | 230  | 
next  | 
| 68422 | 231  | 
case False with Node 1 \<open>x\<noteq>a\<close> show ?thesis by (auto simp add:avl_balR)  | 
| 61232 | 232  | 
qed  | 
233  | 
qed  | 
|
234  | 
case 2  | 
|
| 68422 | 235  | 
show ?case  | 
| 61232 | 236  | 
proof(cases "x = a")  | 
| 68422 | 237  | 
case True with Node 1 show ?thesis by simp  | 
| 61232 | 238  | 
next  | 
239  | 
case False  | 
|
| 68422 | 240  | 
show ?thesis  | 
241  | 
proof(cases "x<a")  | 
|
| 61232 | 242  | 
case True  | 
| 68422 | 243  | 
show ?thesis  | 
| 61232 | 244  | 
proof(cases "height (insert x l) = height r + 2")  | 
| 67406 | 245  | 
case False with Node 2 \<open>x < a\<close> show ?thesis by (auto simp: height_balL2)  | 
| 61232 | 246  | 
next  | 
247  | 
case True  | 
|
| 61581 | 248  | 
hence "(height (balL (insert x l) a r) = height r + 2) \<or>  | 
249  | 
(height (balL (insert x l) a r) = height r + 3)" (is "?A \<or> ?B")  | 
|
250  | 
using Node 2 by (intro height_balL) simp_all  | 
|
| 61232 | 251  | 
thus ?thesis  | 
252  | 
proof  | 
|
| 68422 | 253  | 
assume ?A with 2 \<open>x < a\<close> show ?thesis by (auto)  | 
| 61232 | 254  | 
next  | 
| 68422 | 255  | 
assume ?B with True 1 Node(2) \<open>x < a\<close> show ?thesis by (simp) arith  | 
| 61232 | 256  | 
qed  | 
257  | 
qed  | 
|
258  | 
next  | 
|
259  | 
case False  | 
|
| 68422 | 260  | 
show ?thesis  | 
| 61232 | 261  | 
proof(cases "height (insert x r) = height l + 2")  | 
| 68422 | 262  | 
case False with Node 2 \<open>\<not>x < a\<close> show ?thesis by (auto simp: height_balR2)  | 
| 61232 | 263  | 
next  | 
264  | 
case True  | 
|
| 61581 | 265  | 
hence "(height (balR l a (insert x r)) = height l + 2) \<or>  | 
266  | 
(height (balR l a (insert x r)) = height l + 3)" (is "?A \<or> ?B")  | 
|
267  | 
using Node 2 by (intro height_balR) simp_all  | 
|
| 61232 | 268  | 
thus ?thesis  | 
269  | 
proof  | 
|
| 68422 | 270  | 
assume ?A with 2 \<open>\<not>x < a\<close> show ?thesis by (auto)  | 
| 61232 | 271  | 
next  | 
| 68422 | 272  | 
assume ?B with True 1 Node(4) \<open>\<not>x < a\<close> show ?thesis by (simp) arith  | 
| 61232 | 273  | 
qed  | 
274  | 
qed  | 
|
275  | 
qed  | 
|
276  | 
qed  | 
|
277  | 
qed simp_all  | 
|
278  | 
||
279  | 
||
| 67406 | 280  | 
subsubsection \<open>Deletion maintains AVL balance\<close>  | 
| 61232 | 281  | 
|
| 68023 | 282  | 
lemma avl_split_max:  | 
| 61232 | 283  | 
assumes "avl x" and "x \<noteq> Leaf"  | 
| 68023 | 284  | 
shows "avl (fst (split_max x))" "height x = height(fst (split_max x)) \<or>  | 
285  | 
height x = height(fst (split_max x)) + 1"  | 
|
| 61232 | 286  | 
using assms  | 
| 68023 | 287  | 
proof (induct x rule: split_max_induct)  | 
| 68413 | 288  | 
case (Node l a h r)  | 
| 61232 | 289  | 
case 1  | 
| 61647 | 290  | 
thus ?case using Node  | 
| 68422 | 291  | 
by (auto simp: height_balL height_balL2 avl_balL split:prod.split)  | 
| 61232 | 292  | 
next  | 
| 68413 | 293  | 
case (Node l a h r)  | 
| 61232 | 294  | 
case 2  | 
| 68023 | 295  | 
let ?r' = "fst (split_max r)"  | 
| 67406 | 296  | 
from \<open>avl x\<close> Node 2 have "avl l" and "avl r" by simp_all  | 
| 61581 | 297  | 
thus ?case using Node 2 height_balL[of l ?r' a] height_balL2[of l ?r' a]  | 
| 61232 | 298  | 
apply (auto split:prod.splits simp del:avl.simps) by arith+  | 
299  | 
qed auto  | 
|
300  | 
||
| 61647 | 301  | 
lemma avl_del_root:  | 
| 61232 | 302  | 
assumes "avl t" and "t \<noteq> Leaf"  | 
| 61647 | 303  | 
shows "avl(del_root t)"  | 
| 61232 | 304  | 
using assms  | 
| 61647 | 305  | 
proof (cases t rule:del_root_cases)  | 
| 68413 | 306  | 
case (Node_Node ll ln lh lr n h rl rn rh rr)  | 
307  | 
let ?l = "Node ll ln lh lr"  | 
|
308  | 
let ?r = "Node rl rn rh rr"  | 
|
| 68023 | 309  | 
let ?l' = "fst (split_max ?l)"  | 
| 67406 | 310  | 
from \<open>avl t\<close> and Node_Node have "avl ?r" by simp  | 
311  | 
from \<open>avl t\<close> and Node_Node have "avl ?l" by simp  | 
|
| 61232 | 312  | 
hence "avl(?l')" "height ?l = height(?l') \<or>  | 
| 68023 | 313  | 
height ?l = height(?l') + 1" by (rule avl_split_max,simp)+  | 
| 67406 | 314  | 
with \<open>avl t\<close> Node_Node have "height ?l' = height ?r \<or> height ?l' = height ?r + 1  | 
| 61232 | 315  | 
\<or> height ?r = height ?l' + 1 \<or> height ?r = height ?l' + 2" by fastforce  | 
| 68023 | 316  | 
with \<open>avl ?l'\<close> \<open>avl ?r\<close> have "avl(balR ?l' (snd(split_max ?l)) ?r)"  | 
| 61581 | 317  | 
by (rule avl_balR)  | 
| 61232 | 318  | 
with Node_Node show ?thesis by (auto split:prod.splits)  | 
319  | 
qed simp_all  | 
|
320  | 
||
| 61647 | 321  | 
lemma height_del_root:  | 
| 61232 | 322  | 
assumes "avl t" and "t \<noteq> Leaf"  | 
| 61647 | 323  | 
shows "height t = height(del_root t) \<or> height t = height(del_root t) + 1"  | 
| 61232 | 324  | 
using assms  | 
| 61647 | 325  | 
proof (cases t rule: del_root_cases)  | 
| 68413 | 326  | 
case (Node_Node ll ln lh lr n h rl rn rh rr)  | 
327  | 
let ?l = "Node ll ln lh lr"  | 
|
328  | 
let ?r = "Node rl rn rh rr"  | 
|
| 68023 | 329  | 
let ?l' = "fst (split_max ?l)"  | 
330  | 
let ?t' = "balR ?l' (snd(split_max ?l)) ?r"  | 
|
| 67406 | 331  | 
from \<open>avl t\<close> and Node_Node have "avl ?r" by simp  | 
332  | 
from \<open>avl t\<close> and Node_Node have "avl ?l" by simp  | 
|
| 68023 | 333  | 
hence "avl(?l')" by (rule avl_split_max,simp)  | 
334  | 
have l'_height: "height ?l = height ?l' \<or> height ?l = height ?l' + 1" using \<open>avl ?l\<close> by (intro avl_split_max) auto  | 
|
| 67406 | 335  | 
have t_height: "height t = 1 + max (height ?l) (height ?r)" using \<open>avl t\<close> Node_Node by simp  | 
336  | 
have "height t = height ?t' \<or> height t = height ?t' + 1" using \<open>avl t\<close> Node_Node  | 
|
| 61232 | 337  | 
proof(cases "height ?r = height ?l' + 2")  | 
338  | 
case False  | 
|
| 68422 | 339  | 
show ?thesis using l'_height t_height False  | 
340  | 
by (subst height_balR2[OF \<open>avl ?l'\<close> \<open>avl ?r\<close> False])+ arith  | 
|
| 61232 | 341  | 
next  | 
342  | 
case True  | 
|
343  | 
show ?thesis  | 
|
| 68023 | 344  | 
proof(cases rule: disjE[OF height_balR[OF True \<open>avl ?l'\<close> \<open>avl ?r\<close>, of "snd (split_max ?l)"]])  | 
| 68422 | 345  | 
case 1 thus ?thesis using l'_height t_height True by arith  | 
| 61232 | 346  | 
next  | 
| 68422 | 347  | 
case 2 thus ?thesis using l'_height t_height True by arith  | 
| 61232 | 348  | 
qed  | 
349  | 
qed  | 
|
350  | 
thus ?thesis using Node_Node by (auto split:prod.splits)  | 
|
351  | 
qed simp_all  | 
|
352  | 
||
| 67406 | 353  | 
text\<open>Deletion maintains the AVL property:\<close>  | 
| 61232 | 354  | 
|
| 68422 | 355  | 
theorem avl_delete:  | 
| 61232 | 356  | 
assumes "avl t"  | 
357  | 
shows "avl(delete x t)" and "height t = (height (delete x t)) \<or> height t = height (delete x t) + 1"  | 
|
358  | 
using assms  | 
|
359  | 
proof (induct t)  | 
|
| 68413 | 360  | 
case (Node l n h r)  | 
| 61232 | 361  | 
case 1  | 
| 68422 | 362  | 
show ?case  | 
| 61232 | 363  | 
proof(cases "x = n")  | 
| 68422 | 364  | 
case True with Node 1 show ?thesis by (auto simp:avl_del_root)  | 
| 61232 | 365  | 
next  | 
366  | 
case False  | 
|
| 68422 | 367  | 
show ?thesis  | 
| 61232 | 368  | 
proof(cases "x<n")  | 
| 68422 | 369  | 
case True with Node 1 show ?thesis by (auto simp add:avl_balR)  | 
| 61232 | 370  | 
next  | 
| 68422 | 371  | 
case False with Node 1 \<open>x\<noteq>n\<close> show ?thesis by (auto simp add:avl_balL)  | 
| 61232 | 372  | 
qed  | 
373  | 
qed  | 
|
374  | 
case 2  | 
|
| 68422 | 375  | 
show ?case  | 
| 61232 | 376  | 
proof(cases "x = n")  | 
377  | 
case True  | 
|
| 68413 | 378  | 
with 1 have "height (Node l n h r) = height(del_root (Node l n h r))  | 
379  | 
\<or> height (Node l n h r) = height(del_root (Node l n h r)) + 1"  | 
|
| 61647 | 380  | 
by (subst height_del_root,simp_all)  | 
| 61232 | 381  | 
with True show ?thesis by simp  | 
382  | 
next  | 
|
383  | 
case False  | 
|
| 68422 | 384  | 
show ?thesis  | 
385  | 
proof(cases "x<n")  | 
|
| 61232 | 386  | 
case True  | 
387  | 
show ?thesis  | 
|
388  | 
proof(cases "height r = height (delete x l) + 2")  | 
|
| 67406 | 389  | 
case False with Node 1 \<open>x < n\<close> show ?thesis by(auto simp: balR_def)  | 
| 61232 | 390  | 
next  | 
391  | 
case True  | 
|
| 61581 | 392  | 
hence "(height (balR (delete x l) n r) = height (delete x l) + 2) \<or>  | 
393  | 
height (balR (delete x l) n r) = height (delete x l) + 3" (is "?A \<or> ?B")  | 
|
394  | 
using Node 2 by (intro height_balR) auto  | 
|
| 61232 | 395  | 
thus ?thesis  | 
396  | 
proof  | 
|
| 68422 | 397  | 
assume ?A with \<open>x < n\<close> Node 2 show ?thesis by(auto simp: balR_def)  | 
| 61232 | 398  | 
next  | 
| 68422 | 399  | 
assume ?B with \<open>x < n\<close> Node 2 show ?thesis by(auto simp: balR_def)  | 
| 61232 | 400  | 
qed  | 
401  | 
qed  | 
|
402  | 
next  | 
|
403  | 
case False  | 
|
404  | 
show ?thesis  | 
|
405  | 
proof(cases "height l = height (delete x r) + 2")  | 
|
| 67406 | 406  | 
case False with Node 1 \<open>\<not>x < n\<close> \<open>x \<noteq> n\<close> show ?thesis by(auto simp: balL_def)  | 
| 61232 | 407  | 
next  | 
408  | 
case True  | 
|
| 61581 | 409  | 
hence "(height (balL l n (delete x r)) = height (delete x r) + 2) \<or>  | 
410  | 
height (balL l n (delete x r)) = height (delete x r) + 3" (is "?A \<or> ?B")  | 
|
411  | 
using Node 2 by (intro height_balL) auto  | 
|
| 61232 | 412  | 
thus ?thesis  | 
413  | 
proof  | 
|
| 68422 | 414  | 
assume ?A with \<open>\<not>x < n\<close> \<open>x \<noteq> n\<close> Node 2 show ?thesis by(auto simp: balL_def)  | 
| 61232 | 415  | 
next  | 
| 68422 | 416  | 
assume ?B with \<open>\<not>x < n\<close> \<open>x \<noteq> n\<close> Node 2 show ?thesis by(auto simp: balL_def)  | 
| 61232 | 417  | 
qed  | 
418  | 
qed  | 
|
419  | 
qed  | 
|
420  | 
qed  | 
|
421  | 
qed simp_all  | 
|
422  | 
||
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
423  | 
|
| 68422 | 424  | 
subsection "Overall correctness"  | 
425  | 
||
| 68440 | 426  | 
interpretation S: Set_by_Ordered  | 
| 68431 | 427  | 
where empty = empty and isin = isin and insert = insert and delete = delete  | 
| 68422 | 428  | 
and inorder = inorder and inv = avl  | 
429  | 
proof (standard, goal_cases)  | 
|
| 68431 | 430  | 
case 1 show ?case by (simp add: empty_def)  | 
| 68422 | 431  | 
next  | 
432  | 
case 2 thus ?case by(simp add: isin_set_inorder)  | 
|
433  | 
next  | 
|
434  | 
case 3 thus ?case by(simp add: inorder_insert)  | 
|
435  | 
next  | 
|
436  | 
case 4 thus ?case by(simp add: inorder_delete)  | 
|
437  | 
next  | 
|
| 68431 | 438  | 
case 5 thus ?case by (simp add: empty_def)  | 
| 68422 | 439  | 
next  | 
440  | 
case 6 thus ?case by (simp add: avl_insert(1))  | 
|
441  | 
next  | 
|
442  | 
case 7 thus ?case by (simp add: avl_delete(1))  | 
|
443  | 
qed  | 
|
444  | 
||
445  | 
||
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
446  | 
subsection \<open>Height-Size Relation\<close>  | 
| 
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
447  | 
|
| 68342 | 448  | 
text \<open>Based on theorems by Daniel St\"uwe, Manuel Eberl and Peter Lammich.\<close>  | 
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
449  | 
|
| 68313 | 450  | 
lemma height_invers:  | 
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
451  | 
"(height t = 0) = (t = Leaf)"  | 
| 68413 | 452  | 
"avl t \<Longrightarrow> (height t = Suc h) = (\<exists> l a r . t = Node l a (Suc h) r)"  | 
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
453  | 
by (induction t) auto  | 
| 
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
454  | 
|
| 68313 | 455  | 
text \<open>Any AVL tree of height \<open>h\<close> has at least \<open>fib (h+2)\<close> leaves:\<close>  | 
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
456  | 
|
| 68313 | 457  | 
lemma avl_fib_bound: "avl t \<Longrightarrow> height t = h \<Longrightarrow> fib (h+2) \<le> size1 t"  | 
458  | 
proof (induction h arbitrary: t rule: fib.induct)  | 
|
459  | 
case 1 thus ?case by (simp add: height_invers)  | 
|
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
460  | 
next  | 
| 68313 | 461  | 
case 2 thus ?case by (cases t) (auto simp: height_invers)  | 
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
462  | 
next  | 
| 68313 | 463  | 
case (3 h)  | 
464  | 
from "3.prems" obtain l a r where  | 
|
| 68413 | 465  | 
[simp]: "t = Node l a (Suc(Suc h)) r" "avl l" "avl r"  | 
| 68313 | 466  | 
and C: "  | 
467  | 
height r = Suc h \<and> height l = Suc h  | 
|
468  | 
\<or> height r = Suc h \<and> height l = h  | 
|
469  | 
\<or> height r = h \<and> height l = Suc h" (is "?C1 \<or> ?C2 \<or> ?C3")  | 
|
470  | 
by (cases t) (simp, fastforce)  | 
|
471  | 
  {
 | 
|
472  | 
assume ?C1  | 
|
473  | 
with "3.IH"(1)  | 
|
474  | 
have "fib (h + 3) \<le> size1 l" "fib (h + 3) \<le> size1 r"  | 
|
475  | 
by (simp_all add: eval_nat_numeral)  | 
|
476  | 
hence ?case by (auto simp: eval_nat_numeral)  | 
|
477  | 
  } moreover {
 | 
|
478  | 
assume ?C2  | 
|
479  | 
hence ?case using "3.IH"(1)[of r] "3.IH"(2)[of l] by auto  | 
|
480  | 
  } moreover {
 | 
|
481  | 
assume ?C3  | 
|
482  | 
hence ?case using "3.IH"(1)[of l] "3.IH"(2)[of r] by auto  | 
|
483  | 
} ultimately show ?case using C by blast  | 
|
484  | 
qed  | 
|
485  | 
||
486  | 
lemma fib_alt_induct [consumes 1, case_names 1 2 rec]:  | 
|
487  | 
assumes "n > 0" "P 1" "P 2" "\<And>n. n > 0 \<Longrightarrow> P n \<Longrightarrow> P (Suc n) \<Longrightarrow> P (Suc (Suc n))"  | 
|
488  | 
shows "P n"  | 
|
489  | 
using assms(1)  | 
|
490  | 
proof (induction n rule: fib.induct)  | 
|
491  | 
case (3 n)  | 
|
492  | 
thus ?case using assms by (cases n) (auto simp: eval_nat_numeral)  | 
|
493  | 
qed (insert assms, auto)  | 
|
494  | 
||
495  | 
text \<open>An exponential lower bound for @{const fib}:\<close>
 | 
|
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
496  | 
|
| 68313 | 497  | 
lemma fib_lowerbound:  | 
498  | 
defines "\<phi> \<equiv> (1 + sqrt 5) / 2"  | 
|
499  | 
defines "c \<equiv> 1 / \<phi> ^ 2"  | 
|
500  | 
assumes "n > 0"  | 
|
501  | 
shows "real (fib n) \<ge> c * \<phi> ^ n"  | 
|
502  | 
proof -  | 
|
503  | 
have "\<phi> > 1" by (simp add: \<phi>_def)  | 
|
504  | 
hence "c > 0" by (simp add: c_def)  | 
|
505  | 
from \<open>n > 0\<close> show ?thesis  | 
|
506  | 
proof (induction n rule: fib_alt_induct)  | 
|
507  | 
case (rec n)  | 
|
508  | 
have "c * \<phi> ^ Suc (Suc n) = \<phi> ^ 2 * (c * \<phi> ^ n)"  | 
|
509  | 
by (simp add: field_simps power2_eq_square)  | 
|
510  | 
also have "\<dots> \<le> (\<phi> + 1) * (c * \<phi> ^ n)"  | 
|
511  | 
by (rule mult_right_mono) (insert \<open>c > 0\<close>, simp_all add: \<phi>_def power2_eq_square field_simps)  | 
|
512  | 
also have "\<dots> = c * \<phi> ^ Suc n + c * \<phi> ^ n"  | 
|
513  | 
by (simp add: field_simps)  | 
|
514  | 
also have "\<dots> \<le> real (fib (Suc n)) + real (fib n)"  | 
|
515  | 
by (intro add_mono rec.IH)  | 
|
516  | 
finally show ?case by simp  | 
|
517  | 
qed (insert \<open>\<phi> > 1\<close>, simp_all add: c_def power2_eq_square eval_nat_numeral)  | 
|
518  | 
qed  | 
|
519  | 
||
520  | 
text \<open>The size of an AVL tree is (at least) exponential in its height:\<close>  | 
|
521  | 
||
| 68342 | 522  | 
lemma avl_size_lowerbound:  | 
| 68313 | 523  | 
defines "\<phi> \<equiv> (1 + sqrt 5) / 2"  | 
524  | 
assumes "avl t"  | 
|
| 68342 | 525  | 
shows "\<phi> ^ (height t) \<le> size1 t"  | 
| 68313 | 526  | 
proof -  | 
527  | 
have "\<phi> > 0" by(simp add: \<phi>_def add_pos_nonneg)  | 
|
528  | 
hence "\<phi> ^ height t = (1 / \<phi> ^ 2) * \<phi> ^ (height t + 2)"  | 
|
529  | 
by(simp add: field_simps power2_eq_square)  | 
|
| 68342 | 530  | 
also have "\<dots> \<le> fib (height t + 2)"  | 
| 68313 | 531  | 
using fib_lowerbound[of "height t + 2"] by(simp add: \<phi>_def)  | 
| 68342 | 532  | 
also have "\<dots> \<le> size1 t"  | 
| 68313 | 533  | 
using avl_fib_bound[of t "height t"] assms by simp  | 
534  | 
finally show ?thesis .  | 
|
535  | 
qed  | 
|
| 
63411
 
e051eea34990
got rid of class cmp; added height-size proofs by Daniel Stuewe
 
nipkow 
parents: 
62526 
diff
changeset
 | 
536  | 
|
| 68342 | 537  | 
text \<open>The height of an AVL tree is most @{term "(1/log 2 \<phi>)"} \<open>\<approx> 1.44\<close> times worse
 | 
538  | 
than @{term "log 2 (size1 t)"}:\<close>
 | 
|
539  | 
||
540  | 
lemma avl_height_upperbound:  | 
|
541  | 
defines "\<phi> \<equiv> (1 + sqrt 5) / 2"  | 
|
542  | 
assumes "avl t"  | 
|
543  | 
shows "height t \<le> (1/log 2 \<phi>) * log 2 (size1 t)"  | 
|
544  | 
proof -  | 
|
545  | 
have "\<phi> > 0" "\<phi> > 1" by(auto simp: \<phi>_def pos_add_strict)  | 
|
546  | 
hence "height t = log \<phi> (\<phi> ^ height t)" by(simp add: log_nat_power)  | 
|
547  | 
also have "\<dots> \<le> log \<phi> (size1 t)"  | 
|
548  | 
using avl_size_lowerbound[OF assms(2), folded \<phi>_def] \<open>1 < \<phi>\<close> by simp  | 
|
549  | 
also have "\<dots> = (1/log 2 \<phi>) * log 2 (size1 t)"  | 
|
550  | 
by(simp add: log_base_change[of 2 \<phi>])  | 
|
551  | 
finally show ?thesis .  | 
|
552  | 
qed  | 
|
553  | 
||
| 61232 | 554  | 
end  |