src/HOL/Data_Structures/AVL_Bal_Set.thy
author nipkow
Fri, 11 Dec 2020 17:29:42 +0100
changeset 72883 4e6dc2868d5f
parent 71989 bad75618fb82
permissions -rw-r--r--
tuned
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 (1)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     4
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
     5
theory AVL_Bal_Set
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 detects height increase/decrease from above via the change in balance factors.\<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
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    31
fun is_bal where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    32
"is_bal (Node l (a,b) r) = (b = Bal)"
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    33
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    34
fun incr where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    35
"incr t t' = (t = Leaf \<or> is_bal t \<and> \<not> is_bal t')"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    36
71820
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    37
fun rot2 where
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    38
"rot2 A a B c C = (case B of
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    39
  (Node B\<^sub>1 (b, bb) B\<^sub>2) \<Rightarrow>
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    40
    let b\<^sub>1 = if bb = Rh then Lh else Bal;
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    41
        b\<^sub>2 = if bb = Lh then Rh else Bal
dd5b8072ca90 simplified and tuned
nipkow
parents: 71819
diff changeset
    42
    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
    43
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    44
fun balL :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    45
"balL AB c bc C = (case bc of
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    46
     Bal \<Rightarrow> Node AB (c,Lh) C |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    47
     Rh \<Rightarrow> Node AB (c,Bal) C |
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    48
     Lh \<Rightarrow> (case AB of
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    49
       Node A (a,Lh) B \<Rightarrow> Node A (a,Bal) (Node B (c,Bal) C) |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    50
       Node A (a,Bal) B \<Rightarrow> Node A (a,Rh) (Node B (c,Lh) C) |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    51
       Node A (a,Rh) B \<Rightarrow> rot2 A a B c C))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    52
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    53
fun balR :: "'a tree_bal \<Rightarrow> 'a \<Rightarrow> bal \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    54
"balR A a ba BC = (case ba of
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    55
     Bal \<Rightarrow> Node A (a,Rh) BC |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    56
     Lh \<Rightarrow> Node A (a,Bal) BC |
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
    57
     Rh \<Rightarrow> (case BC of
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    58
       Node B (c,Rh) C \<Rightarrow> Node (Node A (a,Bal) B) (c,Bal) C |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    59
       Node B (c,Bal) C \<Rightarrow> Node (Node A (a,Rh) B) (c,Lh) C |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    60
       Node B (c,Lh) C \<Rightarrow> rot2 A a B c C))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    61
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    62
fun insert :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    63
"insert x Leaf = Node Leaf (x, Bal) Leaf" |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    64
"insert x (Node l (a, b) r) = (case cmp x a of
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    65
   EQ \<Rightarrow> Node l (a, b) r |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    66
   LT \<Rightarrow> let l' = insert x l in if incr l l' then balL l' a b r else Node l' (a,b) r |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    67
   GT \<Rightarrow> let r' = insert x r in if incr r r' then balR l a b r' else Node l (a,b) r')"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    68
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    69
fun decr where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    70
"decr t t' = (t \<noteq> Leaf \<and> (t' = Leaf \<or> \<not> is_bal t \<and> is_bal t'))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    71
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    72
fun split_max :: "'a tree_bal \<Rightarrow> 'a tree_bal * 'a" where
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    73
"split_max (Node l (a, ba) r) =
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    74
  (if r = Leaf then (l,a)
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    75
   else let (r',a') = split_max r;
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    76
            t' = if decr r r' then balL l a ba r' else Node l (a,ba) r'
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    77
        in (t', a'))"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    78
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    79
fun delete :: "'a::linorder \<Rightarrow> 'a tree_bal \<Rightarrow> 'a tree_bal" where
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    80
"delete _ Leaf = Leaf" |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    81
"delete x (Node l (a, ba) r) =
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    82
  (case cmp x a of
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    83
     EQ \<Rightarrow> if l = Leaf then r
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    84
           else let (l', a') = split_max l in
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    85
                if decr l l' then balR l' a' ba r else Node l' (a',ba) r |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    86
     LT \<Rightarrow> let l' = delete x l in if decr l l' then balR l' a ba r else Node l' (a,ba) r |
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    87
     GT \<Rightarrow> let r' = delete x r in if decr r r' then balL l a ba r' else Node l (a,ba) r')"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    88
71846
nipkow
parents: 71844
diff changeset
    89
nipkow
parents: 71844
diff changeset
    90
subsection \<open>Proofs\<close>
nipkow
parents: 71844
diff changeset
    91
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    92
lemmas split_max_induct = split_max.induct[case_names Node Leaf]
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    93
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
    94
lemmas splits = if_splits tree.splits bal.splits
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    95
71846
nipkow
parents: 71844
diff changeset
    96
declare Let_def [simp]
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    97
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
    98
subsubsection "Proofs about insertion"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
    99
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   100
lemma avl_insert: "avl t \<Longrightarrow>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   101
  avl(insert x t) \<and>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   102
  height(insert x t) = height t + (if incr t (insert x t) then 1 else 0)"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   103
apply(induction x t rule: insert.induct)
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   104
apply(auto split!: splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   105
done
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   106
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   107
text \<open>The following two auxiliary lemma merely simplify the proof of \<open>inorder_insert\<close>.\<close>
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   108
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   109
lemma [simp]: "[] \<noteq> ins_list x xs"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   110
by(cases xs) auto
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   111
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   112
lemma [simp]: "avl t \<Longrightarrow> insert x t \<noteq> \<langle>l, (a, Rh), \<langle>\<rangle>\<rangle> \<and> insert x t \<noteq> \<langle>\<langle>\<rangle>, (a, Lh), r\<rangle>"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   113
by(drule avl_insert[of _ x]) (auto split: splits)
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   114
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   115
theorem inorder_insert:
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   116
  "\<lbrakk> avl t;  sorted(inorder t) \<rbrakk> \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   117
apply(induction t)
71846
nipkow
parents: 71844
diff changeset
   118
apply (auto simp: ins_list_simps split!: splits)
71824
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   119
done
95d2d5e60419 avoid hidden undef cases
nipkow
parents: 71820
diff changeset
   120
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
subsubsection "Proofs about deletion"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   123
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   124
lemma inorder_balR:
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   125
  "\<lbrakk> ba = Rh \<longrightarrow> r \<noteq> Leaf; avl r \<rbrakk>
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   126
  \<Longrightarrow> inorder (balR l a ba r) = inorder l @ a # inorder r"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   127
by (auto split: splits)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   128
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   129
lemma inorder_balL:
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   130
  "\<lbrakk> ba = Lh \<longrightarrow> l \<noteq> Leaf; avl l \<rbrakk>
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   131
   \<Longrightarrow> inorder (balL l a ba r) = inorder l @ a # inorder r"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   132
by (auto split: splits)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   133
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   134
lemma height_1_iff: "avl t \<Longrightarrow> height t = Suc 0 \<longleftrightarrow> (\<exists>x. t = Node Leaf (x,Bal) Leaf)"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   135
by(cases t) (auto split: splits prod.splits)
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   136
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   137
lemma avl_split_max:
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   138
  "\<lbrakk> split_max t = (t',a); avl t; t \<noteq> Leaf \<rbrakk> \<Longrightarrow>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   139
   avl t' \<and> height t = height t' + (if decr t t' then 1 else 0)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   140
apply(induction t arbitrary: t' a rule: split_max_induct)
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   141
 apply(auto simp: max_absorb1 max_absorb2 height_1_iff split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   142
done
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   143
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   144
lemma avl_delete: "avl t \<Longrightarrow>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   145
  avl (delete x t) \<and>
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   146
  height t = height (delete x t) + (if decr t (delete x t) then 1 else 0)"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   147
apply(induction x t rule: delete.induct)
71846
nipkow
parents: 71844
diff changeset
   148
 apply(auto simp: max_absorb1 max_absorb2 height_1_iff dest: avl_split_max split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   149
done
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   150
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   151
lemma inorder_split_maxD:
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   152
  "\<lbrakk> split_max t = (t',a); t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow>
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   153
   inorder t' @ [a] = inorder t"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   154
apply(induction t arbitrary: t' rule: split_max.induct)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   155
 apply(fastforce split!: splits prod.splits)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   156
apply simp
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   157
done
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   158
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   159
lemma neq_Leaf_if_height_neq_0: "height t \<noteq> 0 \<Longrightarrow> t \<noteq> Leaf"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   160
by auto
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   161
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   162
lemma split_max_Leaf: "\<lbrakk> t \<noteq> Leaf; avl t \<rbrakk> \<Longrightarrow> split_max t = (\<langle>\<rangle>, x) \<longleftrightarrow> t = Node Leaf (x,Bal) Leaf"
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   163
by(cases t) (auto split: splits prod.splits)
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   164
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   165
theorem inorder_delete:
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   166
  "\<lbrakk> avl t; sorted(inorder t) \<rbrakk>  \<Longrightarrow> inorder (delete x t) = del_list x (inorder t)"
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   167
apply(induction t rule: tree2_induct)
71846
nipkow
parents: 71844
diff changeset
   168
apply(auto simp: del_list_simps inorder_balR inorder_balL avl_delete inorder_split_maxD
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   169
                 split_max_Leaf neq_Leaf_if_height_neq_0
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   170
           simp del: balL.simps balR.simps split!: splits prod.splits)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   171
done
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   172
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   173
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   174
subsubsection \<open>Set Implementation\<close>
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   175
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   176
interpretation S: Set_by_Ordered
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   177
where empty = Leaf and isin = isin
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   178
  and insert = insert
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   179
  and delete = delete
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   180
  and inorder = inorder and inv = avl
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   181
proof (standard, goal_cases)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   182
  case 1 show ?case by (simp)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   183
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   184
  case 2 thus ?case by(simp add: isin_set_inorder)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   185
next
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   186
  case 3 thus ?case by(simp add: inorder_insert)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   187
next
71844
57ace76cbffa another AVL tree version
nipkow
parents: 71828
diff changeset
   188
  case 4 thus ?case by(simp add: inorder_delete)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   189
next
71828
415c38ef38c0 added top-level functions and tuned
nipkow
parents: 71824
diff changeset
   190
  case 5 thus ?case by (simp)
71814
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   191
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   192
  case 6 thus ?case by (simp add: avl_insert)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   193
next
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   194
  case 7 thus ?case by (simp add: avl_delete)
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   195
qed
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   196
a9df6686ed0e AVL trees with balance tags
nipkow
parents:
diff changeset
   197
end