# HG changeset patch # User huffman # Date 1235659146 28800 # Node ID b094999e1d334f6b8ae1503aeaed34fa98b1e879 # Parent 32effb2a8168aedfdabf8ac652cbad1fdb3399a0# Parent e1c714d33c5c922441b7db27eeff910b297376ca merged diff -r e1c714d33c5c -r b094999e1d33 src/HOL/Archimedean_Field.thy --- a/src/HOL/Archimedean_Field.thy Thu Feb 26 11:21:29 2009 +0000 +++ b/src/HOL/Archimedean_Field.thy Thu Feb 26 06:39:06 2009 -0800 @@ -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 diff -r e1c714d33c5c -r b094999e1d33 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Thu Feb 26 11:21:29 2009 +0000 +++ b/src/HOL/Decision_Procs/MIR.thy Thu Feb 26 06:39:06 2009 -0800 @@ -5926,7 +5926,7 @@ apply mir done -lemma "ALL x y. \x\ = \y\ \ 0 \ abs (y - x) \ abs (y - x) \ 1" +lemma "ALL (x::real) (y::real). \x\ = \y\ \ 0 \ abs (y - x) \ abs (y - x) \ 1" apply mir done diff -r e1c714d33c5c -r b094999e1d33 src/HOL/RComplete.thy --- a/src/HOL/RComplete.thy Thu Feb 26 11:21:29 2009 +0000 +++ b/src/HOL/RComplete.thy Thu Feb 26 06:39:06 2009 -0800 @@ -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: " \n::int. real n \ r & r < real (n+1)" unfolding real_of_int_def by (rule floor_exists)