--- a/src/HOL/Deriv.thy Fri Aug 19 17:59:19 2011 -0700
+++ b/src/HOL/Deriv.thy Fri Aug 19 18:06:27 2011 -0700
@@ -304,9 +304,6 @@
==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
by (drule (2) DERIV_divide) (simp add: mult_commute)
-lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"
-by auto
-
text {* @{text "DERIV_intros"} *}
ML {*
structure Deriv_Intros = Named_Thms
--- a/src/HOL/Library/Poly_Deriv.thy Fri Aug 19 17:59:19 2011 -0700
+++ b/src/HOL/Library/Poly_Deriv.thy Fri Aug 19 18:06:27 2011 -0700
@@ -78,11 +78,11 @@
by (simp add: DERIV_cmult mult_commute [of _ c])
lemma DERIV_pow2: "DERIV (%x. x ^ Suc n) x :> real (Suc n) * (x ^ n)"
-by (rule lemma_DERIV_subst, rule DERIV_pow, simp)
+by (rule DERIV_cong, rule DERIV_pow, simp)
declare DERIV_pow2 [simp] DERIV_pow [simp]
lemma DERIV_add_const: "DERIV f x :> D ==> DERIV (%x. a + f x :: 'a::real_normed_field) x :> D"
-by (rule lemma_DERIV_subst, rule DERIV_add, auto)
+by (rule DERIV_cong, rule DERIV_add, auto)
lemma poly_DERIV[simp]: "DERIV (%x. poly p x) x :> poly (pderiv p) x"
by (induct p, auto intro!: DERIV_intros simp add: pderiv_pCons)
--- a/src/HOL/Transcendental.thy Fri Aug 19 17:59:19 2011 -0700
+++ b/src/HOL/Transcendental.thy Fri Aug 19 18:06:27 2011 -0700
@@ -1166,7 +1166,7 @@
lemma DERIV_ln: "0 < x \<Longrightarrow> DERIV ln x :> inverse x"
apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"])
- apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln])
+ apply (erule DERIV_cong [OF DERIV_exp exp_ln])
apply (simp_all add: abs_if isCont_ln)
done
@@ -1309,9 +1309,6 @@
lemma cos_zero [simp]: "cos 0 = 1"
unfolding cos_def cos_coeff_def by (simp add: powser_zero)
-lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"
- by (rule DERIV_cong) (* TODO: delete *)
-
lemma sin_cos_squared_add [simp]: "(sin x)\<twosuperior> + (cos x)\<twosuperior> = 1"
proof -
have "\<forall>x. DERIV (\<lambda>x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0"