| 
61224
 | 
     1  | 
(* Author: Tobias Nipkow *)
  | 
| 
 | 
     2  | 
  | 
| 
 | 
     3  | 
section \<open>Red-Black Tree Implementation of Sets\<close>
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
theory RBT_Set
  | 
| 
 | 
     6  | 
imports
  | 
| 
 | 
     7  | 
  RBT
  | 
| 
61581
 | 
     8  | 
  Cmp
  | 
| 
61224
 | 
     9  | 
  Isin2
  | 
| 
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
61749
 | 
    12  | 
fun ins :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
  | 
| 
 | 
    13  | 
"ins x Leaf = R Leaf x Leaf" |
  | 
| 
 | 
    14  | 
"ins x (B l a r) =
  | 
| 
61678
 | 
    15  | 
  (case cmp x a of
  | 
| 
61749
 | 
    16  | 
     LT \<Rightarrow> bal (ins x l) a r |
  | 
| 
 | 
    17  | 
     GT \<Rightarrow> bal l a (ins x r) |
  | 
| 
61678
 | 
    18  | 
     EQ \<Rightarrow> B l a r)" |
  | 
| 
61749
 | 
    19  | 
"ins x (R l a r) =
  | 
| 
61678
 | 
    20  | 
  (case cmp x a of
  | 
| 
61749
 | 
    21  | 
    LT \<Rightarrow> R (ins x l) a r |
  | 
| 
 | 
    22  | 
    GT \<Rightarrow> R l a (ins x r) |
  | 
| 
61678
 | 
    23  | 
    EQ \<Rightarrow> R l a r)"
  | 
| 
61224
 | 
    24  | 
  | 
| 
61749
 | 
    25  | 
definition insert :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
  | 
| 
 | 
    26  | 
"insert x t = paint Black (ins x t)"
  | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
fun del :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
  | 
| 
 | 
    29  | 
and delL :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
  | 
| 
 | 
    30  | 
and delR :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt"
  | 
| 
61224
 | 
    31  | 
where
  | 
| 
61749
 | 
    32  | 
"del x Leaf = Leaf" |
  | 
| 
 | 
    33  | 
"del x (Node _ l a r) =
  | 
| 
61678
 | 
    34  | 
  (case cmp x a of
  | 
| 
61749
 | 
    35  | 
     LT \<Rightarrow> delL x l a r |
  | 
| 
 | 
    36  | 
     GT \<Rightarrow> delR x l a r |
  | 
| 
61678
 | 
    37  | 
     EQ \<Rightarrow> combine l r)" |
  | 
| 
61749
 | 
    38  | 
"delL x (B t1 a t2) b t3 = balL (del x (B t1 a t2)) b t3" |
  | 
| 
 | 
    39  | 
"delL x l a r = R (del x l) a r" |
  | 
| 
 | 
    40  | 
"delR x t1 a (B t2 b t3) = balR t1 a (del x (B t2 b t3))" | 
  | 
| 
 | 
    41  | 
"delR x l a r = R l a (del x r)"
  | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
definition delete :: "'a::cmp \<Rightarrow> 'a rbt \<Rightarrow> 'a rbt" where
  | 
| 
 | 
    44  | 
"delete x t = paint Black (del x t)"
  | 
| 
61224
 | 
    45  | 
  | 
| 
 | 
    46  | 
  | 
| 
 | 
    47  | 
subsection "Functional Correctness Proofs"
  | 
| 
 | 
    48  | 
  | 
| 
61749
 | 
    49  | 
lemma inorder_paint: "inorder(paint c t) = inorder t"
  | 
| 
62526
 | 
    50  | 
by(cases t) (auto)
  | 
| 
61749
 | 
    51  | 
  | 
| 
61224
 | 
    52  | 
lemma inorder_bal:
  | 
| 
 | 
    53  | 
  "inorder(bal l a r) = inorder l @ a # inorder r"
  | 
| 
62526
 | 
    54  | 
by(cases "(l,a,r)" rule: bal.cases) (auto)
  | 
| 
61224
 | 
    55  | 
  | 
| 
61749
 | 
    56  | 
lemma inorder_ins:
  | 
| 
 | 
    57  | 
  "sorted(inorder t) \<Longrightarrow> inorder(ins x t) = ins_list x (inorder t)"
  | 
| 
 | 
    58  | 
by(induction x t rule: ins.induct) (auto simp: ins_list_simps inorder_bal)
  | 
| 
 | 
    59  | 
  | 
| 
61224
 | 
    60  | 
lemma inorder_insert:
  | 
| 
61749
 | 
    61  | 
  "sorted(inorder t) \<Longrightarrow> inorder(insert x t) = ins_list x (inorder t)"
  | 
| 
 | 
    62  | 
by (simp add: insert_def inorder_ins inorder_paint)
  | 
| 
61224
 | 
    63  | 
  | 
| 
 | 
    64  | 
lemma inorder_balL:
  | 
| 
 | 
    65  | 
  "inorder(balL l a r) = inorder l @ a # inorder r"
  | 
| 
62526
 | 
    66  | 
by(cases "(l,a,r)" rule: balL.cases)(auto simp: inorder_bal inorder_paint)
  | 
| 
61224
 | 
    67  | 
  | 
| 
 | 
    68  | 
lemma inorder_balR:
  | 
| 
 | 
    69  | 
  "inorder(balR l a r) = inorder l @ a # inorder r"
  | 
| 
62526
 | 
    70  | 
by(cases "(l,a,r)" rule: balR.cases) (auto simp: inorder_bal inorder_paint)
  | 
| 
61224
 | 
    71  | 
  | 
| 
 | 
    72  | 
lemma inorder_combine:
  | 
| 
 | 
    73  | 
  "inorder(combine l r) = inorder l @ inorder r"
  | 
| 
 | 
    74  | 
by(induction l r rule: combine.induct)
  | 
| 
61231
 | 
    75  | 
  (auto simp: inorder_balL inorder_balR split: tree.split color.split)
  | 
| 
61224
 | 
    76  | 
  | 
| 
61749
 | 
    77  | 
lemma inorder_del:
  | 
| 
 | 
    78  | 
 "sorted(inorder t) \<Longrightarrow>  inorder(del x t) = del_list x (inorder t)"
  | 
| 
 | 
    79  | 
 "sorted(inorder l) \<Longrightarrow>  inorder(delL x l a r) =
  | 
| 
61678
 | 
    80  | 
    del_list x (inorder l) @ a # inorder r"
  | 
| 
61749
 | 
    81  | 
 "sorted(inorder r) \<Longrightarrow>  inorder(delR x l a r) =
  | 
| 
61224
 | 
    82  | 
    inorder l @ a # del_list x (inorder r)"
  | 
| 
61749
 | 
    83  | 
by(induction x t and x l a r and x l a r rule: del_delL_delR.induct)
  | 
| 
61231
 | 
    84  | 
  (auto simp: del_list_simps inorder_combine inorder_balL inorder_balR)
  | 
| 
61224
 | 
    85  | 
  | 
| 
61749
 | 
    86  | 
lemma inorder_delete:
  | 
| 
 | 
    87  | 
  "sorted(inorder t) \<Longrightarrow> inorder(delete x t) = del_list x (inorder t)"
  | 
| 
 | 
    88  | 
by (auto simp: delete_def inorder_del inorder_paint)
  | 
| 
 | 
    89  | 
  | 
| 
61581
 | 
    90  | 
  | 
| 
61224
 | 
    91  | 
interpretation Set_by_Ordered
  | 
| 
 | 
    92  | 
where empty = Leaf and isin = isin and insert = insert and delete = delete
  | 
| 
61588
 | 
    93  | 
and inorder = inorder and inv = "\<lambda>_. True"
  | 
| 
61224
 | 
    94  | 
proof (standard, goal_cases)
  | 
| 
 | 
    95  | 
  case 1 show ?case by simp
  | 
| 
 | 
    96  | 
next
  | 
| 
 | 
    97  | 
  case 2 thus ?case by(simp add: isin_set)
  | 
| 
 | 
    98  | 
next
  | 
| 
 | 
    99  | 
  case 3 thus ?case by(simp add: inorder_insert)
  | 
| 
 | 
   100  | 
next
  | 
| 
61749
 | 
   101  | 
  case 4 thus ?case by(simp add: inorder_delete)
  | 
| 
 | 
   102  | 
qed auto
  | 
| 
61224
 | 
   103  | 
  | 
| 
61754
 | 
   104  | 
  | 
| 
 | 
   105  | 
subsection \<open>Structural invariants\<close>
  | 
| 
 | 
   106  | 
  | 
| 
 | 
   107  | 
fun color :: "'a rbt \<Rightarrow> color" where
  | 
| 
 | 
   108  | 
"color Leaf = Black" |
  | 
| 
 | 
   109  | 
"color (Node c _ _ _) = c"
  | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
fun bheight :: "'a rbt \<Rightarrow> nat" where
  | 
| 
 | 
   112  | 
"bheight Leaf = 0" |
  | 
| 
 | 
   113  | 
"bheight (Node c l x r) = (if c = Black then Suc(bheight l) else bheight l)"
  | 
| 
 | 
   114  | 
  | 
| 
 | 
   115  | 
fun inv1 :: "'a rbt \<Rightarrow> bool" where
  | 
| 
 | 
   116  | 
"inv1 Leaf = True" |
  | 
| 
 | 
   117  | 
"inv1 (Node c l a r) =
  | 
| 
 | 
   118  | 
  (inv1 l \<and> inv1 r \<and> (c = Black \<or> color l = Black \<and> color r = Black))"
  | 
| 
 | 
   119  | 
  | 
| 
 | 
   120  | 
fun inv1_root :: "'a rbt \<Rightarrow> bool" \<comment> \<open>Weaker version\<close> where
  | 
| 
 | 
   121  | 
"inv1_root Leaf = True" |
  | 
| 
 | 
   122  | 
"inv1_root (Node c l a r) = (inv1 l \<and> inv1 r)"
  | 
| 
 | 
   123  | 
  | 
| 
 | 
   124  | 
fun inv2 :: "'a rbt \<Rightarrow> bool" where
  | 
| 
 | 
   125  | 
"inv2 Leaf = True" |
  | 
| 
 | 
   126  | 
"inv2 (Node c l x r) = (inv2 l \<and> inv2 r \<and> bheight l = bheight r)"
  | 
| 
 | 
   127  | 
  | 
| 
 | 
   128  | 
lemma inv1_rootI[simp]: "inv1 t \<Longrightarrow> inv1_root t"
  | 
| 
 | 
   129  | 
by (cases t) simp+
  | 
| 
 | 
   130  | 
  | 
| 
 | 
   131  | 
definition rbt :: "'a rbt \<Rightarrow> bool" where
  | 
| 
 | 
   132  | 
"rbt t = (inv1 t \<and> inv2 t \<and> color t = Black)"
  | 
| 
 | 
   133  | 
  | 
| 
 | 
   134  | 
lemma color_paint_Black: "color (paint Black t) = Black"
  | 
| 
 | 
   135  | 
by (cases t) auto
  | 
| 
 | 
   136  | 
  | 
| 
 | 
   137  | 
theorem rbt_Leaf: "rbt Leaf"
  | 
| 
 | 
   138  | 
by (simp add: rbt_def)
  | 
| 
 | 
   139  | 
  | 
| 
 | 
   140  | 
lemma paint_inv1_root: "inv1_root t \<Longrightarrow> inv1_root (paint c t)"
  | 
| 
 | 
   141  | 
by (cases t) auto
  | 
| 
 | 
   142  | 
  | 
| 
 | 
   143  | 
lemma inv1_paint_Black: "inv1_root t \<Longrightarrow> inv1 (paint Black t)"
  | 
| 
 | 
   144  | 
by (cases t) auto
  | 
| 
 | 
   145  | 
  | 
| 
 | 
   146  | 
lemma inv2_paint: "inv2 t \<Longrightarrow> inv2 (paint c t)"
  | 
| 
 | 
   147  | 
by (cases t) auto
  | 
| 
 | 
   148  | 
  | 
| 
 | 
   149  | 
lemma inv1_bal: "\<lbrakk>inv1_root l; inv1_root r\<rbrakk> \<Longrightarrow> inv1 (bal l a r)" 
  | 
| 
 | 
   150  | 
by (induct l a r rule: bal.induct) auto
  | 
| 
 | 
   151  | 
  | 
| 
 | 
   152  | 
lemma bheight_bal:
  | 
| 
 | 
   153  | 
  "bheight l = bheight r \<Longrightarrow> bheight (bal l a r) = Suc (bheight l)"
  | 
| 
 | 
   154  | 
by (induct l a r rule: bal.induct) auto
  | 
| 
 | 
   155  | 
  | 
| 
 | 
   156  | 
lemma inv2_bal: 
  | 
| 
 | 
   157  | 
  "\<lbrakk> inv2 l; inv2 r; bheight l = bheight r \<rbrakk> \<Longrightarrow> inv2 (bal l a r)"
  | 
| 
 | 
   158  | 
by (induct l a r rule: bal.induct) auto
  | 
| 
 | 
   159  | 
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
subsubsection \<open>Insertion\<close>
  | 
| 
 | 
   162  | 
  | 
| 
 | 
   163  | 
lemma inv1_ins: assumes "inv1 t"
  | 
| 
 | 
   164  | 
  shows "color t = Black \<Longrightarrow> inv1 (ins x t)" "inv1_root (ins x t)"
  | 
| 
 | 
   165  | 
using assms
  | 
| 
 | 
   166  | 
by (induct x t rule: ins.induct) (auto simp: inv1_bal)
  | 
| 
 | 
   167  | 
  | 
| 
 | 
   168  | 
lemma inv2_ins: assumes "inv2 t"
  | 
| 
 | 
   169  | 
  shows "inv2 (ins x t)" "bheight (ins x t) = bheight t"
  | 
| 
 | 
   170  | 
using assms
  | 
| 
 | 
   171  | 
by (induct x t rule: ins.induct) (auto simp: inv2_bal bheight_bal)
  | 
| 
 | 
   172  | 
  | 
| 
 | 
   173  | 
theorem rbt_ins: "rbt t \<Longrightarrow> rbt (insert x t)"
  | 
| 
 | 
   174  | 
by (simp add: inv1_ins inv2_ins color_paint_Black inv1_paint_Black inv2_paint
  | 
| 
 | 
   175  | 
  rbt_def insert_def)
  | 
| 
 | 
   176  | 
  | 
| 
 | 
   177  | 
(*
  | 
| 
 | 
   178  | 
lemma bheight_paintR'[simp]: "color t = Black \<Longrightarrow> bheight (paint Red t) = bheight t - 1"
  | 
| 
 | 
   179  | 
by (cases t) auto
  | 
| 
 | 
   180  | 
  | 
| 
 | 
   181  | 
lemma balL_inv2_with_inv1:
  | 
| 
 | 
   182  | 
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "inv1 rt"
  | 
| 
 | 
   183  | 
  shows "bheight (balL lt a rt) = bheight lt + 1"  "inv2 (balL lt a rt)"
  | 
| 
 | 
   184  | 
using assms 
  | 
| 
 | 
   185  | 
by (induct lt a rt rule: balL.induct) (auto simp: inv2_bal inv2_paint bheight_bal)
  | 
| 
 | 
   186  | 
  | 
| 
 | 
   187  | 
lemma balL_inv2_app: 
  | 
| 
 | 
   188  | 
  assumes "inv2 lt" "inv2 rt" "bheight lt + 1 = bheight rt" "color rt = Black"
  | 
| 
 | 
   189  | 
  shows "inv2 (balL lt a rt)" 
  | 
| 
 | 
   190  | 
        "bheight (balL lt a rt) = bheight rt"
  | 
| 
 | 
   191  | 
using assms 
  | 
| 
 | 
   192  | 
by (induct lt a rt rule: balL.induct) (auto simp add: inv2_bal bheight_bal) 
  | 
| 
 | 
   193  | 
  | 
| 
 | 
   194  | 
lemma balL_inv1: "\<lbrakk>inv1_root l; inv1 r; color r = Black\<rbrakk> \<Longrightarrow> inv1 (balL l a r)"
  | 
| 
 | 
   195  | 
by (induct l a r rule: balL.induct) (simp_all add: inv1_bal)
  | 
| 
 | 
   196  | 
  | 
| 
 | 
   197  | 
lemma balL_inv1_root: "\<lbrakk> inv1_root lt; inv1 rt \<rbrakk> \<Longrightarrow> inv1_root (balL lt a rt)"
  | 
| 
 | 
   198  | 
by (induct lt a rt rule: balL.induct) (auto simp: inv1_bal paint_inv1_root)
  | 
| 
 | 
   199  | 
  | 
| 
 | 
   200  | 
lemma balR_inv2_with_inv1:
  | 
| 
 | 
   201  | 
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt + 1" "inv1 lt"
  | 
| 
 | 
   202  | 
  shows "inv2 (balR lt a rt) \<and> bheight (balR lt a rt) = bheight lt"
  | 
| 
 | 
   203  | 
using assms
  | 
| 
 | 
   204  | 
by(induct lt a rt rule: balR.induct)(auto simp: inv2_bal bheight_bal inv2_paint)
  | 
| 
 | 
   205  | 
  | 
| 
 | 
   206  | 
lemma balR_inv1: "\<lbrakk>inv1 a; inv1_root b; color a = Black\<rbrakk> \<Longrightarrow> inv1 (balR a x b)"
  | 
| 
 | 
   207  | 
by (induct a x b rule: balR.induct) (simp_all add: inv1_bal)
  | 
| 
 | 
   208  | 
  | 
| 
 | 
   209  | 
lemma balR_inv1_root: "\<lbrakk> inv1 lt; inv1_root rt \<rbrakk> \<Longrightarrow>inv1_root (balR lt x rt)"
  | 
| 
 | 
   210  | 
by (induct lt x rt rule: balR.induct) (auto simp: inv1_bal paint_inv1_root)
  | 
| 
 | 
   211  | 
  | 
| 
 | 
   212  | 
lemma combine_inv2:
  | 
| 
 | 
   213  | 
  assumes "inv2 lt" "inv2 rt" "bheight lt = bheight rt"
  | 
| 
 | 
   214  | 
  shows "bheight (combine lt rt) = bheight lt" "inv2 (combine lt rt)"
  | 
| 
 | 
   215  | 
using assms 
  | 
| 
 | 
   216  | 
by (induct lt rt rule: combine.induct) 
  | 
| 
 | 
   217  | 
   (auto simp: balL_inv2_app split: tree.splits color.splits)
  | 
| 
 | 
   218  | 
  | 
| 
 | 
   219  | 
lemma combine_inv1: 
  | 
| 
 | 
   220  | 
  assumes "inv1 lt" "inv1 rt"
  | 
| 
 | 
   221  | 
  shows "color lt = Black \<Longrightarrow> color rt = Black \<Longrightarrow> inv1 (combine lt rt)"
  | 
| 
 | 
   222  | 
         "inv1_root (combine lt rt)"
  | 
| 
 | 
   223  | 
using assms 
  | 
| 
 | 
   224  | 
by (induct lt rt rule: combine.induct)
  | 
| 
 | 
   225  | 
   (auto simp: balL_inv1 split: tree.splits color.splits)
  | 
| 
 | 
   226  | 
  | 
| 
 | 
   227  | 
  | 
| 
 | 
   228  | 
lemma 
  | 
| 
 | 
   229  | 
  assumes "inv2 lt" "inv1 lt"
  | 
| 
 | 
   230  | 
  shows
  | 
| 
 | 
   231  | 
  "\<lbrakk>inv2 rt; bheight lt = bheight rt; inv1 rt\<rbrakk> \<Longrightarrow>
  | 
| 
 | 
   232  | 
   inv2 (rbt_del_from_left x lt k v rt) \<and> 
  | 
| 
 | 
   233  | 
   bheight (rbt_del_from_left x lt k v rt) = bheight lt \<and> 
  | 
| 
 | 
   234  | 
   (color_of lt = B \<and> color_of rt = B \<and> inv1 (rbt_del_from_left x lt k v rt) \<or> 
  | 
| 
 | 
   235  | 
    (color_of lt \<noteq> B \<or> color_of rt \<noteq> B) \<and> inv1l (rbt_del_from_left x lt k v rt))"
  | 
| 
 | 
   236  | 
  and "\<lbrakk>inv2 rt; bheight lt = bheight rt; inv1 rt\<rbrakk> \<Longrightarrow>
  | 
| 
 | 
   237  | 
  inv2 (rbt_del_from_right x lt k v rt) \<and> 
  | 
| 
 | 
   238  | 
  bheight (rbt_del_from_right x lt k v rt) = bheight lt \<and> 
  | 
| 
 | 
   239  | 
  (color_of lt = B \<and> color_of rt = B \<and> inv1 (rbt_del_from_right x lt k v rt) \<or> 
  | 
| 
 | 
   240  | 
   (color_of lt \<noteq> B \<or> color_of rt \<noteq> B) \<and> inv1l (rbt_del_from_right x lt k v rt))"
  | 
| 
 | 
   241  | 
  and rbt_del_inv1_inv2: "inv2 (rbt_del x lt) \<and> (color_of lt = R \<and> bheight (rbt_del x lt) = bheight lt \<and> inv1 (rbt_del x lt) 
  | 
| 
 | 
   242  | 
  \<or> color_of lt = B \<and> bheight (rbt_del x lt) = bheight lt - 1 \<and> inv1l (rbt_del x lt))"
  | 
| 
 | 
   243  | 
using assms
  | 
| 
 | 
   244  | 
proof (induct x lt k v rt and x lt k v rt and x lt rule: rbt_del_from_left_rbt_del_from_right_rbt_del.induct)
  | 
| 
 | 
   245  | 
case (2 y c _ y')
  | 
| 
 | 
   246  | 
  have "y = y' \<or> y < y' \<or> y > y'" by auto
  | 
| 
 | 
   247  | 
  thus ?case proof (elim disjE)
  | 
| 
 | 
   248  | 
    assume "y = y'"
  | 
| 
 | 
   249  | 
    with 2 show ?thesis by (cases c) (simp add: combine_inv2 combine_inv1)+
  | 
| 
 | 
   250  | 
  next
  | 
| 
 | 
   251  | 
    assume "y < y'"
  | 
| 
 | 
   252  | 
    with 2 show ?thesis by (cases c) auto
  | 
| 
 | 
   253  | 
  next
  | 
| 
 | 
   254  | 
    assume "y' < y"
  | 
| 
 | 
   255  | 
    with 2 show ?thesis by (cases c) auto
  | 
| 
 | 
   256  | 
  qed
  | 
| 
 | 
   257  | 
next
  | 
| 
 | 
   258  | 
  case (3 y lt z v rta y' ss bb) 
  | 
| 
 | 
   259  | 
  thus ?case by (cases "color_of (Branch B lt z v rta) = B \<and> color_of bb = B") (simp add: balance_left_inv2_with_inv1 balance_left_inv1 balance_left_inv1l)+
  | 
| 
 | 
   260  | 
next
  | 
| 
 | 
   261  | 
  case (5 y a y' ss lt z v rta)
  | 
| 
 | 
   262  | 
  thus ?case by (cases "color_of a = B \<and> color_of (Branch B lt z v rta) = B") (simp add: balance_right_inv2_with_inv1 balance_right_inv1 balance_right_inv1l)+
  | 
| 
 | 
   263  | 
next
  | 
| 
 | 
   264  | 
  case ("6_1" y a y' ss) thus ?case by (cases "color_of a = B \<and> color_of Empty = B") simp+
 | 
| 
 | 
   265  | 
qed auto
  | 
| 
 | 
   266  | 
  | 
| 
 | 
   267  | 
theorem rbt_delete_is_rbt [simp]: assumes "rbt t" shows "rbt (delete k t)"
  | 
| 
 | 
   268  | 
proof -
  | 
| 
 | 
   269  | 
  from assms have "inv2 t" and "inv1 t" unfolding rbt_def by auto 
  | 
| 
 | 
   270  | 
  hence "inv2 (del k t) \<and> (color t = Red \<and> bheight (del k t) = bheight t \<and> inv1 (del k t) \<or> color t = Black \<and> bheight (del k t) = bheight t - 1 \<and> inv1_root (del k t))"
  | 
| 
 | 
   271  | 
    by (rule rbt_del_inv1_inv2)
  | 
| 
 | 
   272  | 
  hence "inv2 (del k t) \<and> inv1l (rbt_del k t)" by (cases "color_of t") auto
  | 
| 
 | 
   273  | 
  with assms show ?thesis
  | 
| 
 | 
   274  | 
    unfolding is_rbt_def rbt_delete_def
  | 
| 
 | 
   275  | 
    by (auto intro: paint_rbt_sorted rbt_del_rbt_sorted)
  | 
| 
 | 
   276  | 
qed
  | 
| 
 | 
   277  | 
*)
  | 
| 
61224
 | 
   278  | 
end
  |