diff -r 7e3a026f014f -r aa74ce315bae src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Tue Aug 16 13:07:52 2011 -0700 +++ b/src/HOL/Deriv.thy Tue Aug 16 09:31:23 2011 -0700 @@ -670,7 +670,7 @@ |] ==> \x. a \ x & x \ b & f(x) = y" apply (subgoal_tac "- f a \ -y & -y \ - f b", clarify) apply (drule IVT [where f = "%x. - f x"], assumption) -apply (auto intro: isCont_minus) +apply simp_all done (*HOL style here: object-level formulations*) @@ -750,7 +750,7 @@ with M1 have M3: "\x. a \ x & x \ b --> f x < M" by (fastsimp simp add: linorder_not_le [symmetric]) hence "\x. a \ x & x \ b --> isCont (%x. inverse (M - f x)) x" - by (auto simp add: isCont_inverse isCont_diff con) + by (auto simp add: con) from isCont_bounded [OF le this] obtain k where k: "!!x. a \ x & x \ b --> inverse (M - f x) \ k" by auto have Minv: "!!x. a \ x & x \ b --> 0 < inverse (M - f (x))" @@ -1059,8 +1059,8 @@ (f(b) - f(a) = (b-a) * l)" proof - let ?F = "%x. f x - ((f b - f a) / (b-a)) * x" - have contF: "\x. a \ x \ x \ b \ isCont ?F x" using con - by (fast intro: isCont_diff isCont_const isCont_mult isCont_ident) + have contF: "\x. a \ x \ x \ b \ isCont ?F x" + using con by (fast intro: isCont_intros) have difF: "\x. a < x \ x < b \ ?F differentiable x" proof (clarify) fix x::real @@ -1353,7 +1353,7 @@ ==> \z. \z-x\ \ d & f z < f x" apply (insert lemma_isCont_inj [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d]) -apply (simp add: isCont_minus linorder_not_le) +apply (simp add: linorder_not_le) done text{*Show there's an interval surrounding @{term "f(x)"} in @@ -1509,27 +1509,9 @@ let ?h = "\x. (f b - f a)*(g x) - (g b - g a)*(f x)" from assms have "a < b" by simp moreover have "\x. a \ x \ x \ b \ isCont ?h x" - proof - - have "\x. a <= x \ x <= b \ isCont (\x. f b - f a) x" by simp - with gc have "\x. a <= x \ x <= b \ isCont (\x. (f b - f a) * g x) x" - by (auto intro: isCont_mult) - moreover - have "\x. a <= x \ x <= b \ isCont (\x. g b - g a) x" by simp - with fc have "\x. a <= x \ x <= b \ isCont (\x. (g b - g a) * f x) x" - by (auto intro: isCont_mult) - ultimately show ?thesis - by (fastsimp intro: isCont_diff) - qed - moreover - have "\x. a < x \ x < b \ ?h differentiable x" - proof - - have "\x. a < x \ x < b \ (\x. f b - f a) differentiable x" by simp - with gd have "\x. a < x \ x < b \ (\x. (f b - f a) * g x) differentiable x" by simp - moreover - have "\x. a < x \ x < b \ (\x. g b - g a) differentiable x" by simp - with fd have "\x. a < x \ x < b \ (\x. (g b - g a) * f x) differentiable x" by simp - ultimately show ?thesis by simp - qed + using fc gc by simp + moreover have "\x. a < x \ x < b \ ?h differentiable x" + using fd gd by simp ultimately have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" by (rule MVT) then obtain l where ldef: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where cdef: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" ..