src/HOL/Hyperreal/HDeriv.thy
author wenzelm
Wed, 24 Oct 2007 20:17:48 +0200
changeset 25177 f9ced25685e0
parent 25062 af5ef0d4d655
child 27435 b3f8e9bdf9a7
permissions -rw-r--r--
tuned file names etc.;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
22653
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     1
(*  Title       : Deriv.thy
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     2
    ID          : $Id$
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     5
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     6
*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     7
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     8
header{* Differentiation (Nonstandard) *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
     9
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    10
theory HDeriv
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    11
imports Deriv HLim
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    12
begin
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    13
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    14
text{*Nonstandard Definitions*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    15
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    16
definition
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    17
  nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    18
          ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    19
  "NSDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}.
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    20
      (( *f* f)(star_of x + h)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    21
       - star_of (f x))/h @= star_of D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    22
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    23
definition
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    24
  NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    25
    (infixl "NSdifferentiable" 60) where
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    26
  "f NSdifferentiable x = (\<exists>D. NSDERIV f x :> D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    27
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    28
definition
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    29
  increment :: "[real=>real,real,hypreal] => hypreal" where
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    30
  "increment f x h = (@inc. f NSdifferentiable x &
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    31
           inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    32
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    33
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    34
subsection {* Derivatives *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    35
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    36
lemma DERIV_NS_iff:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    37
      "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    38
by (simp add: deriv_def LIM_NSLIM_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    39
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    40
lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NS> D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    41
by (simp add: deriv_def LIM_NSLIM_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    42
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    43
lemma hnorm_of_hypreal:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    44
  "\<And>r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \<bar>r\<bar>"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    45
by transfer (rule norm_of_real)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    46
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    47
lemma Infinitesimal_of_hypreal:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    48
  "x \<in> Infinitesimal \<Longrightarrow>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    49
   (( *f* of_real) x::'a::real_normed_div_algebra star) \<in> Infinitesimal"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    50
apply (rule InfinitesimalI2)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    51
apply (drule (1) InfinitesimalD2)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    52
apply (simp add: hnorm_of_hypreal)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    53
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    54
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    55
lemma of_hypreal_eq_0_iff:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    56
  "\<And>x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    57
by transfer (rule of_real_eq_0_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    58
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    59
lemma NSDeriv_unique:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    60
     "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    61
apply (subgoal_tac "( *f* of_real) epsilon \<in> Infinitesimal - {0::'a star}")
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    62
apply (simp only: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    63
apply (drule (1) bspec)+
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    64
apply (drule (1) approx_trans3)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    65
apply simp
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    66
apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    67
apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    68
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    69
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    70
text {*First NSDERIV in terms of NSLIM*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    71
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    72
text{*first equivalence *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    73
lemma NSDERIV_NSLIM_iff:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    74
      "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    75
apply (simp add: nsderiv_def NSLIM_def, auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    76
apply (drule_tac x = xa in bspec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    77
apply (rule_tac [3] ccontr)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    78
apply (drule_tac [3] x = h in spec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    79
apply (auto simp add: mem_infmal_iff starfun_lambda_cancel)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    80
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    81
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    82
text{*second equivalence *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    83
lemma NSDERIV_NSLIM_iff2:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    84
     "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    85
by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff  diff_minus [symmetric]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    86
              LIM_NSLIM_iff [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    87
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    88
(* while we're at it! *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    89
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    90
lemma NSDERIV_iff2:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    91
     "(NSDERIV f x :> D) =
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    92
      (\<forall>w.
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    93
        w \<noteq> star_of x & w \<approx> star_of x -->
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    94
        ( *f* (%z. (f z - f x) / (z-x))) w \<approx> star_of D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    95
by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    96
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    97
(*FIXME DELETE*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    98
lemma hypreal_not_eq_minus_iff:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
    99
  "(x \<noteq> a) = (x - a \<noteq> (0::'a::ab_group_add))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   100
by auto
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   101
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   102
lemma NSDERIVD5:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   103
  "(NSDERIV f x :> D) ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   104
   (\<forall>u. u \<approx> hypreal_of_real x -->
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   105
     ( *f* (%z. f z - f x)) u \<approx> hypreal_of_real D * (u - hypreal_of_real x))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   106
apply (auto simp add: NSDERIV_iff2)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   107
apply (case_tac "u = hypreal_of_real x", auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   108
apply (drule_tac x = u in spec, auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   109
apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   110
apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   111
apply (subgoal_tac [2] "( *f* (%z. z-x)) u \<noteq> (0::hypreal) ")
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   112
apply (auto simp add:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   113
         approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   114
         Infinitesimal_subset_HFinite [THEN subsetD])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   115
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   116
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   117
lemma NSDERIVD4:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   118
     "(NSDERIV f x :> D) ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   119
      (\<forall>h \<in> Infinitesimal.
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   120
               (( *f* f)(hypreal_of_real x + h) -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   121
                 hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   122
apply (auto simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   123
apply (case_tac "h = (0::hypreal) ")
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   124
apply (auto simp add: diff_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   125
apply (drule_tac x = h in bspec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   126
apply (drule_tac [2] c = h in approx_mult1)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   127
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   128
            simp add: diff_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   129
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   130
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   131
lemma NSDERIVD3:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   132
     "(NSDERIV f x :> D) ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   133
      (\<forall>h \<in> Infinitesimal - {0}.
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   134
               (( *f* f)(hypreal_of_real x + h) -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   135
                 hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   136
apply (auto simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   137
apply (rule ccontr, drule_tac x = h in bspec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   138
apply (drule_tac [2] c = h in approx_mult1)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   139
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   140
            simp add: mult_assoc diff_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   141
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   142
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   143
text{*Differentiability implies continuity
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   144
         nice and simple "algebraic" proof*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   145
lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   146
apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   147
apply (drule approx_minus_iff [THEN iffD1])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   148
apply (drule hypreal_not_eq_minus_iff [THEN iffD1])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   149
apply (drule_tac x = "xa - star_of x" in bspec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   150
 prefer 2 apply (simp add: add_assoc [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   151
apply (auto simp add: mem_infmal_iff [symmetric] add_commute)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   152
apply (drule_tac c = "xa - star_of x" in approx_mult1)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   153
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   154
            simp add: mult_assoc nonzero_mult_divide_cancel_right)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   155
apply (drule_tac x3=D in
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   156
           HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult,
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   157
             THEN mem_infmal_iff [THEN iffD1]])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   158
apply (auto simp add: mult_commute
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   159
            intro: approx_trans approx_minus_iff [THEN iffD2])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   160
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   161
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   162
text{*Differentiation rules for combinations of functions
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   163
      follow from clear, straightforard, algebraic
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   164
      manipulations*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   165
text{*Constant function*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   166
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   167
(* use simple constant nslimit theorem *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   168
lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   169
by (simp add: NSDERIV_NSLIM_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   170
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   171
text{*Sum of functions- proved easily*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   172
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   173
lemma NSDERIV_add: "[| NSDERIV f x :> Da;  NSDERIV g x :> Db |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   174
      ==> NSDERIV (%x. f x + g x) x :> Da + Db"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   175
apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   176
apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   177
apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   178
apply (auto simp add: diff_def add_ac)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   179
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   180
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   181
text{*Product of functions - Proof is trivial but tedious
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   182
  and long due to rearrangement of terms*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   183
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   184
lemma lemma_nsderiv1:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   185
  fixes a b c d :: "'a::comm_ring star"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   186
  shows "(a*b) - (c*d) = (b*(a - c)) + (c*(b - d))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   187
by (simp add: right_diff_distrib mult_ac)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   188
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   189
lemma lemma_nsderiv2:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   190
  fixes x y z :: "'a::real_normed_field star"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   191
  shows "[| (x - y) / z = star_of D + yb; z \<noteq> 0;
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   192
         z \<in> Infinitesimal; yb \<in> Infinitesimal |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   193
      ==> x - y \<approx> 0"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   194
apply (simp add: nonzero_divide_eq_eq)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   195
apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   196
            simp add: mult_assoc mem_infmal_iff [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   197
apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   198
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   199
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   200
lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   201
      ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   202
apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   203
apply (auto dest!: spec
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   204
      simp add: starfun_lambda_cancel lemma_nsderiv1)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   205
apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   206
apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   207
apply (auto simp add: times_divide_eq_right [symmetric]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   208
            simp del: times_divide_eq)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   209
apply (drule_tac D = Db in lemma_nsderiv2, assumption+)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   210
apply (drule_tac
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   211
     approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   212
apply (auto intro!: approx_add_mono1
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   213
            simp add: left_distrib right_distrib mult_commute add_assoc)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   214
apply (rule_tac b1 = "star_of Db * star_of (f x)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   215
         in add_commute [THEN subst])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   216
apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   217
                    Infinitesimal_add Infinitesimal_mult
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   218
                    Infinitesimal_star_of_mult
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   219
                    Infinitesimal_star_of_mult2
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   220
          simp add: add_assoc [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   221
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   222
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   223
text{*Multiplying by a constant*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   224
lemma NSDERIV_cmult: "NSDERIV f x :> D
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   225
      ==> NSDERIV (%x. c * f x) x :> c*D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   226
apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   227
                  minus_mult_right right_diff_distrib [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   228
apply (erule NSLIM_const [THEN NSLIM_mult])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   229
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   230
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   231
text{*Negation of function*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   232
lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   233
proof (simp add: NSDERIV_NSLIM_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   234
  assume "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   235
  hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   236
    by (rule NSLIM_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   237
  have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   238
    by (simp add: minus_divide_left diff_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   239
  with deriv
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   240
  show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   241
qed
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   242
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   243
text{*Subtraction*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   244
lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   245
by (blast dest: NSDERIV_add NSDERIV_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   246
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   247
lemma NSDERIV_diff:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   248
     "[| NSDERIV f x :> Da; NSDERIV g x :> Db |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   249
      ==> NSDERIV (%x. f x - g x) x :> Da-Db"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   250
apply (simp add: diff_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   251
apply (blast intro: NSDERIV_add_minus)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   252
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   253
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   254
text{*  Similarly to the above, the chain rule admits an entirely
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   255
   straightforward derivation. Compare this with Harrison's
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   256
   HOL proof of the chain rule, which proved to be trickier and
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   257
   required an alternative characterisation of differentiability-
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   258
   the so-called Carathedory derivative. Our main problem is
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   259
   manipulation of terms.*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   260
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   261
(* lemmas *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   262
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   263
lemma NSDERIV_zero:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   264
      "[| NSDERIV g x :> D;
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   265
               ( *f* g) (star_of x + xa) = star_of (g x);
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   266
               xa \<in> Infinitesimal;
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   267
               xa \<noteq> 0
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   268
            |] ==> D = 0"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   269
apply (simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   270
apply (drule bspec, auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   271
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   272
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   273
(* can be proved differently using NSLIM_isCont_iff *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   274
lemma NSDERIV_approx:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   275
     "[| NSDERIV f x :> D;  h \<in> Infinitesimal;  h \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   276
      ==> ( *f* f) (star_of x + h) - star_of (f x) \<approx> 0"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   277
apply (simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   278
apply (simp add: mem_infmal_iff [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   279
apply (rule Infinitesimal_ratio)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   280
apply (rule_tac [3] approx_star_of_HFinite, auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   281
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   282
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   283
(*---------------------------------------------------------------
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   284
   from one version of differentiability
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   285
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   286
                f(x) - f(a)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   287
              --------------- \<approx> Db
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   288
                  x - a
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   289
 ---------------------------------------------------------------*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   290
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   291
lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da;
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   292
         ( *f* g) (star_of(x) + xa) \<noteq> star_of (g x);
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   293
         ( *f* g) (star_of(x) + xa) \<approx> star_of (g x)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   294
      |] ==> (( *f* f) (( *f* g) (star_of(x) + xa))
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   295
                   - star_of (f (g x)))
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   296
              / (( *f* g) (star_of(x) + xa) - star_of (g x))
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   297
             \<approx> star_of(Da)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   298
by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   299
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   300
(*--------------------------------------------------------------
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   301
   from other version of differentiability
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   302
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   303
                f(x + h) - f(x)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   304
               ----------------- \<approx> Db
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   305
                       h
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   306
 --------------------------------------------------------------*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   307
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   308
lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \<in> Infinitesimal; xa \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   309
      ==> (( *f* g) (star_of(x) + xa) - star_of(g x)) / xa
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   310
          \<approx> star_of(Db)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   311
by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   312
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   313
lemma lemma_chain: "(z::'a::real_normed_field star) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   314
proof -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   315
  assume z: "z \<noteq> 0"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   316
  have "x * y = x * (inverse z * z) * y" by (simp add: z)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   317
  thus ?thesis by (simp add: mult_assoc)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   318
qed
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   319
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   320
text{*This proof uses both definitions of differentiability.*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   321
lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   322
      ==> NSDERIV (f o g) x :> Da * Db"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   323
apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   324
                mem_infmal_iff [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   325
apply clarify
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   326
apply (frule_tac f = g in NSDERIV_approx)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   327
apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   328
apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ")
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   329
apply (drule_tac g = g in NSDERIV_zero)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   330
apply (auto simp add: divide_inverse)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   331
apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   332
apply (erule hypreal_not_eq_minus_iff [THEN iffD1])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   333
apply (rule approx_mult_star_of)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   334
apply (simp_all add: divide_inverse [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   335
apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   336
apply (blast intro: NSDERIVD2)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   337
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   338
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   339
text{*Differentiation of natural number powers*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   340
lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1"
23398
0b5a400c7595 made divide_self a simp rule
nipkow
parents: 22653
diff changeset
   341
by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if)
22653
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   342
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   343
lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   344
by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   345
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   346
(*Can't get rid of x \<noteq> 0 because it isn't continuous at zero*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   347
lemma NSDERIV_inverse:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   348
  fixes x :: "'a::{real_normed_field,recpower}"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   349
  shows "x \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   350
apply (simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   351
apply (rule ballI, simp, clarify)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   352
apply (frule (1) Infinitesimal_add_not_zero)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   353
apply (simp add: add_commute)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   354
(*apply (auto simp add: starfun_inverse_inverse realpow_two
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   355
        simp del: minus_mult_left [symmetric] minus_mult_right [symmetric])*)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   356
apply (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric] power_Suc
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   357
              nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_def
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   358
            del: inverse_mult_distrib inverse_minus_eq
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   359
                 minus_mult_left [symmetric] minus_mult_right [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   360
apply (subst mult_commute, simp add: nonzero_mult_divide_cancel_right)
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23398
diff changeset
   361
apply (simp (no_asm_simp) add: mult_assoc [symmetric] left_distrib
22653
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   362
            del: minus_mult_left [symmetric] minus_mult_right [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   363
apply (rule_tac y = "inverse (- (star_of x * star_of x))" in approx_trans)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   364
apply (rule inverse_add_Infinitesimal_approx2)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   365
apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   366
            simp add: inverse_minus_eq [symmetric] HFinite_minus_iff)
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23398
diff changeset
   367
apply (rule Infinitesimal_HFinite_mult, auto)
22653
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   368
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   369
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   370
subsubsection {* Equivalence of NS and Standard definitions *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   371
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23413
diff changeset
   372
lemma divideR_eq_divide: "x /\<^sub>R y = x / y"
22653
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   373
by (simp add: real_scaleR_def divide_inverse mult_commute)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   374
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   375
text{*Now equivalence between NSDERIV and DERIV*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   376
lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   377
by (simp add: deriv_def NSDERIV_NSLIM_iff LIM_NSLIM_iff)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   378
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   379
(* NS version *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   380
lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   381
by (simp add: NSDERIV_DERIV_iff DERIV_pow)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   382
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   383
text{*Derivative of inverse*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   384
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   385
lemma NSDERIV_inverse_fun:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   386
  fixes x :: "'a::{real_normed_field,recpower}"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   387
  shows "[| NSDERIV f x :> d; f(x) \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   388
      ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   389
by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: realpow_Suc)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   390
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   391
text{*Derivative of quotient*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   392
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   393
lemma NSDERIV_quotient:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   394
  fixes x :: "'a::{real_normed_field,recpower}"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   395
  shows "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   396
       ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   397
                            - (e*f(x))) / (g(x) ^ Suc (Suc 0))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   398
by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: realpow_Suc)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   399
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   400
lemma CARAT_NSDERIV: "NSDERIV f x :> l ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   401
      \<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   402
by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   403
                   mult_commute)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   404
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   405
lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   406
by auto
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   407
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   408
lemma CARAT_DERIVD:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   409
  assumes all: "\<forall>z. f z - f x = g z * (z-x)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   410
      and nsc: "isNSCont g x"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   411
  shows "NSDERIV f x :> g x"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   412
proof -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   413
  from nsc
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   414
  have "\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   415
         ( *f* g) w * (w - star_of x) / (w - star_of x) \<approx>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   416
         star_of (g x)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   417
    by (simp add: isNSCont_def nonzero_mult_divide_cancel_right)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   418
  thus ?thesis using all
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   419
    by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   420
qed
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   421
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   422
subsubsection {* Differentiability predicate *}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   423
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   424
lemma NSdifferentiableD: "f NSdifferentiable x ==> \<exists>D. NSDERIV f x :> D"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   425
by (simp add: NSdifferentiable_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   426
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   427
lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   428
by (force simp add: NSdifferentiable_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   429
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   430
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   431
subsection {*(NS) Increment*}
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   432
lemma incrementI:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   433
      "f NSdifferentiable x ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   434
      increment f x h = ( *f* f) (hypreal_of_real(x) + h) -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   435
      hypreal_of_real (f x)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   436
by (simp add: increment_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   437
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   438
lemma incrementI2: "NSDERIV f x :> D ==>
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   439
     increment f x h = ( *f* f) (hypreal_of_real(x) + h) -
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   440
     hypreal_of_real (f x)"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   441
apply (erule NSdifferentiableI [THEN incrementI])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   442
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   443
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   444
(* The Increment theorem -- Keisler p. 65 *)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   445
lemma increment_thm: "[| NSDERIV f x :> D; h \<in> Infinitesimal; h \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   446
      ==> \<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   447
apply (frule_tac h = h in incrementI2, simp add: nsderiv_def)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   448
apply (drule bspec, auto)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   449
apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   450
apply (frule_tac b1 = "hypreal_of_real (D) + y"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   451
        in hypreal_mult_right_cancel [THEN iffD2])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   452
apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   453
apply assumption
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   454
apply (simp add: times_divide_eq_right [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   455
apply (auto simp add: left_distrib)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   456
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   457
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   458
lemma increment_thm2:
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   459
     "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   460
      ==> \<exists>e \<in> Infinitesimal. increment f x h =
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   461
              hypreal_of_real(D)*h + e*h"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   462
by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm)
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   463
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   464
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   465
lemma increment_approx_zero: "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |]
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   466
      ==> increment f x h \<approx> 0"
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   467
apply (drule increment_thm2,
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   468
       auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: left_distrib [symmetric] mem_infmal_iff [symmetric])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   469
apply (erule Infinitesimal_subset_HFinite [THEN subsetD])
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   470
done
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   471
8e016bfdbf2f moved nonstandard derivative stuff from Deriv.thy into new file HDeriv.thy
huffman
parents:
diff changeset
   472
end