src/HOL/ex/Tree23.thy
author huffman
Fri, 04 Nov 2011 07:04:34 +0100
changeset 45333 04b21922ed68
parent 45332 ede9dc025150
child 45334 3f74e041e05c
permissions -rw-r--r--
ex/Tree23.thy: simpler definition of ordered-ness predicate
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33436
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/ex/Tree23.thy
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow, TU Muenchen
nipkow
parents:
diff changeset
     3
*)
nipkow
parents:
diff changeset
     4
nipkow
parents:
diff changeset
     5
header {* 2-3 Trees *}
nipkow
parents:
diff changeset
     6
nipkow
parents:
diff changeset
     7
theory Tree23
nipkow
parents:
diff changeset
     8
imports Main
nipkow
parents:
diff changeset
     9
begin
nipkow
parents:
diff changeset
    10
nipkow
parents:
diff changeset
    11
text{* This is a very direct translation of some of the functions in table.ML
nipkow
parents:
diff changeset
    12
in the Isabelle source code. That source is due to Makarius Wenzel and Stefan
nipkow
parents:
diff changeset
    13
Berghofer.
nipkow
parents:
diff changeset
    14
nipkow
parents:
diff changeset
    15
So far this file contains only data types and functions, but no proofs. Feel
nipkow
parents:
diff changeset
    16
free to have a go at the latter!
nipkow
parents:
diff changeset
    17
nipkow
parents:
diff changeset
    18
Note that because of complicated patterns and mutual recursion, these
nipkow
parents:
diff changeset
    19
function definitions take a few minutes and can also be seen as stress tests
nipkow
parents:
diff changeset
    20
for the function definition facility.  *}
nipkow
parents:
diff changeset
    21
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 33436
diff changeset
    22
type_synonym key = int -- {*for simplicity, should be a type class*}
33436
nipkow
parents:
diff changeset
    23
nipkow
parents:
diff changeset
    24
datatype ord = LESS | EQUAL | GREATER
nipkow
parents:
diff changeset
    25
nipkow
parents:
diff changeset
    26
definition "ord i j = (if i<j then LESS else if i=j then EQUAL else GREATER)"
nipkow
parents:
diff changeset
    27
nipkow
parents:
diff changeset
    28
datatype 'a tree23 =
nipkow
parents:
diff changeset
    29
  Empty |
nipkow
parents:
diff changeset
    30
  Branch2 "'a tree23" "key * 'a" "'a tree23" |
nipkow
parents:
diff changeset
    31
  Branch3 "'a tree23" "key * 'a" "'a tree23" "key * 'a" "'a tree23"
nipkow
parents:
diff changeset
    32
nipkow
parents:
diff changeset
    33
datatype 'a growth =
nipkow
parents:
diff changeset
    34
  Stay "'a tree23" |
nipkow
parents:
diff changeset
    35
  Sprout "'a tree23" "key * 'a" "'a tree23"
nipkow
parents:
diff changeset
    36
nipkow
parents:
diff changeset
    37
fun add :: "key \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a growth" where
nipkow
parents:
diff changeset
    38
"add key y Empty = Sprout Empty (key,y) Empty" |
nipkow
parents:
diff changeset
    39
"add key y (Branch2 left (k,x) right) =
nipkow
parents:
diff changeset
    40
   (case ord key k of
nipkow
parents:
diff changeset
    41
      LESS =>
nipkow
parents:
diff changeset
    42
        (case add key y left of
nipkow
parents:
diff changeset
    43
           Stay left' => Stay (Branch2 left' (k,x) right)
nipkow
parents:
diff changeset
    44
         | Sprout left1 q left2
nipkow
parents:
diff changeset
    45
           => Stay (Branch3 left1 q left2 (k,x) right))
nipkow
parents:
diff changeset
    46
    | EQUAL => Stay (Branch2 left (k,y) right)
nipkow
parents:
diff changeset
    47
    | GREATER =>
nipkow
parents:
diff changeset
    48
        (case add key y right of
nipkow
parents:
diff changeset
    49
           Stay right' => Stay (Branch2 left (k,x) right')
nipkow
parents:
diff changeset
    50
         | Sprout right1 q right2
nipkow
parents:
diff changeset
    51
           => Stay (Branch3 left (k,x) right1 q right2)))" |
nipkow
parents:
diff changeset
    52
"add key y (Branch3 left (k1,x1) mid (k2,x2) right) =
nipkow
parents:
diff changeset
    53
   (case ord key k1 of
nipkow
parents:
diff changeset
    54
      LESS =>
nipkow
parents:
diff changeset
    55
        (case add key y left of
nipkow
parents:
diff changeset
    56
           Stay left' => Stay (Branch3 left' (k1,x1) mid (k2,x2) right)
nipkow
parents:
diff changeset
    57
         | Sprout left1 q left2
nipkow
parents:
diff changeset
    58
           => Sprout (Branch2 left1 q left2) (k1,x1) (Branch2 mid (k2,x2) right))
nipkow
parents:
diff changeset
    59
    | EQUAL => Stay (Branch3 left (k1,y) mid (k2,x2) right)
nipkow
parents:
diff changeset
    60
    | GREATER =>
nipkow
parents:
diff changeset
    61
        (case ord key k2 of
nipkow
parents:
diff changeset
    62
           LESS =>
nipkow
parents:
diff changeset
    63
             (case add key y mid of
nipkow
parents:
diff changeset
    64
                Stay mid' => Stay (Branch3 left (k1,x1) mid' (k2,x2) right)
nipkow
parents:
diff changeset
    65
              | Sprout mid1 q mid2
nipkow
parents:
diff changeset
    66
                => Sprout (Branch2 left (k1,x1) mid1) q (Branch2 mid2 (k2,x2) right))
nipkow
parents:
diff changeset
    67
         | EQUAL => Stay (Branch3 left (k1,x1) mid (k2,y) right)
nipkow
parents:
diff changeset
    68
         | GREATER =>
nipkow
parents:
diff changeset
    69
             (case add key y right of
nipkow
parents:
diff changeset
    70
                Stay right' => Stay (Branch3 left (k1,x1) mid (k2,x2) right')
nipkow
parents:
diff changeset
    71
              | Sprout right1 q right2
nipkow
parents:
diff changeset
    72
                => Sprout (Branch2 left (k1,x1) mid) (k2,x2) (Branch2 right1 q right2))))"
nipkow
parents:
diff changeset
    73
nipkow
parents:
diff changeset
    74
definition add0 :: "key \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
nipkow
parents:
diff changeset
    75
"add0 k y t =
nipkow
parents:
diff changeset
    76
  (case add k y t of Stay t' => t' | Sprout l p r => Branch2 l p r)"
nipkow
parents:
diff changeset
    77
nipkow
parents:
diff changeset
    78
value "add0 5 e (add0 4 d (add0 3 c (add0 2 b (add0 1 a Empty))))"
nipkow
parents:
diff changeset
    79
nipkow
parents:
diff changeset
    80
fun compare where
nipkow
parents:
diff changeset
    81
"compare None (k2, _) = LESS" |
nipkow
parents:
diff changeset
    82
"compare (Some k1) (k2, _) = ord k1 k2"
nipkow
parents:
diff changeset
    83
nipkow
parents:
diff changeset
    84
fun if_eq where
nipkow
parents:
diff changeset
    85
"if_eq EQUAL x y = x" |
nipkow
parents:
diff changeset
    86
"if_eq _ x y = y"
nipkow
parents:
diff changeset
    87
nipkow
parents:
diff changeset
    88
fun del :: "key option \<Rightarrow> 'a tree23 \<Rightarrow> ((key * 'a) * bool * 'a tree23)option"
nipkow
parents:
diff changeset
    89
where
nipkow
parents:
diff changeset
    90
"del (Some k) Empty = None" |
nipkow
parents:
diff changeset
    91
"del None (Branch2 Empty p Empty) = Some(p, (True, Empty))" |
nipkow
parents:
diff changeset
    92
"del None (Branch3 Empty p Empty q Empty) = Some(p, (False, Branch2 Empty q Empty))" |
nipkow
parents:
diff changeset
    93
"del k (Branch2 Empty p Empty) = (case compare k p of
nipkow
parents:
diff changeset
    94
      EQUAL => Some(p, (True, Empty)) | _ => None)" |
nipkow
parents:
diff changeset
    95
"del k (Branch3 Empty p Empty q Empty) = (case compare k p of
nipkow
parents:
diff changeset
    96
      EQUAL => Some(p, (False, Branch2 Empty q Empty))
nipkow
parents:
diff changeset
    97
    | _ => (case compare k q of
nipkow
parents:
diff changeset
    98
        EQUAL => Some(q, (False, Branch2 Empty p Empty))
nipkow
parents:
diff changeset
    99
      | _ => None))" |
nipkow
parents:
diff changeset
   100
"del k (Branch2 l p r) = (case compare k p of
nipkow
parents:
diff changeset
   101
      LESS => (case del k l of None \<Rightarrow> None |
nipkow
parents:
diff changeset
   102
        Some(p', (False, l')) => Some(p', (False, Branch2 l' p r))
nipkow
parents:
diff changeset
   103
      | Some(p', (True, l')) => Some(p', case r of
nipkow
parents:
diff changeset
   104
          Branch2 rl rp rr => (True, Branch3 l' p rl rp rr)
nipkow
parents:
diff changeset
   105
        | Branch3 rl rp rm rq rr => (False, Branch2
nipkow
parents:
diff changeset
   106
            (Branch2 l' p rl) rp (Branch2 rm rq rr))))
nipkow
parents:
diff changeset
   107
    | or => (case del (if_eq or None k) r of None \<Rightarrow> None |
nipkow
parents:
diff changeset
   108
        Some(p', (False, r')) => Some(p', (False, Branch2 l (if_eq or p' p) r'))
nipkow
parents:
diff changeset
   109
      | Some(p', (True, r')) => Some(p', case l of
nipkow
parents:
diff changeset
   110
          Branch2 ll lp lr => (True, Branch3 ll lp lr (if_eq or p' p) r')
nipkow
parents:
diff changeset
   111
        | Branch3 ll lp lm lq lr => (False, Branch2
nipkow
parents:
diff changeset
   112
            (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) r')))))" |
nipkow
parents:
diff changeset
   113
"del k (Branch3 l p m q r) = (case compare k q of
nipkow
parents:
diff changeset
   114
      LESS => (case compare k p of
nipkow
parents:
diff changeset
   115
        LESS => (case del k l of None \<Rightarrow> None |
nipkow
parents:
diff changeset
   116
          Some(p', (False, l')) => Some(p', (False, Branch3 l' p m q r))
nipkow
parents:
diff changeset
   117
        | Some(p', (True, l')) => Some(p', (False, case (m, r) of
nipkow
parents:
diff changeset
   118
            (Branch2 ml mp mr, Branch2 _ _ _) => Branch2 (Branch3 l' p ml mp mr) q r
nipkow
parents:
diff changeset
   119
          | (Branch3 ml mp mm mq mr, _) => Branch3 (Branch2 l' p ml) mp (Branch2 mm mq mr) q r
nipkow
parents:
diff changeset
   120
          | (Branch2 ml mp mr, Branch3 rl rp rm rq rr) =>
nipkow
parents:
diff changeset
   121
              Branch3 (Branch2 l' p ml) mp (Branch2 mr q rl) rp (Branch2 rm rq rr))))
nipkow
parents:
diff changeset
   122
      | or => (case del (if_eq or None k) m of None \<Rightarrow> None |
nipkow
parents:
diff changeset
   123
          Some(p', (False, m')) => Some(p', (False, Branch3 l (if_eq or p' p) m' q r))
nipkow
parents:
diff changeset
   124
        | Some(p', (True, m')) => Some(p', (False, case (l, r) of
nipkow
parents:
diff changeset
   125
            (Branch2 ll lp lr, Branch2 _ _ _) => Branch2 (Branch3 ll lp lr (if_eq or p' p) m') q r
nipkow
parents:
diff changeset
   126
          | (Branch3 ll lp lm lq lr, _) => Branch3 (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) m') q r
nipkow
parents:
diff changeset
   127
          | (_, Branch3 rl rp rm rq rr) => Branch3 l (if_eq or p' p) (Branch2 m' q rl) rp (Branch2 rm rq rr)))))
nipkow
parents:
diff changeset
   128
    | or => (case del (if_eq or None k) r of None \<Rightarrow> None |
nipkow
parents:
diff changeset
   129
        Some(q', (False, r')) => Some(q', (False, Branch3 l p m (if_eq or q' q) r'))
nipkow
parents:
diff changeset
   130
      | Some(q', (True, r')) => Some(q', (False, case (l, m) of
nipkow
parents:
diff changeset
   131
          (Branch2 _ _ _, Branch2 ml mp mr) => Branch2 l p (Branch3 ml mp mr (if_eq or q' q) r')
nipkow
parents:
diff changeset
   132
        | (_, Branch3 ml mp mm mq mr) => Branch3 l p (Branch2 ml mp mm) mq (Branch2 mr (if_eq or q' q) r')
nipkow
parents:
diff changeset
   133
        | (Branch3 ll lp lm lq lr, Branch2 ml mp mr) =>
nipkow
parents:
diff changeset
   134
            Branch3 (Branch2 ll lp lm) lq (Branch2 lr p ml) mp (Branch2 mr (if_eq or q' q) r')))))"
nipkow
parents:
diff changeset
   135
nipkow
parents:
diff changeset
   136
definition del0 :: "key \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
nipkow
parents:
diff changeset
   137
"del0 k t = (case del (Some k) t of None \<Rightarrow> t | Some(_,(_,t')) \<Rightarrow> t')"
nipkow
parents:
diff changeset
   138
nipkow
parents:
diff changeset
   139
45333
04b21922ed68 ex/Tree23.thy: simpler definition of ordered-ness predicate
huffman
parents: 45332
diff changeset
   140
text {* Specifying correctness *}
33436
nipkow
parents:
diff changeset
   141
45325
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   142
definition opt_less :: "key option \<Rightarrow> key option \<Rightarrow> bool" where
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   143
  "opt_less x y = (case x of None \<Rightarrow> True | Some a \<Rightarrow> (case y of None \<Rightarrow> True | Some b \<Rightarrow> a < b))"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   144
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   145
lemma opt_less_simps [simp]:
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   146
  "opt_less None y = True"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   147
  "opt_less x None = True"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   148
  "opt_less (Some a) (Some b) = (a < b)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   149
unfolding opt_less_def by (auto simp add: ord_def split: option.split)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   150
45333
04b21922ed68 ex/Tree23.thy: simpler definition of ordered-ness predicate
huffman
parents: 45332
diff changeset
   151
primrec ord' :: "key option \<Rightarrow> 'a tree23 \<Rightarrow> key option \<Rightarrow> bool" where
45325
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   152
"ord' x Empty y = opt_less x y" |
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   153
"ord' x (Branch2 l p r) y = (ord' x l (Some (fst p)) & ord' (Some (fst p)) r y)" |
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   154
"ord' x (Branch3 l p m q r) y = (ord' x l (Some (fst p)) & ord' (Some (fst p)) m (Some (fst q)) & ord' (Some (fst q)) r y)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   155
45333
04b21922ed68 ex/Tree23.thy: simpler definition of ordered-ness predicate
huffman
parents: 45332
diff changeset
   156
definition ord0 :: "'a tree23 \<Rightarrow> bool" where
04b21922ed68 ex/Tree23.thy: simpler definition of ordered-ness predicate
huffman
parents: 45332
diff changeset
   157
  "ord0 t = ord' None t None"
45325
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   158
33436
nipkow
parents:
diff changeset
   159
fun height :: "'a tree23 \<Rightarrow> nat" where
nipkow
parents:
diff changeset
   160
"height Empty = 0" |
nipkow
parents:
diff changeset
   161
"height (Branch2 l _ r) = Suc(max (height l) (height r))" |
nipkow
parents:
diff changeset
   162
"height (Branch3 l _ m _ r) = Suc(max (height l) (max (height m) (height r)))"
nipkow
parents:
diff changeset
   163
nipkow
parents:
diff changeset
   164
text{* Is a tree balanced? *}
nipkow
parents:
diff changeset
   165
fun bal :: "'a tree23 \<Rightarrow> bool" where
nipkow
parents:
diff changeset
   166
"bal Empty = True" |
nipkow
parents:
diff changeset
   167
"bal (Branch2 l _ r) = (bal l & bal r & height l = height r)" |
nipkow
parents:
diff changeset
   168
"bal (Branch3 l _ m _ r) = (bal l & bal m & bal r & height l = height m & height m = height r)"
nipkow
parents:
diff changeset
   169
45325
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   170
text {* The @{term "add0"} function either preserves the height of the
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   171
tree, or increases it by one. The constructor returned by the @{term
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   172
"add"} function determines which: A return value of the form @{term
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   173
"Stay t"} indicates that the height will be the same. A value of the
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   174
form @{term "Sprout l p r"} indicates an increase in height. *}
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   175
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   176
fun gheight :: "'a growth \<Rightarrow> nat" where
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   177
"gheight (Stay t) = height t" |
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   178
"gheight (Sprout l p r) = max (height l) (height r)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   179
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   180
lemma gheight_add: "gheight (add k y t) = height t"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   181
 apply (induct t)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   182
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   183
  apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   184
  apply (case_tac "ord k a")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   185
    apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   186
    apply (case_tac "add k y t1", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   187
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   188
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   189
  apply (case_tac "add k y t2", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   190
 apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   191
 apply (case_tac "ord k a")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   192
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   193
   apply (case_tac "add k y t1", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   194
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   195
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   196
 apply (case_tac "ord k aa")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   197
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   198
   apply (case_tac "add k y t2", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   199
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   200
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   201
 apply (case_tac "add k y t3", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   202
done
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   203
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   204
lemma add_eq_Stay_dest: "add k y t = Stay t' \<Longrightarrow> height t = height t'"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   205
  using gheight_add [of k y t] by simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   206
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   207
lemma add_eq_Sprout_dest: "add k y t = Sprout l p r \<Longrightarrow> height t = max (height l) (height r)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   208
  using gheight_add [of k y t] by simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   209
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   210
definition gtree :: "'a growth \<Rightarrow> 'a tree23" where
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   211
  "gtree g = (case g of Stay t \<Rightarrow> t | Sprout l p r \<Rightarrow> Branch2 l p r)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   212
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   213
lemma gtree_simps [simp]:
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   214
  "gtree (Stay t) = t"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   215
  "gtree (Sprout l p r) = Branch2 l p r"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   216
unfolding gtree_def by simp_all
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   217
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   218
lemma add0: "add0 k y t = gtree (add k y t)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   219
  unfolding add0_def by (simp split: growth.split)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   220
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   221
text {* The @{term "add0"} operation preserves balance. *}
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   222
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   223
lemma bal_add0: "bal t \<Longrightarrow> bal (add0 k y t)"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   224
unfolding add0
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   225
 apply (induct t)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   226
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   227
  apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   228
  apply (case_tac "ord k a")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   229
    apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   230
    apply (case_tac "add k y t1")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   231
     apply (simp, drule add_eq_Stay_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   232
    apply (simp, drule add_eq_Sprout_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   233
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   234
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   235
  apply (case_tac "add k y t2")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   236
   apply (simp, drule add_eq_Stay_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   237
  apply (simp, drule add_eq_Sprout_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   238
 apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   239
 apply (case_tac "ord k a")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   240
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   241
   apply (case_tac "add k y t1")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   242
    apply (simp, drule add_eq_Stay_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   243
   apply (simp, drule add_eq_Sprout_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   244
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   245
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   246
 apply (case_tac "ord k aa")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   247
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   248
   apply (case_tac "add k y t2")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   249
    apply (simp, drule add_eq_Stay_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   250
   apply (simp, drule add_eq_Sprout_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   251
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   252
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   253
 apply (case_tac "add k y t3")
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   254
  apply (simp, drule add_eq_Stay_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   255
 apply (simp, drule add_eq_Sprout_dest, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   256
done
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   257
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   258
text {* The @{term "add0"} operation preserves order. *}
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   259
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   260
lemma ord_cases:
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   261
  fixes a b :: int obtains
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   262
  "ord a b = LESS" and "a < b" |
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   263
  "ord a b = EQUAL" and "a = b" |
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   264
  "ord a b = GREATER" and "a > b"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   265
unfolding ord_def by (rule linorder_cases [of a b]) auto
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   266
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   267
lemma ord'_add0:
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   268
  "\<lbrakk>ord' k1 t k2; opt_less k1 (Some k); opt_less (Some k) k2\<rbrakk> \<Longrightarrow> ord' k1 (add0 k y t) k2"
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   269
unfolding add0
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   270
 apply (induct t arbitrary: k1 k2)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   271
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   272
  apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   273
  apply (rule_tac a=k and b=a in ord_cases)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   274
    apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   275
    apply (case_tac "add k y t1", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   276
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   277
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   278
  apply (case_tac "add k y t2", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   279
 apply clarsimp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   280
 apply (rule_tac a=k and b=a in ord_cases)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   281
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   282
   apply (case_tac "add k y t1", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   283
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   284
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   285
 apply (rule_tac a=k and b=aa in ord_cases)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   286
   apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   287
   apply (case_tac "add k y t2", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   288
  apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   289
 apply simp
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   290
 apply (case_tac "add k y t3", simp, simp)
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   291
done
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   292
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   293
lemma ord0_add0: "ord0 t \<Longrightarrow> ord0 (add0 k y t)"
45333
04b21922ed68 ex/Tree23.thy: simpler definition of ordered-ness predicate
huffman
parents: 45332
diff changeset
   294
  by (simp add: ord0_def ord'_add0)
45325
26b6179b5a45 ex/Tree23.thy: prove that insertion preserves tree balance and order
huffman
parents: 42463
diff changeset
   295
45332
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   296
text {* The @{term "del"} function preserves balance. *}
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   297
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   298
lemma del_extra_simps:
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   299
"l \<noteq> Empty \<or> r \<noteq> Empty \<Longrightarrow>
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   300
 del k (Branch2 l p r) = (case compare k p of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   301
      LESS => (case del k l of None \<Rightarrow> None |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   302
        Some(p', (False, l')) => Some(p', (False, Branch2 l' p r))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   303
      | Some(p', (True, l')) => Some(p', case r of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   304
          Branch2 rl rp rr => (True, Branch3 l' p rl rp rr)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   305
        | Branch3 rl rp rm rq rr => (False, Branch2
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   306
            (Branch2 l' p rl) rp (Branch2 rm rq rr))))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   307
    | or => (case del (if_eq or None k) r of None \<Rightarrow> None |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   308
        Some(p', (False, r')) => Some(p', (False, Branch2 l (if_eq or p' p) r'))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   309
      | Some(p', (True, r')) => Some(p', case l of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   310
          Branch2 ll lp lr => (True, Branch3 ll lp lr (if_eq or p' p) r')
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   311
        | Branch3 ll lp lm lq lr => (False, Branch2
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   312
            (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) r')))))"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   313
"l \<noteq> Empty \<or> m \<noteq> Empty \<or> r \<noteq> Empty \<Longrightarrow>
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   314
 del k (Branch3 l p m q r) = (case compare k q of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   315
      LESS => (case compare k p of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   316
        LESS => (case del k l of None \<Rightarrow> None |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   317
          Some(p', (False, l')) => Some(p', (False, Branch3 l' p m q r))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   318
        | Some(p', (True, l')) => Some(p', (False, case (m, r) of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   319
            (Branch2 ml mp mr, Branch2 _ _ _) => Branch2 (Branch3 l' p ml mp mr) q r
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   320
          | (Branch3 ml mp mm mq mr, _) => Branch3 (Branch2 l' p ml) mp (Branch2 mm mq mr) q r
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   321
          | (Branch2 ml mp mr, Branch3 rl rp rm rq rr) =>
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   322
              Branch3 (Branch2 l' p ml) mp (Branch2 mr q rl) rp (Branch2 rm rq rr))))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   323
      | or => (case del (if_eq or None k) m of None \<Rightarrow> None |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   324
          Some(p', (False, m')) => Some(p', (False, Branch3 l (if_eq or p' p) m' q r))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   325
        | Some(p', (True, m')) => Some(p', (False, case (l, r) of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   326
            (Branch2 ll lp lr, Branch2 _ _ _) => Branch2 (Branch3 ll lp lr (if_eq or p' p) m') q r
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   327
          | (Branch3 ll lp lm lq lr, _) => Branch3 (Branch2 ll lp lm) lq (Branch2 lr (if_eq or p' p) m') q r
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   328
          | (_, Branch3 rl rp rm rq rr) => Branch3 l (if_eq or p' p) (Branch2 m' q rl) rp (Branch2 rm rq rr)))))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   329
    | or => (case del (if_eq or None k) r of None \<Rightarrow> None |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   330
        Some(q', (False, r')) => Some(q', (False, Branch3 l p m (if_eq or q' q) r'))
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   331
      | Some(q', (True, r')) => Some(q', (False, case (l, m) of
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   332
          (Branch2 _ _ _, Branch2 ml mp mr) => Branch2 l p (Branch3 ml mp mr (if_eq or q' q) r')
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   333
        | (_, Branch3 ml mp mm mq mr) => Branch3 l p (Branch2 ml mp mm) mq (Branch2 mr (if_eq or q' q) r')
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   334
        | (Branch3 ll lp lm lq lr, Branch2 ml mp mr) =>
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   335
            Branch3 (Branch2 ll lp lm) lq (Branch2 lr p ml) mp (Branch2 mr (if_eq or q' q) r')))))"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   336
apply -
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   337
apply (cases l, cases r, simp_all only: del.simps, simp)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   338
apply (cases l, cases m, cases r, simp_all only: del.simps, simp)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   339
done
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   340
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   341
inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   342
"full 0 Empty" |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   343
"\<lbrakk>full n l; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Branch2 l p r)" |
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   344
"\<lbrakk>full n l; full n m; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Branch3 l p m q r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   345
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   346
inductive_cases full_elims:
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   347
  "full n Empty"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   348
  "full n (Branch2 l p r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   349
  "full n (Branch3 l p m q r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   350
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   351
inductive_cases full_0_elim: "full 0 t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   352
inductive_cases full_Suc_elim: "full (Suc n) t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   353
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   354
lemma full_0_iff [simp]: "full 0 t \<longleftrightarrow> t = Empty"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   355
  by (auto elim: full_0_elim intro: full.intros)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   356
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   357
lemma full_Empty_iff [simp]: "full n Empty \<longleftrightarrow> n = 0"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   358
  by (auto elim: full_elims intro: full.intros)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   359
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   360
lemma full_Suc_Branch2_iff [simp]:
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   361
  "full (Suc n) (Branch2 l p r) \<longleftrightarrow> full n l \<and> full n r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   362
  by (auto elim: full_elims intro: full.intros)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   363
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   364
lemma full_Suc_Branch3_iff [simp]:
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   365
  "full (Suc n) (Branch3 l p m q r) \<longleftrightarrow> full n l \<and> full n m \<and> full n r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   366
  by (auto elim: full_elims intro: full.intros)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   367
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   368
definition
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   369
  "full_del n k t = (case del k t of None \<Rightarrow> True | Some (p, b, t') \<Rightarrow> full (if b then n else Suc n) t')"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   370
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   371
lemmas ord_case = ord.exhaust [where y=y and P="ord_case a b c y", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   372
lemmas option_ord_case = ord.exhaust [where y=y and P="option_case a b (ord_case d e f y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   373
lemmas option_option_case = option.exhaust [where y=y and P="option_case a b (option_case d e y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   374
lemmas option_prod_case = prod.exhaust [where y=y and P="option_case a b (prod_case c y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   375
lemmas option_bool_case = bool.exhaust [where y=y and P="option_case a b (bool_case c d y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   376
lemmas prod_tree23_case = tree23.exhaust [where y=y and P="prod_case a (tree23_case b c d y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   377
lemmas option_case = option.exhaust [where y=y and P="option_case a b y", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   378
lemmas full_tree23_case = tree23.exhaust [where y=y and P="full a (tree23_case b c d y)", standard]
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   379
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   380
lemmas case_intros =
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   381
  option_ord_case option_option_case option_prod_case option_bool_case prod_tree23_case option_case
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   382
  full_tree23_case
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   383
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   384
lemma full_del: "full (Suc n) t \<Longrightarrow> full_del n k t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   385
proof -
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   386
  { fix n :: "nat" and p :: "key \<times> 'a" and l r :: "'a tree23" and k
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   387
    assume "\<And>n. \<lbrakk>compare k p = LESS; full (Suc n) l\<rbrakk> \<Longrightarrow> full_del n k l"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   388
    and "\<And>n. \<lbrakk>compare k p = EQUAL; full (Suc n) r\<rbrakk> \<Longrightarrow> full_del n (if_eq EQUAL None k) r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   389
    and "\<And>n. \<lbrakk>compare k p = GREATER; full (Suc n) r\<rbrakk> \<Longrightarrow> full_del n (if_eq GREATER None k) r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   390
    and "full (Suc n) (Branch2 l p r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   391
    hence "full_del n k (Branch2 l p r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   392
     apply clarsimp
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   393
     apply (cases n)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   394
      apply (cases k)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   395
       apply (simp add: full_del_def)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   396
      apply (simp add: full_del_def split: ord.split)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   397
     apply (simp add: full_del_def)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   398
     apply (subst del_extra_simps, force)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   399
     apply (simp | rule case_intros)+
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   400
     done
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   401
  } note A = this
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   402
  { fix n :: "nat" and p q :: "key \<times> 'a" and l m r :: "'a tree23" and k
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   403
    assume "\<And>n. \<lbrakk>compare k q = LESS; compare k p = LESS; full (Suc n) l\<rbrakk> \<Longrightarrow> full_del n k l"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   404
    and "\<And>n. \<lbrakk>compare k q = LESS; compare k p = EQUAL; full (Suc n) m\<rbrakk> \<Longrightarrow> full_del n (if_eq EQUAL None k) m"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   405
    and "\<And>n. \<lbrakk>compare k q = LESS; compare k p = GREATER; full (Suc n) m\<rbrakk> \<Longrightarrow> full_del n (if_eq GREATER None k) m"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   406
    and "\<And>n. \<lbrakk>compare k q = EQUAL; full (Suc n) r\<rbrakk> \<Longrightarrow> full_del n (if_eq EQUAL None k) r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   407
    and "\<And>n. \<lbrakk>compare k q = GREATER; full (Suc n) r\<rbrakk> \<Longrightarrow> full_del n (if_eq GREATER None k) r"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   408
    and "full (Suc n) (Branch3 l p m q r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   409
    hence "full_del n k (Branch3 l p m q r)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   410
     apply clarsimp
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   411
     apply (cases n)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   412
      apply (cases k)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   413
       apply (simp add: full_del_def)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   414
      apply (simp add: full_del_def split: ord.split)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   415
     apply (simp add: full_del_def)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   416
     apply (subst del_extra_simps, force)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   417
     apply (simp | rule case_intros)+
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   418
     done
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   419
  } note B = this
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   420
  show "full (Suc n) t \<Longrightarrow> full_del n k t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   421
  proof (induct k t arbitrary: n rule: del.induct)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   422
    { case goal1 thus "full_del n (Some k) Empty"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   423
        by simp }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   424
    { case goal2 thus "full_del n None (Branch2 Empty p Empty)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   425
        unfolding full_del_def by auto }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   426
    { case goal3 thus "full_del n None (Branch3 Empty p Empty q Empty)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   427
        unfolding full_del_def by auto }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   428
    { case goal4 thus "full_del n (Some v) (Branch2 Empty p Empty)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   429
        unfolding full_del_def by (auto split: ord.split) }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   430
    { case goal5 thus "full_del n (Some v) (Branch3 Empty p Empty q Empty)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   431
        unfolding full_del_def by (auto split: ord.split) }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   432
    { case goal26 thus ?case by simp }
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   433
  qed (fact A | fact B)+
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   434
qed
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   435
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   436
lemma full_imp_height: "full n t \<Longrightarrow> height t = n"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   437
  by (induct set: full, simp_all)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   438
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   439
lemma full_imp_bal: "full n t \<Longrightarrow> bal t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   440
  by (induct set: full, auto dest: full_imp_height)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   441
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   442
lemma bal_imp_full: "bal t \<Longrightarrow> full (height t) t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   443
  by (induct t, simp_all)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   444
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   445
lemma bal_iff_full: "bal t \<longleftrightarrow> full (height t) t"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   446
  using bal_imp_full full_imp_bal ..
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   447
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   448
lemma bal_del0: "bal t \<Longrightarrow> bal (del0 k t)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   449
  unfolding del0_def
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   450
  apply (drule bal_imp_full)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   451
  apply (case_tac "height t", simp, simp)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   452
  apply (frule full_del [where k="Some k"])
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   453
  apply (simp add: full_del_def)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   454
  apply (auto split: option.split elim: full_imp_bal)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   455
  done
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   456
33436
nipkow
parents:
diff changeset
   457
text{* This is a little test harness and should be commented out once the
nipkow
parents:
diff changeset
   458
above functions have been proved correct. *}
nipkow
parents:
diff changeset
   459
nipkow
parents:
diff changeset
   460
datatype 'a act = Add int 'a | Del int
nipkow
parents:
diff changeset
   461
nipkow
parents:
diff changeset
   462
fun exec where
nipkow
parents:
diff changeset
   463
"exec [] t = t" |
nipkow
parents:
diff changeset
   464
"exec (Add k x # as) t = exec as (add0 k x t)" |
nipkow
parents:
diff changeset
   465
"exec (Del k # as) t = exec as (del0 k t)"
nipkow
parents:
diff changeset
   466
nipkow
parents:
diff changeset
   467
text{* Some quick checks: *}
nipkow
parents:
diff changeset
   468
45332
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   469
lemma bal_exec: "bal t \<Longrightarrow> bal (exec as t)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   470
  by (induct as t arbitrary: t rule: exec.induct,
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   471
    simp_all add: bal_add0 bal_del0)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   472
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   473
lemma "bal(exec as Empty)"
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   474
  by (simp add: bal_exec)
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   475
33436
nipkow
parents:
diff changeset
   476
lemma "ord0(exec as Empty)"
nipkow
parents:
diff changeset
   477
quickcheck
nipkow
parents:
diff changeset
   478
oops
nipkow
parents:
diff changeset
   479
45332
ede9dc025150 ex/Tree23.thy: prove that deletion preserves balance
huffman
parents: 45325
diff changeset
   480
end