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