--- a/src/HOL/Archimedean_Field.thy Thu Feb 26 14:16:30 2009 +0100
+++ b/src/HOL/Archimedean_Field.thy Thu Feb 26 15:42:35 2009 +0100
@@ -391,10 +391,10 @@
subsection {* Negation *}
-lemma floor_minus [simp]: "floor (- x) = - ceiling x"
+lemma floor_minus: "floor (- x) = - ceiling x"
unfolding ceiling_def by simp
-lemma ceiling_minus [simp]: "ceiling (- x) = - floor x"
+lemma ceiling_minus: "ceiling (- x) = - floor x"
unfolding ceiling_def by simp
end
--- a/src/HOL/Decision_Procs/MIR.thy Thu Feb 26 14:16:30 2009 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Thu Feb 26 15:42:35 2009 +0100
@@ -5926,7 +5926,7 @@
apply mir
done
-lemma "ALL x y. \<lfloor>x\<rfloor> = \<lfloor>y\<rfloor> \<longrightarrow> 0 \<le> abs (y - x) \<and> abs (y - x) \<le> 1"
+lemma "ALL (x::real) (y::real). \<lfloor>x\<rfloor> = \<lfloor>y\<rfloor> \<longrightarrow> 0 \<le> abs (y - x) \<and> abs (y - x) \<le> 1"
apply mir
done
--- a/src/HOL/RComplete.thy Thu Feb 26 14:16:30 2009 +0100
+++ b/src/HOL/RComplete.thy Thu Feb 26 15:42:35 2009 +0100
@@ -501,13 +501,13 @@
unfolding real_of_nat_def by simp
lemma floor_minus_real_of_nat [simp]: "floor (- real (n::nat)) = - int n"
-unfolding real_of_nat_def by simp
+unfolding real_of_nat_def by (simp add: floor_minus)
lemma floor_real_of_int [simp]: "floor (real (n::int)) = n"
unfolding real_of_int_def by simp
lemma floor_minus_real_of_int [simp]: "floor (- real (n::int)) = - n"
-unfolding real_of_int_def by simp
+unfolding real_of_int_def by (simp add: floor_minus)
lemma real_lb_ub_int: " \<exists>n::int. real n \<le> r & r < real (n+1)"
unfolding real_of_int_def by (rule floor_exists)