fix proofs related to simplification of inequalities on numerals
authorhuffman
Wed, 03 Dec 2008 21:00:39 -0800
changeset 28963 f6d9e0e0b153
parent 28962 f603183f7a5c
child 28964 f0044cdeb945
child 28965 1de908189869
child 28967 3bdb1eae352c
fix proofs related to simplification of inequalities on numerals
src/HOL/Library/Float.thy
--- a/src/HOL/Library/Float.thy	Wed Dec 03 20:45:42 2008 -0800
+++ b/src/HOL/Library/Float.thy	Wed Dec 03 21:00:39 2008 -0800
@@ -514,7 +514,7 @@
   by simp
 
 lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
-  by simp
+  unfolding neg_def number_of_is_id by simp
 
 lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
   by simp
@@ -529,7 +529,7 @@
   by simp
 
 lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
-  by simp
+  unfolding neg_def number_of_is_id by (simp add: not_less)
 
 lemmas intarithrel =
   int_eq_number_of_eq