src/HOL/Deriv.thy
author wenzelm
Fri, 28 Oct 2011 23:41:16 +0200
changeset 46165 3c5d3d286055
parent 45916 c478d1876371
child 46471 1bbbac9a0cb0
permissions -rw-r--r--
tuned Named_Thms: proper binding;
     1 (*  Title       : Deriv.thy
     2     Author      : Jacques D. Fleuriot
     3     Copyright   : 1998  University of Cambridge
     4     Conversion to Isar and new proofs by Lawrence C Paulson, 2004
     5     GMVT by Benjamin Porter, 2005
     6 *)
     7 
     8 header{* Differentiation *}
     9 
    10 theory Deriv
    11 imports Lim
    12 begin
    13 
    14 text{*Standard Definitions*}
    15 
    16 definition
    17   deriv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
    18     --{*Differentiation: D is derivative of function f at x*}
    19           ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
    20   "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)"
    21 
    22 primrec
    23   Bolzano_bisect :: "(real \<times> real \<Rightarrow> bool) \<Rightarrow> real \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real \<times> real" where
    24   "Bolzano_bisect P a b 0 = (a, b)"
    25   | "Bolzano_bisect P a b (Suc n) =
    26       (let (x, y) = Bolzano_bisect P a b n
    27        in if P (x, (x+y) / 2) then ((x+y)/2, y)
    28                               else (x, (x+y)/2))"
    29 
    30 
    31 subsection {* Derivatives *}
    32 
    33 lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)"
    34 by (simp add: deriv_def)
    35 
    36 lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D"
    37 by (simp add: deriv_def)
    38 
    39 lemma DERIV_const [simp]: "DERIV (\<lambda>x. k) x :> 0"
    40   by (simp add: deriv_def tendsto_const)
    41 
    42 lemma DERIV_ident [simp]: "DERIV (\<lambda>x. x) x :> 1"
    43   by (simp add: deriv_def tendsto_const cong: LIM_cong)
    44 
    45 lemma DERIV_add:
    46   "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + g x) x :> D + E"
    47   by (simp only: deriv_def add_diff_add add_divide_distrib tendsto_add)
    48 
    49 lemma DERIV_minus:
    50   "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. - f x) x :> - D"
    51   by (simp only: deriv_def minus_diff_minus divide_minus_left tendsto_minus)
    52 
    53 lemma DERIV_diff:
    54   "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x - g x) x :> D - E"
    55 by (simp only: diff_minus DERIV_add DERIV_minus)
    56 
    57 lemma DERIV_add_minus:
    58   "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + - g x) x :> D + - E"
    59 by (simp only: DERIV_add DERIV_minus)
    60 
    61 lemma DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x"
    62 proof (unfold isCont_iff)
    63   assume "DERIV f x :> D"
    64   hence "(\<lambda>h. (f(x+h) - f(x)) / h) -- 0 --> D"
    65     by (rule DERIV_D)
    66   hence "(\<lambda>h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0"
    67     by (intro tendsto_mult tendsto_ident_at)
    68   hence "(\<lambda>h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0"
    69     by simp
    70   hence "(\<lambda>h. f(x+h) - f(x)) -- 0 --> 0"
    71     by (simp cong: LIM_cong)
    72   thus "(\<lambda>h. f(x+h)) -- 0 --> f(x)"
    73     by (simp add: LIM_def dist_norm)
    74 qed
    75 
    76 lemma DERIV_mult_lemma:
    77   fixes a b c d :: "'a::real_field"
    78   shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d"
    79 by (simp add: diff_minus add_divide_distrib [symmetric] ring_distribs)
    80 
    81 lemma DERIV_mult':
    82   assumes f: "DERIV f x :> D"
    83   assumes g: "DERIV g x :> E"
    84   shows "DERIV (\<lambda>x. f x * g x) x :> f x * E + D * g x"
    85 proof (unfold deriv_def)
    86   from f have "isCont f x"
    87     by (rule DERIV_isCont)
    88   hence "(\<lambda>h. f(x+h)) -- 0 --> f x"
    89     by (simp only: isCont_iff)
    90   hence "(\<lambda>h. f(x+h) * ((g(x+h) - g x) / h) +
    91               ((f(x+h) - f x) / h) * g x)
    92           -- 0 --> f x * E + D * g x"
    93     by (intro tendsto_intros DERIV_D f g)
    94   thus "(\<lambda>h. (f(x+h) * g(x+h) - f x * g x) / h)
    95          -- 0 --> f x * E + D * g x"
    96     by (simp only: DERIV_mult_lemma)
    97 qed
    98 
    99 lemma DERIV_mult:
   100      "[| DERIV f x :> Da; DERIV g x :> Db |]
   101       ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"
   102 by (drule (1) DERIV_mult', simp only: mult_commute add_commute)
   103 
   104 lemma DERIV_unique:
   105       "[| DERIV f x :> D; DERIV f x :> E |] ==> D = E"
   106 apply (simp add: deriv_def)
   107 apply (blast intro: LIM_unique)
   108 done
   109 
   110 text{*Differentiation of finite sum*}
   111 
   112 lemma DERIV_setsum:
   113   assumes "finite S"
   114   and "\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x :> (f' x n)"
   115   shows "DERIV (%x. setsum (f x) S) x :> setsum (f' x) S"
   116   using assms by induct (auto intro!: DERIV_add)
   117 
   118 lemma DERIV_sumr [rule_format (no_asm)]:
   119      "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x))
   120       --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x :> (\<Sum>r=m..<n. f' r x)"
   121   by (auto intro: DERIV_setsum)
   122 
   123 text{*Alternative definition for differentiability*}
   124 
   125 lemma DERIV_LIM_iff:
   126   fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
   127      "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) =
   128       ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
   129 apply (rule iffI)
   130 apply (drule_tac k="- a" in LIM_offset)
   131 apply (simp add: diff_minus)
   132 apply (drule_tac k="a" in LIM_offset)
   133 apply (simp add: add_commute)
   134 done
   135 
   136 lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)"
   137 by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
   138 
   139 lemma DERIV_inverse_lemma:
   140   "\<lbrakk>a \<noteq> 0; b \<noteq> (0::'a::real_normed_field)\<rbrakk>
   141    \<Longrightarrow> (inverse a - inverse b) / h
   142      = - (inverse a * ((a - b) / h) * inverse b)"
   143 by (simp add: inverse_diff_inverse)
   144 
   145 lemma DERIV_inverse':
   146   assumes der: "DERIV f x :> D"
   147   assumes neq: "f x \<noteq> 0"
   148   shows "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))"
   149     (is "DERIV _ _ :> ?E")
   150 proof (unfold DERIV_iff2)
   151   from der have lim_f: "f -- x --> f x"
   152     by (rule DERIV_isCont [unfolded isCont_def])
   153 
   154   from neq have "0 < norm (f x)" by simp
   155   with LIM_D [OF lim_f] obtain s
   156     where s: "0 < s"
   157     and less_fx: "\<And>z. \<lbrakk>z \<noteq> x; norm (z - x) < s\<rbrakk>
   158                   \<Longrightarrow> norm (f z - f x) < norm (f x)"
   159     by fast
   160 
   161   show "(\<lambda>z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E"
   162   proof (rule LIM_equal2 [OF s])
   163     fix z
   164     assume "z \<noteq> x" "norm (z - x) < s"
   165     hence "norm (f z - f x) < norm (f x)" by (rule less_fx)
   166     hence "f z \<noteq> 0" by auto
   167     thus "(inverse (f z) - inverse (f x)) / (z - x) =
   168           - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))"
   169       using neq by (rule DERIV_inverse_lemma)
   170   next
   171     from der have "(\<lambda>z. (f z - f x) / (z - x)) -- x --> D"
   172       by (unfold DERIV_iff2)
   173     thus "(\<lambda>z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x)))
   174           -- x --> ?E"
   175       by (intro tendsto_intros lim_f neq)
   176   qed
   177 qed
   178 
   179 lemma DERIV_divide:
   180   "\<lbrakk>DERIV f x :> D; DERIV g x :> E; g x \<noteq> 0\<rbrakk>
   181    \<Longrightarrow> DERIV (\<lambda>x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)"
   182 apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) +
   183           D * inverse (g x) = (D * g x - f x * E) / (g x * g x)")
   184 apply (erule subst)
   185 apply (unfold divide_inverse)
   186 apply (erule DERIV_mult')
   187 apply (erule (1) DERIV_inverse')
   188 apply (simp add: ring_distribs nonzero_inverse_mult_distrib)
   189 apply (simp add: mult_ac)
   190 done
   191 
   192 lemma DERIV_power_Suc:
   193   fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
   194   assumes f: "DERIV f x :> D"
   195   shows "DERIV (\<lambda>x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)"
   196 proof (induct n)
   197 case 0
   198   show ?case by (simp add: f)
   199 case (Suc k)
   200   from DERIV_mult' [OF f Suc] show ?case
   201     apply (simp only: of_nat_Suc ring_distribs mult_1_left)
   202     apply (simp only: power_Suc algebra_simps)
   203     done
   204 qed
   205 
   206 lemma DERIV_power:
   207   fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
   208   assumes f: "DERIV f x :> D"
   209   shows "DERIV (\<lambda>x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))"
   210 by (cases "n", simp, simp add: DERIV_power_Suc f del: power_Suc)
   211 
   212 text {* Caratheodory formulation of derivative at a point *}
   213 
   214 lemma CARAT_DERIV:
   215      "(DERIV f x :> l) =
   216       (\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isCont g x & g x = l)"
   217       (is "?lhs = ?rhs")
   218 proof
   219   assume der: "DERIV f x :> l"
   220   show "\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) \<and> isCont g x \<and> g x = l"
   221   proof (intro exI conjI)
   222     let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
   223     show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
   224     show "isCont ?g x" using der
   225       by (simp add: isCont_iff DERIV_iff diff_minus
   226                cong: LIM_equal [rule_format])
   227     show "?g x = l" by simp
   228   qed
   229 next
   230   assume "?rhs"
   231   then obtain g where
   232     "(\<forall>z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast
   233   thus "(DERIV f x :> l)"
   234      by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong)
   235 qed
   236 
   237 lemma DERIV_chain':
   238   assumes f: "DERIV f x :> D"
   239   assumes g: "DERIV g (f x) :> E"
   240   shows "DERIV (\<lambda>x. g (f x)) x :> E * D"
   241 proof (unfold DERIV_iff2)
   242   obtain d where d: "\<forall>y. g y - g (f x) = d y * (y - f x)"
   243     and cont_d: "isCont d (f x)" and dfx: "d (f x) = E"
   244     using CARAT_DERIV [THEN iffD1, OF g] by fast
   245   from f have "f -- x --> f x"
   246     by (rule DERIV_isCont [unfolded isCont_def])
   247   with cont_d have "(\<lambda>z. d (f z)) -- x --> d (f x)"
   248     by (rule isCont_tendsto_compose)
   249   hence "(\<lambda>z. d (f z) * ((f z - f x) / (z - x)))
   250           -- x --> d (f x) * D"
   251     by (rule tendsto_mult [OF _ f [unfolded DERIV_iff2]])
   252   thus "(\<lambda>z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D"
   253     by (simp add: d dfx)
   254 qed
   255 
   256 text {*
   257  Let's do the standard proof, though theorem
   258  @{text "LIM_mult2"} follows from a NS proof
   259 *}
   260 
   261 lemma DERIV_cmult:
   262       "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D"
   263 by (drule DERIV_mult' [OF DERIV_const], simp)
   264 
   265 lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c"
   266   apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force)
   267   apply (erule DERIV_cmult)
   268   done
   269 
   270 text {* Standard version *}
   271 lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db"
   272 by (drule (1) DERIV_chain', simp add: o_def mult_commute)
   273 
   274 lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db"
   275 by (auto dest: DERIV_chain simp add: o_def)
   276 
   277 text {* Derivative of linear multiplication *}
   278 lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c"
   279 by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp)
   280 
   281 lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
   282 apply (cut_tac DERIV_power [OF DERIV_ident])
   283 apply (simp add: real_of_nat_def)
   284 done
   285 
   286 text {* Power of @{text "-1"} *}
   287 
   288 lemma DERIV_inverse:
   289   fixes x :: "'a::{real_normed_field}"
   290   shows "x \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))"
   291 by (drule DERIV_inverse' [OF DERIV_ident]) simp
   292 
   293 text {* Derivative of inverse *}
   294 lemma DERIV_inverse_fun:
   295   fixes x :: "'a::{real_normed_field}"
   296   shows "[| DERIV f x :> d; f(x) \<noteq> 0 |]
   297       ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"
   298 by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib)
   299 
   300 text {* Derivative of quotient *}
   301 lemma DERIV_quotient:
   302   fixes x :: "'a::{real_normed_field}"
   303   shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> 0 |]
   304        ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
   305 by (drule (2) DERIV_divide) (simp add: mult_commute)
   306 
   307 text {* @{text "DERIV_intros"} *}
   308 ML {*
   309 structure Deriv_Intros = Named_Thms
   310 (
   311   val name = @{binding DERIV_intros}
   312   val description = "DERIV introduction rules"
   313 )
   314 *}
   315 
   316 setup Deriv_Intros.setup
   317 
   318 lemma DERIV_cong: "\<lbrakk> DERIV f x :> X ; X = Y \<rbrakk> \<Longrightarrow> DERIV f x :> Y"
   319   by simp
   320 
   321 declare
   322   DERIV_const[THEN DERIV_cong, DERIV_intros]
   323   DERIV_ident[THEN DERIV_cong, DERIV_intros]
   324   DERIV_add[THEN DERIV_cong, DERIV_intros]
   325   DERIV_minus[THEN DERIV_cong, DERIV_intros]
   326   DERIV_mult[THEN DERIV_cong, DERIV_intros]
   327   DERIV_diff[THEN DERIV_cong, DERIV_intros]
   328   DERIV_inverse'[THEN DERIV_cong, DERIV_intros]
   329   DERIV_divide[THEN DERIV_cong, DERIV_intros]
   330   DERIV_power[where 'a=real, THEN DERIV_cong,
   331               unfolded real_of_nat_def[symmetric], DERIV_intros]
   332   DERIV_setsum[THEN DERIV_cong, DERIV_intros]
   333 
   334 
   335 subsection {* Differentiability predicate *}
   336 
   337 definition
   338   differentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
   339     (infixl "differentiable" 60) where
   340   "f differentiable x = (\<exists>D. DERIV f x :> D)"
   341 
   342 lemma differentiableE [elim?]:
   343   assumes "f differentiable x"
   344   obtains df where "DERIV f x :> df"
   345   using assms unfolding differentiable_def ..
   346 
   347 lemma differentiableD: "f differentiable x ==> \<exists>D. DERIV f x :> D"
   348 by (simp add: differentiable_def)
   349 
   350 lemma differentiableI: "DERIV f x :> D ==> f differentiable x"
   351 by (force simp add: differentiable_def)
   352 
   353 lemma differentiable_ident [simp]: "(\<lambda>x. x) differentiable x"
   354   by (rule DERIV_ident [THEN differentiableI])
   355 
   356 lemma differentiable_const [simp]: "(\<lambda>z. a) differentiable x"
   357   by (rule DERIV_const [THEN differentiableI])
   358 
   359 lemma differentiable_compose:
   360   assumes f: "f differentiable (g x)"
   361   assumes g: "g differentiable x"
   362   shows "(\<lambda>x. f (g x)) differentiable x"
   363 proof -
   364   from `f differentiable (g x)` obtain df where "DERIV f (g x) :> df" ..
   365   moreover
   366   from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
   367   ultimately
   368   have "DERIV (\<lambda>x. f (g x)) x :> df * dg" by (rule DERIV_chain2)
   369   thus ?thesis by (rule differentiableI)
   370 qed
   371 
   372 lemma differentiable_sum [simp]:
   373   assumes "f differentiable x"
   374   and "g differentiable x"
   375   shows "(\<lambda>x. f x + g x) differentiable x"
   376 proof -
   377   from `f differentiable x` obtain df where "DERIV f x :> df" ..
   378   moreover
   379   from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
   380   ultimately
   381   have "DERIV (\<lambda>x. f x + g x) x :> df + dg" by (rule DERIV_add)
   382   thus ?thesis by (rule differentiableI)
   383 qed
   384 
   385 lemma differentiable_minus [simp]:
   386   assumes "f differentiable x"
   387   shows "(\<lambda>x. - f x) differentiable x"
   388 proof -
   389   from `f differentiable x` obtain df where "DERIV f x :> df" ..
   390   hence "DERIV (\<lambda>x. - f x) x :> - df" by (rule DERIV_minus)
   391   thus ?thesis by (rule differentiableI)
   392 qed
   393 
   394 lemma differentiable_diff [simp]:
   395   assumes "f differentiable x"
   396   assumes "g differentiable x"
   397   shows "(\<lambda>x. f x - g x) differentiable x"
   398   unfolding diff_minus using assms by simp
   399 
   400 lemma differentiable_mult [simp]:
   401   assumes "f differentiable x"
   402   assumes "g differentiable x"
   403   shows "(\<lambda>x. f x * g x) differentiable x"
   404 proof -
   405   from `f differentiable x` obtain df where "DERIV f x :> df" ..
   406   moreover
   407   from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
   408   ultimately
   409   have "DERIV (\<lambda>x. f x * g x) x :> df * g x + dg * f x" by (rule DERIV_mult)
   410   thus ?thesis by (rule differentiableI)
   411 qed
   412 
   413 lemma differentiable_inverse [simp]:
   414   assumes "f differentiable x" and "f x \<noteq> 0"
   415   shows "(\<lambda>x. inverse (f x)) differentiable x"
   416 proof -
   417   from `f differentiable x` obtain df where "DERIV f x :> df" ..
   418   hence "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * df * inverse (f x))"
   419     using `f x \<noteq> 0` by (rule DERIV_inverse')
   420   thus ?thesis by (rule differentiableI)
   421 qed
   422 
   423 lemma differentiable_divide [simp]:
   424   assumes "f differentiable x"
   425   assumes "g differentiable x" and "g x \<noteq> 0"
   426   shows "(\<lambda>x. f x / g x) differentiable x"
   427   unfolding divide_inverse using assms by simp
   428 
   429 lemma differentiable_power [simp]:
   430   fixes f :: "'a::{real_normed_field} \<Rightarrow> 'a"
   431   assumes "f differentiable x"
   432   shows "(\<lambda>x. f x ^ n) differentiable x"
   433   apply (induct n)
   434   apply simp
   435   apply (simp add: assms)
   436   done
   437 
   438 
   439 subsection {* Nested Intervals and Bisection *}
   440 
   441 text{*Lemmas about nested intervals and proof by bisection (cf.Harrison).
   442      All considerably tidied by lcp.*}
   443 
   444 lemma lemma_f_mono_add [rule_format (no_asm)]: "(\<forall>n. (f::nat=>real) n \<le> f (Suc n)) --> f m \<le> f(m + no)"
   445 apply (induct "no")
   446 apply (auto intro: order_trans)
   447 done
   448 
   449 lemma f_inc_g_dec_Beq_f: "[| \<forall>n. f(n) \<le> f(Suc n);
   450          \<forall>n. g(Suc n) \<le> g(n);
   451          \<forall>n. f(n) \<le> g(n) |]
   452       ==> Bseq (f :: nat \<Rightarrow> real)"
   453 apply (rule_tac k = "f 0" and K = "g 0" in BseqI2, rule allI)
   454 apply (rule conjI)
   455 apply (induct_tac "n")
   456 apply (auto intro: order_trans)
   457 apply (rule_tac y = "g n" in order_trans)
   458 apply (induct_tac [2] "n")
   459 apply (auto intro: order_trans)
   460 done
   461 
   462 lemma f_inc_g_dec_Beq_g: "[| \<forall>n. f(n) \<le> f(Suc n);
   463          \<forall>n. g(Suc n) \<le> g(n);
   464          \<forall>n. f(n) \<le> g(n) |]
   465       ==> Bseq (g :: nat \<Rightarrow> real)"
   466 apply (subst Bseq_minus_iff [symmetric])
   467 apply (rule_tac g = "%x. - (f x)" in f_inc_g_dec_Beq_f)
   468 apply auto
   469 done
   470 
   471 lemma f_inc_imp_le_lim:
   472   fixes f :: "nat \<Rightarrow> real"
   473   shows "\<lbrakk>\<forall>n. f n \<le> f (Suc n); convergent f\<rbrakk> \<Longrightarrow> f n \<le> lim f"
   474   by (rule incseq_le, simp add: incseq_SucI, simp add: convergent_LIMSEQ_iff)
   475 
   476 lemma lim_uminus:
   477   fixes g :: "nat \<Rightarrow> 'a::real_normed_vector"
   478   shows "convergent g ==> lim (%x. - g x) = - (lim g)"
   479 apply (rule tendsto_minus [THEN limI])
   480 apply (simp add: convergent_LIMSEQ_iff)
   481 done
   482 
   483 lemma g_dec_imp_lim_le:
   484   fixes g :: "nat \<Rightarrow> real"
   485   shows "\<lbrakk>\<forall>n. g (Suc n) \<le> g(n); convergent g\<rbrakk> \<Longrightarrow> lim g \<le> g n"
   486   by (rule decseq_le, simp add: decseq_SucI, simp add: convergent_LIMSEQ_iff)
   487 
   488 lemma lemma_nest: "[| \<forall>n. f(n) \<le> f(Suc n);
   489          \<forall>n. g(Suc n) \<le> g(n);
   490          \<forall>n. f(n) \<le> g(n) |]
   491       ==> \<exists>l m :: real. l \<le> m &  ((\<forall>n. f(n) \<le> l) & f ----> l) &
   492                             ((\<forall>n. m \<le> g(n)) & g ----> m)"
   493 apply (subgoal_tac "monoseq f & monoseq g")
   494 prefer 2 apply (force simp add: LIMSEQ_iff monoseq_Suc)
   495 apply (subgoal_tac "Bseq f & Bseq g")
   496 prefer 2 apply (blast intro: f_inc_g_dec_Beq_f f_inc_g_dec_Beq_g)
   497 apply (auto dest!: Bseq_monoseq_convergent simp add: convergent_LIMSEQ_iff)
   498 apply (rule_tac x = "lim f" in exI)
   499 apply (rule_tac x = "lim g" in exI)
   500 apply (auto intro: LIMSEQ_le)
   501 apply (auto simp add: f_inc_imp_le_lim g_dec_imp_lim_le convergent_LIMSEQ_iff)
   502 done
   503 
   504 lemma lemma_nest_unique: "[| \<forall>n. f(n) \<le> f(Suc n);
   505          \<forall>n. g(Suc n) \<le> g(n);
   506          \<forall>n. f(n) \<le> g(n);
   507          (%n. f(n) - g(n)) ----> 0 |]
   508       ==> \<exists>l::real. ((\<forall>n. f(n) \<le> l) & f ----> l) &
   509                 ((\<forall>n. l \<le> g(n)) & g ----> l)"
   510 apply (drule lemma_nest, auto)
   511 apply (subgoal_tac "l = m")
   512 apply (drule_tac [2] f = f in tendsto_diff)
   513 apply (auto intro: LIMSEQ_unique)
   514 done
   515 
   516 text{*The universal quantifiers below are required for the declaration
   517   of @{text Bolzano_nest_unique} below.*}
   518 
   519 lemma Bolzano_bisect_le:
   520  "a \<le> b ==> \<forall>n. fst (Bolzano_bisect P a b n) \<le> snd (Bolzano_bisect P a b n)"
   521 apply (rule allI)
   522 apply (induct_tac "n")
   523 apply (auto simp add: Let_def split_def)
   524 done
   525 
   526 lemma Bolzano_bisect_fst_le_Suc: "a \<le> b ==>
   527    \<forall>n. fst(Bolzano_bisect P a b n) \<le> fst (Bolzano_bisect P a b (Suc n))"
   528 apply (rule allI)
   529 apply (induct_tac "n")
   530 apply (auto simp add: Bolzano_bisect_le Let_def split_def)
   531 done
   532 
   533 lemma Bolzano_bisect_Suc_le_snd: "a \<le> b ==>
   534    \<forall>n. snd(Bolzano_bisect P a b (Suc n)) \<le> snd (Bolzano_bisect P a b n)"
   535 apply (rule allI)
   536 apply (induct_tac "n")
   537 apply (auto simp add: Bolzano_bisect_le Let_def split_def)
   538 done
   539 
   540 lemma eq_divide_2_times_iff: "((x::real) = y / (2 * z)) = (2 * x = y/z)"
   541 apply (auto)
   542 apply (drule_tac f = "%u. (1/2) *u" in arg_cong)
   543 apply (simp)
   544 done
   545 
   546 lemma Bolzano_bisect_diff:
   547      "a \<le> b ==>
   548       snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) =
   549       (b-a) / (2 ^ n)"
   550 apply (induct "n")
   551 apply (auto simp add: eq_divide_2_times_iff add_divide_distrib Let_def split_def)
   552 done
   553 
   554 lemmas Bolzano_nest_unique =
   555     lemma_nest_unique
   556     [OF Bolzano_bisect_fst_le_Suc Bolzano_bisect_Suc_le_snd Bolzano_bisect_le]
   557 
   558 
   559 lemma not_P_Bolzano_bisect:
   560   assumes P:    "!!a b c. [| P(a,b); P(b,c); a \<le> b; b \<le> c|] ==> P(a,c)"
   561       and notP: "~ P(a,b)"
   562       and le:   "a \<le> b"
   563   shows "~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))"
   564 proof (induct n)
   565   case 0 show ?case using notP by simp
   566  next
   567   case (Suc n)
   568   thus ?case
   569  by (auto simp del: surjective_pairing [symmetric]
   570              simp add: Let_def split_def Bolzano_bisect_le [OF le]
   571      P [of "fst (Bolzano_bisect P a b n)" _ "snd (Bolzano_bisect P a b n)"])
   572 qed
   573 
   574 (*Now we re-package P_prem as a formula*)
   575 lemma not_P_Bolzano_bisect':
   576      "[| \<forall>a b c. P(a,b) & P(b,c) & a \<le> b & b \<le> c --> P(a,c);
   577          ~ P(a,b);  a \<le> b |] ==>
   578       \<forall>n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))"
   579 by (blast elim!: not_P_Bolzano_bisect [THEN [2] rev_notE])
   580 
   581 
   582 
   583 lemma lemma_BOLZANO:
   584      "[| \<forall>a b c. P(a,b) & P(b,c) & a \<le> b & b \<le> c --> P(a,c);
   585          \<forall>x. \<exists>d::real. 0 < d &
   586                 (\<forall>a b. a \<le> x & x \<le> b & (b-a) < d --> P(a,b));
   587          a \<le> b |]
   588       ==> P(a,b)"
   589 apply (rule Bolzano_nest_unique [where P1=P, THEN exE], assumption+)
   590 apply (rule tendsto_minus_cancel)
   591 apply (simp (no_asm_simp) add: Bolzano_bisect_diff LIMSEQ_divide_realpow_zero)
   592 apply (rule ccontr)
   593 apply (drule not_P_Bolzano_bisect', assumption+)
   594 apply (rename_tac "l")
   595 apply (drule_tac x = l in spec, clarify)
   596 apply (simp add: LIMSEQ_iff)
   597 apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec)
   598 apply (drule_tac P = "%r. 0<r --> ?Q r" and x = "d/2" in spec)
   599 apply (drule real_less_half_sum, auto)
   600 apply (drule_tac x = "fst (Bolzano_bisect P a b (no + noa))" in spec)
   601 apply (drule_tac x = "snd (Bolzano_bisect P a b (no + noa))" in spec)
   602 apply safe
   603 apply (simp_all (no_asm_simp))
   604 apply (rule_tac y = "abs (fst (Bolzano_bisect P a b (no + noa)) - l) + abs (snd (Bolzano_bisect P a b (no + noa)) - l)" in order_le_less_trans)
   605 apply (simp (no_asm_simp) add: abs_if)
   606 apply (rule real_sum_of_halves [THEN subst])
   607 apply (rule add_strict_mono)
   608 apply (simp_all add: diff_minus [symmetric])
   609 done
   610 
   611 
   612 lemma lemma_BOLZANO2: "((\<forall>a b c. (a \<le> b & b \<le> c & P(a,b) & P(b,c)) --> P(a,c)) &
   613        (\<forall>x. \<exists>d::real. 0 < d &
   614                 (\<forall>a b. a \<le> x & x \<le> b & (b-a) < d --> P(a,b))))
   615       --> (\<forall>a b. a \<le> b --> P(a,b))"
   616 apply clarify
   617 apply (blast intro: lemma_BOLZANO)
   618 done
   619 
   620 
   621 subsection {* Intermediate Value Theorem *}
   622 
   623 text {*Prove Contrapositive by Bisection*}
   624 
   625 lemma IVT: "[| f(a::real) \<le> (y::real); y \<le> f(b);
   626          a \<le> b;
   627          (\<forall>x. a \<le> x & x \<le> b --> isCont f x) |]
   628       ==> \<exists>x. a \<le> x & x \<le> b & f(x) = y"
   629 apply (rule contrapos_pp, assumption)
   630 apply (cut_tac P = "% (u,v) . a \<le> u & u \<le> v & v \<le> b --> ~ (f (u) \<le> y & y \<le> f (v))" in lemma_BOLZANO2)
   631 apply safe
   632 apply simp_all
   633 apply (simp add: isCont_iff LIM_eq)
   634 apply (rule ccontr)
   635 apply (subgoal_tac "a \<le> x & x \<le> b")
   636  prefer 2
   637  apply simp
   638  apply (drule_tac P = "%d. 0<d --> ?P d" and x = 1 in spec, arith)
   639 apply (drule_tac x = x in spec)+
   640 apply simp
   641 apply (drule_tac P = "%r. ?P r --> (\<exists>s>0. ?Q r s) " and x = "\<bar>y - f x\<bar>" in spec)
   642 apply safe
   643 apply simp
   644 apply (drule_tac x = s in spec, clarify)
   645 apply (cut_tac x = "f x" and y = y in linorder_less_linear, safe)
   646 apply (drule_tac x = "ba-x" in spec)
   647 apply (simp_all add: abs_if)
   648 apply (drule_tac x = "aa-x" in spec)
   649 apply (case_tac "x \<le> aa", simp_all)
   650 done
   651 
   652 lemma IVT2: "[| f(b::real) \<le> (y::real); y \<le> f(a);
   653          a \<le> b;
   654          (\<forall>x. a \<le> x & x \<le> b --> isCont f x)
   655       |] ==> \<exists>x. a \<le> x & x \<le> b & f(x) = y"
   656 apply (subgoal_tac "- f a \<le> -y & -y \<le> - f b", clarify)
   657 apply (drule IVT [where f = "%x. - f x"], assumption)
   658 apply simp_all
   659 done
   660 
   661 (*HOL style here: object-level formulations*)
   662 lemma IVT_objl: "(f(a::real) \<le> (y::real) & y \<le> f(b) & a \<le> b &
   663       (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
   664       --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
   665 apply (blast intro: IVT)
   666 done
   667 
   668 lemma IVT2_objl: "(f(b::real) \<le> (y::real) & y \<le> f(a) & a \<le> b &
   669       (\<forall>x. a \<le> x & x \<le> b --> isCont f x))
   670       --> (\<exists>x. a \<le> x & x \<le> b & f(x) = y)"
   671 apply (blast intro: IVT2)
   672 done
   673 
   674 
   675 subsection {* Boundedness of continuous functions *}
   676 
   677 text{*By bisection, function continuous on closed interval is bounded above*}
   678 
   679 lemma isCont_bounded:
   680      "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
   681       ==> \<exists>M::real. \<forall>x::real. a \<le> x & x \<le> b --> f(x) \<le> M"
   682 apply (cut_tac P = "% (u,v) . a \<le> u & u \<le> v & v \<le> b --> (\<exists>M. \<forall>x. u \<le> x & x \<le> v --> f x \<le> M)" in lemma_BOLZANO2)
   683 apply safe
   684 apply simp_all
   685 apply (rename_tac x xa ya M Ma)
   686 apply (metis linorder_not_less order_le_less order_trans)
   687 apply (case_tac "a \<le> x & x \<le> b")
   688  prefer 2
   689  apply (rule_tac x = 1 in exI, force)
   690 apply (simp add: LIM_eq isCont_iff)
   691 apply (drule_tac x = x in spec, auto)
   692 apply (erule_tac V = "\<forall>M. \<exists>x. a \<le> x & x \<le> b & ~ f x \<le> M" in thin_rl)
   693 apply (drule_tac x = 1 in spec, auto)
   694 apply (rule_tac x = s in exI, clarify)
   695 apply (rule_tac x = "\<bar>f x\<bar> + 1" in exI, clarify)
   696 apply (drule_tac x = "xa-x" in spec)
   697 apply (auto simp add: abs_ge_self)
   698 done
   699 
   700 text{*Refine the above to existence of least upper bound*}
   701 
   702 lemma lemma_reals_complete: "((\<exists>x. x \<in> S) & (\<exists>y. isUb UNIV S (y::real))) -->
   703       (\<exists>t. isLub UNIV S t)"
   704 by (blast intro: reals_complete)
   705 
   706 lemma isCont_has_Ub: "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
   707          ==> \<exists>M::real. (\<forall>x::real. a \<le> x & x \<le> b --> f(x) \<le> M) &
   708                    (\<forall>N. N < M --> (\<exists>x. a \<le> x & x \<le> b & N < f(x)))"
   709 apply (cut_tac S = "Collect (%y. \<exists>x. a \<le> x & x \<le> b & y = f x)"
   710         in lemma_reals_complete)
   711 apply auto
   712 apply (drule isCont_bounded, assumption)
   713 apply (auto simp add: isUb_def leastP_def isLub_def setge_def setle_def)
   714 apply (rule exI, auto)
   715 apply (auto dest!: spec simp add: linorder_not_less)
   716 done
   717 
   718 text{*Now show that it attains its upper bound*}
   719 
   720 lemma isCont_eq_Ub:
   721   assumes le: "a \<le> b"
   722       and con: "\<forall>x::real. a \<le> x & x \<le> b --> isCont f x"
   723   shows "\<exists>M::real. (\<forall>x. a \<le> x & x \<le> b --> f(x) \<le> M) &
   724              (\<exists>x. a \<le> x & x \<le> b & f(x) = M)"
   725 proof -
   726   from isCont_has_Ub [OF le con]
   727   obtain M where M1: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M"
   728              and M2: "!!N. N<M ==> \<exists>x. a \<le> x \<and> x \<le> b \<and> N < f x"  by blast
   729   show ?thesis
   730   proof (intro exI, intro conjI)
   731     show " \<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> f x \<le> M" by (rule M1)
   732     show "\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M"
   733     proof (rule ccontr)
   734       assume "\<not> (\<exists>x. a \<le> x \<and> x \<le> b \<and> f x = M)"
   735       with M1 have M3: "\<forall>x. a \<le> x & x \<le> b --> f x < M"
   736         by (fastforce simp add: linorder_not_le [symmetric])
   737       hence "\<forall>x. a \<le> x & x \<le> b --> isCont (%x. inverse (M - f x)) x"
   738         by (auto simp add: con)
   739       from isCont_bounded [OF le this]
   740       obtain k where k: "!!x. a \<le> x & x \<le> b --> inverse (M - f x) \<le> k" by auto
   741       have Minv: "!!x. a \<le> x & x \<le> b --> 0 < inverse (M - f (x))"
   742         by (simp add: M3 algebra_simps)
   743       have "!!x. a \<le> x & x \<le> b --> inverse (M - f x) < k+1" using k
   744         by (auto intro: order_le_less_trans [of _ k])
   745       with Minv
   746       have "!!x. a \<le> x & x \<le> b --> inverse(k+1) < inverse(inverse(M - f x))"
   747         by (intro strip less_imp_inverse_less, simp_all)
   748       hence invlt: "!!x. a \<le> x & x \<le> b --> inverse(k+1) < M - f x"
   749         by simp
   750       have "M - inverse (k+1) < M" using k [of a] Minv [of a] le
   751         by (simp, arith)
   752       from M2 [OF this]
   753       obtain x where ax: "a \<le> x & x \<le> b & M - inverse(k+1) < f x" ..
   754       thus False using invlt [of x] by force
   755     qed
   756   qed
   757 qed
   758 
   759 
   760 text{*Same theorem for lower bound*}
   761 
   762 lemma isCont_eq_Lb: "[| a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
   763          ==> \<exists>M::real. (\<forall>x::real. a \<le> x & x \<le> b --> M \<le> f(x)) &
   764                    (\<exists>x. a \<le> x & x \<le> b & f(x) = M)"
   765 apply (subgoal_tac "\<forall>x. a \<le> x & x \<le> b --> isCont (%x. - (f x)) x")
   766 prefer 2 apply (blast intro: isCont_minus)
   767 apply (drule_tac f = "(%x. - (f x))" in isCont_eq_Ub)
   768 apply safe
   769 apply auto
   770 done
   771 
   772 
   773 text{*Another version.*}
   774 
   775 lemma isCont_Lb_Ub: "[|a \<le> b; \<forall>x. a \<le> x & x \<le> b --> isCont f x |]
   776       ==> \<exists>L M::real. (\<forall>x::real. a \<le> x & x \<le> b --> L \<le> f(x) & f(x) \<le> M) &
   777           (\<forall>y. L \<le> y & y \<le> M --> (\<exists>x. a \<le> x & x \<le> b & (f(x) = y)))"
   778 apply (frule isCont_eq_Lb)
   779 apply (frule_tac [2] isCont_eq_Ub)
   780 apply (assumption+, safe)
   781 apply (rule_tac x = "f x" in exI)
   782 apply (rule_tac x = "f xa" in exI, simp, safe)
   783 apply (cut_tac x = x and y = xa in linorder_linear, safe)
   784 apply (cut_tac f = f and a = x and b = xa and y = y in IVT_objl)
   785 apply (cut_tac [2] f = f and a = xa and b = x and y = y in IVT2_objl, safe)
   786 apply (rule_tac [2] x = xb in exI)
   787 apply (rule_tac [4] x = xb in exI, simp_all)
   788 done
   789 
   790 
   791 subsection {* Local extrema *}
   792 
   793 text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
   794 
   795 lemma DERIV_pos_inc_right:
   796   fixes f :: "real => real"
   797   assumes der: "DERIV f x :> l"
   798       and l:   "0 < l"
   799   shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x + h)"
   800 proof -
   801   from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
   802   have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l)"
   803     by (simp add: diff_minus)
   804   then obtain s
   805         where s:   "0 < s"
   806           and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
   807     by auto
   808   thus ?thesis
   809   proof (intro exI conjI strip)
   810     show "0<s" using s .
   811     fix h::real
   812     assume "0 < h" "h < s"
   813     with all [of h] show "f x < f (x+h)"
   814     proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
   815     split add: split_if_asm)
   816       assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
   817       with l
   818       have "0 < (f (x+h) - f x) / h" by arith
   819       thus "f x < f (x+h)"
   820   by (simp add: pos_less_divide_eq h)
   821     qed
   822   qed
   823 qed
   824 
   825 lemma DERIV_neg_dec_left:
   826   fixes f :: "real => real"
   827   assumes der: "DERIV f x :> l"
   828       and l:   "l < 0"
   829   shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x-h)"
   830 proof -
   831   from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
   832   have "\<exists>s > 0. (\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l)"
   833     by (simp add: diff_minus)
   834   then obtain s
   835         where s:   "0 < s"
   836           and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
   837     by auto
   838   thus ?thesis
   839   proof (intro exI conjI strip)
   840     show "0<s" using s .
   841     fix h::real
   842     assume "0 < h" "h < s"
   843     with all [of "-h"] show "f x < f (x-h)"
   844     proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
   845     split add: split_if_asm)
   846       assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
   847       with l
   848       have "0 < (f (x-h) - f x) / h" by arith
   849       thus "f x < f (x-h)"
   850   by (simp add: pos_less_divide_eq h)
   851     qed
   852   qed
   853 qed
   854 
   855 
   856 lemma DERIV_pos_inc_left:
   857   fixes f :: "real => real"
   858   shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)"
   859   apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified])
   860   apply (auto simp add: DERIV_minus)
   861   done
   862 
   863 lemma DERIV_neg_dec_right:
   864   fixes f :: "real => real"
   865   shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x) > f(x + h)"
   866   apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified])
   867   apply (auto simp add: DERIV_minus)
   868   done
   869 
   870 lemma DERIV_local_max:
   871   fixes f :: "real => real"
   872   assumes der: "DERIV f x :> l"
   873       and d:   "0 < d"
   874       and le:  "\<forall>y. \<bar>x-y\<bar> < d --> f(y) \<le> f(x)"
   875   shows "l = 0"
   876 proof (cases rule: linorder_cases [of l 0])
   877   case equal thus ?thesis .
   878 next
   879   case less
   880   from DERIV_neg_dec_left [OF der less]
   881   obtain d' where d': "0 < d'"
   882              and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x-h)" by blast
   883   from real_lbound_gt_zero [OF d d']
   884   obtain e where "0 < e \<and> e < d \<and> e < d'" ..
   885   with lt le [THEN spec [where x="x-e"]]
   886   show ?thesis by (auto simp add: abs_if)
   887 next
   888   case greater
   889   from DERIV_pos_inc_right [OF der greater]
   890   obtain d' where d': "0 < d'"
   891              and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" by blast
   892   from real_lbound_gt_zero [OF d d']
   893   obtain e where "0 < e \<and> e < d \<and> e < d'" ..
   894   with lt le [THEN spec [where x="x+e"]]
   895   show ?thesis by (auto simp add: abs_if)
   896 qed
   897 
   898 
   899 text{*Similar theorem for a local minimum*}
   900 lemma DERIV_local_min:
   901   fixes f :: "real => real"
   902   shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0"
   903 by (drule DERIV_minus [THEN DERIV_local_max], auto)
   904 
   905 
   906 text{*In particular, if a function is locally flat*}
   907 lemma DERIV_local_const:
   908   fixes f :: "real => real"
   909   shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0"
   910 by (auto dest!: DERIV_local_max)
   911 
   912 
   913 subsection {* Rolle's Theorem *}
   914 
   915 text{*Lemma about introducing open ball in open interval*}
   916 lemma lemma_interval_lt:
   917      "[| a < x;  x < b |]
   918       ==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)"
   919 
   920 apply (simp add: abs_less_iff)
   921 apply (insert linorder_linear [of "x-a" "b-x"], safe)
   922 apply (rule_tac x = "x-a" in exI)
   923 apply (rule_tac [2] x = "b-x" in exI, auto)
   924 done
   925 
   926 lemma lemma_interval: "[| a < x;  x < b |] ==>
   927         \<exists>d::real. 0 < d &  (\<forall>y. \<bar>x-y\<bar> < d --> a \<le> y & y \<le> b)"
   928 apply (drule lemma_interval_lt, auto)
   929 apply force
   930 done
   931 
   932 text{*Rolle's Theorem.
   933    If @{term f} is defined and continuous on the closed interval
   934    @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
   935    and @{term "f(a) = f(b)"},
   936    then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
   937 theorem Rolle:
   938   assumes lt: "a < b"
   939       and eq: "f(a) = f(b)"
   940       and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
   941       and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
   942   shows "\<exists>z::real. a < z & z < b & DERIV f z :> 0"
   943 proof -
   944   have le: "a \<le> b" using lt by simp
   945   from isCont_eq_Ub [OF le con]
   946   obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x"
   947              and alex: "a \<le> x" and xleb: "x \<le> b"
   948     by blast
   949   from isCont_eq_Lb [OF le con]
   950   obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z"
   951               and alex': "a \<le> x'" and x'leb: "x' \<le> b"
   952     by blast
   953   show ?thesis
   954   proof cases
   955     assume axb: "a < x & x < b"
   956         --{*@{term f} attains its maximum within the interval*}
   957     hence ax: "a<x" and xb: "x<b" by arith + 
   958     from lemma_interval [OF ax xb]
   959     obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
   960       by blast
   961     hence bound': "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> f y \<le> f x" using x_max
   962       by blast
   963     from differentiableD [OF dif [OF axb]]
   964     obtain l where der: "DERIV f x :> l" ..
   965     have "l=0" by (rule DERIV_local_max [OF der d bound'])
   966         --{*the derivative at a local maximum is zero*}
   967     thus ?thesis using ax xb der by auto
   968   next
   969     assume notaxb: "~ (a < x & x < b)"
   970     hence xeqab: "x=a | x=b" using alex xleb by arith
   971     hence fb_eq_fx: "f b = f x" by (auto simp add: eq)
   972     show ?thesis
   973     proof cases
   974       assume ax'b: "a < x' & x' < b"
   975         --{*@{term f} attains its minimum within the interval*}
   976       hence ax': "a<x'" and x'b: "x'<b" by arith+ 
   977       from lemma_interval [OF ax' x'b]
   978       obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
   979   by blast
   980       hence bound': "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> f x' \<le> f y" using x'_min
   981   by blast
   982       from differentiableD [OF dif [OF ax'b]]
   983       obtain l where der: "DERIV f x' :> l" ..
   984       have "l=0" by (rule DERIV_local_min [OF der d bound'])
   985         --{*the derivative at a local minimum is zero*}
   986       thus ?thesis using ax' x'b der by auto
   987     next
   988       assume notax'b: "~ (a < x' & x' < b)"
   989         --{*@{term f} is constant througout the interval*}
   990       hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith
   991       hence fb_eq_fx': "f b = f x'" by (auto simp add: eq)
   992       from dense [OF lt]
   993       obtain r where ar: "a < r" and rb: "r < b" by blast
   994       from lemma_interval [OF ar rb]
   995       obtain d where d: "0<d" and bound: "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
   996   by blast
   997       have eq_fb: "\<forall>z. a \<le> z --> z \<le> b --> f z = f b"
   998       proof (clarify)
   999         fix z::real
  1000         assume az: "a \<le> z" and zb: "z \<le> b"
  1001         show "f z = f b"
  1002         proof (rule order_antisym)
  1003           show "f z \<le> f b" by (simp add: fb_eq_fx x_max az zb)
  1004           show "f b \<le> f z" by (simp add: fb_eq_fx' x'_min az zb)
  1005         qed
  1006       qed
  1007       have bound': "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> f r = f y"
  1008       proof (intro strip)
  1009         fix y::real
  1010         assume lt: "\<bar>r-y\<bar> < d"
  1011         hence "f y = f b" by (simp add: eq_fb bound)
  1012         thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le)
  1013       qed
  1014       from differentiableD [OF dif [OF conjI [OF ar rb]]]
  1015       obtain l where der: "DERIV f r :> l" ..
  1016       have "l=0" by (rule DERIV_local_const [OF der d bound'])
  1017         --{*the derivative of a constant function is zero*}
  1018       thus ?thesis using ar rb der by auto
  1019     qed
  1020   qed
  1021 qed
  1022 
  1023 
  1024 subsection{*Mean Value Theorem*}
  1025 
  1026 lemma lemma_MVT:
  1027      "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)"
  1028 proof cases
  1029   assume "a=b" thus ?thesis by simp
  1030 next
  1031   assume "a\<noteq>b"
  1032   hence ba: "b-a \<noteq> 0" by arith
  1033   show ?thesis
  1034     by (rule real_mult_left_cancel [OF ba, THEN iffD1],
  1035         simp add: right_diff_distrib,
  1036         simp add: left_diff_distrib)
  1037 qed
  1038 
  1039 theorem MVT:
  1040   assumes lt:  "a < b"
  1041       and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
  1042       and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
  1043   shows "\<exists>l z::real. a < z & z < b & DERIV f z :> l &
  1044                    (f(b) - f(a) = (b-a) * l)"
  1045 proof -
  1046   let ?F = "%x. f x - ((f b - f a) / (b-a)) * x"
  1047   have contF: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?F x"
  1048     using con by (fast intro: isCont_intros)
  1049   have difF: "\<forall>x. a < x \<and> x < b \<longrightarrow> ?F differentiable x"
  1050   proof (clarify)
  1051     fix x::real
  1052     assume ax: "a < x" and xb: "x < b"
  1053     from differentiableD [OF dif [OF conjI [OF ax xb]]]
  1054     obtain l where der: "DERIV f x :> l" ..
  1055     show "?F differentiable x"
  1056       by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"],
  1057           blast intro: DERIV_diff DERIV_cmult_Id der)
  1058   qed
  1059   from Rolle [where f = ?F, OF lt lemma_MVT contF difF]
  1060   obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0"
  1061     by blast
  1062   have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)"
  1063     by (rule DERIV_cmult_Id)
  1064   hence derF: "DERIV (\<lambda>x. ?F x + (f b - f a) / (b - a) * x) z
  1065                    :> 0 + (f b - f a) / (b - a)"
  1066     by (rule DERIV_add [OF der])
  1067   show ?thesis
  1068   proof (intro exI conjI)
  1069     show "a < z" using az .
  1070     show "z < b" using zb .
  1071     show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp)
  1072     show "DERIV f z :> ((f b - f a)/(b-a))"  using derF by simp
  1073   qed
  1074 qed
  1075 
  1076 lemma MVT2:
  1077      "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |]
  1078       ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))"
  1079 apply (drule MVT)
  1080 apply (blast intro: DERIV_isCont)
  1081 apply (force dest: order_less_imp_le simp add: differentiable_def)
  1082 apply (blast dest: DERIV_unique order_less_imp_le)
  1083 done
  1084 
  1085 
  1086 text{*A function is constant if its derivative is 0 over an interval.*}
  1087 
  1088 lemma DERIV_isconst_end:
  1089   fixes f :: "real => real"
  1090   shows "[| a < b;
  1091          \<forall>x. a \<le> x & x \<le> b --> isCont f x;
  1092          \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
  1093         ==> f b = f a"
  1094 apply (drule MVT, assumption)
  1095 apply (blast intro: differentiableI)
  1096 apply (auto dest!: DERIV_unique simp add: diff_eq_eq)
  1097 done
  1098 
  1099 lemma DERIV_isconst1:
  1100   fixes f :: "real => real"
  1101   shows "[| a < b;
  1102          \<forall>x. a \<le> x & x \<le> b --> isCont f x;
  1103          \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
  1104         ==> \<forall>x. a \<le> x & x \<le> b --> f x = f a"
  1105 apply safe
  1106 apply (drule_tac x = a in order_le_imp_less_or_eq, safe)
  1107 apply (drule_tac b = x in DERIV_isconst_end, auto)
  1108 done
  1109 
  1110 lemma DERIV_isconst2:
  1111   fixes f :: "real => real"
  1112   shows "[| a < b;
  1113          \<forall>x. a \<le> x & x \<le> b --> isCont f x;
  1114          \<forall>x. a < x & x < b --> DERIV f x :> 0;
  1115          a \<le> x; x \<le> b |]
  1116         ==> f x = f a"
  1117 apply (blast dest: DERIV_isconst1)
  1118 done
  1119 
  1120 lemma DERIV_isconst3: fixes a b x y :: real
  1121   assumes "a < b" and "x \<in> {a <..< b}" and "y \<in> {a <..< b}"
  1122   assumes derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
  1123   shows "f x = f y"
  1124 proof (cases "x = y")
  1125   case False
  1126   let ?a = "min x y"
  1127   let ?b = "max x y"
  1128   
  1129   have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0"
  1130   proof (rule allI, rule impI)
  1131     fix z :: real assume "?a \<le> z \<and> z \<le> ?b"
  1132     hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
  1133     hence "z \<in> {a<..<b}" by auto
  1134     thus "DERIV f z :> 0" by (rule derivable)
  1135   qed
  1136   hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z"
  1137     and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto
  1138 
  1139   have "?a < ?b" using `x \<noteq> y` by auto
  1140   from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y]
  1141   show ?thesis by auto
  1142 qed auto
  1143 
  1144 lemma DERIV_isconst_all:
  1145   fixes f :: "real => real"
  1146   shows "\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)"
  1147 apply (rule linorder_cases [of x y])
  1148 apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+
  1149 done
  1150 
  1151 lemma DERIV_const_ratio_const:
  1152   fixes f :: "real => real"
  1153   shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k"
  1154 apply (rule linorder_cases [of a b], auto)
  1155 apply (drule_tac [!] f = f in MVT)
  1156 apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
  1157 apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
  1158 done
  1159 
  1160 lemma DERIV_const_ratio_const2:
  1161   fixes f :: "real => real"
  1162   shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k"
  1163 apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1])
  1164 apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc)
  1165 done
  1166 
  1167 lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)"
  1168 by (simp)
  1169 
  1170 lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)"
  1171 by (simp)
  1172 
  1173 text{*Gallileo's "trick": average velocity = av. of end velocities*}
  1174 
  1175 lemma DERIV_const_average:
  1176   fixes v :: "real => real"
  1177   assumes neq: "a \<noteq> (b::real)"
  1178       and der: "\<forall>x. DERIV v x :> k"
  1179   shows "v ((a + b)/2) = (v a + v b)/2"
  1180 proof (cases rule: linorder_cases [of a b])
  1181   case equal with neq show ?thesis by simp
  1182 next
  1183   case less
  1184   have "(v b - v a) / (b - a) = k"
  1185     by (rule DERIV_const_ratio_const2 [OF neq der])
  1186   hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
  1187   moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k"
  1188     by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
  1189   ultimately show ?thesis using neq by force
  1190 next
  1191   case greater
  1192   have "(v b - v a) / (b - a) = k"
  1193     by (rule DERIV_const_ratio_const2 [OF neq der])
  1194   hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
  1195   moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k"
  1196     by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
  1197   ultimately show ?thesis using neq by (force simp add: add_commute)
  1198 qed
  1199 
  1200 (* A function with positive derivative is increasing. 
  1201    A simple proof using the MVT, by Jeremy Avigad. And variants.
  1202 *)
  1203 lemma DERIV_pos_imp_increasing:
  1204   fixes a::real and b::real and f::"real => real"
  1205   assumes "a < b" and "\<forall>x. a \<le> x & x \<le> b --> (EX y. DERIV f x :> y & y > 0)"
  1206   shows "f a < f b"
  1207 proof (rule ccontr)
  1208   assume f: "~ f a < f b"
  1209   have "EX l z. a < z & z < b & DERIV f z :> l
  1210       & f b - f a = (b - a) * l"
  1211     apply (rule MVT)
  1212       using assms
  1213       apply auto
  1214       apply (metis DERIV_isCont)
  1215      apply (metis differentiableI less_le)
  1216     done
  1217   then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
  1218       and "f b - f a = (b - a) * l"
  1219     by auto
  1220   with assms f have "~(l > 0)"
  1221     by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le)
  1222   with assms z show False
  1223     by (metis DERIV_unique less_le)
  1224 qed
  1225 
  1226 lemma DERIV_nonneg_imp_nonincreasing:
  1227   fixes a::real and b::real and f::"real => real"
  1228   assumes "a \<le> b" and
  1229     "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<ge> 0)"
  1230   shows "f a \<le> f b"
  1231 proof (rule ccontr, cases "a = b")
  1232   assume "~ f a \<le> f b" and "a = b"
  1233   then show False by auto
  1234 next
  1235   assume A: "~ f a \<le> f b"
  1236   assume B: "a ~= b"
  1237   with assms have "EX l z. a < z & z < b & DERIV f z :> l
  1238       & f b - f a = (b - a) * l"
  1239     apply -
  1240     apply (rule MVT)
  1241       apply auto
  1242       apply (metis DERIV_isCont)
  1243      apply (metis differentiableI less_le)
  1244     done
  1245   then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
  1246       and C: "f b - f a = (b - a) * l"
  1247     by auto
  1248   with A have "a < b" "f b < f a" by auto
  1249   with C have "\<not> l \<ge> 0" by (auto simp add: not_le algebra_simps)
  1250     (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl)
  1251   with assms z show False
  1252     by (metis DERIV_unique order_less_imp_le)
  1253 qed
  1254 
  1255 lemma DERIV_neg_imp_decreasing:
  1256   fixes a::real and b::real and f::"real => real"
  1257   assumes "a < b" and
  1258     "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y < 0)"
  1259   shows "f a > f b"
  1260 proof -
  1261   have "(%x. -f x) a < (%x. -f x) b"
  1262     apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"])
  1263     using assms
  1264     apply auto
  1265     apply (metis DERIV_minus neg_0_less_iff_less)
  1266     done
  1267   thus ?thesis
  1268     by simp
  1269 qed
  1270 
  1271 lemma DERIV_nonpos_imp_nonincreasing:
  1272   fixes a::real and b::real and f::"real => real"
  1273   assumes "a \<le> b" and
  1274     "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<le> 0)"
  1275   shows "f a \<ge> f b"
  1276 proof -
  1277   have "(%x. -f x) a \<le> (%x. -f x) b"
  1278     apply (rule DERIV_nonneg_imp_nonincreasing [of a b "%x. -f x"])
  1279     using assms
  1280     apply auto
  1281     apply (metis DERIV_minus neg_0_le_iff_le)
  1282     done
  1283   thus ?thesis
  1284     by simp
  1285 qed
  1286 
  1287 subsection {* Continuous injective functions *}
  1288 
  1289 text{*Dull lemma: an continuous injection on an interval must have a
  1290 strict maximum at an end point, not in the middle.*}
  1291 
  1292 lemma lemma_isCont_inj:
  1293   fixes f :: "real \<Rightarrow> real"
  1294   assumes d: "0 < d"
  1295       and inj [rule_format]: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z"
  1296       and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z"
  1297   shows "\<exists>z. \<bar>z-x\<bar> \<le> d & f x < f z"
  1298 proof (rule ccontr)
  1299   assume  "~ (\<exists>z. \<bar>z-x\<bar> \<le> d & f x < f z)"
  1300   hence all [rule_format]: "\<forall>z. \<bar>z - x\<bar> \<le> d --> f z \<le> f x" by auto
  1301   show False
  1302   proof (cases rule: linorder_le_cases [of "f(x-d)" "f(x+d)"])
  1303     case le
  1304     from d cont all [of "x+d"]
  1305     have flef: "f(x+d) \<le> f x"
  1306      and xlex: "x - d \<le> x"
  1307      and cont': "\<forall>z. x - d \<le> z \<and> z \<le> x \<longrightarrow> isCont f z"
  1308        by (auto simp add: abs_if)
  1309     from IVT [OF le flef xlex cont']
  1310     obtain x' where "x-d \<le> x'" "x' \<le> x" "f x' = f(x+d)" by blast
  1311     moreover
  1312     hence "g(f x') = g (f(x+d))" by simp
  1313     ultimately show False using d inj [of x'] inj [of "x+d"]
  1314       by (simp add: abs_le_iff)
  1315   next
  1316     case ge
  1317     from d cont all [of "x-d"]
  1318     have flef: "f(x-d) \<le> f x"
  1319      and xlex: "x \<le> x+d"
  1320      and cont': "\<forall>z. x \<le> z \<and> z \<le> x+d \<longrightarrow> isCont f z"
  1321        by (auto simp add: abs_if)
  1322     from IVT2 [OF ge flef xlex cont']
  1323     obtain x' where "x \<le> x'" "x' \<le> x+d" "f x' = f(x-d)" by blast
  1324     moreover
  1325     hence "g(f x') = g (f(x-d))" by simp
  1326     ultimately show False using d inj [of x'] inj [of "x-d"]
  1327       by (simp add: abs_le_iff)
  1328   qed
  1329 qed
  1330 
  1331 
  1332 text{*Similar version for lower bound.*}
  1333 
  1334 lemma lemma_isCont_inj2:
  1335   fixes f g :: "real \<Rightarrow> real"
  1336   shows "[|0 < d; \<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z;
  1337         \<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z |]
  1338       ==> \<exists>z. \<bar>z-x\<bar> \<le> d & f z < f x"
  1339 apply (insert lemma_isCont_inj
  1340           [where f = "%x. - f x" and g = "%y. g(-y)" and x = x and d = d])
  1341 apply (simp add: linorder_not_le)
  1342 done
  1343 
  1344 text{*Show there's an interval surrounding @{term "f(x)"} in
  1345 @{text "f[[x - d, x + d]]"} .*}
  1346 
  1347 lemma isCont_inj_range:
  1348   fixes f :: "real \<Rightarrow> real"
  1349   assumes d: "0 < d"
  1350       and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z"
  1351       and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z"
  1352   shows "\<exists>e>0. \<forall>y. \<bar>y - f x\<bar> \<le> e --> (\<exists>z. \<bar>z-x\<bar> \<le> d & f z = y)"
  1353 proof -
  1354   have "x-d \<le> x+d" "\<forall>z. x-d \<le> z \<and> z \<le> x+d \<longrightarrow> isCont f z" using cont d
  1355     by (auto simp add: abs_le_iff)
  1356   from isCont_Lb_Ub [OF this]
  1357   obtain L M
  1358   where all1 [rule_format]: "\<forall>z. x-d \<le> z \<and> z \<le> x+d \<longrightarrow> L \<le> f z \<and> f z \<le> M"
  1359     and all2 [rule_format]:
  1360            "\<forall>y. L \<le> y \<and> y \<le> M \<longrightarrow> (\<exists>z. x-d \<le> z \<and> z \<le> x+d \<and> f z = y)"
  1361     by auto
  1362   with d have "L \<le> f x & f x \<le> M" by simp
  1363   moreover have "L \<noteq> f x"
  1364   proof -
  1365     from lemma_isCont_inj2 [OF d inj cont]
  1366     obtain u where "\<bar>u - x\<bar> \<le> d" "f u < f x"  by auto
  1367     thus ?thesis using all1 [of u] by arith
  1368   qed
  1369   moreover have "f x \<noteq> M"
  1370   proof -
  1371     from lemma_isCont_inj [OF d inj cont]
  1372     obtain u where "\<bar>u - x\<bar> \<le> d" "f x < f u"  by auto
  1373     thus ?thesis using all1 [of u] by arith
  1374   qed
  1375   ultimately have "L < f x & f x < M" by arith
  1376   hence "0 < f x - L" "0 < M - f x" by arith+
  1377   from real_lbound_gt_zero [OF this]
  1378   obtain e where e: "0 < e" "e < f x - L" "e < M - f x" by auto
  1379   thus ?thesis
  1380   proof (intro exI conjI)
  1381     show "0<e" using e(1) .
  1382     show "\<forall>y. \<bar>y - f x\<bar> \<le> e \<longrightarrow> (\<exists>z. \<bar>z - x\<bar> \<le> d \<and> f z = y)"
  1383     proof (intro strip)
  1384       fix y::real
  1385       assume "\<bar>y - f x\<bar> \<le> e"
  1386       with e have "L \<le> y \<and> y \<le> M" by arith
  1387       from all2 [OF this]
  1388       obtain z where "x - d \<le> z" "z \<le> x + d" "f z = y" by blast
  1389       thus "\<exists>z. \<bar>z - x\<bar> \<le> d \<and> f z = y" 
  1390         by (force simp add: abs_le_iff)
  1391     qed
  1392   qed
  1393 qed
  1394 
  1395 
  1396 text{*Continuity of inverse function*}
  1397 
  1398 lemma isCont_inverse_function:
  1399   fixes f g :: "real \<Rightarrow> real"
  1400   assumes d: "0 < d"
  1401       and inj: "\<forall>z. \<bar>z-x\<bar> \<le> d --> g(f z) = z"
  1402       and cont: "\<forall>z. \<bar>z-x\<bar> \<le> d --> isCont f z"
  1403   shows "isCont g (f x)"
  1404 proof (simp add: isCont_iff LIM_eq)
  1405   show "\<forall>r. 0 < r \<longrightarrow>
  1406          (\<exists>s>0. \<forall>z. z\<noteq>0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>g(f x + z) - g(f x)\<bar> < r)"
  1407   proof (intro strip)
  1408     fix r::real
  1409     assume r: "0<r"
  1410     from real_lbound_gt_zero [OF r d]
  1411     obtain e where e: "0 < e" and e_lt: "e < r \<and> e < d" by blast
  1412     with inj cont
  1413     have e_simps: "\<forall>z. \<bar>z-x\<bar> \<le> e --> g (f z) = z"
  1414                   "\<forall>z. \<bar>z-x\<bar> \<le> e --> isCont f z"   by auto
  1415     from isCont_inj_range [OF e this]
  1416     obtain e' where e': "0 < e'"
  1417         and all: "\<forall>y. \<bar>y - f x\<bar> \<le> e' \<longrightarrow> (\<exists>z. \<bar>z - x\<bar> \<le> e \<and> f z = y)"
  1418           by blast
  1419     show "\<exists>s>0. \<forall>z. z\<noteq>0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>g(f x + z) - g(f x)\<bar> < r"
  1420     proof (intro exI conjI)
  1421       show "0<e'" using e' .
  1422       show "\<forall>z. z \<noteq> 0 \<and> \<bar>z\<bar> < e' \<longrightarrow> \<bar>g (f x + z) - g (f x)\<bar> < r"
  1423       proof (intro strip)
  1424         fix z::real
  1425         assume z: "z \<noteq> 0 \<and> \<bar>z\<bar> < e'"
  1426         with e e_lt e_simps all [rule_format, of "f x + z"]
  1427         show "\<bar>g (f x + z) - g (f x)\<bar> < r" by force
  1428       qed
  1429     qed
  1430   qed
  1431 qed
  1432 
  1433 text {* Derivative of inverse function *}
  1434 
  1435 lemma DERIV_inverse_function:
  1436   fixes f g :: "real \<Rightarrow> real"
  1437   assumes der: "DERIV f (g x) :> D"
  1438   assumes neq: "D \<noteq> 0"
  1439   assumes a: "a < x" and b: "x < b"
  1440   assumes inj: "\<forall>y. a < y \<and> y < b \<longrightarrow> f (g y) = y"
  1441   assumes cont: "isCont g x"
  1442   shows "DERIV g x :> inverse D"
  1443 unfolding DERIV_iff2
  1444 proof (rule LIM_equal2)
  1445   show "0 < min (x - a) (b - x)"
  1446     using a b by arith 
  1447 next
  1448   fix y
  1449   assume "norm (y - x) < min (x - a) (b - x)"
  1450   hence "a < y" and "y < b" 
  1451     by (simp_all add: abs_less_iff)
  1452   thus "(g y - g x) / (y - x) =
  1453         inverse ((f (g y) - x) / (g y - g x))"
  1454     by (simp add: inj)
  1455 next
  1456   have "(\<lambda>z. (f z - f (g x)) / (z - g x)) -- g x --> D"
  1457     by (rule der [unfolded DERIV_iff2])
  1458   hence 1: "(\<lambda>z. (f z - x) / (z - g x)) -- g x --> D"
  1459     using inj a b by simp
  1460   have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x"
  1461   proof (safe intro!: exI)
  1462     show "0 < min (x - a) (b - x)"
  1463       using a b by simp
  1464   next
  1465     fix y
  1466     assume "norm (y - x) < min (x - a) (b - x)"
  1467     hence y: "a < y" "y < b"
  1468       by (simp_all add: abs_less_iff)
  1469     assume "g y = g x"
  1470     hence "f (g y) = f (g x)" by simp
  1471     hence "y = x" using inj y a b by simp
  1472     also assume "y \<noteq> x"
  1473     finally show False by simp
  1474   qed
  1475   have "(\<lambda>y. (f (g y) - x) / (g y - g x)) -- x --> D"
  1476     using cont 1 2 by (rule isCont_LIM_compose2)
  1477   thus "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x)))
  1478         -- x --> inverse D"
  1479     using neq by (rule tendsto_inverse)
  1480 qed
  1481 
  1482 
  1483 subsection {* Generalized Mean Value Theorem *}
  1484 
  1485 theorem GMVT:
  1486   fixes a b :: real
  1487   assumes alb: "a < b"
  1488     and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
  1489     and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable x"
  1490     and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x"
  1491     and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable x"
  1492   shows "\<exists>g'c f'c c. DERIV g c :> g'c \<and> DERIV f c :> f'c \<and> a < c \<and> c < b \<and> ((f b - f a) * g'c) = ((g b - g a) * f'c)"
  1493 proof -
  1494   let ?h = "\<lambda>x. (f b - f a)*(g x) - (g b - g a)*(f x)"
  1495   from assms have "a < b" by simp
  1496   moreover have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?h x"
  1497     using fc gc by simp
  1498   moreover have "\<forall>x. a < x \<and> x < b \<longrightarrow> ?h differentiable x"
  1499     using fd gd by simp
  1500   ultimately have "\<exists>l z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" by (rule MVT)
  1501   then obtain l where ldef: "\<exists>z. a < z \<and> z < b \<and> DERIV ?h z :> l \<and> ?h b - ?h a = (b - a) * l" ..
  1502   then obtain c where cdef: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" ..
  1503 
  1504   from cdef have cint: "a < c \<and> c < b" by auto
  1505   with gd have "g differentiable c" by simp
  1506   hence "\<exists>D. DERIV g c :> D" by (rule differentiableD)
  1507   then obtain g'c where g'cdef: "DERIV g c :> g'c" ..
  1508 
  1509   from cdef have "a < c \<and> c < b" by auto
  1510   with fd have "f differentiable c" by simp
  1511   hence "\<exists>D. DERIV f c :> D" by (rule differentiableD)
  1512   then obtain f'c where f'cdef: "DERIV f c :> f'c" ..
  1513 
  1514   from cdef have "DERIV ?h c :> l" by auto
  1515   moreover have "DERIV ?h c :>  g'c * (f b - f a) - f'c * (g b - g a)"
  1516     using g'cdef f'cdef by (auto intro!: DERIV_intros)
  1517   ultimately have leq: "l =  g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique)
  1518 
  1519   {
  1520     from cdef have "?h b - ?h a = (b - a) * l" by auto
  1521     also with leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
  1522     finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
  1523   }
  1524   moreover
  1525   {
  1526     have "?h b - ?h a =
  1527          ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) -
  1528           ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))"
  1529       by (simp add: algebra_simps)
  1530     hence "?h b - ?h a = 0" by auto
  1531   }
  1532   ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto
  1533   with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp
  1534   hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp
  1535   hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac)
  1536 
  1537   with g'cdef f'cdef cint show ?thesis by auto
  1538 qed
  1539 
  1540 
  1541 subsection {* Theorems about Limits *}
  1542 
  1543 (* need to rename second isCont_inverse *)
  1544 
  1545 lemma isCont_inv_fun:
  1546   fixes f g :: "real \<Rightarrow> real"
  1547   shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;  
  1548          \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]  
  1549       ==> isCont g (f x)"
  1550 by (rule isCont_inverse_function)
  1551 
  1552 lemma isCont_inv_fun_inv:
  1553   fixes f g :: "real \<Rightarrow> real"
  1554   shows "[| 0 < d;  
  1555          \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z;  
  1556          \<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |]  
  1557        ==> \<exists>e. 0 < e &  
  1558              (\<forall>y. 0 < \<bar>y - f(x)\<bar> & \<bar>y - f(x)\<bar> < e --> f(g(y)) = y)"
  1559 apply (drule isCont_inj_range)
  1560 prefer 2 apply (assumption, assumption, auto)
  1561 apply (rule_tac x = e in exI, auto)
  1562 apply (rotate_tac 2)
  1563 apply (drule_tac x = y in spec, auto)
  1564 done
  1565 
  1566 
  1567 text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*}
  1568 lemma LIM_fun_gt_zero:
  1569      "[| f -- c --> (l::real); 0 < l |]  
  1570          ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> 0 < f x)"
  1571 apply (drule (1) LIM_D, clarify)
  1572 apply (rule_tac x = s in exI)
  1573 apply (simp add: abs_less_iff)
  1574 done
  1575 
  1576 lemma LIM_fun_less_zero:
  1577      "[| f -- c --> (l::real); l < 0 |]  
  1578       ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x < 0)"
  1579 apply (drule LIM_D [where r="-l"], simp, clarify)
  1580 apply (rule_tac x = s in exI)
  1581 apply (simp add: abs_less_iff)
  1582 done
  1583 
  1584 lemma LIM_fun_not_zero:
  1585      "[| f -- c --> (l::real); l \<noteq> 0 |] 
  1586       ==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x \<noteq> 0)"
  1587 apply (rule linorder_cases [of l 0])
  1588 apply (drule (1) LIM_fun_less_zero, force)
  1589 apply simp
  1590 apply (drule (1) LIM_fun_gt_zero, force)
  1591 done
  1592 
  1593 end