# HG changeset patch # User nipkow # Date 1228982162 -3600 # Node ID 12ca66b887a0747430ab391eb403b53ec8e0c33d # Parent e70b9c2bee145958606e4d5df078c58e6c32dbdb codegen diff -r e70b9c2bee14 -r 12ca66b887a0 src/HOL/Divides.thy --- a/src/HOL/Divides.thy Thu Dec 11 08:53:53 2008 +0100 +++ b/src/HOL/Divides.thy Thu Dec 11 08:56:02 2008 +0100 @@ -127,7 +127,7 @@ note that ultimately show thesis by blast qed -lemma dvd_eq_mod_eq_0 [code]: "a dvd b \ b mod a = 0" +lemma dvd_eq_mod_eq_0 [code unfold]: "a dvd b \ b mod a = 0" proof assume "b mod a = 0" with mod_div_equality [of b a] have "b div a * a = b" by simp diff -r e70b9c2bee14 -r 12ca66b887a0 src/HOL/Real.thy --- a/src/HOL/Real.thy Thu Dec 11 08:53:53 2008 +0100 +++ b/src/HOL/Real.thy Thu Dec 11 08:56:02 2008 +0100 @@ -1,5 +1,5 @@ theory Real -imports "~~/src/HOL/Real/RealVector" +imports RComplete "~~/src/HOL/Real/RealVector" begin end