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