src/HOL/Data_Structures/Brother12_Set.thy
author paulson <lp15@cam.ac.uk>
Sat, 26 Aug 2023 11:36:25 +0100
changeset 78653 7ed1759fe1bd
parent 76063 24c9f56aa035
permissions -rw-r--r--
tidying up old apply-style proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
     1
(* Author: Tobias Nipkow, Daniel Stüwe *)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
     2
62130
90a3016a6c12 added AA_Map; tuned titles
nipkow
parents: 61809
diff changeset
     3
section \<open>1-2 Brother Tree Implementation of Sets\<close>
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
     4
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
     5
theory Brother12_Set
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
     6
  imports
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
     7
    Cmp
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
     8
    Set_Specs
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
     9
    "HOL-Number_Theory.Fib"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    10
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    11
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    12
subsection \<open>Data Type and Operations\<close>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    13
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    14
datatype 'a bro =
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    15
  N0 |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    16
  N1 "'a bro" |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    17
  N2 "'a bro" 'a "'a bro" |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    18
  (* auxiliary constructors: *)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    19
  L2 'a |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    20
  N3 "'a bro" 'a "'a bro" 'a "'a bro"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    21
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    22
definition empty :: "'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    23
  "empty = N0"
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
    24
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    25
fun inorder :: "'a bro \<Rightarrow> 'a list" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    26
  "inorder N0 = []" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    27
  "inorder (N1 t) = inorder t" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    28
  "inorder (N2 l a r) = inorder l @ a # inorder r" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    29
  "inorder (L2 a) = [a]" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    30
  "inorder (N3 t1 a1 t2 a2 t3) = inorder t1 @ a1 # inorder t2 @ a2 # inorder t3"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    31
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
    32
fun isin :: "'a bro \<Rightarrow> 'a::linorder \<Rightarrow> bool" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    33
  "isin N0 x = False" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    34
  "isin (N1 t) x = isin t x" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    35
  "isin (N2 l a r) x =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    36
  (case cmp x a of
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    37
     LT \<Rightarrow> isin l x |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    38
     EQ \<Rightarrow> True |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    39
     GT \<Rightarrow> isin r x)"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    40
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    41
fun n1 :: "'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    42
  "n1 (L2 a) = N2 N0 a N0" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    43
  "n1 (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    44
  "n1 t = N1 t"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    45
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    46
hide_const (open) insert
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    47
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    48
locale insert
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    49
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    50
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    51
fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    52
  "n2 (L2 a1) a2 t = N3 N0 a1 N0 a2 t" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    53
  "n2 (N3 t1 a1 t2 a2 t3) a3 (N1 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    54
  "n2 (N3 t1 a1 t2 a2 t3) a3 t4 = N3 (N2 t1 a1 t2) a2 (N1 t3) a3 t4" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    55
  "n2 t1 a1 (L2 a2) = N3 t1 a1 N0 a2 N0" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    56
  "n2 (N1 t1) a1 (N3 t2 a2 t3 a3 t4) = N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    57
  "n2 t1 a1 (N3 t2 a2 t3 a3 t4) = N3 t1 a1 (N1 t2) a2 (N2 t3 a3 t4)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    58
  "n2 t1 a t2 = N2 t1 a t2"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    59
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
    60
fun ins :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    61
  "ins x N0 = L2 x" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    62
  "ins x (N1 t) = n1 (ins x t)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    63
  "ins x (N2 l a r) =
61789
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
    64
  (case cmp x a of
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
    65
     LT \<Rightarrow> n2 (ins x l) a r |
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
    66
     EQ \<Rightarrow> N2 l a r |
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
    67
     GT \<Rightarrow> n2 l a (ins x r))"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    68
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    69
fun tree :: "'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    70
  "tree (L2 a) = N2 N0 a N0" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    71
  "tree (N3 t1 a1 t2 a2 t3) = N2 (N2 t1 a1 t2) a2 (N1 t3)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    72
  "tree t = t"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    73
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
    74
definition insert :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    75
  "insert x t = tree(ins x t)"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    76
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    77
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    78
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    79
locale delete
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    80
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    81
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    82
fun n2 :: "'a bro \<Rightarrow> 'a \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    83
  "n2 (N1 t1) a1 (N1 t2) = N1 (N2 t1 a1 t2)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    84
  "n2 (N1 (N1 t1)) a1 (N2 (N1 t2) a2 (N2 t3 a3 t4)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    85
  N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    86
  "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N1 t4)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    87
  N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    88
  "n2 (N1 (N1 t1)) a1 (N2 (N2 t2 a2 t3) a3 (N2 t4 a4 t5)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    89
  N2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N2 t4 a4 t5))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    90
  "n2 (N2 (N1 t1) a1 (N2 t2 a2 t3)) a3 (N1 (N1 t4)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    91
  N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    92
  "n2 (N2 (N2 t1 a1 t2) a2 (N1 t3)) a3 (N1 (N1 t4)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    93
  N1 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    94
  "n2 (N2 (N2 t1 a1 t2) a2 (N2 t3 a3 t4)) a5 (N1 (N1 t5)) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    95
  N2 (N1 (N2 t1 a1 t2)) a2 (N2 (N2 t3 a3 t4) a5 (N1 t5))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    96
  "n2 t1 a1 t2 = N2 t1 a1 t2"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
    97
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
    98
fun split_min :: "'a bro \<Rightarrow> ('a \<times> 'a bro) option" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
    99
  "split_min N0 = None" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   100
  "split_min (N1 t) =
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   101
  (case split_min t of
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   102
     None \<Rightarrow> None |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   103
     Some (a, t') \<Rightarrow> Some (a, N1 t'))" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   104
  "split_min (N2 t1 a t2) =
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   105
  (case split_min t1 of
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   106
     None \<Rightarrow> Some (a, N1 t2) |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   107
     Some (b, t1') \<Rightarrow> Some (b, n2 t1' a t2))"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   108
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   109
fun del :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   110
  "del _ N0         = N0" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   111
  "del x (N1 t)     = N1 (del x t)" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   112
  "del x (N2 l a r) =
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   113
  (case cmp x a of
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   114
     LT \<Rightarrow> n2 (del x l) a r |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   115
     GT \<Rightarrow> n2 l a (del x r) |
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   116
     EQ \<Rightarrow> (case split_min r of
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   117
              None \<Rightarrow> N1 l |
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   118
              Some (b, r') \<Rightarrow> n2 l b r'))"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   119
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   120
fun tree :: "'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   121
  "tree (N1 t) = t" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   122
  "tree t = t"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   123
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   124
definition delete :: "'a::linorder \<Rightarrow> 'a bro \<Rightarrow> 'a bro" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   125
  "delete a t = tree (del a t)"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   126
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   127
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   128
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   129
subsection \<open>Invariants\<close>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   130
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   131
fun B :: "nat \<Rightarrow> 'a bro set"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   132
  and U :: "nat \<Rightarrow> 'a bro set" where
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   133
  "B 0 = {N0}" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   134
  "B (Suc h) = { N2 t1 a t2 | t1 a t2. 
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   135
  t1 \<in> B h \<union> U h \<and> t2 \<in> B h \<or> t1 \<in> B h \<and> t2 \<in> B h \<union> U h}" |
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   136
  "U 0 = {}" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   137
  "U (Suc h) = N1 ` B h"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   138
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   139
abbreviation "T h \<equiv> B h \<union> U h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   140
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   141
fun Bp :: "nat \<Rightarrow> 'a bro set" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   142
  "Bp 0 = B 0 \<union> L2 ` UNIV" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   143
  "Bp (Suc 0) = B (Suc 0) \<union> {N3 N0 a N0 b N0|a b. True}" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   144
  "Bp (Suc(Suc h)) = B (Suc(Suc h)) \<union>
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   145
  {N3 t1 a t2 b t3 | t1 a t2 b t3. t1 \<in> B (Suc h) \<and> t2 \<in> U (Suc h) \<and> t3 \<in> B (Suc h)}"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   146
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   147
fun Um :: "nat \<Rightarrow> 'a bro set" where
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   148
  "Um 0 = {}" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   149
  "Um (Suc h) = N1 ` T h"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   150
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   151
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   152
subsection "Functional Correctness Proofs"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   153
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   154
subsubsection "Proofs for isin"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   155
67929
30486b96274d eliminated "elems"
nipkow
parents: 67613
diff changeset
   156
lemma isin_set:
30486b96274d eliminated "elems"
nipkow
parents: 67613
diff changeset
   157
  "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> set(inorder t))"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   158
  by(induction h arbitrary: t) (fastforce simp: isin_simps split: if_splits)+
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   159
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   160
subsubsection "Proofs for insertion"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   161
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   162
lemma inorder_n1: "inorder(n1 t) = inorder t"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   163
  by(cases t rule: n1.cases) (auto simp: sorted_lems)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   164
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   165
context insert
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   166
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   167
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   168
lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   169
  by(cases "(l,a,r)" rule: n2.cases) (auto simp: sorted_lems)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   170
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   171
lemma inorder_tree: "inorder(tree t) = inorder t"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   172
  by(cases t) auto
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   173
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   174
lemma inorder_ins: "t \<in> T h \<Longrightarrow>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   175
  sorted(inorder t) \<Longrightarrow> inorder(ins a t) = ins_list a (inorder t)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   176
  by(induction h arbitrary: t) (auto simp: ins_list_simps inorder_n1 inorder_n2)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   177
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   178
lemma inorder_insert: "t \<in> T h \<Longrightarrow>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   179
  sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   180
  by(simp add: insert_def inorder_ins inorder_tree)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   181
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   182
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   183
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   184
subsubsection \<open>Proofs for deletion\<close>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   185
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   186
context delete
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   187
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   188
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   189
lemma inorder_tree: "inorder(tree t) = inorder t"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   190
  by(cases t) auto
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   191
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   192
lemma inorder_n2: "inorder(n2 l a r) = inorder l @ a # inorder r"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   193
  by(cases "(l,a,r)" rule: n2.cases) (auto)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   194
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   195
lemma inorder_split_min:
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   196
  "t \<in> T h \<Longrightarrow> (split_min t = None \<longleftrightarrow> inorder t = []) \<and>
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   197
  (split_min t = Some(a,t') \<longrightarrow> inorder t = a # inorder t')"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   198
  by(induction h arbitrary: t a t') (auto simp: inorder_n2 split: option.splits)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   199
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   200
lemma inorder_del:
61792
nipkow
parents: 61789
diff changeset
   201
  "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(del x t) = del_list x (inorder t)"
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 68431
diff changeset
   202
  apply (induction h arbitrary: t) 
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 68431
diff changeset
   203
  apply (auto simp: del_list_simps inorder_n2 split: option.splits)
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 68431
diff changeset
   204
  apply (auto simp: del_list_simps inorder_n2
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   205
      inorder_split_min[OF UnI1] inorder_split_min[OF UnI2] split: option.splits)
73526
a3cc9fa1295d new automatic order prover: stateless, complete, verified
nipkow
parents: 68431
diff changeset
   206
  done
61792
nipkow
parents: 61789
diff changeset
   207
nipkow
parents: 61789
diff changeset
   208
lemma inorder_delete:
nipkow
parents: 61789
diff changeset
   209
  "t \<in> T h \<Longrightarrow> sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   210
  by(simp add: delete_def inorder_del inorder_tree)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   211
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   212
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   213
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   214
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   215
subsection \<open>Invariant Proofs\<close>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   216
61789
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
   217
subsubsection \<open>Proofs for insertion\<close>
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   218
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   219
lemma n1_type: "t \<in> Bp h \<Longrightarrow> n1 t \<in> T (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   220
  by(cases h rule: Bp.cases) auto
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   221
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   222
context insert
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   223
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   224
61809
81d34cf268d8 tightened invariant
nipkow
parents: 61792
diff changeset
   225
lemma tree_type: "t \<in> Bp h \<Longrightarrow> tree t \<in> B h \<union> B (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   226
  by(cases h rule: Bp.cases) auto
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   227
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   228
lemma n2_type:
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   229
  "(t1 \<in> Bp h \<and> t2 \<in> T h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h)) \<and>
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   230
   (t1 \<in> T h \<and> t2 \<in> Bp h \<longrightarrow> n2 t1 a t2 \<in> Bp (Suc h))"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   231
  apply(cases h rule: Bp.cases)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   232
    apply (auto)[2]
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   233
  apply(rule conjI impI | erule conjE exE imageE | simp | erule disjE)+
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   234
  done
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   235
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   236
lemma Bp_if_B: "t \<in> B h \<Longrightarrow> t \<in> Bp h"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   237
  by (cases h rule: Bp.cases) simp_all
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   238
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   239
text\<open>An automatic proof:\<close>
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   240
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   241
lemma
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   242
  "(t \<in> B h \<longrightarrow> ins x t \<in> Bp h) \<and> (t \<in> U h \<longrightarrow> ins x t \<in> T h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   243
proof (induction h arbitrary: t)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   244
  case 0
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   245
  then show ?case by simp
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   246
next
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   247
  case (Suc h)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   248
  then show ?case by (fastforce simp: Bp_if_B n2_type dest: n1_type)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   249
qed
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   250
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   251
text\<open>A detailed proof:\<close>
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   252
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   253
lemma ins_type:
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   254
  shows "t \<in> B h \<Longrightarrow> ins x t \<in> Bp h" and "t \<in> U h \<Longrightarrow> ins x t \<in> T h"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   255
proof(induction h arbitrary: t)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   256
  case 0
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   257
  { case 1 thus ?case by simp
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   258
  next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   259
    case 2 thus ?case by simp }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   260
next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   261
  case (Suc h)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   262
  { case 1
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   263
    then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   264
      t1: "t1 \<in> T h" and t2: "t2 \<in> T h" and t12: "t1 \<in> B h \<or> t2 \<in> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   265
      by auto
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   266
    have ?case if "x < a"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   267
    proof -
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   268
      have "n2 (ins x t1) a t2 \<in> Bp (Suc h)"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   269
      proof cases
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   270
        assume "t1 \<in> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   271
        with t2 show ?thesis by (simp add: Suc.IH(1) n2_type)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   272
      next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   273
        assume "t1 \<notin> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   274
        hence 1: "t1 \<in> U h" and 2: "t2 \<in> B h" using t1 t12 by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   275
        show ?thesis by (metis Suc.IH(2)[OF 1] Bp_if_B[OF 2] n2_type)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   276
      qed
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   277
      with \<open>x < a\<close> show ?case by simp
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   278
    qed
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   279
    moreover
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   280
    have ?case if "a < x"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   281
    proof -
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   282
      have "n2 t1 a (ins x t2) \<in> Bp (Suc h)"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   283
      proof cases
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   284
        assume "t2 \<in> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   285
        with t1 show ?thesis by (simp add: Suc.IH(1) n2_type)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   286
      next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   287
        assume "t2 \<notin> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   288
        hence 1: "t1 \<in> B h" and 2: "t2 \<in> U h" using t2 t12 by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   289
        show ?thesis by (metis Bp_if_B[OF 1] Suc.IH(2)[OF 2] n2_type)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   290
      qed
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   291
      with \<open>a < x\<close> show ?case by simp
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   292
    qed
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   293
    moreover
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   294
    have ?case if "x = a"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   295
    proof -
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   296
      from 1 have "t \<in> Bp (Suc h)" by(rule Bp_if_B)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   297
      thus "?case" using \<open>x = a\<close> by simp
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   298
    qed
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   299
    ultimately show ?case by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   300
  next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   301
    case 2 thus ?case using Suc(1) n1_type by fastforce }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   302
qed
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   303
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   304
lemma insert_type:
61809
81d34cf268d8 tightened invariant
nipkow
parents: 61792
diff changeset
   305
  "t \<in> B h \<Longrightarrow> insert x t \<in> B h \<union> B (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   306
  unfolding insert_def by (metis ins_type(1) tree_type)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   307
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   308
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   309
61789
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
   310
subsubsection "Proofs for deletion"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   311
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   312
lemma B_simps[simp]: 
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   313
  "N1 t \<in> B h = False"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   314
  "L2 y \<in> B h = False"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   315
  "(N3 t1 a1 t2 a2 t3) \<in> B h = False"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   316
  "N0 \<in> B h \<longleftrightarrow> h = 0"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   317
  by (cases h, auto)+
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   318
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   319
context delete
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   320
begin
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   321
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   322
lemma n2_type1:
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   323
  "\<lbrakk>t1 \<in> Um h; t2 \<in> B h\<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   324
  apply(cases h rule: Bp.cases)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   325
    apply auto[2]
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   326
  apply(erule exE bexE conjE imageE | simp | erule disjE)+
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   327
  done
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   328
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   329
lemma n2_type2:
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   330
  "\<lbrakk>t1 \<in> B h ; t2 \<in> Um h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   331
  apply(cases h rule: Bp.cases)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   332
  using Um.simps(1) apply blast
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   333
   apply force
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   334
  apply(erule exE bexE conjE imageE | simp | erule disjE)+
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   335
  done
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   336
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   337
lemma n2_type3:
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   338
  "\<lbrakk>t1 \<in> T h ; t2 \<in> T h \<rbrakk> \<Longrightarrow> n2 t1 a t2 \<in> T (Suc h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   339
  apply(cases h rule: Bp.cases)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   340
    apply auto[2]
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   341
  apply(erule exE bexE conjE imageE | simp | erule disjE)+
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   342
  done
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   343
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   344
lemma split_minNoneN0: "\<lbrakk>t \<in> B h; split_min t = None\<rbrakk> \<Longrightarrow>  t = N0"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   345
  by (cases t) (auto split: option.splits)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   346
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   347
lemma split_minNoneN1 : "\<lbrakk>t \<in> U h; split_min t = None\<rbrakk> \<Longrightarrow> t = N1 N0"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   348
  by (cases h) (auto simp: split_minNoneN0  split: option.splits)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   349
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   350
lemma split_min_type:
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   351
  "t \<in> B h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> T h"
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   352
  "t \<in> U h \<Longrightarrow> split_min t = Some (a, t') \<Longrightarrow> t' \<in> Um h"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   353
proof (induction h arbitrary: t a t')
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   354
  case (Suc h)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   355
  { case 1
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   356
    then obtain t1 a t2 where [simp]: "t = N2 t1 a t2" and
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   357
      t12: "t1 \<in> T h" "t2 \<in> T h" "t1 \<in> B h \<or> t2 \<in> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   358
      by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   359
    show ?case
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   360
    proof (cases "split_min t1")
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   361
      case None
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   362
      show ?thesis
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   363
      proof cases
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   364
        assume "t1 \<in> B h"
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   365
        with split_minNoneN0[OF this None] 1 show ?thesis by(auto)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   366
      next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   367
        assume "t1 \<notin> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   368
        thus ?thesis using 1 None by (auto)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   369
      qed
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   370
    next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   371
      case [simp]: (Some bt')
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   372
      obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   373
      show ?thesis
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   374
      proof cases
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   375
        assume "t1 \<in> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   376
        from Suc.IH(1)[OF this] 1 have "t1' \<in> T h" by simp
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   377
        from n2_type3[OF this t12(2)] 1 show ?thesis by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   378
      next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   379
        assume "t1 \<notin> B h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   380
        hence t1: "t1 \<in> U h" and t2: "t2 \<in> B h" using t12 by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   381
        from Suc.IH(2)[OF t1] have "t1' \<in> Um h" by simp
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   382
        from n2_type1[OF this t2] 1 show ?thesis by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   383
      qed
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   384
    qed
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   385
  }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   386
  { case 2
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   387
    then obtain t1 where [simp]: "t = N1 t1" and t1: "t1 \<in> B h" by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   388
    show ?case
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   389
    proof (cases "split_min t1")
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   390
      case None
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   391
      with split_minNoneN0[OF t1 None] 2 show ?thesis by(auto)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   392
    next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   393
      case [simp]: (Some bt')
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   394
      obtain b t1' where [simp]: "bt' = (b,t1')" by fastforce
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   395
      from Suc.IH(1)[OF t1] have "t1' \<in> T h" by simp
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   396
      thus ?thesis using 2 by auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   397
    qed
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   398
  }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   399
qed auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   400
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   401
lemma del_type:
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   402
  "t \<in> B h \<Longrightarrow> del x t \<in> T h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   403
  "t \<in> U h \<Longrightarrow> del x t \<in> Um h"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   404
proof (induction h arbitrary: x t)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   405
  case (Suc h)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   406
  { case 1
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   407
    then obtain l a r where [simp]: "t = N2 l a r" and
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   408
      lr: "l \<in> T h" "r \<in> T h" "l \<in> B h \<or> r \<in> B h" by auto
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   409
    have ?case if "x < a"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   410
    proof cases
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   411
      assume "l \<in> B h"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   412
      from n2_type3[OF Suc.IH(1)[OF this] lr(2)]
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   413
      show ?thesis using \<open>x<a\<close> by(simp)
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   414
    next
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   415
      assume "l \<notin> B h"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   416
      hence "l \<in> U h" "r \<in> B h" using lr by auto
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   417
      from n2_type1[OF Suc.IH(2)[OF this(1)] this(2)]
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   418
      show ?thesis using \<open>x<a\<close> by(simp)
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   419
    qed
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   420
    moreover
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   421
    have ?case if "x > a"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   422
    proof cases
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   423
      assume "r \<in> B h"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   424
      from n2_type3[OF lr(1) Suc.IH(1)[OF this]]
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   425
      show ?thesis using \<open>x>a\<close> by(simp)
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   426
    next
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   427
      assume "r \<notin> B h"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   428
      hence "l \<in> B h" "r \<in> U h" using lr by auto
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   429
      from n2_type2[OF this(1) Suc.IH(2)[OF this(2)]]
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67040
diff changeset
   430
      show ?thesis using \<open>x>a\<close> by(simp)
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   431
    qed
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   432
    moreover
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   433
    have ?case if [simp]: "x=a"
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   434
    proof (cases "split_min r")
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   435
      case None
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   436
      show ?thesis
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   437
      proof cases
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   438
        assume "r \<in> B h"
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   439
        with split_minNoneN0[OF this None] lr show ?thesis by(simp)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   440
      next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   441
        assume "r \<notin> B h"
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   442
        hence "r \<in> U h" using lr by auto
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   443
        with split_minNoneN1[OF this None] lr(3) show ?thesis by (simp)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   444
      qed
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   445
    next
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   446
      case [simp]: (Some br')
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   447
      obtain b r' where [simp]: "br' = (b,r')" by fastforce
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   448
      show ?thesis
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   449
      proof cases
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   450
        assume "r \<in> B h"
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   451
        from split_min_type(1)[OF this] n2_type3[OF lr(1)]
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   452
        show ?thesis by simp
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   453
      next
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   454
        assume "r \<notin> B h"
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   455
        hence "l \<in> B h" and "r \<in> U h" using lr by auto
68020
6aade817bee5 del_min -> split_min
nipkow
parents: 67965
diff changeset
   456
        from split_min_type(2)[OF this(2)] n2_type2[OF this(1)]
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   457
        show ?thesis by simp
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   458
      qed
67040
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   459
    qed
c1b87d15774a replaced raw proof blocks by local lemmas
nipkow
parents: 66453
diff changeset
   460
    ultimately show ?case by auto
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   461
  }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   462
  { case 2 with Suc.IH(1) show ?case by auto }
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   463
qed auto
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   464
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   465
lemma tree_type: "t \<in> T (h+1) \<Longrightarrow> tree t \<in> B (h+1) \<union> B h"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   466
  by(auto)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   467
61809
81d34cf268d8 tightened invariant
nipkow
parents: 61792
diff changeset
   468
lemma delete_type: "t \<in> B h \<Longrightarrow> delete x t \<in> B h \<union> B(h-1)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   469
  unfolding delete_def
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   470
  by (cases h) (simp, metis del_type(1) tree_type Suc_eq_plus1 diff_Suc_1)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   471
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   472
end
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   473
61789
9ce1a397410a added Brother12_Map
nipkow
parents: 61784
diff changeset
   474
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   475
subsection "Overall correctness"
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   476
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   477
interpretation Set_by_Ordered
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   478
  where empty = empty and isin = isin and insert = insert.insert
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   479
    and delete = delete.delete and inorder = inorder and inv = "\<lambda>t. \<exists>h. t \<in> B h"
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   480
proof (standard, goal_cases)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   481
  case 2 thus ?case by(auto intro!: isin_set)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   482
next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   483
  case 3 thus ?case by(auto intro!: insert.inorder_insert)
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   484
next
61792
nipkow
parents: 61789
diff changeset
   485
  case 4 thus ?case by(auto intro!: delete.inorder_delete)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   486
next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   487
  case 6 thus ?case using insert.insert_type by blast
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   488
next
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   489
  case 7 thus ?case using delete.delete_type by blast
68431
b294e095f64c more abstract naming
nipkow
parents: 68020
diff changeset
   490
qed (auto simp: empty_def)
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   491
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   492
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   493
subsection \<open>Height-Size Relation\<close>
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   494
76063
24c9f56aa035 proper umlauts;
wenzelm
parents: 73526
diff changeset
   495
text \<open>By Daniel Stüwe\<close>
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   496
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   497
fun fib_tree :: "nat \<Rightarrow> unit bro" where
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   498
  "fib_tree 0 = N0" 
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   499
| "fib_tree (Suc 0) = N2 N0 () N0"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   500
| "fib_tree (Suc(Suc h)) = N2 (fib_tree (h+1)) () (N1 (fib_tree h))"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   501
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   502
fun fib' :: "nat \<Rightarrow> nat" where
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   503
  "fib' 0 = 0" 
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   504
| "fib' (Suc 0) = 1"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   505
| "fib' (Suc(Suc h)) = 1 + fib' (Suc h) + fib' h"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   506
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   507
fun size :: "'a bro \<Rightarrow> nat" where
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   508
  "size N0 = 0" 
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   509
| "size (N1 t) = size t"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   510
| "size (N2 t1 _ t2) = 1 + size t1 + size t2"
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   511
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   512
lemma fib_tree_B: "fib_tree h \<in> B h"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   513
  by (induction h rule: fib_tree.induct) auto
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   514
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   515
declare [[names_short]]
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   516
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   517
lemma size_fib': "size (fib_tree h) = fib' h"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   518
  by (induction h rule: fib_tree.induct) auto
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   519
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   520
lemma fibfib: "fib' h + 1 = fib (Suc(Suc h))"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   521
  by (induction h rule: fib_tree.induct) auto
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   522
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   523
lemma B_N2_cases[consumes 1]:
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   524
  assumes "N2 t1 a t2 \<in> B (Suc n)"
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   525
  obtains 
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   526
    (BB) "t1 \<in> B n" and "t2 \<in> B n" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   527
    (UB) "t1 \<in> U n" and "t2 \<in> B n" |
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   528
    (BU) "t1 \<in> B n" and "t2 \<in> U n"
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   529
  using assms by auto
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   530
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   531
lemma size_bounded: "t \<in> B h \<Longrightarrow> size t \<ge> size (fib_tree h)"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   532
  unfolding size_fib' proof (induction h arbitrary: t rule: fib'.induct)
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   533
  case (3 h t')
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   534
  note main = 3
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   535
  then obtain t1 a t2 where t': "t' = N2 t1 a t2" by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   536
  with main have "N2 t1 a t2 \<in> B (Suc (Suc h))" by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   537
  thus ?case proof (cases rule: B_N2_cases)
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   538
    case BB
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   539
    then obtain x y z where t2: "t2 = N2 x y z \<or> t2 = N2 z y x" "x \<in> B h" by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   540
    show ?thesis unfolding t' using main(1)[OF BB(1)] main(2)[OF t2(2)] t2(1) by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   541
  next
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   542
    case UB
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   543
    then obtain t11 where t1: "t1 = N1 t11" "t11 \<in> B h" by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   544
    show ?thesis unfolding t' t1(1) using main(2)[OF t1(2)] main(1)[OF UB(2)] by simp
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   545
  next
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   546
    case BU
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   547
    then obtain t22 where t2: "t2 = N1 t22" "t22 \<in> B h" by auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   548
    show ?thesis unfolding t' t2(1) using main(2)[OF t2(2)] main(1)[OF BU(1)] by simp
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   549
  qed
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   550
qed auto
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   551
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   552
theorem "t \<in> B h \<Longrightarrow> fib (h + 2) \<le> size t + 1"
78653
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   553
  using size_bounded
7ed1759fe1bd tidying up old apply-style proofs
paulson <lp15@cam.ac.uk>
parents: 76063
diff changeset
   554
  by (simp add: size_fib' fibfib[symmetric] del: fib.simps)
63411
e051eea34990 got rid of class cmp; added height-size proofs by Daniel Stuewe
nipkow
parents: 62526
diff changeset
   555
61784
21b34a2269e5 added 1-2 brother trees
nipkow
parents:
diff changeset
   556
end