src/HOL/Nonstandard_Analysis/CLim.thy
author paulson <lp15@cam.ac.uk>
Mon, 15 Aug 2022 21:57:55 +0100
changeset 75866 9eeed5c424f9
parent 70723 4e39d87c9737
permissions -rw-r--r--
A bit of cleaning up
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61976
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/CLim.thy
716336f19aa9 clarified session;
wenzelm
parents: 61976
diff changeset
     2
    Author:     Jacques D. Fleuriot
716336f19aa9 clarified session;
wenzelm
parents: 61976
diff changeset
     3
    Copyright:  2001 University of Edinburgh
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
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
     7
section \<open>Limits, Continuity and Differentiation for Complex Functions\<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 CLim
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    10
  imports CStar
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
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
(*not in simpset?*)
70723
4e39d87c9737 imported new material mostly due to Sébastien Gouëzel
paulson <lp15@cam.ac.uk>
parents: 70228
diff changeset
    14
declare epsilon_not_zero [simp]
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    15
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
(*??generalize*)
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    17
lemma lemma_complex_mult_inverse_squared [simp]: "x \<noteq> 0 \<Longrightarrow> x * (inverse x)\<^sup>2 = inverse x"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    18
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    19
  by (simp add: numeral_2_eq_2)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    20
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    21
text \<open>Changing the quantified variable. Install earlier?\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    22
lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) \<longleftrightarrow> (\<forall>x. P (x - a))"
75866
9eeed5c424f9 A bit of cleaning up
paulson <lp15@cam.ac.uk>
parents: 70723
diff changeset
    23
  by (metis add_diff_cancel)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    25
subsection \<open>Limit of Complex to Complex Function\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    26
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    27
lemma NSLIM_Re: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    28
  by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hRe_hcomplex_of_complex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    30
lemma NSLIM_Im: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    31
  by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hIm_hcomplex_of_complex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    33
lemma LIM_Re: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    34
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    35
  by (simp add: LIM_NSLIM_iff NSLIM_Re)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    36
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    37
lemma LIM_Im: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    38
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    39
  by (simp add: LIM_NSLIM_iff NSLIM_Im)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    41
lemma LIM_cnj: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    42
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    43
  by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    44
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    45
lemma LIM_cnj_iff: "((\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L) \<longleftrightarrow> f \<midarrow>a\<rightarrow> L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    46
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    47
  by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    48
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))"
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    50
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    51
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)"
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    53
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    54
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)"
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    56
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    57
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    58
text \<open>Another equivalence result.\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    59
lemma NSCLIM_NSCRLIM_iff: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    60
  by (simp add: NSLIM_def starfun_norm
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    61
      approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    63
text \<open>Much, much easier standard proof.\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    64
lemma CLIM_CRLIM_iff: "f \<midarrow>x\<rightarrow> L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow> 0"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    65
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    66
  by (simp add: LIM_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    68
text \<open>So this is nicer nonstandard proof.\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    69
lemma NSCLIM_NSCRLIM_iff2: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    70
  by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    72
lemma NSLIM_NSCRLIM_Re_Im_iff:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    73
  "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    74
  apply (auto intro: NSLIM_Re NSLIM_Im)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    75
  apply (auto simp add: NSLIM_def starfun_Re starfun_Im)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    76
  apply (auto dest!: spec)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    77
  apply (simp add: hcomplex_approx_iff)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    78
  done
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    79
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    80
lemma LIM_CRLIM_Re_Im_iff: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    81
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    82
  by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    83
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    84
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    85
subsection \<open>Continuity\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    86
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    87
lemma NSLIM_isContc_iff: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a"
70228
2d5b122aa0ff De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    88
  by (rule NSLIM_at0_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    90
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    91
subsection \<open>Functions from Complex to Reals\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    92
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    93
lemma isNSContCR_cmod [simp]: "isNSCont cmod a"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    94
  by (auto intro: approx_hnorm
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    95
      simp: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] isNSCont_def)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    96
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    97
lemma isContCR_cmod [simp]: "isCont cmod a"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    98
  by (simp add: isNSCont_isCont_iff [symmetric])
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    99
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   100
lemma isCont_Re: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Re (f x)) a"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   101
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   102
  by (simp add: isCont_def LIM_Re)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   103
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   104
lemma isCont_Im: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Im (f x)) a"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   105
  for f :: "'a::real_normed_vector \<Rightarrow> complex"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   106
  by (simp add: isCont_def LIM_Im)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   107
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   108
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   109
subsection \<open>Differentiation of Natural Number Powers\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   110
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   111
lemma CDERIV_pow [simp]: "DERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - Suc 0))"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   112
  apply (induct n)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   113
   apply (drule_tac [2] DERIV_ident [THEN DERIV_mult])
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   114
   apply (auto simp add: distrib_right of_nat_Suc)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   115
  apply (case_tac "n")
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   116
   apply (auto simp add: ac_simps)
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   117
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   119
text \<open>Nonstandard version.\<close>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   120
lemma NSCDERIV_pow: "NSDERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   121
  by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   122
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64601
diff changeset
   123
text \<open>Can't relax the premise \<^term>\<open>x \<noteq> 0\<close>: it isn't continuous at zero.\<close>
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   124
lemma NSCDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   125
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   126
  unfolding numeral_2_eq_2 by (rule NSDERIV_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   127
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   128
lemma CDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   129
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   130
  unfolding numeral_2_eq_2 by (rule DERIV_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   131
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   132
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 64601
diff changeset
   133
subsection \<open>Derivative of Reciprocals (Function \<^term>\<open>inverse\<close>)\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   134
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
lemma CDERIV_inverse_fun:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   136
  "DERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   137
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   138
  unfolding numeral_2_eq_2 by (rule DERIV_inverse_fun)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
lemma NSCDERIV_inverse_fun:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   141
  "NSDERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   142
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   143
  unfolding numeral_2_eq_2 by (rule NSDERIV_inverse_fun)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   144
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   145
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   146
subsection \<open>Derivative of Quotient\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   147
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   148
lemma CDERIV_quotient:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   149
  "DERIV f x :> d \<Longrightarrow> DERIV g x :> e \<Longrightarrow> g(x) \<noteq> 0 \<Longrightarrow>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   150
    DERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   151
  for x :: complex
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   152
  unfolding numeral_2_eq_2 by (rule DERIV_quotient)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   153
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
lemma NSCDERIV_quotient:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   155
  "NSDERIV f x :> d \<Longrightarrow> NSDERIV g x :> e \<Longrightarrow> g x \<noteq> (0::complex) \<Longrightarrow>
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   156
    NSDERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   157
  unfolding numeral_2_eq_2 by (rule NSDERIV_quotient)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   158
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   159
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   160
subsection \<open>Caratheodory Formulation of Derivative at a Point: Standard Proof\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   162
lemma CARAT_CDERIVD:
64601
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   163
  "(\<forall>z. f z - f x = g z * (z - x)) \<and> isNSCont g x \<and> g x = l \<Longrightarrow> NSDERIV f x :> l"
37ce6ceacbb7 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   164
  by clarify (rule CARAT_DERIVD)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   165
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   166
end