src/HOL/Data_Structures/Tree23_Set.thy
changeset 61469 cd82b1023932
child 61513 c0126c001b3d
equal deleted inserted replaced
61468:7d1127ac2251 61469:cd82b1023932
       
     1 (* Author: Tobias Nipkow *)
       
     2 
       
     3 section \<open>2-3 Tree Implementation of Sets\<close>
       
     4 
       
     5 theory Tree23_Set
       
     6 imports
       
     7   Tree23
       
     8   Set_by_Ordered
       
     9 begin
       
    10 
       
    11 fun isin :: "'a::linorder tree23 \<Rightarrow> 'a \<Rightarrow> bool" where
       
    12 "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 (Node3 l a m b r) x =
       
    15   (x < a \<and> isin l x \<or> x = a \<or> (x < b \<and> isin m x \<or> x = b \<or> isin r x))"
       
    16 
       
    17 datatype 'a up\<^sub>i = T\<^sub>i "'a tree23" | Up\<^sub>i "'a tree23" 'a "'a tree23"
       
    18 
       
    19 fun tree\<^sub>i :: "'a up\<^sub>i \<Rightarrow> 'a tree23" where
       
    20 "tree\<^sub>i (T\<^sub>i t) = t" |
       
    21 "tree\<^sub>i (Up\<^sub>i l p r) = Node2 l p r"
       
    22 
       
    23 fun ins :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>i" where
       
    24 "ins a Leaf = Up\<^sub>i Leaf a Leaf" |
       
    25 "ins a (Node2 l x r) =
       
    26    (if a < x then
       
    27       case ins a l of
       
    28          T\<^sub>i l' => T\<^sub>i (Node2 l' x r)
       
    29        | Up\<^sub>i l1 q l2 => T\<^sub>i (Node3 l1 q l2 x r)
       
    30     else if a=x then T\<^sub>i (Node2 l x r)
       
    31     else
       
    32       case ins a r of
       
    33         T\<^sub>i r' => T\<^sub>i (Node2 l x r')
       
    34       | Up\<^sub>i r1 q r2 => T\<^sub>i (Node3 l x r1 q r2))" |
       
    35 "ins a (Node3 l x1 m x2 r) =
       
    36    (if a < x1 then
       
    37       case ins a l of
       
    38         T\<^sub>i l' => T\<^sub>i (Node3 l' x1 m x2 r)
       
    39       | Up\<^sub>i l1 q l2 => Up\<^sub>i (Node2 l1 q l2) x1 (Node2 m x2 r)
       
    40     else if a=x1 then T\<^sub>i (Node3 l x1 m x2 r)
       
    41     else if a < x2 then
       
    42            case ins a m of
       
    43              T\<^sub>i m' => T\<^sub>i (Node3 l x1 m' x2 r)
       
    44            | Up\<^sub>i m1 q m2 => Up\<^sub>i (Node2 l x1 m1) q (Node2 m2 x2 r)
       
    45          else if a=x2 then T\<^sub>i (Node3 l x1 m x2 r)
       
    46          else
       
    47            case ins a r of
       
    48              T\<^sub>i r' => T\<^sub>i (Node3 l x1 m x2 r')
       
    49            | Up\<^sub>i r1 q r2 => Up\<^sub>i (Node2 l x1 m) x2 (Node2 r1 q r2))"
       
    50 
       
    51 hide_const insert
       
    52 
       
    53 definition insert :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
       
    54 "insert a t = tree\<^sub>i(ins a t)"
       
    55 
       
    56 datatype 'a up\<^sub>d = T\<^sub>d "'a tree23" | Up\<^sub>d "'a tree23"
       
    57 
       
    58 fun tree\<^sub>d :: "'a up\<^sub>d \<Rightarrow> 'a tree23" where
       
    59 "tree\<^sub>d (T\<^sub>d x) = x" |
       
    60 "tree\<^sub>d (Up\<^sub>d x) = x"
       
    61 
       
    62 (* Variation: return None to signal no-change *)
       
    63 
       
    64 fun node21 :: "'a up\<^sub>d \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d" where
       
    65 "node21 (T\<^sub>d t1) a t2 = T\<^sub>d(Node2 t1 a t2)" |
       
    66 "node21 (Up\<^sub>d t1) a (Node2 t2 b t3) = Up\<^sub>d(Node3 t1 a t2 b t3)" |
       
    67 "node21 (Up\<^sub>d t1) a (Node3 t2 b t3 c t4) = T\<^sub>d(Node2 (Node2 t1 a t2) b (Node2 t3 c t4))"
       
    68 
       
    69 fun node22 :: "'a tree23 \<Rightarrow> 'a \<Rightarrow> 'a up\<^sub>d \<Rightarrow> 'a up\<^sub>d" where
       
    70 "node22 t1 a (T\<^sub>d t2) = T\<^sub>d(Node2 t1 a t2)" |
       
    71 "node22 (Node2 t1 b t2) a (Up\<^sub>d t3) = Up\<^sub>d(Node3 t1 b t2 a t3)" |
       
    72 "node22 (Node3 t1 b t2 c t3) a (Up\<^sub>d t4) = T\<^sub>d(Node2 (Node2 t1 b t2) c (Node2 t3 a t4))"
       
    73 
       
    74 fun node31 :: "'a up\<^sub>d \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d" where
       
    75 "node31 (T\<^sub>d t1) a t2 b t3 = T\<^sub>d(Node3 t1 a t2 b t3)" |
       
    76 "node31 (Up\<^sub>d t1) a (Node2 t2 b t3) c t4 = T\<^sub>d(Node2 (Node3 t1 a t2 b t3) c t4)" |
       
    77 "node31 (Up\<^sub>d t1) a (Node3 t2 b t3 c t4) d t5 = T\<^sub>d(Node3 (Node2 t1 a t2) b (Node2 t3 c t4) d t5)"
       
    78 
       
    79 fun node32 :: "'a tree23 \<Rightarrow> 'a \<Rightarrow> 'a up\<^sub>d \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d" where
       
    80 "node32 t1 a (T\<^sub>d t2) b t3 = T\<^sub>d(Node3 t1 a t2 b t3)" |
       
    81 "node32 t1 a (Up\<^sub>d t2) b (Node2 t3 c t4) = T\<^sub>d(Node2 t1 a (Node3 t2 b t3 c t4))" |
       
    82 "node32 t1 a (Up\<^sub>d t2) b (Node3 t3 c t4 d t5) = T\<^sub>d(Node3 t1 a (Node2 t2 b t3) c (Node2 t4 d t5))"
       
    83 
       
    84 fun node33 :: "'a tree23 \<Rightarrow> 'a \<Rightarrow> 'a tree23 \<Rightarrow> 'a \<Rightarrow> 'a up\<^sub>d \<Rightarrow> 'a up\<^sub>d" where
       
    85 "node33 l a m b (T\<^sub>d r) = T\<^sub>d(Node3 l a m b r)" |
       
    86 "node33 t1 a (Node2 t2 b t3) c (Up\<^sub>d t4) = T\<^sub>d(Node2 t1 a (Node3 t2 b t3 c t4))" |
       
    87 "node33 t1 a (Node3 t2 b t3 c t4) d (Up\<^sub>d t5) = T\<^sub>d(Node3 t1 a (Node2 t2 b t3) c (Node2 t4 d t5))"
       
    88 
       
    89 fun del_min :: "'a tree23 \<Rightarrow> 'a * 'a up\<^sub>d" where
       
    90 "del_min (Node2 Leaf a Leaf) = (a, Up\<^sub>d Leaf)" |
       
    91 "del_min (Node3 Leaf a Leaf b Leaf) = (a, T\<^sub>d(Node2 Leaf b Leaf))" |
       
    92 "del_min (Node2 l a r) = (let (x,l') = del_min l in (x, node21 l' a r))" |
       
    93 "del_min (Node3 l a m b r) = (let (x,l') = del_min l in (x, node31 l' a m b r))"
       
    94 
       
    95 fun del :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a up\<^sub>d"
       
    96 where
       
    97 "del k Leaf = T\<^sub>d Leaf" |
       
    98 "del k (Node2 Leaf p Leaf) = (if k=p then Up\<^sub>d Leaf else T\<^sub>d(Node2 Leaf p Leaf))" |
       
    99 "del k (Node3 Leaf p Leaf q Leaf) = T\<^sub>d(if k=p then Node2 Leaf q Leaf
       
   100   else if k=q then Node2 Leaf p Leaf else Node3 Leaf p Leaf q Leaf)" |
       
   101 "del k (Node2 l a r) = (if k<a then node21 (del k l) a r else
       
   102   if k > a then node22 l a (del k r) else
       
   103   let (a',t) = del_min r in node22 l a' t)" |
       
   104 "del k (Node3 l a m b r) = (if k<a then node31 (del k l) a m b r else
       
   105   if k = a then let (a',m') = del_min m in node32 l a' m' b r else
       
   106   if k < b then node32 l a (del k m) b r else
       
   107   if k = b then let (b',r') = del_min r in node33 l a m b' r'
       
   108   else node33 l a m b (del k r))"
       
   109 
       
   110 definition delete :: "'a::linorder \<Rightarrow> 'a tree23 \<Rightarrow> 'a tree23" where
       
   111 "delete k t = tree\<^sub>d(del k t)"
       
   112 
       
   113 
       
   114 declare prod.splits [split]
       
   115 
       
   116 subsection "Functional Correctness"
       
   117 
       
   118 
       
   119 subsubsection "Proofs for isin"
       
   120 
       
   121 lemma "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
       
   122 by (induction t) (auto simp: elems_simps1)
       
   123 
       
   124 lemma isin_set: "sorted(inorder t) \<Longrightarrow> isin t x = (x \<in> elems (inorder t))"
       
   125 by (induction t) (auto simp: elems_simps2)
       
   126 
       
   127 
       
   128 subsubsection "Proofs for insert"
       
   129 
       
   130 lemma inorder_ins:
       
   131   "sorted(inorder t) \<Longrightarrow> inorder(tree\<^sub>i(ins x t)) = ins_list x (inorder t)"
       
   132 by(induction t) (auto simp: ins_list_simps split: up\<^sub>i.splits)
       
   133 
       
   134 lemma inorder_insert:
       
   135   "sorted(inorder t) \<Longrightarrow> inorder(insert a t) = ins_list a (inorder t)"
       
   136 by(simp add: insert_def inorder_ins)
       
   137 
       
   138 
       
   139 subsubsection "Proofs for delete"
       
   140 
       
   141 lemma inorder_node21: "height r > 0 \<Longrightarrow>
       
   142   inorder (tree\<^sub>d (node21 l' a r)) = inorder (tree\<^sub>d l') @ a # inorder r"
       
   143 by(induct l' a r rule: node21.induct) auto
       
   144 
       
   145 lemma inorder_node22: "height l > 0 \<Longrightarrow>
       
   146   inorder (tree\<^sub>d (node22 l a r')) = inorder l @ a # inorder (tree\<^sub>d r')"
       
   147 by(induct l a r' rule: node22.induct) auto
       
   148 
       
   149 lemma inorder_node31: "height m > 0 \<Longrightarrow>
       
   150   inorder (tree\<^sub>d (node31 l' a m b r)) = inorder (tree\<^sub>d l') @ a # inorder m @ b # inorder r"
       
   151 by(induct l' a m b r rule: node31.induct) auto
       
   152 
       
   153 lemma inorder_node32: "height r > 0 \<Longrightarrow>
       
   154   inorder (tree\<^sub>d (node32 l a m' b r)) = inorder l @ a # inorder (tree\<^sub>d m') @ b # inorder r"
       
   155 by(induct l a m' b r rule: node32.induct) auto
       
   156 
       
   157 lemma inorder_node33: "height m > 0 \<Longrightarrow>
       
   158   inorder (tree\<^sub>d (node33 l a m b r')) = inorder l @ a # inorder m @ b # inorder (tree\<^sub>d r')"
       
   159 by(induct l a m b r' rule: node33.induct) auto
       
   160 
       
   161 lemmas inorder_nodes = inorder_node21 inorder_node22
       
   162   inorder_node31 inorder_node32 inorder_node33
       
   163 
       
   164 lemma del_minD:
       
   165   "del_min t = (x,t') \<Longrightarrow> bal t \<Longrightarrow> height t > 0 \<Longrightarrow>
       
   166   x # inorder(tree\<^sub>d t') = inorder t"
       
   167 by(induction t arbitrary: t' rule: del_min.induct)
       
   168   (auto simp: inorder_nodes)
       
   169 
       
   170 lemma inorder_del: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
       
   171   inorder(tree\<^sub>d (del x t)) = del_list x (inorder t)"
       
   172 by(induction t rule: del.induct)
       
   173   (auto simp: del_list_simps inorder_nodes del_minD)
       
   174 
       
   175 lemma inorder_delete: "\<lbrakk> bal t ; sorted(inorder t) \<rbrakk> \<Longrightarrow>
       
   176   inorder(delete x t) = del_list x (inorder t)"
       
   177 by(simp add: delete_def inorder_del)
       
   178 
       
   179 
       
   180 subsection \<open>Balancedness\<close>
       
   181 
       
   182 
       
   183 subsubsection "Proofs for insert"
       
   184 
       
   185 text{* First a standard proof that @{const ins} preserves @{const bal}. *}
       
   186 
       
   187 instantiation up\<^sub>i :: (type)height
       
   188 begin
       
   189 
       
   190 fun height_up\<^sub>i :: "'a up\<^sub>i \<Rightarrow> nat" where
       
   191 "height (T\<^sub>i t) = height t" |
       
   192 "height (Up\<^sub>i l a r) = height l"
       
   193 
       
   194 instance ..
       
   195 
       
   196 end
       
   197 
       
   198 lemma bal_ins: "bal t \<Longrightarrow> bal (tree\<^sub>i(ins a t)) \<and> height(ins a t) = height t"
       
   199 by (induct t) (auto split: up\<^sub>i.split)
       
   200 
       
   201 text{* Now an alternative proof (by Brian Huffman) that runs faster because
       
   202 two properties (balance and height) are combined in one predicate. *}
       
   203 
       
   204 inductive full :: "nat \<Rightarrow> 'a tree23 \<Rightarrow> bool" where
       
   205 "full 0 Leaf" |
       
   206 "\<lbrakk>full n l; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Node2 l p r)" |
       
   207 "\<lbrakk>full n l; full n m; full n r\<rbrakk> \<Longrightarrow> full (Suc n) (Node3 l p m q r)"
       
   208 
       
   209 inductive_cases full_elims:
       
   210   "full n Leaf"
       
   211   "full n (Node2 l p r)"
       
   212   "full n (Node3 l p m q r)"
       
   213 
       
   214 inductive_cases full_0_elim: "full 0 t"
       
   215 inductive_cases full_Suc_elim: "full (Suc n) t"
       
   216 
       
   217 lemma full_0_iff [simp]: "full 0 t \<longleftrightarrow> t = Leaf"
       
   218   by (auto elim: full_0_elim intro: full.intros)
       
   219 
       
   220 lemma full_Leaf_iff [simp]: "full n Leaf \<longleftrightarrow> n = 0"
       
   221   by (auto elim: full_elims intro: full.intros)
       
   222 
       
   223 lemma full_Suc_Node2_iff [simp]:
       
   224   "full (Suc n) (Node2 l p r) \<longleftrightarrow> full n l \<and> full n r"
       
   225   by (auto elim: full_elims intro: full.intros)
       
   226 
       
   227 lemma full_Suc_Node3_iff [simp]:
       
   228   "full (Suc n) (Node3 l p m q r) \<longleftrightarrow> full n l \<and> full n m \<and> full n r"
       
   229   by (auto elim: full_elims intro: full.intros)
       
   230 
       
   231 lemma full_imp_height: "full n t \<Longrightarrow> height t = n"
       
   232   by (induct set: full, simp_all)
       
   233 
       
   234 lemma full_imp_bal: "full n t \<Longrightarrow> bal t"
       
   235   by (induct set: full, auto dest: full_imp_height)
       
   236 
       
   237 lemma bal_imp_full: "bal t \<Longrightarrow> full (height t) t"
       
   238   by (induct t, simp_all)
       
   239 
       
   240 lemma bal_iff_full: "bal t \<longleftrightarrow> (\<exists>n. full n t)"
       
   241   by (auto elim!: bal_imp_full full_imp_bal)
       
   242 
       
   243 text {* The @{const "insert"} function either preserves the height of the
       
   244 tree, or increases it by one. The constructor returned by the @{term
       
   245 "insert"} function determines which: A return value of the form @{term
       
   246 "T\<^sub>i t"} indicates that the height will be the same. A value of the
       
   247 form @{term "Up\<^sub>i l p r"} indicates an increase in height. *}
       
   248 
       
   249 fun full\<^sub>i :: "nat \<Rightarrow> 'a up\<^sub>i \<Rightarrow> bool" where
       
   250 "full\<^sub>i n (T\<^sub>i t) \<longleftrightarrow> full n t" |
       
   251 "full\<^sub>i n (Up\<^sub>i l p r) \<longleftrightarrow> full n l \<and> full n r"
       
   252 
       
   253 lemma full\<^sub>i_ins: "full n t \<Longrightarrow> full\<^sub>i n (ins a t)"
       
   254 by (induct rule: full.induct) (auto split: up\<^sub>i.split)
       
   255 
       
   256 text {* The @{const insert} operation preserves balance. *}
       
   257 
       
   258 lemma bal_insert: "bal t \<Longrightarrow> bal (insert a t)"
       
   259 unfolding bal_iff_full insert_def
       
   260 apply (erule exE)
       
   261 apply (drule full\<^sub>i_ins [of _ _ a])
       
   262 apply (cases "ins a t")
       
   263 apply (auto intro: full.intros)
       
   264 done
       
   265 
       
   266 
       
   267 subsection "Proofs for delete"
       
   268 
       
   269 instantiation up\<^sub>d :: (type)height
       
   270 begin
       
   271 
       
   272 fun height_up\<^sub>d :: "'a up\<^sub>d \<Rightarrow> nat" where
       
   273 "height (T\<^sub>d t) = height t" |
       
   274 "height (Up\<^sub>d t) = height t + 1"
       
   275 
       
   276 instance ..
       
   277 
       
   278 end
       
   279 
       
   280 lemma bal_tree\<^sub>d_node21:
       
   281   "\<lbrakk>bal r; bal (tree\<^sub>d l'); height r = height l' \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d (node21 l' a r))"
       
   282 by(induct l' a r rule: node21.induct) auto
       
   283 
       
   284 lemma bal_tree\<^sub>d_node22:
       
   285   "\<lbrakk>bal(tree\<^sub>d r'); bal l; height r' = height l \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d (node22 l a r'))"
       
   286 by(induct l a r' rule: node22.induct) auto
       
   287 
       
   288 lemma bal_tree\<^sub>d_node31:
       
   289   "\<lbrakk> bal (tree\<^sub>d l'); bal m; bal r; height l' = height r; height m = height r \<rbrakk>
       
   290   \<Longrightarrow> bal (tree\<^sub>d (node31 l' a m b r))"
       
   291 by(induct l' a m b r rule: node31.induct) auto
       
   292 
       
   293 lemma bal_tree\<^sub>d_node32:
       
   294   "\<lbrakk> bal l; bal (tree\<^sub>d m'); bal r; height l = height r; height m' = height r \<rbrakk>
       
   295   \<Longrightarrow> bal (tree\<^sub>d (node32 l a m' b r))"
       
   296 by(induct l a m' b r rule: node32.induct) auto
       
   297 
       
   298 lemma bal_tree\<^sub>d_node33:
       
   299   "\<lbrakk> bal l; bal m; bal(tree\<^sub>d r'); height l = height r'; height m = height r' \<rbrakk>
       
   300   \<Longrightarrow> bal (tree\<^sub>d (node33 l a m b r'))"
       
   301 by(induct l a m b r' rule: node33.induct) auto
       
   302 
       
   303 lemmas bals = bal_tree\<^sub>d_node21 bal_tree\<^sub>d_node22
       
   304   bal_tree\<^sub>d_node31 bal_tree\<^sub>d_node32 bal_tree\<^sub>d_node33
       
   305 
       
   306 lemma height'_node21:
       
   307    "height r > 0 \<Longrightarrow> height(node21 l' a r) = max (height l') (height r) + 1"
       
   308 by(induct l' a r rule: node21.induct)(simp_all)
       
   309 
       
   310 lemma height'_node22:
       
   311    "height l > 0 \<Longrightarrow> height(node22 l a r') = max (height l) (height r') + 1"
       
   312 by(induct l a r' rule: node22.induct)(simp_all)
       
   313 
       
   314 lemma height'_node31:
       
   315   "height m > 0 \<Longrightarrow> height(node31 l a m b r) =
       
   316    max (height l) (max (height m) (height r)) + 1"
       
   317 by(induct l a m b r rule: node31.induct)(simp_all add: max_def)
       
   318 
       
   319 lemma height'_node32:
       
   320   "height r > 0 \<Longrightarrow> height(node32 l a m b r) =
       
   321    max (height l) (max (height m) (height r)) + 1"
       
   322 by(induct l a m b r rule: node32.induct)(simp_all add: max_def)
       
   323 
       
   324 lemma height'_node33:
       
   325   "height m > 0 \<Longrightarrow> height(node33 l a m b r) =
       
   326    max (height l) (max (height m) (height r)) + 1"
       
   327 by(induct l a m b r rule: node33.induct)(simp_all add: max_def)
       
   328 
       
   329 lemmas heights = height'_node21 height'_node22
       
   330   height'_node31 height'_node32 height'_node33
       
   331 
       
   332 lemma height_del_min:
       
   333   "del_min t = (x, t') \<Longrightarrow> height t > 0 \<Longrightarrow> bal t \<Longrightarrow> height t' = height t"
       
   334 by(induct t arbitrary: x t' rule: del_min.induct)
       
   335   (auto simp: heights split: prod.splits)
       
   336 
       
   337 lemma height_del: "bal t \<Longrightarrow> height(del x t) = height t"
       
   338 by(induction x t rule: del.induct)
       
   339   (auto simp add: heights max_def height_del_min)
       
   340 
       
   341 lemma bal_del_min:
       
   342   "\<lbrakk> del_min t = (x, t'); bal t; height t > 0 \<rbrakk> \<Longrightarrow> bal (tree\<^sub>d t')"
       
   343 by(induct t arbitrary: x t' rule: del_min.induct)
       
   344   (auto simp: heights height_del_min bals)
       
   345 
       
   346 lemma bal_tree\<^sub>d_del: "bal t \<Longrightarrow> bal(tree\<^sub>d(del x t))"
       
   347 by(induction x t rule: del.induct)
       
   348   (auto simp: bals bal_del_min height_del height_del_min)
       
   349 corollary bal_delete: "bal t \<Longrightarrow> bal(delete x t)"
       
   350 by(simp add: delete_def bal_tree\<^sub>d_del)
       
   351 
       
   352 
       
   353 subsection \<open>Overall Correctness\<close>
       
   354 
       
   355 interpretation Set_by_Ordered
       
   356 where empty = Leaf and isin = isin and insert = insert and delete = delete
       
   357 and inorder = inorder and wf = bal
       
   358 proof (standard, goal_cases)
       
   359   case 2 thus ?case by(simp add: isin_set)
       
   360 next
       
   361   case 3 thus ?case by(simp add: inorder_insert)
       
   362 next
       
   363   case 4 thus ?case by(simp add: inorder_delete)
       
   364 next
       
   365   case 6 thus ?case by(simp add: bal_insert)
       
   366 next
       
   367   case 7 thus ?case by(simp add: bal_delete)
       
   368 qed simp+
       
   369 
       
   370 end