# HG changeset patch # User huffman # Date 1329992442 -3600 # Node ID fd0ac848140e724ddccd9e5771ba7ca549496e14 # Parent 7fc239ebece2ac043ac977c089d03c31887eeb9d simplify proof diff -r 7fc239ebece2 -r fd0ac848140e src/HOL/Word/Bit_Representation.thy --- a/src/HOL/Word/Bit_Representation.thy Thu Feb 23 08:59:55 2012 +0100 +++ b/src/HOL/Word/Bit_Representation.thy Thu Feb 23 11:20:42 2012 +0100 @@ -166,7 +166,7 @@ lemma bin_abs_lem: "bin = (w BIT b) ==> ~ bin = Int.Min --> ~ bin = Int.Pls --> nat (abs w) < nat (abs bin)" - apply (clarsimp simp add: bin_rl_char) + apply clarsimp apply (unfold Pls_def Min_def Bit_def) apply (cases b) apply (clarsimp, arith)