src/HOL/Complex/CLim.thy
author haftmann
Sat, 19 May 2007 11:33:30 +0200
changeset 23024 70435ffe077d
parent 22979 d95580341be5
child 23069 cdfff0241c12
permissions -rw-r--r--
fixed text
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     1
(*  Title       : CLim.thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Copyright   : 2001 University of Edinburgh
14469
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
14469
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
     7
header{*Limits, Continuity and Differentiation for Complex Functions*}
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15013
diff changeset
     9
theory CLim
22979
d95580341be5 minimize imports
huffman
parents: 22883
diff changeset
    10
imports CStar
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15013
diff changeset
    11
begin
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    12
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    13
(*not in simpset?*)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    14
declare hypreal_epsilon_not_zero [simp]
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    15
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    16
(*??generalize*)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    17
lemma lemma_complex_mult_inverse_squared [simp]:
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    18
     "x \<noteq> (0::complex) \<Longrightarrow> (x * inverse(x) ^ 2) = inverse x"
14469
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
    19
by (simp add: numeral_2_eq_2)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    20
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    21
text{*Changing the quantified variable. Install earlier?*}
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14469
diff changeset
    22
lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))";
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    23
apply auto 
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    24
apply (drule_tac x="x+a" in spec) 
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    25
apply (simp add: diff_minus add_assoc) 
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    26
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    27
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    28
lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)"
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    29
by (simp add: diff_eq_eq diff_minus [symmetric])
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    30
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    31
lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)"
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    32
apply auto
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    33
apply (drule sym [THEN diff_eq_eq [THEN iffD2]], auto)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    34
done
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    36
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    37
subsection{*Limit of Complex to Complex Function*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    38
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    39
lemma NSLIM_Re: "f -- a --NS> L ==> (%x. Re(f x)) -- a --NS> Re(L)"
20659
66b8455090b8 changed constants into abbreviations; shortened proofs
huffman
parents: 20563
diff changeset
    40
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff 
14469
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
    41
              hRe_hcomplex_of_complex)
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
    42
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    43
lemma NSLIM_Im: "f -- a --NS> L ==> (%x. Im(f x)) -- a --NS> Im(L)"
20659
66b8455090b8 changed constants into abbreviations; shortened proofs
huffman
parents: 20563
diff changeset
    44
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff 
14469
c7674b7034f5 heavy tidying
paulson
parents: 14430
diff changeset
    45
              hIm_hcomplex_of_complex)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    46
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    47
(** get this result easily now **)
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    48
lemma LIM_Re: "f -- a --> L ==> (%x. Re(f x)) -- a --> Re(L)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    49
by (simp add: LIM_NSLIM_iff NSLIM_Re)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    50
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    51
lemma LIM_Im: "f -- a --> L ==> (%x. Im(f x)) -- a --> Im(L)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    52
by (simp add: LIM_NSLIM_iff NSLIM_Im)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    53
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    54
lemma LIM_cnj: "f -- a --> L ==> (%x. cnj (f x)) -- a --> cnj L"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    55
by (simp add: LIM_def complex_cnj_diff [symmetric])
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    56
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    57
lemma LIM_cnj_iff: "((%x. cnj (f x)) -- a --> cnj L) = (f -- a --> L)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    58
by (simp add: LIM_def complex_cnj_diff [symmetric])
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    59
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    60
(*** NSLIM_not zero and hence LIM_not_zero ***)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    61
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    62
lemma NSCLIM_not_zero: "k \<noteq> 0 ==> ~ ((%x::complex. k) -- x --NS> 0)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    63
apply (auto simp del: star_of_zero simp add: NSLIM_def)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    64
apply (rule_tac x = "hcomplex_of_complex x + hcomplex_of_hypreal epsilon" in exI)
20559
2116b7a371c7 removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents: 20552
diff changeset
    65
apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym]
17373
27509e72f29e removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents: 17318
diff changeset
    66
            simp del: star_of_zero)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    67
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    68
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    69
(* [| k \<noteq> 0; (%x. k) -- x --NS> 0 |] ==> R *)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    70
lemmas NSCLIM_not_zeroE = NSCLIM_not_zero [THEN notE, standard]
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    71
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    72
(*** NSLIM_const hence LIM_const ***)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    73
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    74
lemma NSCLIM_const_eq: "(%x::complex. k) -- x --NS> L ==> k = L"
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    75
apply (rule ccontr)
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    76
apply (drule NSLIM_zero)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    77
apply (rule NSCLIM_not_zeroE [of "k-L"], auto)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    78
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    79
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    80
(*** NSLIM and hence LIM are unique ***)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    81
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    82
lemma NSCLIM_unique: "[| f -- (x::complex) --NS> L; f -- x --NS> M |] ==> L = M"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    83
apply (drule (1) NSLIM_diff)
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
    84
apply (drule NSLIM_minus)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    85
apply (auto dest!: NSCLIM_const_eq [symmetric])
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    86
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    87
21831
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    88
lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))"
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    89
by transfer (rule refl)
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    90
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    91
lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)"
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    92
by transfer (rule refl)
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    93
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    94
lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)"
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    95
by transfer (rule refl)
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    96
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
    97
text""
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    98
(** another equivalence result **)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
    99
lemma NSCLIM_NSCRLIM_iff:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   100
   "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)"
21831
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
   101
by (simp add: NSLIM_def starfun_norm
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
   102
    approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric])
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   103
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   104
(** much, much easier standard proof **)
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   105
lemma CLIM_CRLIM_iff: "(f -- x --> L) = ((%y. cmod(f y - L)) -- x --> 0)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   106
by (simp add: LIM_def)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   107
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   108
(* so this is nicer nonstandard proof *)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   109
lemma NSCLIM_NSCRLIM_iff2:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   110
     "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   111
by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   112
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   113
lemma NSLIM_NSCRLIM_Re_Im_iff:
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   114
     "(f -- a --NS> L) = ((%x. Re(f x)) -- a --NS> Re(L) &
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   115
                            (%x. Im(f x)) -- a --NS> Im(L))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   116
apply (auto intro: NSLIM_Re NSLIM_Im)
21831
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
   117
apply (auto simp add: NSLIM_def starfun_Re starfun_Im)
1897fe3d72d5 remove references to star_n
huffman
parents: 21792
diff changeset
   118
apply (auto dest!: spec)
21839
54018ed3b99d added lemmas about hRe, hIm, HComplex; removed all uses of star_n
huffman
parents: 21831
diff changeset
   119
apply (simp add: hcomplex_approx_iff)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   120
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   121
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   122
lemma LIM_CRLIM_Re_Im_iff:
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   123
     "(f -- a --> L) = ((%x. Re(f x)) -- a --> Re(L) &
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   124
                         (%x. Im(f x)) -- a --> Im(L))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   125
by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   126
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   127
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   128
subsection{*Continuity*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   129
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   130
lemma NSLIM_isContc_iff:
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   131
     "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)"
20659
66b8455090b8 changed constants into abbreviations; shortened proofs
huffman
parents: 20563
diff changeset
   132
by (rule NSLIM_h_iff)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   133
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   134
subsection{*Functions from Complex to Reals*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   135
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   136
lemma isNSContCR_cmod [simp]: "isNSCont cmod (a)"
22883
005be8dafce0 hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
huffman
parents: 21839
diff changeset
   137
by (auto intro: approx_hnorm
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   138
         simp add: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] 
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   139
                    isNSCont_def)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   140
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   141
lemma isContCR_cmod [simp]: "isCont cmod (a)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   142
by (simp add: isNSCont_isCont_iff [symmetric])
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   143
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   144
lemma isCont_Re: "isCont f a ==> isCont (%x. Re (f x)) a"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   145
by (simp add: isCont_def LIM_Re)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   146
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   147
lemma isCont_Im: "isCont f a ==> isCont (%x. Im (f x)) a"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   148
by (simp add: isCont_def LIM_Im)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   149
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   150
subsection{* Differentiation of Natural Number Powers*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   151
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   152
lemma CDERIV_pow [simp]:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   153
     "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))"
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   154
apply (induct_tac "n")
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   155
apply (drule_tac [2] DERIV_Id [THEN DERIV_mult])
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 14738
diff changeset
   156
apply (auto simp add: left_distrib real_of_nat_Suc)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   157
apply (case_tac "n")
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   158
apply (auto simp add: mult_ac add_commute)
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   159
done
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   160
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   161
text{*Nonstandard version*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   162
lemma NSCDERIV_pow:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   163
     "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   164
by (simp add: NSDERIV_DERIV_iff)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   165
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   166
text{*Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   167
lemma NSCDERIV_inverse:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   168
     "(x::complex) \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   169
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   170
by (rule NSDERIV_inverse)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   171
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   172
lemma CDERIV_inverse:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   173
     "(x::complex) \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   174
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   175
by (rule DERIV_inverse)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   176
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   177
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   178
subsection{*Derivative of Reciprocals (Function @{term inverse})*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   179
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   180
lemma CDERIV_inverse_fun:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   181
     "[| DERIV f x :> d; f(x) \<noteq> (0::complex) |]
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   182
      ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   183
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   184
by (rule DERIV_inverse_fun)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   185
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   186
lemma NSCDERIV_inverse_fun:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   187
     "[| NSDERIV f x :> d; f(x) \<noteq> (0::complex) |]
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   188
      ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   189
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   190
by (rule NSDERIV_inverse_fun)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   191
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   192
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   193
subsection{* Derivative of Quotient*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   194
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   195
lemma CDERIV_quotient:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   196
     "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> (0::complex) |]
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   197
       ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   198
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   199
by (rule DERIV_quotient)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   200
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   201
lemma NSCDERIV_quotient:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   202
     "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> (0::complex) |]
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   203
       ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   204
unfolding numeral_2_eq_2
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   205
by (rule NSDERIV_quotient)
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   206
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   207
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   208
subsection{*Caratheodory Formulation of Derivative at a Point: Standard Proof*}
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   209
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   210
lemma CARAT_CDERIVD:
21792
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   211
     "(\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   212
      ==> NSDERIV f x :> l"
266a1056a2a3 removed redundant constants and lemmas
huffman
parents: 21404
diff changeset
   213
by clarify (rule CARAT_DERIVD)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15228
diff changeset
   214
14405
534de3572a65 conversion of Complex/CLim to Isar script
paulson
parents: 13957
diff changeset
   215
end