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