# HG changeset patch # User nipkow # Date 1497517896 -7200 # Node ID c9c9438cfc0f5bf1498e0ae4fec183ec1191bfd3 # Parent 6e0c330f4051e13f2229261d307df657d7212440 tuned diff -r 6e0c330f4051 -r c9c9438cfc0f src/HOL/Data_Structures/RBT_Set.thy --- a/src/HOL/Data_Structures/RBT_Set.thy Wed Jun 14 19:39:12 2017 +0200 +++ b/src/HOL/Data_Structures/RBT_Set.thy Thu Jun 15 11:11:36 2017 +0200 @@ -228,23 +228,23 @@ lemma neq_LeafD: "t \ Leaf \ \c l x r. t = Node c l x r" by(cases t) auto -lemma del_invc_invh: "invh l \ invc l \ invh (del x l) \ - (color l = Red \ bheight (del x l) = bheight l \ invc (del x l) \ - color l = Black \ bheight (del x l) = bheight l - 1 \ invc2 (del x l))" -proof (induct x l rule: del.induct) -case (2 y c _ y' W) - have "y = y' \ y < y' \ y > y'" by auto +lemma del_invc_invh: "invh t \ invc t \ invh (del x t) \ + (color t = Red \ bheight (del x t) = bheight t \ invc (del x t) \ + color t = Black \ bheight (del x t) = bheight t - 1 \ invc2 (del x t))" +proof (induct x t rule: del.induct) +case (2 x c _ y) + have "x = y \ x < y \ x > y" by auto thus ?case proof (elim disjE) - assume "y = y'" + assume "x = y" with 2 show ?thesis by (cases c) (simp_all add: invh_combine invc_combine) next - assume "y < y'" + assume "x < y" with 2 show ?thesis by(cases c) (auto simp: invh_baldL_invc invc_baldL invc2_baldL dest: neq_LeafD) next - assume "y' < y" + assume "y < x" with 2 show ?thesis by(cases c) (auto simp: invh_baldR_invc invc_baldR invc2_baldR dest: neq_LeafD)