diff -r 80a1aa30e24d -r 408e15cbd2a6 src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Fri Jun 14 08:34:27 2019 +0000 +++ b/src/HOL/Deriv.thy Fri Jun 14 08:34:28 2019 +0000 @@ -1409,12 +1409,14 @@ and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - - obtain f' where derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" + obtain f' :: "real \ real \ real" + where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis - by (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative linordered_field_class.sign_simps(5) real_differentiable_def) + by (simp add: ac_simps) + (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: