src/HOL/Data_Structures/Tree23_Set.thy
changeset 61581 00d9682e8dd7
parent 61534 a88e07c8d0d5
child 61588 1d2907d0ed73
equal deleted inserted replaced
61569:947ce60a06e1 61581:00d9682e8dd7
     3 section \<open>A 2-3 Tree Implementation of Sets\<close>
     3 section \<open>A 2-3 Tree Implementation of Sets\<close>
     4 
     4 
     5 theory Tree23_Set
     5 theory Tree23_Set
     6 imports
     6 imports
     7   Tree23
     7   Tree23
       
     8   Cmp
     8   Set_by_Ordered
     9   Set_by_Ordered
     9 begin
    10 begin
    10 
    11 
    11 fun isin :: "'a::linorder tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
    12 fun isin :: "'a::cmp tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
    12 "isin Leaf x = False" |
    13 "isin Leaf x = False" |
    13 "isin (Node2 l a r) x = (x < a \<and> isin l x \<or> x=a \<or> isin r x)" |
    14 "isin (Node2 l a r) x =
       
    15   (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x)" |
    14 "isin (Node3 l a m b r) x =
    16 "isin (Node3 l a m b r) x =
    15   (x < a \<and> isin l x \<or> x > b \<and> isin r x \<or> x = a \<or> x = b \<or> isin m x)"
    17   (case cmp x a of LT \<Rightarrow> isin l x | EQ \<Rightarrow> True | GT \<Rightarrow> (case cmp x b of
       
    18    LT \<Rightarrow> isin m x | EQ \<Rightarrow> True | GT \<Rightarrow> isin r x))"
    16 
    19 
    17 datatype 'a up\<^sub>i = T\<^sub>i "'a tree23" | Up\<^sub>i "'a tree23" 'a "'a tree23"
    20 datatype 'a up\<^sub>i = T\<^sub>i "'a tree23" | Up\<^sub>i "'a tree23" 'a "'a tree23"
    18 
    21 
    19 fun tree\<^sub>i :: "'a up\<^sub>i \<Rightarrow> 'a tree23" where
    22 fun tree\<^sub>i :: "'a up\<^sub>i \<Rightarrow> 'a tree23" where
    20 "tree\<^sub>i (T\<^sub>i t) = t" |
    23 "tree\<^sub>i (T\<^sub>i t) = t" |
    21 "tree\<^sub>i (Up\<^sub>i l p r) = Node2 l p r"
    24 "tree\<^sub>i (Up\<^sub>i l p r) = Node2 l p r"
    22 
    25 
    23 fun ins :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>i" where
    26 fun ins :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>i" where
    24 "ins x Leaf = Up\<^sub>i Leaf x Leaf" |
    27 "ins x Leaf = Up\<^sub>i Leaf x Leaf" |
    25 "ins x (Node2 l a r) =
    28 "ins x (Node2 l a r) =
    26    (if x < a then
    29    (case cmp x a of
    27       case ins x l of
    30       LT \<Rightarrow> (case ins x l of
    28          T\<^sub>i l' => T\<^sub>i (Node2 l' a r)
    31               T\<^sub>i l' => T\<^sub>i (Node2 l' a r)
    29        | Up\<^sub>i l1 b l2 => T\<^sub>i (Node3 l1 b l2 a r)
    32             | Up\<^sub>i l1 b l2 => T\<^sub>i (Node3 l1 b l2 a r)) |
    30     else if x=a then T\<^sub>i (Node2 l x r)
    33       EQ \<Rightarrow> T\<^sub>i (Node2 l x r) |
    31     else
    34       GT \<Rightarrow> (case ins x r of
    32       case ins x r of
    35               T\<^sub>i r' => T\<^sub>i (Node2 l a r')
    33         T\<^sub>i r' => T\<^sub>i (Node2 l a r')
    36             | Up\<^sub>i r1 b r2 => T\<^sub>i (Node3 l a r1 b r2)))" |
    34       | Up\<^sub>i r1 b r2 => T\<^sub>i (Node3 l a r1 b r2))" |
       
    35 "ins x (Node3 l a m b r) =
    37 "ins x (Node3 l a m b r) =
    36    (if x < a then
    38    (case cmp x a of
    37       case ins x l of
    39       LT \<Rightarrow> (case ins x l of
    38         T\<^sub>i l' => T\<^sub>i (Node3 l' a m b r)
    40               T\<^sub>i l' => T\<^sub>i (Node3 l' a m b r)
    39       | Up\<^sub>i l1 c l2 => Up\<^sub>i (Node2 l1 c l2) a (Node2 m b r)
    41             | Up\<^sub>i l1 c l2 => Up\<^sub>i (Node2 l1 c l2) a (Node2 m b r)) |
    40     else
    42       EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
    41     if x > b then
    43       GT \<Rightarrow> (case cmp x b of
    42       case ins x r of
    44                GT \<Rightarrow> (case ins x r of
    43         T\<^sub>i r' => T\<^sub>i (Node3 l a m b r')
    45                        T\<^sub>i r' => T\<^sub>i (Node3 l a m b r')
    44       | Up\<^sub>i r1 c r2 => Up\<^sub>i (Node2 l a m) b (Node2 r1 c r2)
    46                      | Up\<^sub>i r1 c r2 => Up\<^sub>i (Node2 l a m) b (Node2 r1 c r2)) |
    45     else
    47                EQ \<Rightarrow> T\<^sub>i (Node3 l a m b r) |
    46     if x=a \<or> x = b then T\<^sub>i (Node3 l a m b r)
    48                LT \<Rightarrow> (case ins x m of
    47     else
    49                        T\<^sub>i m' => T\<^sub>i (Node3 l a m' b r)
    48       case ins x m of
    50                      | Up\<^sub>i m1 c m2 => Up\<^sub>i (Node2 l a m1) c (Node2 m2 b r))))"
    49         T\<^sub>i m' => T\<^sub>i (Node3 l a m' b r)
       
    50       | Up\<^sub>i m1 c m2 => Up\<^sub>i (Node2 l a m1) c (Node2 m2 b r))"
       
    51 
    51 
    52 hide_const insert
    52 hide_const insert
    53 
    53 
    54 definition insert :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
    54 definition insert :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
    55 "insert x t = tree\<^sub>i(ins x t)"
    55 "insert x t = tree\<^sub>i(ins x t)"
    56 
    56 
    57 datatype 'a up\<^sub>d = T\<^sub>d "'a tree23" | Up\<^sub>d "'a tree23"
    57 datatype 'a up\<^sub>d = T\<^sub>d "'a tree23" | Up\<^sub>d "'a tree23"
    58 
    58 
    59 fun tree\<^sub>d :: "'a up\<^sub>d \<Rightarrow> 'a tree23" where
    59 fun tree\<^sub>d :: "'a up\<^sub>d \<Rightarrow> 'a tree23" where
    91 "del_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
    91 "del_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
    92 "del_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
    92 "del_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
    93 "del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
    93 "del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
    94 "del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))"
    94 "del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))"
    95 
    95 
    96 fun del :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d"
    96 fun del :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d"
    97 where
    97 where
    98 "del x Leaf = T\<^sub>d Leaf" |
    98 "del x Leaf = T\<^sub>d Leaf" |
    99 "del x (Node2 Leaf a Leaf) = (if x = a then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf a Leaf))" |
    99 "del x (Node2 Leaf a Leaf) = (if x = a then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf a Leaf))" |
   100 "del x (Node3 Leaf a Leaf b Leaf) = T\<^sub>d(if x = a then Node2 Leaf b Leaf
   100 "del x (Node3 Leaf a Leaf b Leaf) = T\<^sub>d(if x = a then Node2 Leaf b Leaf
   101   else if x = b then Node2 Leaf a Leaf else Node3 Leaf a Leaf b Leaf)" |
   101   else if x = b then Node2 Leaf a Leaf else Node3 Leaf a Leaf b Leaf)" |
   102 "del x (Node2 l a r) = (if x<a then node21 (del x l) a r else
   102 "del x (Node2 l a r) = (case cmp x a of
   103   if x > a then node22 l a (del x r) else
   103   LT \<Rightarrow> node21 (del x l) a r |
   104   let (a',t) = del_min r in node22 l a' t)" |
   104   GT \<Rightarrow> node22 l a (del x r) |
   105 "del x (Node3 l a m b r) = (if x<a then node31 (del x l) a m b r else
   105   EQ \<Rightarrow> let (a',t) = del_min r in node22 l a' t)" |
   106   if x = a then let (a',m') = del_min m in node32 l a' m' b r else
   106 "del x (Node3 l a m b r) = (case cmp x a of
   107   if x < b then node32 l a (del x m) b r else
   107   LT \<Rightarrow> node31 (del x l) a m b r |
   108   if x = b then let (b',r') = del_min r in node33 l a m b' r'
   108   EQ \<Rightarrow> let (a',m') = del_min m in node32 l a' m' b r |
   109   else node33 l a m b (del x r))"
   109   GT \<Rightarrow> (case cmp x b of
   110 
   110           LT \<Rightarrow> node32 l a (del x m) b r |
   111 definition delete :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
   111           EQ \<Rightarrow> let (b',r') = del_min r in node33 l a m b' r' |
       
   112           GT \<Rightarrow> node33 l a m b (del x r)))"
       
   113 
       
   114 definition delete :: "'a::cmp \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
   112 "delete x t = tree\<^sub>d(del x t)"
   115 "delete x t = tree\<^sub>d(del x t)"
   113 
   116 
   114 
   117 
   115 subsection "Functional Correctness"
   118 subsection "Functional Correctness"
   116 
   119 
   117 
       
   118 subsubsection "Proofs for isin"
   120 subsubsection "Proofs for isin"
   119 
   121 
   120 lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
   122 lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
   121 by (induction t) (auto simp: elems_simps1)
   123 by (induction t) (auto simp: elems_simps1 ball_Un)
   122 
   124 
   123 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
   125 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
   124 by (induction t) (auto simp: elems_simps2)
   126 by (induction t) (auto simp: elems_simps2)
   125 
   127 
   126 
   128 
   127 subsubsection "Proofs for insert"
   129 subsubsection "Proofs for insert"
   128 
   130 
   129 lemma inorder_ins:
   131 lemma inorder_ins:
   130   "sorted(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(ins x t)) = ins_list x (inorder t)"
   132   "sorted(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(ins x t)) = ins_list x (inorder t)"
   131 by(induction t) (auto simp: ins_list_simps split: up\<^sub>i.splits) (* 38 secs in 2015 *)
   133 by(induction t) (auto simp: ins_list_simps split: up\<^sub>i.splits)
   132 
   134 
   133 lemma inorder_insert:
   135 lemma inorder_insert:
   134   "sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
   136   "sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
   135 by(simp add: insert_def inorder_ins)
   137 by(simp add: insert_def inorder_ins)
   136 
   138 
   193 instance ..
   195 instance ..
   194 
   196 
   195 end
   197 end
   196 
   198 
   197 lemma bal_ins: "bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t)) \<and> height(ins a t) = height t"
   199 lemma bal_ins: "bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t)) \<and> height(ins a t) = height t"
   198 by (induct t) (auto split: up\<^sub>i.split) (* 87 secs in 2015 *)
   200 by (induct t) (auto split: up\<^sub>i.split) (* 15 secs in 2015 *)
   199 
   201 
   200 text{* Now an alternative proof (by Brian Huffman) that runs faster because
   202 text{* Now an alternative proof (by Brian Huffman) that runs faster because
   201 two properties (balance and height) are combined in one predicate. *}
   203 two properties (balance and height) are combined in one predicate. *}
   202 
   204 
   203 inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where
   205 inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where