src/HOL/Deriv.thy
author hoelzl
Tue, 09 Apr 2013 14:04:41 +0200
changeset 51641 cd05e9fcc63d
parent 51529 2d2f59e6055a
child 51642 400ec5ae7f8f
permissions -rw-r--r--
remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     1
(*  Title       : Deriv.thy
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     5
    GMVT by Benjamin Porter, 2005
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     6
*)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     7
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     8
header{* Differentiation *}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
     9
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    10
theory Deriv
51526
155263089e7b move SEQ.thy and Lim.thy to Limits.thy
hoelzl
parents: 51481
diff changeset
    11
imports Limits
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    12
begin
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    13
22984
67d434ad9ef8 section labels
huffman
parents: 22653
diff changeset
    14
text{*Standard Definitions*}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    15
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    16
definition
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    17
  deriv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    18
    --{*Differentiation: D is derivative of function f at x*}
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21239
diff changeset
    19
          ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    20
  "DERIV f x :> D = ((%h. (f(x + h) - f x) / h) -- 0 --> D)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    21
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    22
subsection {* Derivatives *}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    23
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    24
lemma DERIV_iff: "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --> D)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    25
by (simp add: deriv_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    26
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    27
lemma DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --> D"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    28
by (simp add: deriv_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    29
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    30
lemma DERIV_const [simp]: "DERIV (\<lambda>x. k) x :> 0"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    31
  by (simp add: deriv_def tendsto_const)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    32
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23044
diff changeset
    33
lemma DERIV_ident [simp]: "DERIV (\<lambda>x. x) x :> 1"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    34
  by (simp add: deriv_def tendsto_const cong: LIM_cong)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    35
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    36
lemma DERIV_add:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    37
  "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + g x) x :> D + E"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    38
  by (simp only: deriv_def add_diff_add add_divide_distrib tendsto_add)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    39
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    40
lemma DERIV_minus:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    41
  "DERIV f x :> D \<Longrightarrow> DERIV (\<lambda>x. - f x) x :> - D"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    42
  by (simp only: deriv_def minus_diff_minus divide_minus_left tendsto_minus)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    43
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    44
lemma DERIV_diff:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    45
  "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x - g x) x :> D - E"
37887
2ae085b07f2f diff_minus subsumes diff_def
haftmann
parents: 36777
diff changeset
    46
by (simp only: diff_minus DERIV_add DERIV_minus)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    47
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    48
lemma DERIV_add_minus:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    49
  "\<lbrakk>DERIV f x :> D; DERIV g x :> E\<rbrakk> \<Longrightarrow> DERIV (\<lambda>x. f x + - g x) x :> D + - E"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    50
by (simp only: DERIV_add DERIV_minus)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    51
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    52
lemma DERIV_isCont: "DERIV f x :> D \<Longrightarrow> isCont f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    53
proof (unfold isCont_iff)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    54
  assume "DERIV f x :> D"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    55
  hence "(\<lambda>h. (f(x+h) - f(x)) / h) -- 0 --> D"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    56
    by (rule DERIV_D)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    57
  hence "(\<lambda>h. (f(x+h) - f(x)) / h * h) -- 0 --> D * 0"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    58
    by (intro tendsto_mult tendsto_ident_at)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    59
  hence "(\<lambda>h. (f(x+h) - f(x)) * (h / h)) -- 0 --> 0"
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    60
    by simp
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    61
  hence "(\<lambda>h. f(x+h) - f(x)) -- 0 --> 0"
23398
0b5a400c7595 made divide_self a simp rule
nipkow
parents: 23255
diff changeset
    62
    by (simp cong: LIM_cong)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    63
  thus "(\<lambda>h. f(x+h)) -- 0 --> f(x)"
31338
d41a8ba25b67 generalize constants from Lim.thy to class metric_space
huffman
parents: 31336
diff changeset
    64
    by (simp add: LIM_def dist_norm)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    65
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    66
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    67
lemma DERIV_mult_lemma:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    68
  fixes a b c d :: "'a::real_field"
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    69
  shows "(a * b - c * d) / h = a * ((b - d) / h) + ((a - c) / h) * d"
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
    70
  by (simp add: field_simps diff_divide_distrib)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    71
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    72
lemma DERIV_mult':
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    73
  assumes f: "DERIV f x :> D"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    74
  assumes g: "DERIV g x :> E"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    75
  shows "DERIV (\<lambda>x. f x * g x) x :> f x * E + D * g x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    76
proof (unfold deriv_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    77
  from f have "isCont f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    78
    by (rule DERIV_isCont)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    79
  hence "(\<lambda>h. f(x+h)) -- 0 --> f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    80
    by (simp only: isCont_iff)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    81
  hence "(\<lambda>h. f(x+h) * ((g(x+h) - g x) / h) +
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    82
              ((f(x+h) - f x) / h) * g x)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    83
          -- 0 --> f x * E + D * g x"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
    84
    by (intro tendsto_intros DERIV_D f g)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
    85
  thus "(\<lambda>h. (f(x+h) * g(x+h) - f x * g x) / h)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    86
         -- 0 --> f x * E + D * g x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    87
    by (simp only: DERIV_mult_lemma)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    88
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    89
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    90
lemma DERIV_mult:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
    91
    "DERIV f x :> Da \<Longrightarrow> DERIV g x :> Db \<Longrightarrow> DERIV (\<lambda>x. f x * g x) x :> Da * g x + Db * f x"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
    92
  by (drule (1) DERIV_mult', simp only: mult_commute add_commute)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    93
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    94
lemma DERIV_unique:
50331
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
    95
    "DERIV f x :> D \<Longrightarrow> DERIV f x :> E \<Longrightarrow> D = E"
4b6dc5077e98 use filterlim in Lim and SEQ; tuned proofs
hoelzl
parents: 50330
diff changeset
    96
  unfolding deriv_def by (rule LIM_unique) 
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    97
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    98
text{*Differentiation of finite sum*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
    99
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   100
lemma DERIV_setsum:
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   101
  assumes "finite S"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   102
  and "\<And> n. n \<in> S \<Longrightarrow> DERIV (%x. f x n) x :> (f' x n)"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   103
  shows "DERIV (%x. setsum (f x) S) x :> setsum (f' x) S"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   104
  using assms by induct (auto intro!: DERIV_add)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   105
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   106
lemma DERIV_sumr [rule_format (no_asm)]:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   107
     "(\<forall>r. m \<le> r & r < (m + n) --> DERIV (%x. f r x) x :> (f' r x))
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   108
      --> DERIV (%x. \<Sum>n=m..<n::nat. f n x :: real) x :> (\<Sum>r=m..<n. f' r x)"
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   109
  by (auto intro: DERIV_setsum)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   110
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   111
text{*Alternative definition for differentiability*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   112
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   113
lemma DERIV_LIM_iff:
31338
d41a8ba25b67 generalize constants from Lim.thy to class metric_space
huffman
parents: 31336
diff changeset
   114
  fixes f :: "'a::{real_normed_vector,inverse} \<Rightarrow> 'a" shows
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   115
     "((%h. (f(a + h) - f(a)) / h) -- 0 --> D) =
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   116
      ((%x. (f(x)-f(a)) / (x-a)) -- a --> D)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   117
apply (rule iffI)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   118
apply (drule_tac k="- a" in LIM_offset)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   119
apply (simp add: diff_minus)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   120
apply (drule_tac k="a" in LIM_offset)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   121
apply (simp add: add_commute)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   122
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   123
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   124
lemma DERIV_iff2: "(DERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --> D)"
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   125
by (simp add: deriv_def diff_minus [symmetric] DERIV_LIM_iff)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   126
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   127
lemma DERIV_inverse_lemma:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   128
  "\<lbrakk>a \<noteq> 0; b \<noteq> (0::'a::real_normed_field)\<rbrakk>
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   129
   \<Longrightarrow> (inverse a - inverse b) / h
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   130
     = - (inverse a * ((a - b) / h) * inverse b)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   131
by (simp add: inverse_diff_inverse)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   132
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   133
lemma DERIV_inverse':
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   134
  assumes der: "DERIV f x :> D"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   135
  assumes neq: "f x \<noteq> 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   136
  shows "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * D * inverse (f x))"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   137
    (is "DERIV _ _ :> ?E")
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   138
proof (unfold DERIV_iff2)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   139
  from der have lim_f: "f -- x --> f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   140
    by (rule DERIV_isCont [unfolded isCont_def])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   141
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   142
  from neq have "0 < norm (f x)" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   143
  with LIM_D [OF lim_f] obtain s
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   144
    where s: "0 < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   145
    and less_fx: "\<And>z. \<lbrakk>z \<noteq> x; norm (z - x) < s\<rbrakk>
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   146
                  \<Longrightarrow> norm (f z - f x) < norm (f x)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   147
    by fast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   148
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   149
  show "(\<lambda>z. (inverse (f z) - inverse (f x)) / (z - x)) -- x --> ?E"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   150
  proof (rule LIM_equal2 [OF s])
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   151
    fix z
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   152
    assume "z \<noteq> x" "norm (z - x) < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   153
    hence "norm (f z - f x) < norm (f x)" by (rule less_fx)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   154
    hence "f z \<noteq> 0" by auto
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   155
    thus "(inverse (f z) - inverse (f x)) / (z - x) =
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   156
          - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x))"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   157
      using neq by (rule DERIV_inverse_lemma)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   158
  next
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   159
    from der have "(\<lambda>z. (f z - f x) / (z - x)) -- x --> D"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   160
      by (unfold DERIV_iff2)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   161
    thus "(\<lambda>z. - (inverse (f z) * ((f z - f x) / (z - x)) * inverse (f x)))
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   162
          -- x --> ?E"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
   163
      by (intro tendsto_intros lim_f neq)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   164
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   165
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   166
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   167
lemma DERIV_divide:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   168
  "\<lbrakk>DERIV f x :> D; DERIV g x :> E; g x \<noteq> 0\<rbrakk>
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   169
   \<Longrightarrow> DERIV (\<lambda>x. f x / g x) x :> (D * g x - f x * E) / (g x * g x)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   170
apply (subgoal_tac "f x * - (inverse (g x) * E * inverse (g x)) +
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   171
          D * inverse (g x) = (D * g x - f x * E) / (g x * g x)")
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   172
apply (erule subst)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   173
apply (unfold divide_inverse)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   174
apply (erule DERIV_mult')
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   175
apply (erule (1) DERIV_inverse')
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23441
diff changeset
   176
apply (simp add: ring_distribs nonzero_inverse_mult_distrib)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   177
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   178
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   179
lemma DERIV_power_Suc:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   180
  fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   181
  assumes f: "DERIV f x :> D"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23413
diff changeset
   182
  shows "DERIV (\<lambda>x. f x ^ Suc n) x :> (1 + of_nat n) * (D * f x ^ n)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   183
proof (induct n)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   184
case 0
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   185
  show ?case by (simp add: f)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   186
case (Suc k)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   187
  from DERIV_mult' [OF f Suc] show ?case
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23441
diff changeset
   188
    apply (simp only: of_nat_Suc ring_distribs mult_1_left)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29472
diff changeset
   189
    apply (simp only: power_Suc algebra_simps)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   190
    done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   191
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   192
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   193
lemma DERIV_power:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   194
  fixes f :: "'a \<Rightarrow> 'a::{real_normed_field}"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   195
  assumes f: "DERIV f x :> D"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   196
  shows "DERIV (\<lambda>x. f x ^ n) x :> of_nat n * (D * f x ^ (n - Suc 0))"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   197
by (cases "n", simp, simp add: DERIV_power_Suc f del: power_Suc)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   198
29975
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   199
text {* Caratheodory formulation of derivative at a point *}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   200
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   201
lemma CARAT_DERIV:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   202
     "(DERIV f x :> l) =
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   203
      (\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isCont g x & g x = l)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   204
      (is "?lhs = ?rhs")
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   205
proof
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   206
  assume der: "DERIV f x :> l"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   207
  show "\<exists>g. (\<forall>z. f z - f x = g z * (z-x)) \<and> isCont g x \<and> g x = l"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   208
  proof (intro exI conjI)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   209
    let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23412
diff changeset
   210
    show "\<forall>z. f z - f x = ?g z * (z-x)" by simp
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   211
    show "isCont ?g x" using der
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   212
      by (simp add: isCont_iff DERIV_iff diff_minus
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   213
               cong: LIM_equal [rule_format])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   214
    show "?g x = l" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   215
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   216
next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   217
  assume "?rhs"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   218
  then obtain g where
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   219
    "(\<forall>z. f z - f x = g z * (z-x))" and "isCont g x" and "g x = l" by blast
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   220
  thus "(DERIV f x :> l)"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23412
diff changeset
   221
     by (auto simp add: isCont_iff DERIV_iff cong: LIM_cong)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   222
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   223
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   224
lemma DERIV_chain':
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   225
  assumes f: "DERIV f x :> D"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   226
  assumes g: "DERIV g (f x) :> E"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   227
  shows "DERIV (\<lambda>x. g (f x)) x :> E * D"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   228
proof (unfold DERIV_iff2)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   229
  obtain d where d: "\<forall>y. g y - g (f x) = d y * (y - f x)"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   230
    and cont_d: "isCont d (f x)" and dfx: "d (f x) = E"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   231
    using CARAT_DERIV [THEN iffD1, OF g] by fast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   232
  from f have "f -- x --> f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   233
    by (rule DERIV_isCont [unfolded isCont_def])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   234
  with cont_d have "(\<lambda>z. d (f z)) -- x --> d (f x)"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
   235
    by (rule isCont_tendsto_compose)
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   236
  hence "(\<lambda>z. d (f z) * ((f z - f x) / (z - x)))
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   237
          -- x --> d (f x) * D"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
   238
    by (rule tendsto_mult [OF _ f [unfolded DERIV_iff2]])
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   239
  thus "(\<lambda>z. (g (f z) - g (f x)) / (z - x)) -- x --> E * D"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 34941
diff changeset
   240
    by (simp add: d dfx)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   241
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   242
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   243
text {*
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   244
 Let's do the standard proof, though theorem
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   245
 @{text "LIM_mult2"} follows from a NS proof
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   246
*}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   247
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   248
lemma DERIV_cmult:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   249
      "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   250
by (drule DERIV_mult' [OF DERIV_const], simp)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   251
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   252
lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   253
  apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   254
  apply (erule DERIV_cmult)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   255
  done
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   256
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   257
text {* Standard version *}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   258
lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 34941
diff changeset
   259
by (drule (1) DERIV_chain', simp add: o_def mult_commute)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   260
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   261
lemma DERIV_chain2: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (%x. f (g x)) x :> Da * Db"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   262
by (auto dest: DERIV_chain simp add: o_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   263
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   264
text {* Derivative of linear multiplication *}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   265
lemma DERIV_cmult_Id [simp]: "DERIV (op * c) x :> c"
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23044
diff changeset
   266
by (cut_tac c = c and x = x in DERIV_ident [THEN DERIV_cmult], simp)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   267
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   268
lemma DERIV_pow: "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23044
diff changeset
   269
apply (cut_tac DERIV_power [OF DERIV_ident])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 34941
diff changeset
   270
apply (simp add: real_of_nat_def)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   271
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   272
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   273
text {* Power of @{text "-1"} *}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   274
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   275
lemma DERIV_inverse:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   276
  fixes x :: "'a::{real_normed_field}"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   277
  shows "x \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   278
by (drule DERIV_inverse' [OF DERIV_ident]) simp
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   279
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   280
text {* Derivative of inverse *}
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   281
lemma DERIV_inverse_fun:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   282
  fixes x :: "'a::{real_normed_field}"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   283
  shows "[| DERIV f x :> d; f(x) \<noteq> 0 |]
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   284
      ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   285
by (drule (1) DERIV_inverse') (simp add: mult_ac nonzero_inverse_mult_distrib)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   286
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   287
text {* Derivative of quotient *}
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   288
lemma DERIV_quotient:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   289
  fixes x :: "'a::{real_normed_field}"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   290
  shows "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> 0 |]
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   291
       ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30242
diff changeset
   292
by (drule (2) DERIV_divide) (simp add: mult_commute)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   293
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   294
text {* @{text "DERIV_intros"} *}
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   295
ML {*
31902
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31899
diff changeset
   296
structure Deriv_Intros = Named_Thms
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   297
(
45294
3c5d3d286055 tuned Named_Thms: proper binding;
wenzelm
parents: 45051
diff changeset
   298
  val name = @{binding DERIV_intros}
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   299
  val description = "DERIV introduction rules"
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   300
)
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   301
*}
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   302
31902
862ae16a799d renamed NamedThmsFun to Named_Thms;
wenzelm
parents: 31899
diff changeset
   303
setup Deriv_Intros.setup
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   304
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   305
lemma DERIV_cong: "\<lbrakk> DERIV f x :> X ; X = Y \<rbrakk> \<Longrightarrow> DERIV f x :> Y"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   306
  by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   307
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   308
declare
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   309
  DERIV_const[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   310
  DERIV_ident[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   311
  DERIV_add[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   312
  DERIV_minus[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   313
  DERIV_mult[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   314
  DERIV_diff[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   315
  DERIV_inverse'[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   316
  DERIV_divide[THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   317
  DERIV_power[where 'a=real, THEN DERIV_cong,
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   318
              unfolded real_of_nat_def[symmetric], DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31404
diff changeset
   319
  DERIV_setsum[THEN DERIV_cong, DERIV_intros]
22984
67d434ad9ef8 section labels
huffman
parents: 22653
diff changeset
   320
31899
1a7ade46061b fixed document (DERIV_intros);
wenzelm
parents: 31880
diff changeset
   321
22984
67d434ad9ef8 section labels
huffman
parents: 22653
diff changeset
   322
subsection {* Differentiability predicate *}
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   323
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   324
definition
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   325
  differentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   326
    (infixl "differentiable" 60) where
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   327
  "f differentiable x = (\<exists>D. DERIV f x :> D)"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   328
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   329
lemma differentiableE [elim?]:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   330
  assumes "f differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   331
  obtains df where "DERIV f x :> df"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   332
  using assms unfolding differentiable_def ..
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   333
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   334
lemma differentiableD: "f differentiable x ==> \<exists>D. DERIV f x :> D"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   335
by (simp add: differentiable_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   336
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   337
lemma differentiableI: "DERIV f x :> D ==> f differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   338
by (force simp add: differentiable_def)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   339
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   340
lemma differentiable_ident [simp]: "(\<lambda>x. x) differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   341
  by (rule DERIV_ident [THEN differentiableI])
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   342
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   343
lemma differentiable_const [simp]: "(\<lambda>z. a) differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   344
  by (rule DERIV_const [THEN differentiableI])
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   345
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   346
lemma differentiable_compose:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   347
  assumes f: "f differentiable (g x)"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   348
  assumes g: "g differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   349
  shows "(\<lambda>x. f (g x)) differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   350
proof -
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   351
  from `f differentiable (g x)` obtain df where "DERIV f (g x) :> df" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   352
  moreover
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   353
  from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   354
  ultimately
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   355
  have "DERIV (\<lambda>x. f (g x)) x :> df * dg" by (rule DERIV_chain2)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   356
  thus ?thesis by (rule differentiableI)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   357
qed
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   358
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   359
lemma differentiable_sum [simp]:
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   360
  assumes "f differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   361
  and "g differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   362
  shows "(\<lambda>x. f x + g x) differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   363
proof -
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   364
  from `f differentiable x` obtain df where "DERIV f x :> df" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   365
  moreover
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   366
  from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   367
  ultimately
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   368
  have "DERIV (\<lambda>x. f x + g x) x :> df + dg" by (rule DERIV_add)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   369
  thus ?thesis by (rule differentiableI)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   370
qed
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   371
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   372
lemma differentiable_minus [simp]:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   373
  assumes "f differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   374
  shows "(\<lambda>x. - f x) differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   375
proof -
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   376
  from `f differentiable x` obtain df where "DERIV f x :> df" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   377
  hence "DERIV (\<lambda>x. - f x) x :> - df" by (rule DERIV_minus)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   378
  thus ?thesis by (rule differentiableI)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   379
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   380
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   381
lemma differentiable_diff [simp]:
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   382
  assumes "f differentiable x"
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   383
  assumes "g differentiable x"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   384
  shows "(\<lambda>x. f x - g x) differentiable x"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   385
  unfolding diff_minus using assms by simp
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   386
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   387
lemma differentiable_mult [simp]:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   388
  assumes "f differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   389
  assumes "g differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   390
  shows "(\<lambda>x. f x * g x) differentiable x"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   391
proof -
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   392
  from `f differentiable x` obtain df where "DERIV f x :> df" ..
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   393
  moreover
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   394
  from `g differentiable x` obtain dg where "DERIV g x :> dg" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   395
  ultimately
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   396
  have "DERIV (\<lambda>x. f x * g x) x :> df * g x + dg * f x" by (rule DERIV_mult)
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   397
  thus ?thesis by (rule differentiableI)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   398
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   399
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   400
lemma differentiable_inverse [simp]:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   401
  assumes "f differentiable x" and "f x \<noteq> 0"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   402
  shows "(\<lambda>x. inverse (f x)) differentiable x"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   403
proof -
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   404
  from `f differentiable x` obtain df where "DERIV f x :> df" ..
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   405
  hence "DERIV (\<lambda>x. inverse (f x)) x :> - (inverse (f x) * df * inverse (f x))"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   406
    using `f x \<noteq> 0` by (rule DERIV_inverse')
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   407
  thus ?thesis by (rule differentiableI)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   408
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   409
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   410
lemma differentiable_divide [simp]:
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   411
  assumes "f differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   412
  assumes "g differentiable x" and "g x \<noteq> 0"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   413
  shows "(\<lambda>x. f x / g x) differentiable x"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   414
  unfolding divide_inverse using assms by simp
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   415
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   416
lemma differentiable_power [simp]:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   417
  fixes f :: "'a::{real_normed_field} \<Rightarrow> 'a"
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   418
  assumes "f differentiable x"
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   419
  shows "(\<lambda>x. f x ^ n) differentiable x"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   420
  apply (induct n)
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   421
  apply simp
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   422
  apply (simp add: assms)
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   423
  done
29169
6a5f1d8d7344 more proofs about differentiable
huffman
parents: 29166
diff changeset
   424
29975
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   425
subsection {* Local extrema *}
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   426
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   427
text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   428
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   429
lemma DERIV_pos_inc_right:
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   430
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   431
  assumes der: "DERIV f x :> l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   432
      and l:   "0 < l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   433
  shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x + h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   434
proof -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   435
  from l der [THEN DERIV_D, THEN LIM_D [where r = "l"]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   436
  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)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   437
    by (simp add: diff_minus)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   438
  then obtain s
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   439
        where s:   "0 < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   440
          and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   441
    by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   442
  thus ?thesis
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   443
  proof (intro exI conjI strip)
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23431
diff changeset
   444
    show "0<s" using s .
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   445
    fix h::real
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   446
    assume "0 < h" "h < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   447
    with all [of h] show "f x < f (x+h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   448
    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   449
    split add: split_if_asm)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   450
      assume "~ (f (x+h) - f x) / h < l" and h: "0 < h"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   451
      with l
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   452
      have "0 < (f (x+h) - f x) / h" by arith
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   453
      thus "f x < f (x+h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   454
  by (simp add: pos_less_divide_eq h)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   455
    qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   456
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   457
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   458
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   459
lemma DERIV_neg_dec_left:
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   460
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   461
  assumes der: "DERIV f x :> l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   462
      and l:   "l < 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   463
  shows "\<exists>d > 0. \<forall>h > 0. h < d --> f(x) < f(x-h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   464
proof -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   465
  from l der [THEN DERIV_D, THEN LIM_D [where r = "-l"]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   466
  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)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   467
    by (simp add: diff_minus)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   468
  then obtain s
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   469
        where s:   "0 < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   470
          and all: "!!z. z \<noteq> 0 \<and> \<bar>z\<bar> < s \<longrightarrow> \<bar>(f(x+z) - f x) / z - l\<bar> < -l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   471
    by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   472
  thus ?thesis
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   473
  proof (intro exI conjI strip)
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23431
diff changeset
   474
    show "0<s" using s .
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   475
    fix h::real
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   476
    assume "0 < h" "h < s"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   477
    with all [of "-h"] show "f x < f (x-h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   478
    proof (simp add: abs_if pos_less_divide_eq diff_minus [symmetric]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   479
    split add: split_if_asm)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   480
      assume " - ((f (x-h) - f x) / h) < l" and h: "0 < h"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   481
      with l
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   482
      have "0 < (f (x-h) - f x) / h" by arith
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   483
      thus "f x < f (x-h)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   484
  by (simp add: pos_less_divide_eq h)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   485
    qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   486
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   487
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   488
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   489
lemma DERIV_pos_inc_left:
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   490
  fixes f :: "real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   491
  shows "DERIV f x :> l \<Longrightarrow> 0 < l \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x - h) < f(x)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   492
  apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified])
41368
8afa26855137 use DERIV_intros
hoelzl
parents: 37891
diff changeset
   493
  apply (auto simp add: DERIV_minus)
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   494
  done
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   495
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   496
lemma DERIV_neg_dec_right:
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   497
  fixes f :: "real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   498
  shows "DERIV f x :> l \<Longrightarrow> l < 0 \<Longrightarrow> \<exists>d > 0. \<forall>h > 0. h < d --> f(x) > f(x + h)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   499
  apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified])
41368
8afa26855137 use DERIV_intros
hoelzl
parents: 37891
diff changeset
   500
  apply (auto simp add: DERIV_minus)
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   501
  done
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   502
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   503
lemma DERIV_local_max:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   504
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   505
  assumes der: "DERIV f x :> l"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   506
      and d:   "0 < d"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   507
      and le:  "\<forall>y. \<bar>x-y\<bar> < d --> f(y) \<le> f(x)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   508
  shows "l = 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   509
proof (cases rule: linorder_cases [of l 0])
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23431
diff changeset
   510
  case equal thus ?thesis .
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   511
next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   512
  case less
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   513
  from DERIV_neg_dec_left [OF der less]
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   514
  obtain d' where d': "0 < d'"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   515
             and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x-h)" by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   516
  from real_lbound_gt_zero [OF d d']
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   517
  obtain e where "0 < e \<and> e < d \<and> e < d'" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   518
  with lt le [THEN spec [where x="x-e"]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   519
  show ?thesis by (auto simp add: abs_if)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   520
next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   521
  case greater
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   522
  from DERIV_pos_inc_right [OF der greater]
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   523
  obtain d' where d': "0 < d'"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   524
             and lt: "\<forall>h > 0. h < d' \<longrightarrow> f x < f (x + h)" by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   525
  from real_lbound_gt_zero [OF d d']
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   526
  obtain e where "0 < e \<and> e < d \<and> e < d'" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   527
  with lt le [THEN spec [where x="x+e"]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   528
  show ?thesis by (auto simp add: abs_if)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   529
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   530
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   531
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   532
text{*Similar theorem for a local minimum*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   533
lemma DERIV_local_min:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   534
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   535
  shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) \<le> f(y) |] ==> l = 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   536
by (drule DERIV_minus [THEN DERIV_local_max], auto)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   537
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   538
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   539
text{*In particular, if a function is locally flat*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   540
lemma DERIV_local_const:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   541
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   542
  shows "[| DERIV f x :> l; 0 < d; \<forall>y. \<bar>x-y\<bar> < d --> f(x) = f(y) |] ==> l = 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   543
by (auto dest!: DERIV_local_max)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   544
29975
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   545
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   546
subsection {* Rolle's Theorem *}
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   547
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   548
text{*Lemma about introducing open ball in open interval*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   549
lemma lemma_interval_lt:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   550
     "[| a < x;  x < b |]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   551
      ==> \<exists>d::real. 0 < d & (\<forall>y. \<bar>x-y\<bar> < d --> a < y & y < b)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 26120
diff changeset
   552
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22984
diff changeset
   553
apply (simp add: abs_less_iff)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   554
apply (insert linorder_linear [of "x-a" "b-x"], safe)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   555
apply (rule_tac x = "x-a" in exI)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   556
apply (rule_tac [2] x = "b-x" in exI, auto)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   557
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   558
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   559
lemma lemma_interval: "[| a < x;  x < b |] ==>
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   560
        \<exists>d::real. 0 < d &  (\<forall>y. \<bar>x-y\<bar> < d --> a \<le> y & y \<le> b)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   561
apply (drule lemma_interval_lt, auto)
44921
58eef4843641 tuned proofs
huffman
parents: 44890
diff changeset
   562
apply force
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   563
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   564
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   565
text{*Rolle's Theorem.
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   566
   If @{term f} is defined and continuous on the closed interval
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   567
   @{text "[a,b]"} and differentiable on the open interval @{text "(a,b)"},
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   568
   and @{term "f(a) = f(b)"},
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   569
   then there exists @{text "x0 \<in> (a,b)"} such that @{term "f'(x0) = 0"}*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   570
theorem Rolle:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   571
  assumes lt: "a < b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   572
      and eq: "f(a) = f(b)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   573
      and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   574
      and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   575
  shows "\<exists>z::real. a < z & z < b & DERIV f z :> 0"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   576
proof -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   577
  have le: "a \<le> b" using lt by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   578
  from isCont_eq_Ub [OF le con]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   579
  obtain x where x_max: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f z \<le> f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   580
             and alex: "a \<le> x" and xleb: "x \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   581
    by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   582
  from isCont_eq_Lb [OF le con]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   583
  obtain x' where x'_min: "\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> f x' \<le> f z"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   584
              and alex': "a \<le> x'" and x'leb: "x' \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   585
    by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   586
  show ?thesis
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   587
  proof cases
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   588
    assume axb: "a < x & x < b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   589
        --{*@{term f} attains its maximum within the interval*}
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 26120
diff changeset
   590
    hence ax: "a<x" and xb: "x<b" by arith + 
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   591
    from lemma_interval [OF ax xb]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   592
    obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   593
      by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   594
    hence bound': "\<forall>y. \<bar>x-y\<bar> < d \<longrightarrow> f y \<le> f x" using x_max
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   595
      by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   596
    from differentiableD [OF dif [OF axb]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   597
    obtain l where der: "DERIV f x :> l" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   598
    have "l=0" by (rule DERIV_local_max [OF der d bound'])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   599
        --{*the derivative at a local maximum is zero*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   600
    thus ?thesis using ax xb der by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   601
  next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   602
    assume notaxb: "~ (a < x & x < b)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   603
    hence xeqab: "x=a | x=b" using alex xleb by arith
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   604
    hence fb_eq_fx: "f b = f x" by (auto simp add: eq)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   605
    show ?thesis
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   606
    proof cases
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   607
      assume ax'b: "a < x' & x' < b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   608
        --{*@{term f} attains its minimum within the interval*}
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 26120
diff changeset
   609
      hence ax': "a<x'" and x'b: "x'<b" by arith+ 
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   610
      from lemma_interval [OF ax' x'b]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   611
      obtain d where d: "0<d" and bound: "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   612
  by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   613
      hence bound': "\<forall>y. \<bar>x'-y\<bar> < d \<longrightarrow> f x' \<le> f y" using x'_min
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   614
  by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   615
      from differentiableD [OF dif [OF ax'b]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   616
      obtain l where der: "DERIV f x' :> l" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   617
      have "l=0" by (rule DERIV_local_min [OF der d bound'])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   618
        --{*the derivative at a local minimum is zero*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   619
      thus ?thesis using ax' x'b der by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   620
    next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   621
      assume notax'b: "~ (a < x' & x' < b)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   622
        --{*@{term f} is constant througout the interval*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   623
      hence x'eqab: "x'=a | x'=b" using alex' x'leb by arith
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   624
      hence fb_eq_fx': "f b = f x'" by (auto simp add: eq)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   625
      from dense [OF lt]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   626
      obtain r where ar: "a < r" and rb: "r < b" by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   627
      from lemma_interval [OF ar rb]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   628
      obtain d where d: "0<d" and bound: "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> a \<le> y \<and> y \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   629
  by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   630
      have eq_fb: "\<forall>z. a \<le> z --> z \<le> b --> f z = f b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   631
      proof (clarify)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   632
        fix z::real
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   633
        assume az: "a \<le> z" and zb: "z \<le> b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   634
        show "f z = f b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   635
        proof (rule order_antisym)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   636
          show "f z \<le> f b" by (simp add: fb_eq_fx x_max az zb)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   637
          show "f b \<le> f z" by (simp add: fb_eq_fx' x'_min az zb)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   638
        qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   639
      qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   640
      have bound': "\<forall>y. \<bar>r-y\<bar> < d \<longrightarrow> f r = f y"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   641
      proof (intro strip)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   642
        fix y::real
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   643
        assume lt: "\<bar>r-y\<bar> < d"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   644
        hence "f y = f b" by (simp add: eq_fb bound)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   645
        thus "f r = f y" by (simp add: eq_fb ar rb order_less_imp_le)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   646
      qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   647
      from differentiableD [OF dif [OF conjI [OF ar rb]]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   648
      obtain l where der: "DERIV f r :> l" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   649
      have "l=0" by (rule DERIV_local_const [OF der d bound'])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   650
        --{*the derivative of a constant function is zero*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   651
      thus ?thesis using ar rb der by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   652
    qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   653
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   654
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   655
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   656
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   657
subsection{*Mean Value Theorem*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   658
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   659
lemma lemma_MVT:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   660
     "f a - (f b - f a)/(b-a) * a = f b - (f b - f a)/(b-a) * (b::real)"
51481
ef949192e5d6 move continuous_on_inv to HOL image (simplifies isCont_inverse_function)
hoelzl
parents: 51480
diff changeset
   661
  by (cases "a = b") (simp_all add: field_simps)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   662
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   663
theorem MVT:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   664
  assumes lt:  "a < b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   665
      and con: "\<forall>x. a \<le> x & x \<le> b --> isCont f x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   666
      and dif [rule_format]: "\<forall>x. a < x & x < b --> f differentiable x"
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   667
  shows "\<exists>l z::real. a < z & z < b & DERIV f z :> l &
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   668
                   (f(b) - f(a) = (b-a) * l)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   669
proof -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   670
  let ?F = "%x. f x - ((f b - f a) / (b-a)) * x"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44209
diff changeset
   671
  have contF: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?F x"
aa74ce315bae add simp rules for isCont
huffman
parents: 44209
diff changeset
   672
    using con by (fast intro: isCont_intros)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   673
  have difF: "\<forall>x. a < x \<and> x < b \<longrightarrow> ?F differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   674
  proof (clarify)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   675
    fix x::real
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   676
    assume ax: "a < x" and xb: "x < b"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   677
    from differentiableD [OF dif [OF conjI [OF ax xb]]]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   678
    obtain l where der: "DERIV f x :> l" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   679
    show "?F differentiable x"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   680
      by (rule differentiableI [where D = "l - (f b - f a)/(b-a)"],
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   681
          blast intro: DERIV_diff DERIV_cmult_Id der)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   682
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   683
  from Rolle [where f = ?F, OF lt lemma_MVT contF difF]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   684
  obtain z where az: "a < z" and zb: "z < b" and der: "DERIV ?F z :> 0"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   685
    by blast
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   686
  have "DERIV (%x. ((f b - f a)/(b-a)) * x) z :> (f b - f a)/(b-a)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   687
    by (rule DERIV_cmult_Id)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   688
  hence derF: "DERIV (\<lambda>x. ?F x + (f b - f a) / (b - a) * x) z
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   689
                   :> 0 + (f b - f a) / (b - a)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   690
    by (rule DERIV_add [OF der])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   691
  show ?thesis
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   692
  proof (intro exI conjI)
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23431
diff changeset
   693
    show "a < z" using az .
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23431
diff changeset
   694
    show "z < b" using zb .
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   695
    show "f b - f a = (b - a) * ((f b - f a)/(b-a))" by (simp)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   696
    show "DERIV f z :> ((f b - f a)/(b-a))"  using derF by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   697
  qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   698
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   699
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   700
lemma MVT2:
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   701
     "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   702
      ==> \<exists>z::real. a < z & z < b & (f b - f a = (b - a) * f'(z))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   703
apply (drule MVT)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   704
apply (blast intro: DERIV_isCont)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   705
apply (force dest: order_less_imp_le simp add: differentiable_def)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   706
apply (blast dest: DERIV_unique order_less_imp_le)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   707
done
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   708
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   709
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   710
text{*A function is constant if its derivative is 0 over an interval.*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   711
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   712
lemma DERIV_isconst_end:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   713
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   714
  shows "[| a < b;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   715
         \<forall>x. a \<le> x & x \<le> b --> isCont f x;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   716
         \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   717
        ==> f b = f a"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   718
apply (drule MVT, assumption)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   719
apply (blast intro: differentiableI)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   720
apply (auto dest!: DERIV_unique simp add: diff_eq_eq)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   721
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   722
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   723
lemma DERIV_isconst1:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   724
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   725
  shows "[| a < b;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   726
         \<forall>x. a \<le> x & x \<le> b --> isCont f x;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   727
         \<forall>x. a < x & x < b --> DERIV f x :> 0 |]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   728
        ==> \<forall>x. a \<le> x & x \<le> b --> f x = f a"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   729
apply safe
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   730
apply (drule_tac x = a in order_le_imp_less_or_eq, safe)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   731
apply (drule_tac b = x in DERIV_isconst_end, auto)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   732
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   733
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   734
lemma DERIV_isconst2:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   735
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   736
  shows "[| a < b;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   737
         \<forall>x. a \<le> x & x \<le> b --> isCont f x;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   738
         \<forall>x. a < x & x < b --> DERIV f x :> 0;
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   739
         a \<le> x; x \<le> b |]
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   740
        ==> f x = f a"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   741
apply (blast dest: DERIV_isconst1)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   742
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   743
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   744
lemma DERIV_isconst3: fixes a b x y :: real
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   745
  assumes "a < b" and "x \<in> {a <..< b}" and "y \<in> {a <..< b}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   746
  assumes derivable: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> DERIV f x :> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   747
  shows "f x = f y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   748
proof (cases "x = y")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   749
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   750
  let ?a = "min x y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   751
  let ?b = "max x y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   752
  
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   753
  have "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> DERIV f z :> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   754
  proof (rule allI, rule impI)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   755
    fix z :: real assume "?a \<le> z \<and> z \<le> ?b"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   756
    hence "a < z" and "z < b" using `x \<in> {a <..< b}` and `y \<in> {a <..< b}` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   757
    hence "z \<in> {a<..<b}" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   758
    thus "DERIV f z :> 0" by (rule derivable)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   759
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   760
  hence isCont: "\<forall>z. ?a \<le> z \<and> z \<le> ?b \<longrightarrow> isCont f z"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   761
    and DERIV: "\<forall>z. ?a < z \<and> z < ?b \<longrightarrow> DERIV f z :> 0" using DERIV_isCont by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   762
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   763
  have "?a < ?b" using `x \<noteq> y` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   764
  from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   765
  show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   766
qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29667
diff changeset
   767
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   768
lemma DERIV_isconst_all:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   769
  fixes f :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   770
  shows "\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   771
apply (rule linorder_cases [of x y])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   772
apply (blast intro: sym DERIV_isCont DERIV_isconst_end)+
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   773
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   774
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   775
lemma DERIV_const_ratio_const:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   776
  fixes f :: "real => real"
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   777
  shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b-a) * k"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   778
apply (rule linorder_cases [of a b], auto)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   779
apply (drule_tac [!] f = f in MVT)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   780
apply (auto dest: DERIV_isCont DERIV_unique simp add: differentiable_def)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23441
diff changeset
   781
apply (auto dest: DERIV_unique simp add: ring_distribs diff_minus)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   782
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   783
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   784
lemma DERIV_const_ratio_const2:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   785
  fixes f :: "real => real"
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   786
  shows "[|a \<noteq> b; \<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b-a) = k"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   787
apply (rule_tac c1 = "b-a" in real_mult_right_cancel [THEN iffD1])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   788
apply (auto dest!: DERIV_const_ratio_const simp add: mult_assoc)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   789
done
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   790
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   791
lemma real_average_minus_first [simp]: "((a + b) /2 - a) = (b-a)/(2::real)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   792
by (simp)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   793
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   794
lemma real_average_minus_second [simp]: "((b + a)/2 - a) = (b-a)/(2::real)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   795
by (simp)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   796
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   797
text{*Gallileo's "trick": average velocity = av. of end velocities*}
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   798
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   799
lemma DERIV_const_average:
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   800
  fixes v :: "real => real"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   801
  assumes neq: "a \<noteq> (b::real)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   802
      and der: "\<forall>x. DERIV v x :> k"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   803
  shows "v ((a + b)/2) = (v a + v b)/2"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   804
proof (cases rule: linorder_cases [of a b])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   805
  case equal with neq show ?thesis by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   806
next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   807
  case less
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   808
  have "(v b - v a) / (b - a) = k"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   809
    by (rule DERIV_const_ratio_const2 [OF neq der])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   810
  hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   811
  moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   812
    by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   813
  ultimately show ?thesis using neq by force
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   814
next
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   815
  case greater
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   816
  have "(v b - v a) / (b - a) = k"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   817
    by (rule DERIV_const_ratio_const2 [OF neq der])
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   818
  hence "(b-a) * ((v b - v a) / (b-a)) = (b-a) * k" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   819
  moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   820
    by (rule DERIV_const_ratio_const2 [OF _ der], simp add: neq)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   821
  ultimately show ?thesis using neq by (force simp add: add_commute)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   822
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   823
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   824
(* A function with positive derivative is increasing. 
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   825
   A simple proof using the MVT, by Jeremy Avigad. And variants.
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   826
*)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   827
lemma DERIV_pos_imp_increasing:
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   828
  fixes a::real and b::real and f::"real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   829
  assumes "a < b" and "\<forall>x. a \<le> x & x \<le> b --> (EX y. DERIV f x :> y & y > 0)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   830
  shows "f a < f b"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   831
proof (rule ccontr)
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   832
  assume f: "~ f a < f b"
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   833
  have "EX l z. a < z & z < b & DERIV f z :> l
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   834
      & f b - f a = (b - a) * l"
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   835
    apply (rule MVT)
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   836
      using assms
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   837
      apply auto
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   838
      apply (metis DERIV_isCont)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35216
diff changeset
   839
     apply (metis differentiableI less_le)
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   840
    done
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   841
  then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   842
      and "f b - f a = (b - a) * l"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   843
    by auto
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   844
  with assms f have "~(l > 0)"
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35216
diff changeset
   845
    by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le)
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   846
  with assms z show False
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35216
diff changeset
   847
    by (metis DERIV_unique less_le)
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   848
qed
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   849
45791
d985ec974815 more systematic lemma name
noschinl
parents: 45600
diff changeset
   850
lemma DERIV_nonneg_imp_nondecreasing:
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   851
  fixes a::real and b::real and f::"real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   852
  assumes "a \<le> b" and
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   853
    "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<ge> 0)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   854
  shows "f a \<le> f b"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   855
proof (rule ccontr, cases "a = b")
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   856
  assume "~ f a \<le> f b" and "a = b"
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   857
  then show False by auto
37891
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   858
next
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   859
  assume A: "~ f a \<le> f b"
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   860
  assume B: "a ~= b"
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   861
  with assms have "EX l z. a < z & z < b & DERIV f z :> l
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   862
      & f b - f a = (b - a) * l"
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   863
    apply -
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   864
    apply (rule MVT)
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   865
      apply auto
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   866
      apply (metis DERIV_isCont)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 35216
diff changeset
   867
     apply (metis differentiableI less_le)
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   868
    done
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   869
  then obtain l z where z: "a < z" "z < b" "DERIV f z :> l"
37891
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   870
      and C: "f b - f a = (b - a) * l"
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   871
    by auto
37891
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   872
  with A have "a < b" "f b < f a" by auto
c26f9d06e82c robustified metis proof
haftmann
parents: 37888
diff changeset
   873
  with C have "\<not> l \<ge> 0" by (auto simp add: not_le algebra_simps)
45051
c478d1876371 discontinued legacy theorem names from RealDef.thy
huffman
parents: 44921
diff changeset
   874
    (metis A add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl)
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   875
  with assms z show False
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   876
    by (metis DERIV_unique order_less_imp_le)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   877
qed
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   878
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   879
lemma DERIV_neg_imp_decreasing:
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   880
  fixes a::real and b::real and f::"real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   881
  assumes "a < b" and
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   882
    "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y < 0)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   883
  shows "f a > f b"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   884
proof -
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   885
  have "(%x. -f x) a < (%x. -f x) b"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   886
    apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"])
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   887
    using assms
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   888
    apply auto
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   889
    apply (metis DERIV_minus neg_0_less_iff_less)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   890
    done
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   891
  thus ?thesis
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   892
    by simp
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   893
qed
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   894
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   895
lemma DERIV_nonpos_imp_nonincreasing:
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   896
  fixes a::real and b::real and f::"real => real"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   897
  assumes "a \<le> b" and
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   898
    "\<forall>x. a \<le> x & x \<le> b --> (\<exists>y. DERIV f x :> y & y \<le> 0)"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   899
  shows "f a \<ge> f b"
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   900
proof -
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   901
  have "(%x. -f x) a \<le> (%x. -f x) b"
45791
d985ec974815 more systematic lemma name
noschinl
parents: 45600
diff changeset
   902
    apply (rule DERIV_nonneg_imp_nondecreasing [of a b "%x. -f x"])
33690
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   903
    using assms
889d06128608 simplified bulky metis proofs;
wenzelm
parents: 33659
diff changeset
   904
    apply auto
33654
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   905
    apply (metis DERIV_minus neg_0_le_iff_le)
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   906
    done
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   907
  thus ?thesis
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   908
    by simp
abf780db30ea A number of theorems contributed by Jeremy Avigad
paulson
parents: 31902
diff changeset
   909
qed
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   910
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   911
text {* Derivative of inverse function *}
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   912
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   913
lemma DERIV_inverse_function:
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   914
  fixes f g :: "real \<Rightarrow> real"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   915
  assumes der: "DERIV f (g x) :> D"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   916
  assumes neq: "D \<noteq> 0"
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   917
  assumes a: "a < x" and b: "x < b"
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   918
  assumes inj: "\<forall>y. a < y \<and> y < b \<longrightarrow> f (g y) = y"
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   919
  assumes cont: "isCont g x"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   920
  shows "DERIV g x :> inverse D"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   921
unfolding DERIV_iff2
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   922
proof (rule LIM_equal2)
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   923
  show "0 < min (x - a) (b - x)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 26120
diff changeset
   924
    using a b by arith 
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   925
next
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   926
  fix y
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   927
  assume "norm (y - x) < min (x - a) (b - x)"
27668
6eb20b2cecf8 Tuned and simplified proofs
chaieb
parents: 26120
diff changeset
   928
  hence "a < y" and "y < b" 
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   929
    by (simp_all add: abs_less_iff)
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   930
  thus "(g y - g x) / (y - x) =
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   931
        inverse ((f (g y) - x) / (g y - g x))"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   932
    by (simp add: inj)
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   933
next
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   934
  have "(\<lambda>z. (f z - f (g x)) / (z - g x)) -- g x --> D"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   935
    by (rule der [unfolded DERIV_iff2])
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   936
  hence 1: "(\<lambda>z. (f z - x) / (z - g x)) -- g x --> D"
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   937
    using inj a b by simp
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   938
  have 2: "\<exists>d>0. \<forall>y. y \<noteq> x \<and> norm (y - x) < d \<longrightarrow> g y \<noteq> g x"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   939
  proof (safe intro!: exI)
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   940
    show "0 < min (x - a) (b - x)"
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   941
      using a b by simp
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   942
  next
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   943
    fix y
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   944
    assume "norm (y - x) < min (x - a) (b - x)"
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   945
    hence y: "a < y" "y < b"
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   946
      by (simp_all add: abs_less_iff)
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   947
    assume "g y = g x"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   948
    hence "f (g y) = f (g x)" by simp
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23041
diff changeset
   949
    hence "y = x" using inj y a b by simp
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   950
    also assume "y \<noteq> x"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   951
    finally show False by simp
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   952
  qed
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   953
  have "(\<lambda>y. (f (g y) - x) / (g y - g x)) -- x --> D"
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   954
    using cont 1 2 by (rule isCont_LIM_compose2)
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   955
  thus "(\<lambda>y. inverse ((f (g y) - x) / (g y - g x)))
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   956
        -- x --> inverse D"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44317
diff changeset
   957
    using neq by (rule tendsto_inverse)
23041
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   958
qed
a0f26d47369b add lemma DERIV_inverse_function
huffman
parents: 22998
diff changeset
   959
29975
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   960
subsection {* Generalized Mean Value Theorem *}
28c5322f0df3 more subsection headings
huffman
parents: 29803
diff changeset
   961
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   962
theorem GMVT:
21784
e76faa6e65fd changed (ns)deriv to take functions of type 'a::real_normed_field => 'a
huffman
parents: 21404
diff changeset
   963
  fixes a b :: real
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   964
  assumes alb: "a < b"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   965
    and fc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont f x"
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   966
    and fd: "\<forall>x. a < x \<and> x < b \<longrightarrow> f differentiable x"
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   967
    and gc: "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont g x"
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   968
    and gd: "\<forall>x. a < x \<and> x < b \<longrightarrow> g differentiable x"
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   969
  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)"
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   970
proof -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   971
  let ?h = "\<lambda>x. (f b - f a)*(g x) - (g b - g a)*(f x)"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 41368
diff changeset
   972
  from assms have "a < b" by simp
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   973
  moreover have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> isCont ?h x"
44233
aa74ce315bae add simp rules for isCont
huffman
parents: 44209
diff changeset
   974
    using fc gc by simp
aa74ce315bae add simp rules for isCont
huffman
parents: 44209
diff changeset
   975
  moreover have "\<forall>x. a < x \<and> x < b \<longrightarrow> ?h differentiable x"
aa74ce315bae add simp rules for isCont
huffman
parents: 44209
diff changeset
   976
    using fd gd by simp
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   977
  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)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   978
  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" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   979
  then obtain c where cdef: "a < c \<and> c < b \<and> DERIV ?h c :> l \<and> ?h b - ?h a = (b - a) * l" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   980
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   981
  from cdef have cint: "a < c \<and> c < b" by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   982
  with gd have "g differentiable c" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   983
  hence "\<exists>D. DERIV g c :> D" by (rule differentiableD)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   984
  then obtain g'c where g'cdef: "DERIV g c :> g'c" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   985
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   986
  from cdef have "a < c \<and> c < b" by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   987
  with fd have "f differentiable c" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   988
  hence "\<exists>D. DERIV f c :> D" by (rule differentiableD)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   989
  then obtain f'c where f'cdef: "DERIV f c :> f'c" ..
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   990
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   991
  from cdef have "DERIV ?h c :> l" by auto
41368
8afa26855137 use DERIV_intros
hoelzl
parents: 37891
diff changeset
   992
  moreover have "DERIV ?h c :>  g'c * (f b - f a) - f'c * (g b - g a)"
8afa26855137 use DERIV_intros
hoelzl
parents: 37891
diff changeset
   993
    using g'cdef f'cdef by (auto intro!: DERIV_intros)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   994
  ultimately have leq: "l =  g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   995
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   996
  {
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   997
    from cdef have "?h b - ?h a = (b - a) * l" by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   998
    also with leq have "\<dots> = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
   999
    finally have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1000
  }
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1001
  moreover
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1002
  {
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1003
    have "?h b - ?h a =
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1004
         ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) -
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1005
          ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29472
diff changeset
  1006
      by (simp add: algebra_simps)
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1007
    hence "?h b - ?h a = 0" by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1008
  }
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1009
  ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1010
  with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1011
  hence "g'c * (f b - f a) = f'c * (g b - g a)" by simp
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1012
  hence "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: mult_ac)
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1013
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1014
  with g'cdef f'cdef cint show ?thesis by auto
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1015
qed
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1016
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1017
lemma GMVT':
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1018
  fixes f g :: "real \<Rightarrow> real"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1019
  assumes "a < b"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1020
  assumes isCont_f: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont f z"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1021
  assumes isCont_g: "\<And>z. a \<le> z \<Longrightarrow> z \<le> b \<Longrightarrow> isCont g z"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1022
  assumes DERIV_g: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV g z :> (g' z)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1023
  assumes DERIV_f: "\<And>z. a < z \<Longrightarrow> z < b \<Longrightarrow> DERIV f z :> (f' z)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1024
  shows "\<exists>c. a < c \<and> c < b \<and> (f b - f a) * g' c = (g b - g a) * f' c"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1025
proof -
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1026
  have "\<exists>g'c f'c c. DERIV g c :> g'c \<and> DERIV f c :> f'c \<and>
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1027
    a < c \<and> c < b \<and> (f b - f a) * g'c = (g b - g a) * f'c"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1028
    using assms by (intro GMVT) (force simp: differentiable_def)+
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1029
  then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1030
    using DERIV_f DERIV_g by (force dest: DERIV_unique)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1031
  then show ?thesis
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1032
    by auto
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1033
qed
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1034
51529
2d2f59e6055a move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents: 51526
diff changeset
  1035
2d2f59e6055a move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents: 51526
diff changeset
  1036
subsection {* L'Hopitals rule *}
2d2f59e6055a move theorems about compactness of real closed intervals, the intermediate value theorem, and lemmas about continuity of bijective functions from Deriv.thy to Limits.thy
hoelzl
parents: 51526
diff changeset
  1037
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1038
lemma DERIV_cong_ev: "x = y \<Longrightarrow> eventually (\<lambda>x. f x = g x) (nhds x) \<Longrightarrow> u = v \<Longrightarrow>
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1039
    DERIV f x :> u \<longleftrightarrow> DERIV g y :> v"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1040
  unfolding DERIV_iff2
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1041
proof (rule filterlim_cong)
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1042
  assume "eventually (\<lambda>x. f x = g x) (nhds x)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1043
  moreover then have "f x = g x" by (auto simp: eventually_nhds)
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1044
  moreover assume "x = y" "u = v"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1045
  ultimately show "eventually (\<lambda>xa. (f xa - f x) / (xa - x) = (g xa - g y) / (xa - y)) (at x)"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1046
    by (auto simp: eventually_at_filter elim: eventually_elim1)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1047
qed simp_all
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1048
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1049
lemma DERIV_shift:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1050
  "(DERIV f (x + z) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (x + z)) x :> y)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1051
  by (simp add: DERIV_iff field_simps)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1052
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1053
lemma DERIV_mirror:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1054
  "(DERIV f (- x) :> y) \<longleftrightarrow> (DERIV (\<lambda>x. f (- x::real) :: real) x :> - y)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1055
  by (simp add: deriv_def filterlim_at_split filterlim_at_left_to_right
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1056
                tendsto_minus_cancel_left field_simps conj_commute)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1057
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1058
lemma isCont_If_ge:
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1059
  fixes a :: "'a :: linorder_topology"
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1060
  shows "continuous (at_left a) g \<Longrightarrow> (f ---> g a) (at_right a) \<Longrightarrow> isCont (\<lambda>x. if x \<le> a then g x else f x) a"
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1061
  unfolding isCont_def continuous_within
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1062
  apply (intro filterlim_split_at)
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1063
  apply (subst filterlim_cong[OF refl refl, where g=g])
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1064
  apply (simp_all add: eventually_at_filter less_le)
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1065
  apply (subst filterlim_cong[OF refl refl, where g=f])
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1066
  apply (simp_all add: eventually_at_filter less_le)
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1067
  done
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1068
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1069
lemma lhopital_right_0:
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1070
  fixes f0 g0 :: "real \<Rightarrow> real"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1071
  assumes f_0: "(f0 ---> 0) (at_right 0)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1072
  assumes g_0: "(g0 ---> 0) (at_right 0)"
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1073
  assumes ev:
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1074
    "eventually (\<lambda>x. g0 x \<noteq> 0) (at_right 0)"
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1075
    "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)"
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1076
    "eventually (\<lambda>x. DERIV f0 x :> f' x) (at_right 0)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1077
    "eventually (\<lambda>x. DERIV g0 x :> g' x) (at_right 0)"
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1078
  assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)"
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1079
  shows "((\<lambda> x. f0 x / g0 x) ---> x) (at_right 0)"
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1080
proof -
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1081
  def f \<equiv> "\<lambda>x. if x \<le> 0 then 0 else f0 x"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1082
  then have "f 0 = 0" by simp
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1083
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1084
  def g \<equiv> "\<lambda>x. if x \<le> 0 then 0 else g0 x"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1085
  then have "g 0 = 0" by simp
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1086
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1087
  have "eventually (\<lambda>x. g0 x \<noteq> 0 \<and> g' x \<noteq> 0 \<and>
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1088
      DERIV f0 x :> (f' x) \<and> DERIV g0 x :> (g' x)) (at_right 0)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1089
    using ev by eventually_elim auto
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1090
  then obtain a where [arith]: "0 < a"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1091
    and g0_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g0 x \<noteq> 0"
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1092
    and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0"
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1093
    and f0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV f0 x :> (f' x)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1094
    and g0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> DERIV g0 x :> (g' x)"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1095
    unfolding eventually_at eventually_at by (auto simp: dist_real_def)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1096
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1097
  have g_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g x \<noteq> 0"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1098
    using g0_neq_0 by (simp add: g_def)
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1099
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1100
  { fix x assume x: "0 < x" "x < a" then have "DERIV f x :> (f' x)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1101
      by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]])
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1102
         (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) }
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1103
  note f = this
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1104
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1105
  { fix x assume x: "0 < x" "x < a" then have "DERIV g x :> (g' x)"
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1106
      by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]])
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1107
         (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) }
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1108
  note g = this
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1109
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1110
  have "isCont f 0"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1111
    unfolding f_def by (intro isCont_If_ge f_0 continuous_const)
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1112
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1113
  have "isCont g 0"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1114
    unfolding g_def by (intro isCont_If_ge g_0 continuous_const)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1115
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1116
  have "\<exists>\<zeta>. \<forall>x\<in>{0 <..< a}. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1117
  proof (rule bchoice, rule)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1118
    fix x assume "x \<in> {0 <..< a}"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1119
    then have x[arith]: "0 < x" "x < a" by auto
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1120
    with g'_neq_0 g_neq_0 `g 0 = 0` have g': "\<And>x. 0 < x \<Longrightarrow> x < a  \<Longrightarrow> 0 \<noteq> g' x" "g 0 \<noteq> g x"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1121
      by auto
50328
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1122
    have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont f x"
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1123
      using `isCont f 0` f by (auto intro: DERIV_isCont simp: le_less)
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1124
    moreover have "\<And>x. 0 \<le> x \<Longrightarrow> x < a \<Longrightarrow> isCont g x"
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1125
      using `isCont g 0` g by (auto intro: DERIV_isCont simp: le_less)
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1126
    ultimately have "\<exists>c. 0 < c \<and> c < x \<and> (f x - f 0) * g' c = (g x - g 0) * f' c"
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1127
      using f g `x < a` by (intro GMVT') auto
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1128
    then guess c ..
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1129
    moreover
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1130
    with g'(1)[of c] g'(2) have "(f x - f 0)  / (g x - g 0) = f' c / g' c"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1131
      by (simp add: field_simps)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1132
    ultimately show "\<exists>y. 0 < y \<and> y < x \<and> f x / g x = f' y / g' y"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1133
      using `f 0 = 0` `g 0 = 0` by (auto intro!: exI[of _ c])
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1134
  qed
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1135
  then guess \<zeta> ..
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1136
  then have \<zeta>: "eventually (\<lambda>x. 0 < \<zeta> x \<and> \<zeta> x < x \<and> f x / g x = f' (\<zeta> x) / g' (\<zeta> x)) (at_right 0)"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1137
    unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1138
  moreover
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1139
  from \<zeta> have "eventually (\<lambda>x. norm (\<zeta> x) \<le> x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1140
    by eventually_elim auto
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1141
  then have "((\<lambda>x. norm (\<zeta> x)) ---> 0) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1142
    by (rule_tac real_tendsto_sandwich[where f="\<lambda>x. 0" and h="\<lambda>x. x"])
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1143
       (auto intro: tendsto_const tendsto_ident_at)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1144
  then have "(\<zeta> ---> 0) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1145
    by (rule tendsto_norm_zero_cancel)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1146
  with \<zeta> have "filterlim \<zeta> (at_right 0) (at_right 0)"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1147
    by (auto elim!: eventually_elim1 simp: filterlim_at)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1148
  from this lim have "((\<lambda>t. f' (\<zeta> t) / g' (\<zeta> t)) ---> x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1149
    by (rule_tac filterlim_compose[of _ _ _ \<zeta>])
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1150
  ultimately have "((\<lambda>t. f t / g t) ---> x) (at_right 0)" (is ?P)
50328
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1151
    by (rule_tac filterlim_cong[THEN iffD1, OF refl refl])
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1152
       (auto elim: eventually_elim1)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1153
  also have "?P \<longleftrightarrow> ?thesis"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1154
    by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter)
50329
9bd6b6b8a554 weakened assumptions for lhopital_right_0
hoelzl
parents: 50328
diff changeset
  1155
  finally show ?thesis .
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1156
qed
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1157
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1158
lemma lhopital_right:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1159
  "((f::real \<Rightarrow> real) ---> 0) (at_right x) \<Longrightarrow> (g ---> 0) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1160
    eventually (\<lambda>x. g x \<noteq> 0) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1161
    eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1162
    eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1163
    eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1164
    ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1165
  ((\<lambda> x. f x / g x) ---> y) (at_right x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1166
  unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1167
  by (rule lhopital_right_0)
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1168
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1169
lemma lhopital_left:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1170
  "((f::real \<Rightarrow> real) ---> 0) (at_left x) \<Longrightarrow> (g ---> 0) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1171
    eventually (\<lambda>x. g x \<noteq> 0) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1172
    eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1173
    eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1174
    eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1175
    ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1176
  ((\<lambda> x. f x / g x) ---> y) (at_left x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1177
  unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1178
  by (rule lhopital_right[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror)
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1179
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1180
lemma lhopital:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1181
  "((f::real \<Rightarrow> real) ---> 0) (at x) \<Longrightarrow> (g ---> 0) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1182
    eventually (\<lambda>x. g x \<noteq> 0) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1183
    eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1184
    eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1185
    eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1186
    ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1187
  ((\<lambda> x. f x / g x) ---> y) (at x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1188
  unfolding eventually_at_split filterlim_at_split
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1189
  by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f'])
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1190
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1191
lemma lhopital_right_0_at_top:
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1192
  fixes f g :: "real \<Rightarrow> real"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1193
  assumes g_0: "LIM x at_right 0. g x :> at_top"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1194
  assumes ev:
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1195
    "eventually (\<lambda>x. g' x \<noteq> 0) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1196
    "eventually (\<lambda>x. DERIV f x :> f' x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1197
    "eventually (\<lambda>x. DERIV g x :> g' x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1198
  assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1199
  shows "((\<lambda> x. f x / g x) ---> x) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1200
  unfolding tendsto_iff
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1201
proof safe
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1202
  fix e :: real assume "0 < e"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1203
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1204
  with lim[unfolded tendsto_iff, rule_format, of "e / 4"]
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1205
  have "eventually (\<lambda>t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1206
  from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]]
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1207
  obtain a where [arith]: "0 < a"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1208
    and g'_neq_0: "\<And>x. 0 < x \<Longrightarrow> x < a \<Longrightarrow> g' x \<noteq> 0"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1209
    and f0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV f x :> (f' x)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1210
    and g0: "\<And>x. 0 < x \<Longrightarrow> x \<le> a \<Longrightarrow> DERIV g x :> (g' x)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1211
    and Df: "\<And>t. 0 < t \<Longrightarrow> t < a \<Longrightarrow> dist (f' t / g' t) x < e / 4"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1212
    unfolding eventually_at_le by (auto simp: dist_real_def)
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1213
    
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1214
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1215
  from Df have
50328
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1216
    "eventually (\<lambda>t. t < a) (at_right 0)" "eventually (\<lambda>t::real. 0 < t) (at_right 0)"
51641
cd05e9fcc63d remove the within-filter, replace "at" by "at _ within UNIV" (This allows to remove a couple of redundant lemmas)
hoelzl
parents: 51529
diff changeset
  1217
    unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1218
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1219
  moreover
50328
25b1e8686ce0 tuned proof
hoelzl
parents: 50327
diff changeset
  1220
  have "eventually (\<lambda>t. 0 < g t) (at_right 0)" "eventually (\<lambda>t. g a < g t) (at_right 0)"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50331
diff changeset
  1221
    using g_0 by (auto elim: eventually_elim1 simp: filterlim_at_top_dense)
50327
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1222
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1223
  moreover
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1224
  have inv_g: "((\<lambda>x. inverse (g x)) ---> 0) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1225
    using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl]
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1226
    by (rule filterlim_compose)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1227
  then have "((\<lambda>x. norm (1 - g a * inverse (g x))) ---> norm (1 - g a * 0)) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1228
    by (intro tendsto_intros)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1229
  then have "((\<lambda>x. norm (1 - g a / g x)) ---> 1) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1230
    by (simp add: inverse_eq_divide)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1231
  from this[unfolded tendsto_iff, rule_format, of 1]
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1232
  have "eventually (\<lambda>x. norm (1 - g a / g x) < 2) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1233
    by (auto elim!: eventually_elim1 simp: dist_real_def)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1234
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1235
  moreover
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1236
  from inv_g have "((\<lambda>t. norm ((f a - x * g a) * inverse (g t))) ---> norm ((f a - x * g a) * 0)) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1237
    by (intro tendsto_intros)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1238
  then have "((\<lambda>t. norm (f a - x * g a) / norm (g t)) ---> 0) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1239
    by (simp add: inverse_eq_divide)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1240
  from this[unfolded tendsto_iff, rule_format, of "e / 2"] `0 < e`
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1241
  have "eventually (\<lambda>t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1242
    by (auto simp: dist_real_def)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1243
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1244
  ultimately show "eventually (\<lambda>t. dist (f t / g t) x < e) (at_right 0)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1245
  proof eventually_elim
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1246
    fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1247
    assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1248
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1249
    have "\<exists>y. t < y \<and> y < a \<and> (g a - g t) * f' y = (f a - f t) * g' y"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1250
      using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1251
    then guess y ..
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1252
    from this
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1253
    have [arith]: "t < y" "y < a" and D_eq: "(f t - f a) / (g t - g a) = f' y / g' y"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1254
      using `g a < g t` g'_neq_0[of y] by (auto simp add: field_simps)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1255
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1256
    have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1257
      by (simp add: field_simps)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1258
    have "norm (f t / g t - x) \<le>
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1259
        norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1260
      unfolding * by (rule norm_triangle_ineq)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1261
    also have "\<dots> = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1262
      by (simp add: abs_mult D_eq dist_real_def)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1263
    also have "\<dots> < (e / 4) * 2 + e / 2"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1264
      using ineq Df[of y] `0 < e` by (intro add_le_less_mono mult_mono) auto
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1265
    finally show "dist (f t / g t) x < e"
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1266
      by (simp add: dist_real_def)
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1267
  qed
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1268
qed
bbea2e82871c add L'Hôpital's rule
hoelzl
parents: 47108
diff changeset
  1269
50330
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1270
lemma lhopital_right_at_top:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1271
  "LIM x at_right x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1272
    eventually (\<lambda>x. g' x \<noteq> 0) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1273
    eventually (\<lambda>x. DERIV f x :> f' x) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1274
    eventually (\<lambda>x. DERIV g x :> g' x) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1275
    ((\<lambda> x. (f' x / g' x)) ---> y) (at_right x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1276
    ((\<lambda> x. f x / g x) ---> y) (at_right x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1277
  unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1278
  by (rule lhopital_right_0_at_top)
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1279
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1280
lemma lhopital_left_at_top:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1281
  "LIM x at_left x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1282
    eventually (\<lambda>x. g' x \<noteq> 0) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1283
    eventually (\<lambda>x. DERIV f x :> f' x) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1284
    eventually (\<lambda>x. DERIV g x :> g' x) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1285
    ((\<lambda> x. (f' x / g' x)) ---> y) (at_left x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1286
    ((\<lambda> x. f x / g x) ---> y) (at_left x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1287
  unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1288
  by (rule lhopital_right_at_top[where f'="\<lambda>x. - f' (- x)"]) (auto simp: DERIV_mirror)
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1289
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1290
lemma lhopital_at_top:
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1291
  "LIM x at x. (g::real \<Rightarrow> real) x :> at_top \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1292
    eventually (\<lambda>x. g' x \<noteq> 0) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1293
    eventually (\<lambda>x. DERIV f x :> f' x) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1294
    eventually (\<lambda>x. DERIV g x :> g' x) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1295
    ((\<lambda> x. (f' x / g' x)) ---> y) (at x) \<Longrightarrow>
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1296
    ((\<lambda> x. f x / g x) ---> y) (at x)"
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1297
  unfolding eventually_at_split filterlim_at_split
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1298
  by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f'])
d0b12171118e conversion rules for at, at_left and at_right; applied to l'Hopital's rules.
hoelzl
parents: 50329
diff changeset
  1299
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1300
lemma lhospital_at_top_at_top:
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1301
  fixes f g :: "real \<Rightarrow> real"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1302
  assumes g_0: "LIM x at_top. g x :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1303
  assumes g': "eventually (\<lambda>x. g' x \<noteq> 0) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1304
  assumes Df: "eventually (\<lambda>x. DERIV f x :> f' x) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1305
  assumes Dg: "eventually (\<lambda>x. DERIV g x :> g' x) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1306
  assumes lim: "((\<lambda> x. (f' x / g' x)) ---> x) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1307
  shows "((\<lambda> x. f x / g x) ---> x) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1308
  unfolding filterlim_at_top_to_right
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1309
proof (rule lhopital_right_0_at_top)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1310
  let ?F = "\<lambda>x. f (inverse x)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1311
  let ?G = "\<lambda>x. g (inverse x)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1312
  let ?R = "at_right (0::real)"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1313
  let ?D = "\<lambda>f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1314
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1315
  show "LIM x ?R. ?G x :> at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1316
    using g_0 unfolding filterlim_at_top_to_right .
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1317
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1318
  show "eventually (\<lambda>x. DERIV ?G x  :> ?D g' x) ?R"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1319
    unfolding eventually_at_right_to_top
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1320
    using Dg eventually_ge_at_top[where c="1::real"]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1321
    apply eventually_elim
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1322
    apply (rule DERIV_cong)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1323
    apply (rule DERIV_chain'[where f=inverse])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1324
    apply (auto intro!:  DERIV_inverse)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1325
    done
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1326
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1327
  show "eventually (\<lambda>x. DERIV ?F x  :> ?D f' x) ?R"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1328
    unfolding eventually_at_right_to_top
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1329
    using Df eventually_ge_at_top[where c="1::real"]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1330
    apply eventually_elim
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1331
    apply (rule DERIV_cong)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1332
    apply (rule DERIV_chain'[where f=inverse])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1333
    apply (auto intro!:  DERIV_inverse)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1334
    done
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1335
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1336
  show "eventually (\<lambda>x. ?D g' x \<noteq> 0) ?R"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1337
    unfolding eventually_at_right_to_top
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1338
    using g' eventually_ge_at_top[where c="1::real"]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1339
    by eventually_elim auto
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1340
    
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1341
  show "((\<lambda>x. ?D f' x / ?D g' x) ---> x) ?R"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1342
    unfolding filterlim_at_right_to_top
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1343
    apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1344
    using eventually_ge_at_top[where c="1::real"]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1345
    by eventually_elim simp
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1346
qed
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1347
21164
0742fc979c67 new Deriv.thy contains stuff from Lim.thy
huffman
parents:
diff changeset
  1348
end