src/HOL/Data_Structures/AVL_Bal2_Set.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 82308 3529946fca19
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
     1
(* Author: Tobias Nipkow *)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     2
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
     3
section "AVL Tree with Balance Factors (2)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     4
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
     5
theory AVL_Bal2_Set
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     6
imports
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     7
  Cmp
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     8
  Isin2
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     9
begin
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    10
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    11
text \<open>This version passes a flag (\<open>Same\<close>/\<open>Diff\<close>) back up to signal if the height changed.\<close>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    12
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    13
datatype bal = Lh | Bal | Rh
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    14
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    15
type_synonym 'a tree_bal = "('a * bal) tree"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    16
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    17
text \<open>Invariant:\<close>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    18
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    19
fun avl :: "'a tree_bal \<Rightarrow> bool" where
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    20
"avl Leaf = True" |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    21
"avl (Node l (a,b) r) =
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    22
  ((case b of
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    23
    Bal \<Rightarrow> height r = height l |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    24
    Lh \<Rightarrow> height l = height r + 1 |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    25
    Rh \<Rightarrow> height r = height l + 1)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    26
  \<and> avl l \<and> avl r)"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    27
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    28
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    29
subsection \<open>Code\<close>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    30
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    31
datatype 'a alt = Same 'a | Diff 'a
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    32
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    33
type_synonym 'a tree_bal2 = "'a tree_bal alt"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    34
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    35
fun tree :: "'a alt \<Rightarrow> 'a" where
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    36
"tree(Same t) = t" |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    37
"tree(Diff t) = t"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    38
71820
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    39
fun rot2 where
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    40
"rot2 A a B c C = (case B of
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    41
  (Node B\<^sub>1 (b, bb) B\<^sub>2) \<Rightarrow>
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    42
    let b\<^sub>1 = if bb = Rh then Lh else Bal;
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    43
        b\<^sub>2 = if bb = Lh then Rh else Bal
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    44
    in Node (Node A (a,b\<^sub>1) B\<^sub>1) (b,Bal) (Node B\<^sub>2 (c,b\<^sub>2) C))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    45
71818
nipkow
parents: 71816
diff changeset
    46
fun balL :: "'a tree_bal2 \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
71815
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    47
"balL AB' c bc C = (case AB' of
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    48
   Same AB \<Rightarrow> Same (Node AB (c,bc) C) |
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    49
   Diff AB \<Rightarrow> (case bc of
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    50
     Bal \<Rightarrow> Diff (Node AB (c,Lh) C) |
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    51
     Rh \<Rightarrow> Same (Node AB (c,Bal) C) |
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    52
     Lh \<Rightarrow> (case AB of
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    53
       Node A (a,Lh) B \<Rightarrow> Same(Node A (a,Bal) (Node B (c,Bal) C)) |
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    54
       Node A (a,Rh) B \<Rightarrow> Same(rot2 A a B c C))))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    55
71818
nipkow
parents: 71816
diff changeset
    56
fun balR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal2 \<Rightarrow> 'a tree_bal2" where
71815
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    57
"balR A a ba BC' = (case BC' of
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    58
   Same BC \<Rightarrow> Same (Node A (a,ba) BC) |
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    59
   Diff BC \<Rightarrow> (case ba of
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    60
     Bal \<Rightarrow> Diff (Node A (a,Rh) BC) |
a86e37f4ad60 tuned var. names
nipkow
parents: 71814
diff changeset
    61
     Lh \<Rightarrow> Same (Node A (a,Bal) BC) |
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    62
     Rh \<Rightarrow> (case BC of
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    63
       Node B (c,Rh) C \<Rightarrow> Same(Node (Node A (a,Bal) B) (c,Bal) C) |
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    64
       Node B (c,Lh) C \<Rightarrow> Same(rot2 A a B c C))))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    65
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    66
fun ins :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    67
"ins x Leaf = Diff(Node Leaf (x, Bal) Leaf)" |
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    68
"ins x (Node l (a, b) r) = (case cmp x a of
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    69
   EQ \<Rightarrow> Same(Node l (a, b) r) |
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    70
   LT \<Rightarrow> balL (ins x l) a b r |
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    71
   GT \<Rightarrow> balR l a b (ins x r))"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    72
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    73
definition insert :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    74
"insert x t = tree(ins x t)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    75
71818
nipkow
parents: 71816
diff changeset
    76
fun baldR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal2 \<Rightarrow> 'a tree_bal2" where
71816
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    77
"baldR AB c bc C' = (case C' of
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    78
   Same C \<Rightarrow> Same (Node AB (c,bc) C) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    79
   Diff C \<Rightarrow> (case bc of
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    80
     Bal \<Rightarrow> Same (Node AB (c,Lh) C) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    81
     Rh \<Rightarrow> Diff (Node AB (c,Bal) C) |
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    82
     Lh \<Rightarrow> (case AB of
71816
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    83
       Node A (a,Lh) B \<Rightarrow> Diff(Node A (a,Bal) (Node B (c,Bal) C)) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    84
       Node A (a,Bal) B \<Rightarrow> Same(Node A (a,Rh) (Node B (c,Lh) C)) |
71820
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    85
       Node A (a,Rh) B \<Rightarrow> Diff(rot2 A a B c C))))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    86
71818
nipkow
parents: 71816
diff changeset
    87
fun baldL :: "'a tree_bal2 \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
71816
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    88
"baldL A' a ba BC = (case A' of
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    89
   Same A \<Rightarrow> Same (Node A (a,ba) BC) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    90
   Diff A \<Rightarrow> (case ba of
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    91
     Bal \<Rightarrow> Same (Node A (a,Rh) BC) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    92
     Lh \<Rightarrow> Diff (Node A (a,Bal) BC) |
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    93
     Rh \<Rightarrow> (case BC of
71816
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    94
       Node B (c,Rh) C \<Rightarrow> Diff(Node (Node A (a,Bal) B) (c,Bal) C) |
489c907b9e05 tuned var. names
nipkow
parents: 71815
diff changeset
    95
       Node B (c,Bal) C \<Rightarrow> Same(Node (Node A (a,Rh) B) (c,Lh) C) |
71820
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    96
       Node B (c,Lh) C \<Rightarrow> Diff(rot2 A a B c C))))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    97
71818
nipkow
parents: 71816
diff changeset
    98
fun split_max :: "'a tree_bal \<Rightarrow> 'a tree_bal2 * 'a" where
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    99
"split_max (Node l (a, ba) r) =
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   100
  (if r = Leaf then (Diff l,a) else let (r',a') = split_max r in (baldR l a ba r', a'))"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   101
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   102
fun del :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal2" where
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   103
"del _ Leaf = Same Leaf" |
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   104
"del x (Node l (a, ba) r) =
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   105
  (case cmp x a of
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   106
     EQ \<Rightarrow> if l = Leaf then Diff r
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   107
           else let (l', a') = split_max l in baldL l' a' ba r |
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   108
     LT \<Rightarrow> baldL (del x l) a ba r |
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   109
     GT \<Rightarrow> baldR l a ba (del x r))"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   110
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   111
definition delete :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   112
"delete x t = tree(del x t)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   113
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   114
lemmas split_max_induct = split_max.induct[case_names Node Leaf]
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   115
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   116
lemmas splits = if_splits tree.splits alt.splits bal.splits
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   117
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   118
subsection \<open>Proofs\<close>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   119
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   120
subsubsection "Proofs about insertion"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   121
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   122
lemma avl_ins_case: "avl t \<Longrightarrow> case ins x t of
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   123
   Same t' \<Rightarrow> avl t' \<and> height t' = height t |
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   124
   Diff t' \<Rightarrow> avl t' \<and> height t' = height t + 1 \<and>
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   125
      (\<forall>l a r. t' = Node l (a,Bal) r \<longrightarrow> a = x \<and> l = Leaf \<and> r = Leaf)"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   126
  by (induction x t rule: ins.induct) (auto simp: max_absorb1 split!: splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   127
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   128
corollary avl_insert: "avl t \<Longrightarrow> avl(insert x t)"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   129
using avl_ins_case[of t x] by (simp add: insert_def split: splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   130
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   131
(* The following aux lemma simplifies the inorder_ins proof *)
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   132
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   133
lemma ins_Diff[simp]: "avl t \<Longrightarrow>
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   134
  ins x t \<noteq> Diff Leaf \<and>
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   135
  (ins x t = Diff (Node l (a,Bal) r) \<longleftrightarrow> t = Leaf \<and> a = x \<and> l=Leaf \<and> r=Leaf) \<and>
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   136
  ins x t \<noteq> Diff (Node l (a,Rh) Leaf) \<and>
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   137
  ins x t \<noteq> Diff (Node Leaf (a,Lh) r)"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   138
by(drule avl_ins_case[of _ x]) (auto split: splits)
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   139
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   140
theorem inorder_ins:
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   141
  "\<lbrakk> avl t;  sorted(inorder t) \<rbrakk> \<Longrightarrow> inorder(tree(ins x t)) = ins_list x (inorder t)"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   142
  by (induction t) (auto simp: ins_list_simps  split!: splits)
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   143
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   144
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   145
subsubsection "Proofs about deletion"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   146
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   147
lemma inorder_baldL:
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   148
  "\<lbrakk> ba = Rh \<longrightarrow> r \<noteq> Leaf; avl r \<rbrakk>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   149
  \<Longrightarrow> inorder (tree(baldL l a ba r)) = inorder (tree l) @ a # inorder r"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   150
by (auto split: splits)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   151
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   152
lemma inorder_baldR:
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   153
  "\<lbrakk> ba = Lh \<longrightarrow> l \<noteq> Leaf; avl l \<rbrakk>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   154
   \<Longrightarrow> inorder (tree(baldR l a ba r)) = inorder l @ a # inorder (tree r)"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   155
by (auto split: splits)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   156
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   157
lemma avl_split_max:
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   158
  "\<lbrakk> split_max t = (t',a); avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow> case t' of
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   159
   Same t' \<Rightarrow> avl t' \<and> height t = height t' |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   160
   Diff t' \<Rightarrow> avl t' \<and> height t = height t' + 1"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   161
proof (induction t arbitrary: t' a rule: split_max_induct)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   162
qed (auto simp: max_def split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   163
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   164
lemma avl_del_case: "avl t \<Longrightarrow> case del x t of
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   165
   Same t' \<Rightarrow> avl t' \<and> height t = height t' |
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   166
   Diff t' \<Rightarrow> avl t' \<and> height t = height t' + 1"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   167
proof (induction x t rule: del.induct)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   168
qed (auto simp: max_absorb1 max_absorb2 dest: avl_split_max split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   169
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   170
corollary avl_delete: "avl t \<Longrightarrow> avl(delete x t)"
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   171
using avl_del_case[of t x] by(simp add: delete_def split: splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   172
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   173
lemma inorder_split_maxD:
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   174
  "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   175
   inorder (tree t') @ [a] = inorder t"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   176
proof (induction t arbitrary: t' rule: split_max.induct)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   177
qed (auto  split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   178
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   179
lemma neq_Leaf_if_height_neq_0[simp]: "height t \<noteq> 0 \<Longrightarrow> t \<noteq> Leaf"
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   180
by auto
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   181
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   182
theorem inorder_del:
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   183
  "\<lbrakk> avl t; sorted(inorder t) \<rbrakk>  \<Longrightarrow> inorder (tree(del x t)) = del_list x (inorder t)"
82308
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   184
proof (induction t rule: tree2_induct)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   185
  case Leaf
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   186
  then show ?case by simp
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   187
next
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   188
  case (Node x1 a b x3)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   189
  then show ?case 
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   190
    by (auto simp: del_list_simps inorder_baldL inorder_baldR avl_delete inorder_split_maxD
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   191
           simp del:  baldL.simps split!: splits prod.splits)
3529946fca19 tidied old proofs
paulson <lp15@cam.ac.uk>
parents: 71844
diff changeset
   192
qed
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   193
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   194
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   195
subsubsection \<open>Set Implementation\<close>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   196
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   197
interpretation S: Set_by_Ordered
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   198
where empty = Leaf and isin = isin
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   199
  and insert = insert
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   200
  and delete = delete
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   201
  and inorder = inorder and inv = avl
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   202
proof (standard, goal_cases)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   203
  case 1 show ?case by (simp)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   204
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   205
  case 2 thus ?case by(simp add: isin_set_inorder)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   206
next
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   207
  case 3 thus ?case by(simp add: inorder_ins insert_def)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   208
next
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   209
  case 4 thus ?case by(simp add: inorder_del delete_def)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   210
next
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   211
  case 5 thus ?case by (simp)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   212
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   213
  case 6 thus ?case by (simp add: avl_insert)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   214
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   215
  case 7 thus ?case by (simp add: avl_delete)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   216
qed
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   217
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   218
end