# HG changeset patch # User paulson # Date 1258112013 0 # Node ID abf780db30eadfb1c50a8e346d2cd347e13f458d # Parent 0d82107dc07a939ebd972665cf3fcb34b6ff1d93 A number of theorems contributed by Jeremy Avigad diff -r 0d82107dc07a -r abf780db30ea src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Thu Nov 12 17:21:51 2009 +0100 +++ b/src/HOL/Deriv.thy Fri Nov 13 11:33:33 2009 +0000 @@ -273,6 +273,11 @@ "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" by (drule DERIV_mult' [OF DERIV_const], simp) +lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c" + apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force) + apply (erule DERIV_cmult) + done + text {* Standard version *} lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) @@ -702,14 +707,10 @@ apply safe apply simp_all apply (rename_tac x xa ya M Ma) -apply (cut_tac x = M and y = Ma in linorder_linear, safe) -apply (rule_tac x = Ma in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) -apply (rule_tac x = M in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (metis linorder_not_less order_le_less real_le_trans) apply (case_tac "a \ x & x \ b") -apply (rule_tac [2] x = 1 in exI) -prefer 2 apply force + prefer 2 + apply (rule_tac x = 1 in exI, force) apply (simp add: LIM_eq isCont_iff) apply (drule_tac x = x in spec, auto) apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) @@ -815,7 +816,7 @@ text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} -lemma DERIV_left_inc: +lemma DERIV_pos_inc_right: fixes f :: "real => real" assumes der: "DERIV f x :> l" and l: "0 < l" @@ -845,7 +846,7 @@ qed qed -lemma DERIV_left_dec: +lemma DERIV_neg_dec_left: fixes f :: "real => real" assumes der: "DERIV f x :> l" and l: "l < 0" @@ -875,6 +876,21 @@ qed qed + +lemma DERIV_pos_inc_left: + fixes f :: "real => real" + shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d --> f(x - h) < f(x)" + apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified]) + apply (auto simp add: DERIV_minus) + done + +lemma DERIV_neg_dec_right: + fixes f :: "real => real" + shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d --> f(x) > f(x + h)" + apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified]) + apply (auto simp add: DERIV_minus) + done + lemma DERIV_local_max: fixes f :: "real => real" assumes der: "DERIV f x :> l" @@ -885,7 +901,7 @@ case equal thus ?thesis . next case less - from DERIV_left_dec [OF der less] + from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast from real_lbound_gt_zero [OF d d'] @@ -894,7 +910,7 @@ show ?thesis by (auto simp add: abs_if) next case greater - from DERIV_left_inc [OF der greater] + from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast from real_lbound_gt_zero [OF d d'] @@ -1205,6 +1221,87 @@ ultimately show ?thesis using neq by (force simp add: add_commute) qed +(* A function with positive derivative is increasing. + A simple proof using the MVT, by Jeremy Avigad. And variants. +*) + +lemma DERIV_pos_imp_increasing: + fixes a::real and b::real and f::"real => real" + assumes "a < b" and "\x. a \ x & x \ b --> (EX y. DERIV f x :> y & y > 0)" + shows "f a < f b" +proof (rule ccontr) + assume "~ f a < f b" + from assms have "EX l z. a < z & z < b & DERIV f z :> l + & f b - f a = (b - a) * l" + by (metis MVT DERIV_isCont differentiableI real_less_def) + then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + and "f b - f a = (b - a) * l" + by auto + + from prems have "~(l > 0)" + by (metis assms(1) linorder_not_le mult_le_0_iff real_le_eq_diff) + with prems show False + by (metis DERIV_unique real_less_def) +qed + + +lemma DERIV_nonneg_imp_nonincreasing: + fixes a::real and b::real and f::"real => real" + assumes "a \ b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" + shows "f a \ f b" +proof (rule ccontr, cases "a = b") + assume "~ f a \ f b" + assume "a = b" + with prems show False by auto + next assume "~ f a \ f b" + assume "a ~= b" + with assms have "EX l z. a < z & z < b & DERIV f z :> l + & f b - f a = (b - a) * l" + apply (intro MVT) + apply auto + apply (metis DERIV_isCont) + apply (metis differentiableI real_less_def) + done + then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + and "f b - f a = (b - a) * l" + by auto + from prems have "~(l >= 0)" + by (metis diff_self le_eqI le_iff_diff_le_0 real_le_anti_sym real_le_linear + split_mult_pos_le) + with prems show False + by (metis DERIV_unique order_less_imp_le) +qed + +lemma DERIV_neg_imp_decreasing: + fixes a::real and b::real and f::"real => real" + assumes "a < b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y < 0)" + shows "f a > f b" +proof - + have "(%x. -f x) a < (%x. -f x) b" + apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"]) + apply (insert prems, auto) + apply (metis DERIV_minus neg_0_less_iff_less) + done + thus ?thesis + by simp +qed + +lemma DERIV_nonpos_imp_nonincreasing: + fixes a::real and b::real and f::"real => real" + assumes "a \ b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" + shows "f a \ f b" +proof - + have "(%x. -f x) a \ (%x. -f x) b" + apply (rule DERIV_nonneg_imp_nonincreasing [of a b "%x. -f x"]) + apply (insert prems, auto) + apply (metis DERIV_minus neg_0_le_iff_le) + done + thus ?thesis + by simp +qed subsection {* Continuous injective functions *}