src/HOL/Complex/CLim.ML
author paulson
Sun, 15 Feb 2004 10:46:37 +0100
changeset 14387 e96d5c42c4b0
parent 14373 67a628beb981
permissions -rw-r--r--
Polymorphic treatment of binary arithmetic using axclasses
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.ML
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Description : A first theory of limits, continuity and 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
                  differentiation for complex functions
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
     8
(*FIXME: MOVE these two to Complex.thy*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
     9
Goal "(x + - a = (0::complex)) = (x=a)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    10
by (simp_tac (simpset() addsimps [diff_eq_eq,symmetric complex_diff_def]) 1);
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    11
qed "complex_add_minus_iff";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    12
Addsimps [complex_add_minus_iff];
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    13
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    14
Goal "(x+y = (0::complex)) = (y = -x)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    15
by Auto_tac;
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    16
by (dtac (sym RS (diff_eq_eq RS iffD2)) 1);
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    17
by Auto_tac;  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    18
qed "complex_add_eq_0_iff";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    19
AddIffs [complex_add_eq_0_iff];
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    20
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
    21
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
    22
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
(* Limit of complex to complex function                                               *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
    24
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
Goalw [NSCLIM_def,NSCRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
   "f -- a --NSC> L ==> (%x. Re(f x)) -- a --NSCR> Re(L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
by (auto_tac (claset(),simpset() addsimps [starfunC_approx_Re_Im_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
    hRe_hcomplex_of_complex]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
qed "NSCLIM_NSCRLIM_Re";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
Goalw [NSCLIM_def,NSCRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
   "f -- a --NSC> L ==> (%x. Im(f x)) -- a --NSCR> Im(L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
by (auto_tac (claset(),simpset() addsimps [starfunC_approx_Re_Im_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
    hIm_hcomplex_of_complex]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
qed "NSCLIM_NSCRLIM_Im";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
Goalw [CLIM_def,NSCLIM_def,capprox_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
      "f -- x --C> L ==> f -- x --NSC> L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
    starfunC,hcomplex_diff,CInfinitesimal_hcmod_iff,hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
    Infinitesimal_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
by (dres_inst_tac [("x","s")] spec 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
by (dtac sym 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
qed "CLIM_NSCLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
Goal "(ALL t. P t) = (ALL X. P (Abs_hcomplex(hcomplexrel `` {X})))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
qed "eq_Abs_hcomplex_ALL";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
Goal "ALL s. 0 < s --> (EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
\        cmod (xa - x) < s  & r <= cmod (f xa - L)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
\     ==> ALL (n::nat). EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
\             cmod(xa - x) < inverse(real(Suc n)) & r <= cmod(f xa - L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
by (Clarify_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
by (cut_inst_tac [("n1","n")]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
    67
    (real_of_nat_Suc_gt_zero RS positive_imp_inverse_positive) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
val lemma_CLIM = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
(* not needed? *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    72
Goal "ALL x z. EX y. Q x z y ==> EX f. ALL x z. Q x z (f x z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
by (rtac choice 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
by (blast_tac (claset() addIs [choice]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
qed "choice2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
Goal "ALL s. 0 < s --> (EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
\        cmod (xa - x) < s  & r <= cmod (f xa - L)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
\     ==> EX X. ALL (n::nat). X n ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
\               cmod(X n - x) < inverse(real(Suc n)) & r <= cmod(f (X n) - L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
by (dtac lemma_CLIM 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
by (dtac choice 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
val lemma_skolemize_CLIM2 = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
Goal "ALL n. X n ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
\         cmod (X n - x) < inverse (real(Suc n)) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
\         r <= cmod (f (X n) - L) ==> \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
\         ALL n. cmod (X n - x) < inverse (real(Suc n))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
by (Auto_tac );
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
val lemma_csimp = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
Goalw [CLIM_def,NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
     "f -- x --NSC> L ==> f -- x --C> L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
by (auto_tac (claset(),simpset() addsimps [eq_Abs_hcomplex_ALL,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
    starfunC,CInfinitesimal_capprox_minus RS sym,hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
    CInfinitesimal_hcmod_iff,hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    98
    Infinitesimal_FreeUltrafilterNat_iff,hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    99
by (EVERY1[rtac ccontr, Asm_full_simp_tac]);
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14354
diff changeset
   100
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
by (dtac lemma_skolemize_CLIM2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
by (dres_inst_tac [("x","X")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
by (dtac (lemma_csimp RS complex_seq_to_hcomplex_CInfinitesimal) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
    hcomplex_of_complex_def,Infinitesimal_FreeUltrafilterNat_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
    hcomplex_diff,hcmod]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
by (Blast_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
by (dres_inst_tac [("x","r")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
by (Clarify_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
by (dtac FreeUltrafilterNat_all 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
by (arith_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
qed "NSCLIM_CLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
(**** First key result ****)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   119
Goal "(f -- x --C> L) = (f -- x --NSC> L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   120
by (blast_tac (claset() addIs [CLIM_NSCLIM,NSCLIM_CLIM]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
qed "CLIM_NSCLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   123
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
(* Limit of complex to real function                                                  *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   125
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
Goalw [CRLIM_def,NSCRLIM_def,capprox_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
      "f -- x --CR> L ==> f -- x --NSCR> L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   129
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   130
by (res_inst_tac [("z","xa")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
    starfunCR,hcomplex_diff,CInfinitesimal_hcmod_iff,hcmod,hypreal_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
    Infinitesimal_FreeUltrafilterNat_iff,Infinitesimal_approx_minus RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   134
    hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   135
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   137
by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   138
by (dres_inst_tac [("x","s")] spec 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   139
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
by (dtac sym 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
qed "CRLIM_NSCRLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   142
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   143
Goal "ALL s. 0 < s --> (EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
\        cmod (xa - x) < s  & r <= abs (f xa - L)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
\     ==> ALL (n::nat). EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
\             cmod(xa - x) < inverse(real(Suc n)) & r <= abs (f xa - L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   147
by (Clarify_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   148
by (cut_inst_tac [("n1","n")]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   149
    (real_of_nat_Suc_gt_zero RS positive_imp_inverse_positive) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   150
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   151
val lemma_CRLIM = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   152
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   153
Goal "ALL s. 0 < s --> (EX xa.  xa ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   154
\        cmod (xa - x) < s  & r <= abs (f xa - L)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   155
\     ==> EX X. ALL (n::nat). X n ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   156
\               cmod(X n - x) < inverse(real(Suc n)) & r <= abs (f (X n) - L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   157
by (dtac lemma_CRLIM 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   158
by (dtac choice 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   159
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   160
val lemma_skolemize_CRLIM2 = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   161
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   162
Goal "ALL n. X n ~= x & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   163
\         cmod (X n - x) < inverse (real(Suc n)) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   164
\         r <= abs (f (X n) - L) ==> \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   165
\         ALL n. cmod (X n - x) < inverse (real(Suc n))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   166
by (Auto_tac );
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   167
val lemma_crsimp = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   168
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   169
Goalw [CRLIM_def,NSCRLIM_def,capprox_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   170
      "f -- x --NSCR> L ==> f -- x --CR> L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   171
by (auto_tac (claset(),simpset() addsimps [eq_Abs_hcomplex_ALL,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   172
    starfunCR,hcomplex_diff,hcomplex_of_complex_def,hypreal_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   173
    CInfinitesimal_hcmod_iff,hcmod,Infinitesimal_approx_minus RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   174
    Infinitesimal_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   175
by (EVERY1[rtac ccontr, Asm_full_simp_tac]);
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14354
diff changeset
   176
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   177
by (dtac lemma_skolemize_CRLIM2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   178
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   179
by (dres_inst_tac [("x","X")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   180
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   181
by (dtac (lemma_crsimp RS complex_seq_to_hcomplex_CInfinitesimal) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   182
by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   183
    hcomplex_of_complex_def,Infinitesimal_FreeUltrafilterNat_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   184
    hcomplex_diff,hcmod]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   185
by (Blast_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   186
by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   187
    hypreal_diff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   188
by (dres_inst_tac [("x","r")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   189
by (Clarify_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   190
by (dtac FreeUltrafilterNat_all 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   191
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   192
qed "NSCRLIM_CRLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   193
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   194
(** second key result **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   195
Goal "(f -- x --CR> L) = (f -- x --NSCR> L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   196
by (blast_tac (claset() addIs [CRLIM_NSCRLIM,NSCRLIM_CRLIM]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   197
qed "CRLIM_NSCRLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   198
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
(** get this result easily now **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   200
Goal "f -- a --C> L ==> (%x. Re(f x)) -- a --CR> Re(L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   201
by (auto_tac (claset() addDs [NSCLIM_NSCRLIM_Re],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   202
    addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   203
qed "CLIM_CRLIM_Re";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   204
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   205
Goal "f -- a --C> L ==> (%x. Im(f x)) -- a --CR> Im(L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   206
by (auto_tac (claset() addDs [NSCLIM_NSCRLIM_Im],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   207
    addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   208
qed "CLIM_CRLIM_Im";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   209
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   210
Goal "f -- a --C> L ==> (%x. cnj (f x)) -- a --C> cnj L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   211
by (auto_tac (claset(),simpset() addsimps [CLIM_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   212
    complex_cnj_diff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   213
qed "CLIM_cnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   214
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   215
Goal "((%x. cnj (f x)) -- a --C> cnj L) = (f -- a --C> L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   216
by (auto_tac (claset(),simpset() addsimps [CLIM_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   217
    complex_cnj_diff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   218
qed "CLIM_cnj_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   219
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   220
(*** NSLIM_add hence CLIM_add *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   221
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   222
Goalw [NSCLIM_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   223
     "[| f -- x --NSC> l; g -- x --NSC> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   224
\     ==> (%x. f(x) + g(x)) -- x --NSC> (l + m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   225
by (auto_tac (claset() addSIs [capprox_add], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   226
qed "NSCLIM_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   227
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   228
Goal "[| f -- x --C> l; g -- x --C> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   229
\     ==> (%x. f(x) + g(x)) -- x --C> (l + m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   230
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_add]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   231
qed "CLIM_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   232
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   233
(*** NSLIM_mult hence CLIM_mult *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   234
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   235
Goalw [NSCLIM_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   236
     "[| f -- x --NSC> l; g -- x --NSC> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   237
\     ==> (%x. f(x) * g(x)) -- x --NSC> (l * m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   238
by (auto_tac (claset() addSIs [capprox_mult_CFinite],  simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   239
qed "NSCLIM_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   240
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   241
Goal "[| f -- x --C> l; g -- x --C> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   242
\     ==> (%x. f(x) * g(x)) -- x --C> (l * m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   243
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   244
qed "CLIM_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   245
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   246
(*** NSCLIM_const and CLIM_const ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   247
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   248
Goalw [NSCLIM_def] "(%x. k) -- x --NSC> k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   249
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   250
qed "NSCLIM_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   251
Addsimps [NSCLIM_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   252
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   253
Goalw [CLIM_def] "(%x. k) -- x --C> k";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   254
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   255
qed "CLIM_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   256
Addsimps [CLIM_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   257
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   258
(*** NSCLIM_minus and CLIM_minus ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   259
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   260
Goalw [NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   261
      "f -- a --NSC> L ==> (%x. -f(x)) -- a --NSC> -L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   262
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   263
qed "NSCLIM_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   264
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   265
Goal "f -- a --C> L ==> (%x. -f(x)) -- a --C> -L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   266
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   267
qed "CLIM_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   268
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   269
(*** NSCLIM_diff hence CLIM_diff ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   270
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   271
Goalw [complex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   272
     "[| f -- x --NSC> l; g -- x --NSC> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   273
\     ==> (%x. f(x) - g(x)) -- x --NSC> (l - m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   274
by (auto_tac (claset(), simpset() addsimps [NSCLIM_add,NSCLIM_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   275
qed "NSCLIM_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   276
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   277
Goal "[| f -- x --C> l; g -- x --C> m |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   278
\     ==> (%x. f(x) - g(x)) -- x --C> (l - m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   279
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_diff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   280
qed "CLIM_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   281
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   282
(*** NSCLIM_inverse and hence CLIM_inverse *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   283
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   284
Goalw [NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   285
     "[| f -- a --NSC> L;  L ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   286
\     ==> (%x. inverse(f(x))) -- a --NSC> (inverse L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   287
by (Clarify_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   288
by (dtac spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   289
by (auto_tac (claset(), 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   290
              simpset() addsimps [hcomplex_of_complex_capprox_inverse]));  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   291
qed "NSCLIM_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   292
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   293
Goal "[| f -- a --C> L;  L ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   294
\     ==> (%x. inverse(f(x))) -- a --C> (inverse L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   295
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   296
qed "CLIM_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   297
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   298
(*** NSCLIM_zero, CLIM_zero, etc. ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   299
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   300
Goal "f -- a --NSC> l ==> (%x. f(x) - l) -- a --NSC> 0";
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   301
by (res_inst_tac [("a1","l")] (right_minus RS subst) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   302
by (rewtac complex_diff_def);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   303
by (rtac NSCLIM_add 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   304
qed "NSCLIM_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   305
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   306
Goal "f -- a --C> l ==> (%x. f(x) - l) -- a --C> 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   307
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   308
qed "CLIM_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   309
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   310
Goal "(%x. f(x) - l) -- x --NSC> 0 ==> f -- x --NSC> l";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   311
by (dres_inst_tac [("g","%x. l"),("m","l")] NSCLIM_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   312
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   313
qed "NSCLIM_zero_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   314
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   315
Goal "(%x. f(x) - l) -- x --C> 0 ==> f -- x --C> l";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   316
by (dres_inst_tac [("g","%x. l"),("m","l")] CLIM_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   317
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   318
qed "CLIM_zero_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   319
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   320
(*** NSCLIM_not zero and hence CLIM_not_zero ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   321
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   322
(*not in simpset?*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   323
Addsimps [hypreal_epsilon_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   324
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   325
Goalw [NSCLIM_def] "k ~= 0 ==> ~ ((%x. k) -- x --NSC> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   326
by (auto_tac (claset(),simpset() delsimps [hcomplex_of_complex_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   327
by (res_inst_tac [("x","hcomplex_of_complex x + hcomplex_of_hypreal epsilon")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   328
by (auto_tac (claset() addIs [CInfinitesimal_add_capprox_self RS capprox_sym],simpset()
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   329
    delsimps [hcomplex_of_complex_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   330
qed "NSCLIM_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   331
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   332
(* [| k ~= 0; (%x. k) -- x --NSC> 0 |] ==> R *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   333
bind_thm("NSCLIM_not_zeroE", NSCLIM_not_zero RS notE);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   334
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   335
Goal "k ~= 0 ==> ~ ((%x. k) -- x --C> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   336
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_not_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   337
qed "CLIM_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   338
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   339
(*** NSCLIM_const hence CLIM_const ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   340
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   341
Goal "(%x. k) -- x --NSC> L ==> k = L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   342
by (rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   343
by (dtac NSCLIM_zero 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   344
by (rtac NSCLIM_not_zeroE 1 THEN assume_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   345
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   346
qed "NSCLIM_const_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   347
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   348
Goal "(%x. k) -- x --C> L ==> k = L";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   349
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff,NSCLIM_const_eq]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   350
qed "CLIM_const_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   351
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   352
(*** NSCLIM and hence CLIM are unique ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   353
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   354
Goal "[| f -- x --NSC> L; f -- x --NSC> M |] ==> L = M";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   355
by (dtac NSCLIM_minus 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   356
by (dtac NSCLIM_add 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   357
by (auto_tac (claset() addSDs [NSCLIM_const_eq RS sym], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   358
qed "NSCLIM_unique";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   359
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   360
Goal "[| f -- x --C> L; f -- x --C> M |] ==> L = M";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   361
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_unique]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   362
qed "CLIM_unique";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   363
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   364
(***  NSCLIM_mult_zero and CLIM_mult_zero ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   365
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   366
Goal "[| f -- x --NSC> 0; g -- x --NSC> 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   367
\         ==> (%x. f(x)*g(x)) -- x --NSC> 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   368
by (dtac NSCLIM_mult 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   369
qed "NSCLIM_mult_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   370
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   371
Goal "[| f -- x --C> 0; g -- x --C> 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   372
\     ==> (%x. f(x)*g(x)) -- x --C> 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   373
by (dtac CLIM_mult 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   374
qed "CLIM_mult_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   375
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   376
(*** NSCLIM_self hence CLIM_self ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   377
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   378
Goalw [NSCLIM_def] "(%x. x) -- a --NSC> a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   379
by (auto_tac (claset() addIs [starfunC_Idfun_capprox],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   380
qed "NSCLIM_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   381
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   382
Goal "(%x. x) -- a --C> a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   383
by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff,NSCLIM_self]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   384
qed "CLIM_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   385
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   386
(** another equivalence result **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   387
Goalw [NSCLIM_def,NSCRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   388
   "(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   389
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_capprox_minus 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   390
    RS sym,CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   391
by (ALLGOALS(dtac spec) THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   392
by (ALLGOALS(res_inst_tac [("z","xa")] eq_Abs_hcomplex));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   393
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   394
    starfunC,starfunCR,hcomplex_of_complex_def,hcmod,mem_infmal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   395
qed "NSCLIM_NSCRLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   396
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   397
(** much, much easier standard proof **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   398
Goalw [CLIM_def,CRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   399
   "(f -- x --C> L) = ((%y. cmod(f y - L)) -- x --CR> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   400
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   401
qed "CLIM_CRLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   402
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   403
(* so this is nicer nonstandard proof *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   404
Goal "(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   405
by (auto_tac (claset(),simpset() addsimps [CRLIM_NSCRLIM_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   406
    CLIM_CRLIM_iff,CLIM_NSCLIM_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   407
qed "NSCLIM_NSCRLIM_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   408
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   409
Goal "(f -- a --NSC> L) = ((%x. Re(f x)) -- a --NSCR> Re(L) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   410
\                           (%x. Im(f x)) -- a --NSCR> Im(L))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   411
by (auto_tac (claset() addIs [NSCLIM_NSCRLIM_Re,NSCLIM_NSCRLIM_Im],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   412
by (auto_tac (claset(),simpset() addsimps [NSCLIM_def,NSCRLIM_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   413
by (REPEAT(dtac spec 1) THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   414
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   415
by (auto_tac (claset(),simpset() addsimps [capprox_approx_iff,starfunC,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   416
    hcomplex_of_complex_def,starfunCR,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   417
qed "NSCLIM_NSCRLIM_Re_Im_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   418
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   419
Goal "(f -- a --C> L) = ((%x. Re(f x)) -- a --CR> Re(L) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   420
\                        (%x. Im(f x)) -- a --CR> Im(L))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   421
by (auto_tac (claset(),simpset() addsimps [CLIM_NSCLIM_iff,CRLIM_NSCRLIM_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   422
    NSCLIM_NSCRLIM_Re_Im_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   423
qed "CLIM_CRLIM_Re_Im_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   424
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   425
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   426
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   427
(* Continuity                                                                         *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   428
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   429
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   430
Goalw [isNSContc_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
      "[| isNSContc f a; y @c= hcomplex_of_complex a |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   432
\           ==> ( *fc* f) y @c= hcomplex_of_complex (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   433
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   434
qed "isNSContcD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   435
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   436
Goalw [isNSContc_def,NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   437
      "isNSContc f a ==> f -- a --NSC> (f a) ";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   438
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   439
qed "isNSContc_NSCLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   440
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   441
Goalw [isNSContc_def,NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   442
      "f -- a --NSC> (f a) ==> isNSContc f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   443
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   444
by (res_inst_tac [("Q","y = hcomplex_of_complex a")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   445
    (excluded_middle RS disjE) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   446
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   447
qed "NSCLIM_isNSContc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   448
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   449
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   450
(* NS continuity can be defined using NS Limit in   *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   451
(* similar fashion to standard def of continuity    *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   452
(* -------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   453
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   454
Goal "(isNSContc f a) = (f -- a --NSC> (f a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   455
by (blast_tac (claset() addIs [isNSContc_NSCLIM,NSCLIM_isNSContc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   456
qed "isNSContc_NSCLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   457
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   458
Goal "(isNSContc f a) = (f -- a --C> (f a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   459
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   460
    [CLIM_NSCLIM_iff,isNSContc_NSCLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   461
qed "isNSContc_CLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   462
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   463
(*** key result for continuity ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   464
Goalw [isContc_def] "(isNSContc f a) = (isContc f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   465
by (rtac isNSContc_CLIM_iff 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   466
qed "isNSContc_isContc_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   467
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   468
Goal "isContc f a ==> isNSContc f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   469
by (etac (isNSContc_isContc_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   470
qed "isContc_isNSContc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   471
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   472
Goal "isNSContc f a ==> isContc f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   473
by (etac (isNSContc_isContc_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   474
qed "isNSContc_isContc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   475
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   476
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   477
(* Alternative definition of continuity             *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   478
(* -------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   479
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   480
Goalw [NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   481
     "(f -- a --NSC> L) = ((%h. f(a + h)) -- 0 --NSC> L)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   482
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   483
by (dres_inst_tac [("x","hcomplex_of_complex a + x")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   484
by (dres_inst_tac [("x","- hcomplex_of_complex a + x")] spec 2);
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
   485
by Safe_tac;
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   486
by (Asm_full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   487
by (rtac ((mem_cinfmal_iff RS iffD2) RS 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   488
    (CInfinitesimal_add_capprox_self RS capprox_sym)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   489
by (rtac (capprox_minus_iff2 RS iffD1) 4);
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
   490
by (asm_full_simp_tac (simpset() addsimps compare_rls@[hcomplex_add_commute]) 3);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   491
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   492
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 4);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   493
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   494
       simpset() addsimps [starfunC, hcomplex_of_complex_def, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   495
              hcomplex_minus, hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   496
qed "NSCLIM_h_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   497
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   498
Goal "(f -- a --NSC> f a) = ((%h. f(a + h)) -- 0 --NSC> f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   499
by (rtac NSCLIM_h_iff 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   500
qed "NSCLIM_isContc_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   501
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   502
Goal "(f -- a --C> f a) = ((%h. f(a + h)) -- 0 --C> f(a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   503
by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff, NSCLIM_isContc_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   504
qed "CLIM_isContc_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   505
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   506
Goalw [isContc_def] "(isContc f x) = ((%h. f(x + h)) -- 0 --C> f(x))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   507
by (simp_tac (simpset() addsimps [CLIM_isContc_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   508
qed "isContc_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   509
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   510
Goal "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) + g(x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   511
by (auto_tac (claset() addIs [capprox_add],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   512
              simpset() addsimps [isNSContc_isContc_iff RS sym, isNSContc_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   513
qed "isContc_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   514
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   515
Goal "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) * g(x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   516
by (auto_tac (claset() addSIs [starfunC_mult_CFinite_capprox],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   517
              simpset() delsimps [starfunC_mult RS sym]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   518
			addsimps [isNSContc_isContc_iff RS sym, isNSContc_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   519
qed "isContc_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   520
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   521
(*** more theorems: note simple proofs ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   522
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   523
Goal "[| isContc f a; isContc g (f a) |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   524
\     ==> isContc (g o f) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   525
by (auto_tac (claset(),simpset() addsimps [isNSContc_isContc_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   526
              isNSContc_def,starfunC_o RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   527
qed "isContc_o";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   528
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   529
Goal "[| isContc f a; isContc g (f a) |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   530
\     ==> isContc (%x. g (f x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   531
by (auto_tac (claset() addDs [isContc_o],simpset() addsimps [o_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   532
qed "isContc_o2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   533
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   534
Goalw [isNSContc_def] "isNSContc f a ==> isNSContc (%x. - f x) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   535
by Auto_tac; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   536
qed "isNSContc_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   537
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   538
Goal "isContc f a ==> isContc (%x. - f x) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   539
by (auto_tac (claset(),simpset() addsimps [isNSContc_isContc_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   540
              isNSContc_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   541
qed "isContc_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   542
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   543
Goalw [isContc_def]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   544
      "[| isContc f x; f x ~= 0 |] ==> isContc (%x. inverse (f x)) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   545
by (blast_tac (claset() addIs [CLIM_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   546
qed "isContc_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   547
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   548
Goal "[| isNSContc f x; f x ~= 0 |] ==> isNSContc (%x. inverse (f x)) x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   549
by (auto_tac (claset() addIs [isContc_inverse],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   550
    [isNSContc_isContc_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   551
qed "isNSContc_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   552
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   553
Goalw [complex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   554
      "[| isContc f a; isContc g a |] ==> isContc (%x. f(x) - g(x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   555
by (auto_tac (claset() addIs [isContc_add,isContc_minus],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   556
qed "isContc_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   557
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   558
Goalw [isContc_def]  "isContc (%x. k) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   559
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   560
qed "isContc_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
Addsimps [isContc_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   563
Goalw [isNSContc_def]  "isNSContc (%x. k) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   564
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   565
qed "isNSContc_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   566
Addsimps [isNSContc_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   567
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   568
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   569
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   570
(* functions from complex to reals                                                    *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   571
(* ----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   572
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   573
Goalw [isNSContCR_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   574
      "[| isNSContCR f a; y @c= hcomplex_of_complex a |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   575
\           ==> ( *fcR* f) y @= hypreal_of_real (f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   576
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   577
qed "isNSContCRD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   578
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   579
Goalw [isNSContCR_def,NSCRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   580
      "isNSContCR f a ==> f -- a --NSCR> (f a) ";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   581
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   582
qed "isNSContCR_NSCRLIM";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   583
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   584
Goalw [isNSContCR_def,NSCRLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   585
      "f -- a --NSCR> (f a) ==> isNSContCR f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   586
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   587
by (res_inst_tac [("Q","y = hcomplex_of_complex a")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   588
    (excluded_middle RS disjE) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   589
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   590
qed "NSCRLIM_isNSContCR";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   591
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   592
Goal "(isNSContCR f a) = (f -- a --NSCR> (f a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   593
by (blast_tac (claset() addIs [isNSContCR_NSCRLIM,NSCRLIM_isNSContCR]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   594
qed "isNSContCR_NSCRLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   595
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   596
Goal "(isNSContCR f a) = (f -- a --CR> (f a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   597
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   598
    [CRLIM_NSCRLIM_iff,isNSContCR_NSCRLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   599
qed "isNSContCR_CRLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   600
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   601
(*** another key result for continuity ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   602
Goalw [isContCR_def] "(isNSContCR f a) = (isContCR f a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   603
by (rtac isNSContCR_CRLIM_iff 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   604
qed "isNSContCR_isContCR_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   605
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   606
Goal "isContCR f a ==> isNSContCR f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   607
by (etac (isNSContCR_isContCR_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   608
qed "isContCR_isNSContCR";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   609
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   610
Goal "isNSContCR f a ==> isContCR f a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   611
by (etac (isNSContCR_isContCR_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   612
qed "isNSContCR_isContCR";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   613
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   614
Goalw [isNSContCR_def]  "isNSContCR cmod (a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   615
by (auto_tac (claset() addIs [capprox_hcmod_approx],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   616
    simpset() addsimps [starfunCR_cmod,hcmod_hcomplex_of_complex
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   617
    RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   618
qed "isNSContCR_cmod";    
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   619
Addsimps [isNSContCR_cmod];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   620
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   621
Goal "isContCR cmod (a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   622
by (auto_tac (claset(),simpset() addsimps [isNSContCR_isContCR_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   623
qed "isContCR_cmod";    
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   624
Addsimps [isContCR_cmod];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   625
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   626
Goalw [isContc_def,isContCR_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   627
  "isContc f a ==> isContCR (%x. Re (f x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   628
by (etac CLIM_CRLIM_Re 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   629
qed "isContc_isContCR_Re"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   630
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   631
Goalw [isContc_def,isContCR_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   632
  "isContc f a ==> isContCR (%x. Im (f x)) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   633
by (etac CLIM_CRLIM_Im 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   634
qed "isContc_isContCR_Im"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   635
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   636
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   637
(* Derivatives                                                                        *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   638
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   639
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   640
Goalw [cderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   641
      "(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   642
by (Blast_tac 1);        
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   643
qed "CDERIV_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   644
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   645
Goalw [cderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   646
      "(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   647
by (simp_tac (simpset() addsimps [CLIM_NSCLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   648
qed "CDERIV_NSC_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   649
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   650
Goalw [cderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   651
      "CDERIV f x :> D \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   652
\      ==> (%h. (f(x + h) - f(x))/h) -- 0 --C> D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   653
by (Blast_tac 1);        
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   654
qed "CDERIVD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   655
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   656
Goalw [cderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   657
      "CDERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NSC> D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   658
by (asm_full_simp_tac (simpset() addsimps [CLIM_NSCLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   659
qed "NSC_DERIVD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   660
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   661
(*** Uniqueness ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   662
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   663
Goalw [cderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   664
      "[| CDERIV f x :> D; CDERIV f x :> E |] ==> D = E";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   665
by (blast_tac (claset() addIs [CLIM_unique]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   666
qed "CDERIV_unique";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   667
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   668
(*** uniqueness: a nonstandard proof ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   669
Goalw [nscderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   670
     "[| NSCDERIV f x :> D; NSCDERIV f x :> E |] ==> D = E";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   671
by (auto_tac (claset() addSDs [inst "x" "hcomplex_of_hypreal epsilon" bspec] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   672
                       addSIs [inj_hcomplex_of_complex RS injD] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   673
                       addDs [capprox_trans3],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   674
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   675
qed "NSCDeriv_unique";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   676
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   677
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   678
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   679
(* Differentiability                                                                  *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   680
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   681
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   682
Goalw [cdifferentiable_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   683
      "f cdifferentiable x ==> EX D. CDERIV f x :> D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   684
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   685
qed "cdifferentiableD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   686
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   687
Goalw [cdifferentiable_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   688
      "CDERIV f x :> D ==> f cdifferentiable x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   689
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   690
qed "cdifferentiableI";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   691
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   692
Goalw [NSCdifferentiable_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   693
      "f NSCdifferentiable x ==> EX D. NSCDERIV f x :> D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   694
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   695
qed "NSCdifferentiableD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   696
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   697
Goalw [NSCdifferentiable_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   698
      "NSCDERIV f x :> D ==> f NSCdifferentiable x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   699
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   700
qed "NSCdifferentiableI";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   701
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   702
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   703
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   704
(* Alternative definition for differentiability                                       *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   705
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   706
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   707
Goalw [CLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   708
 "((%h. (f(a + h) - f(a))/h) -- 0 --C> D) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   709
\ ((%x. (f(x) - f(a)) / (x - a)) -- a --C> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   710
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   711
by (ALLGOALS(dtac spec));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   712
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   713
by (Blast_tac 1 THEN Blast_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   714
by (ALLGOALS(res_inst_tac [("x","s")] exI));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   715
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   716
by (dres_inst_tac [("x","x - a")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   717
by (dres_inst_tac [("x","x + a")] spec 2);
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   718
by (auto_tac (claset(), simpset() addsimps add_ac));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   719
qed "CDERIV_CLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   720
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   721
Goalw [cderiv_def] "(CDERIV f x :> D) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   722
\         ((%z. (f(z) - f(x)) / (z - x)) -- x --C> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   723
by (simp_tac (simpset() addsimps [CDERIV_CLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   724
qed "CDERIV_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   725
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   726
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   727
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   728
(* Equivalence of NS and standard defs of differentiation                             *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   729
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   730
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   731
(*** first equivalence ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   732
Goalw [nscderiv_def,NSCLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   733
      "(NSCDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   734
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   735
by (dres_inst_tac [("x","xa")] bspec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   736
by (rtac ccontr 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   737
by (dres_inst_tac [("x","h")] spec 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   738
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   739
              simpset() addsimps [mem_cinfmal_iff, starfunC_lambda_cancel]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   740
qed "NSCDERIV_NSCLIM_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   741
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   742
(*** 2nd equivalence ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   743
Goal "(NSCDERIV f x :> D) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   744
\         ((%z. (f(z) - f(x)) / (z - x)) -- x --NSC> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   745
by (full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   746
     [NSCDERIV_NSCLIM_iff, CDERIV_CLIM_iff, CLIM_NSCLIM_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   747
qed "NSCDERIV_NSCLIM_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   748
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   749
Goal "(NSCDERIV f x :> D) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   750
\     (ALL xa. xa ~= hcomplex_of_complex x & xa @c= hcomplex_of_complex x --> \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   751
\       ( *fc* (%z. (f z - f x) / (z - x))) xa @c= hcomplex_of_complex D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   752
by (auto_tac (claset(), simpset() addsimps [NSCDERIV_NSCLIM_iff2, NSCLIM_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   753
qed "NSCDERIV_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   754
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   755
Goalw [cderiv_def] "(NSCDERIV f x :> D) = (CDERIV f x :> D)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   756
by (simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,CLIM_NSCLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   757
qed "NSCDERIV_CDERIV_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   758
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   759
Goalw [nscderiv_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   760
      "NSCDERIV f x :> D ==> isNSContc f x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   761
by (auto_tac (claset(),simpset() addsimps [isNSContc_NSCLIM_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   762
    NSCLIM_def,hcomplex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   763
by (dtac (capprox_minus_iff RS iffD1) 1);
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
   764
by (subgoal_tac "xa + - (hcomplex_of_complex x) ~= 0" 1);
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
   765
 by (asm_full_simp_tac (simpset() addsimps compare_rls) 2);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   766
by (dres_inst_tac [("x","- hcomplex_of_complex x + xa")] bspec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   767
by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   768
by (auto_tac (claset(),simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   769
    [mem_cinfmal_iff RS sym,hcomplex_add_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   770
by (dres_inst_tac [("c","xa + - hcomplex_of_complex x")] capprox_mult1 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   771
by (auto_tac (claset() addIs [CInfinitesimal_subset_CFinite
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   772
    RS subsetD],simpset() addsimps [mult_assoc]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   773
by (dres_inst_tac [("x3","D")] (CFinite_hcomplex_of_complex RSN
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   774
    (2,CInfinitesimal_CFinite_mult) RS (mem_cinfmal_iff RS iffD1)) 1);
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   775
by (blast_tac (claset() addIs [capprox_trans,mult_commute RS subst,
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   776
    (capprox_minus_iff RS iffD2)]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   777
qed "NSCDERIV_isNSContc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   778
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   779
Goal "CDERIV f x :> D ==> isContc f x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   780
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   781
    isNSContc_isContc_iff RS sym,NSCDERIV_isNSContc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   782
qed "CDERIV_isContc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   783
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   784
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   785
(* Differentiation rules for combinations of functions follow from clear,             *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   786
(* straightforard, algebraic manipulations                                            *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   787
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   788
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   789
(* use simple constant nslimit theorem *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   790
Goal "(NSCDERIV (%x. k) x :> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   791
by (simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   792
qed "NSCDERIV_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   793
Addsimps [NSCDERIV_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   794
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   795
Goal "(CDERIV (%x. k) x :> 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   796
by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   797
qed "CDERIV_const";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   798
Addsimps [CDERIV_const];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   799
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   800
Goal "[| NSCDERIV f x :> Da;  NSCDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   801
\     ==> NSCDERIV (%x. f x + g x) x :> Da + Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   802
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   803
           NSCLIM_def]) 1 THEN REPEAT(Step_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   804
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   805
       simpset() addsimps [hcomplex_add_divide_distrib,hcomplex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   806
by (dres_inst_tac [("b","hcomplex_of_complex Da"),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   807
                   ("d","hcomplex_of_complex Db")] capprox_add 1);
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   808
by (auto_tac (claset(), simpset() addsimps add_ac));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   809
qed "NSCDERIV_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   810
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   811
Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   812
\     ==> CDERIV (%x. f x + g x) x :> Da + Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   813
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   814
                                     NSCDERIV_CDERIV_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   815
qed "CDERIV_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   816
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   817
(*** lemmas for multiplication ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   818
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   819
Goal "((a::hcomplex)*b) - (c*d) = (b*(a - c)) + (c*(b - d))";
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   820
by (simp_tac (simpset() addsimps [right_diff_distrib]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   821
val lemma_nscderiv1 = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   822
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   823
Goal "[| (x + y) / z = hcomplex_of_complex D + yb; z ~= 0; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   824
\        z : CInfinitesimal; yb : CInfinitesimal |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   825
\     ==> x + y @c= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   826
by (forw_inst_tac [("c1","z")] (hcomplex_mult_right_cancel RS iffD2) 1 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   827
    THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   828
by (thin_tac "(x + y) / z = hcomplex_of_complex D + yb" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   829
by (auto_tac (claset() addSIs [CInfinitesimal_CFinite_mult2, CFinite_add],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   830
              simpset() addsimps [mem_cinfmal_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   831
by (etac (CInfinitesimal_subset_CFinite RS subsetD) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   832
val lemma_nscderiv2 = result();
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   833
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   834
Goal "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   835
\     ==> NSCDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   836
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff, NSCLIM_def]) 1 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   837
    THEN REPEAT(Step_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   838
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   839
       simpset() addsimps [starfunC_lambda_cancel, lemma_nscderiv1,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   840
       hcomplex_of_complex_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   841
by (simp_tac (simpset() addsimps [hcomplex_add_divide_distrib]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   842
by (REPEAT(dtac (bex_CInfinitesimal_iff2 RS iffD2) 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   843
by (auto_tac (claset(),
14318
7dbd3988b15b type hcomplex is now in class field
paulson
parents: 13957
diff changeset
   844
        simpset() delsimps [times_divide_eq_right]
7dbd3988b15b type hcomplex is now in class field
paulson
parents: 13957
diff changeset
   845
		  addsimps [times_divide_eq_right RS sym]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   846
by (rewtac hcomplex_diff_def);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   847
by (dres_inst_tac [("D","Db")] lemma_nscderiv2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   848
by (dtac (capprox_minus_iff RS iffD2 RS (bex_CInfinitesimal_iff2 RS iffD2)) 4);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   849
by (auto_tac (claset() addSIs [capprox_add_mono1],
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   850
     simpset() addsimps [left_distrib, right_distrib, mult_commute, add_assoc]));
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   851
by (res_inst_tac [("b1","hcomplex_of_complex Db * hcomplex_of_complex (f x)")]
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   852
    (add_commute RS subst) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   853
by (auto_tac (claset() addSIs [CInfinitesimal_add_capprox_self2 RS capprox_sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   854
			       CInfinitesimal_add, CInfinitesimal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   855
			       CInfinitesimal_hcomplex_of_complex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   856
			       CInfinitesimal_hcomplex_of_complex_mult2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   857
	      simpset() addsimps [hcomplex_add_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   858
qed "NSCDERIV_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   859
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   860
Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   861
\     ==> CDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   862
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   863
                                           NSCDERIV_CDERIV_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   864
qed "CDERIV_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   865
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   866
Goal "NSCDERIV f x :> D ==> NSCDERIV (%x. c * f x) x :> c*D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   867
by (asm_full_simp_tac 
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14335
diff changeset
   868
    (simpset() addsimps [times_divide_eq_right RS sym, NSCDERIV_NSCLIM_iff,
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   869
                         minus_mult_right, right_distrib RS sym,
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   870
                         complex_diff_def] 
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14335
diff changeset
   871
             delsimps [times_divide_eq_right, minus_mult_right RS sym]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   872
by (etac (NSCLIM_const RS NSCLIM_mult) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   873
qed "NSCDERIV_cmult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   874
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   875
Goal "CDERIV f x :> D ==> CDERIV (%x. c * f x) x :> c*D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   876
by (auto_tac (claset(),simpset() addsimps [NSCDERIV_cmult,NSCDERIV_CDERIV_iff
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   877
    RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   878
qed "CDERIV_cmult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   879
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   880
Goal "NSCDERIV f x :> D ==> NSCDERIV (%x. -(f x)) x :> -D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   881
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,complex_diff_def]) 1);
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   882
by (res_inst_tac [("t","f x")] (minus_minus RS subst) 1);
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
   883
by (asm_simp_tac (simpset() addsimps [minus_add_distrib RS sym] 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   884
                   delsimps [minus_add_distrib, minus_minus]
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   885
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   886
) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   887
by (etac NSCLIM_minus 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   888
qed "NSCDERIV_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   889
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   890
Goal "CDERIV f x :> D ==> CDERIV (%x. -(f x)) x :> -D";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   891
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_minus,NSCDERIV_CDERIV_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   892
qed "CDERIV_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   893
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   894
Goal "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   895
\     ==> NSCDERIV (%x. f x + -g x) x :> Da + -Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   896
by (blast_tac (claset() addDs [NSCDERIV_add,NSCDERIV_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   897
qed "NSCDERIV_add_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   898
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   899
Goal "[| CDERIV f x :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   900
\     ==> CDERIV (%x. f x + -g x) x :> Da + -Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   901
by (blast_tac (claset() addDs [CDERIV_add,CDERIV_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   902
qed "CDERIV_add_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   903
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   904
Goalw [complex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   905
     "[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   906
\     ==> NSCDERIV (%x. f x - g x) x :> Da - Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   907
by (blast_tac (claset() addIs [NSCDERIV_add_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   908
qed "NSCDERIV_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   909
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   910
Goalw [complex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   911
     "[| CDERIV f x :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   912
\      ==> CDERIV (%x. f x - g x) x :> Da - Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   913
by (blast_tac (claset() addIs [CDERIV_add_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   914
qed "CDERIV_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   915
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   916
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   917
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   918
(* Chain rule                                       *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   919
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   920
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   921
(* lemmas *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   922
Goalw [nscderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   923
      "[| NSCDERIV g x :> D; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   924
\         ( *fc* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex(g x);\
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   925
\         xa : CInfinitesimal; xa ~= 0 \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   926
\      |] ==> D = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   927
by (dtac bspec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   928
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   929
qed "NSCDERIV_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   930
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   931
Goalw [nscderiv_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   932
     "[| NSCDERIV f x :> D;  h: CInfinitesimal;  h ~= 0 |]  \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   933
\     ==> ( *fc* f) (hcomplex_of_complex(x) + h) - hcomplex_of_complex(f x) @c= 0";    
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   934
by (asm_full_simp_tac (simpset() addsimps [mem_cinfmal_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   935
by (rtac CInfinitesimal_ratio 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   936
by (rtac capprox_hcomplex_of_complex_CFinite 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   937
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   938
qed "NSCDERIV_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   939
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   940
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   941
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   942
(* from one version of differentiability            *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   943
(*                                                  *)                                   
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   944
(*   f(x) - f(a)                                    *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   945
(* --------------- @= Db                            *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   946
(*     x - a                                        *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   947
(* -------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   948
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   949
Goal "[| NSCDERIV f (g x) :> Da; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   950
\        ( *fc* g) (hcomplex_of_complex(x) + xa) ~= hcomplex_of_complex (g x); \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   951
\        ( *fc* g) (hcomplex_of_complex(x) + xa) @c= hcomplex_of_complex (g x) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   952
\     |] ==> (( *fc* f) (( *fc* g) (hcomplex_of_complex(x) + xa)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   953
\                     - hcomplex_of_complex (f (g x))) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   954
\             / (( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x)) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   955
\            @c= hcomplex_of_complex (Da)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   956
by (auto_tac (claset(),simpset() addsimps [NSCDERIV_NSCLIM_iff2, NSCLIM_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   957
qed "NSCDERIVD1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   958
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   959
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   960
(* from other version of differentiability          *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   961
(*                                                  *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   962
(*  f(x + h) - f(x)                                 *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   963
(* ----------------- @= Db                          *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   964
(*         h                                        *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   965
(*--------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   966
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   967
Goal "[| NSCDERIV g x :> Db; xa: CInfinitesimal; xa ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   968
\     ==> (( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex(g x)) / xa \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   969
\         @c= hcomplex_of_complex (Db)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   970
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   971
    simpset() addsimps [NSCDERIV_NSCLIM_iff, NSCLIM_def, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   972
		mem_cinfmal_iff, starfunC_lambda_cancel]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   973
qed "NSCDERIVD2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   974
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   975
Goal "(z::hcomplex) ~= 0 ==> x*y = (x*inverse(z))*(z*y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   976
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   977
qed "lemma_complex_chain";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   978
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   979
(*** chain rule ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   980
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   981
Goal "[| NSCDERIV f (g x) :> Da; NSCDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   982
\     ==> NSCDERIV (f o g) x :> Da * Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   983
by (asm_simp_tac (simpset() addsimps [NSCDERIV_NSCLIM_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   984
    NSCLIM_def,mem_cinfmal_iff RS sym]) 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   985
by (forw_inst_tac [("f","g")] NSCDERIV_capprox 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   986
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   987
              simpset() addsimps [starfunC_lambda_cancel2, starfunC_o RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   988
by (case_tac "( *fc* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex (g x)" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   989
by (dres_inst_tac [("g","g")] NSCDERIV_zero 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   990
by (auto_tac (claset(),simpset() addsimps [hcomplex_divide_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   991
by (res_inst_tac [("z1","( *fc* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x)"),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   992
    ("y1","inverse xa")] (lemma_complex_chain RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   993
by (Asm_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   994
by (rtac capprox_mult_hcomplex_of_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   995
by (fold_tac [hcomplex_divide_def]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   996
by (blast_tac (claset() addIs [NSCDERIVD2]) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   997
by (auto_tac (claset() addSIs [NSCDERIVD1] addIs [capprox_minus_iff RS iffD2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   998
    simpset() addsimps [symmetric hcomplex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   999
qed "NSCDERIV_chain";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1000
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1001
(* standard version *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1002
Goal "[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1003
\     ==> CDERIV (f o g) x :> Da * Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1004
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1005
    NSCDERIV_chain]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1006
qed "CDERIV_chain";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1007
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1008
Goal "[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1009
\     ==> CDERIV (%x. f (g x)) x :> Da * Db";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1010
by (auto_tac (claset() addDs [CDERIV_chain], simpset() addsimps [o_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1011
qed "CDERIV_chain2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1012
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1013
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1014
(* Differentiation of natural number powers                                           *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1015
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1016
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1017
Goal "NSCDERIV (%x. x) x :> 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1018
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1019
     simpset() addsimps [NSCDERIV_NSCLIM_iff,NSCLIM_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1020
qed "NSCDERIV_Id";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1021
Addsimps [NSCDERIV_Id];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1022
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1023
Goal "CDERIV (%x. x) x :> 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1024
by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1025
qed "CDERIV_Id";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1026
Addsimps [CDERIV_Id];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1027
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1028
bind_thm ("isContc_Id", CDERIV_Id RS CDERIV_isContc);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1029
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1030
(*derivative of linear multiplication*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1031
Goal "CDERIV (op * c) x :> c";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1032
by (cut_inst_tac [("c","c"),("x","x")] (CDERIV_Id RS CDERIV_cmult) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1033
by (Asm_full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1034
qed "CDERIV_cmult_Id";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1035
Addsimps [CDERIV_cmult_Id];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1036
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1037
Goal "NSCDERIV (op * c) x :> c";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1038
by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1039
qed "NSCDERIV_cmult_Id";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1040
Addsimps [NSCDERIV_cmult_Id];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1041
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1042
Goal "CDERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - 1))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1043
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1044
by (dtac (CDERIV_Id RS CDERIV_mult) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1045
by (auto_tac (claset(), 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1046
              simpset() addsimps [complex_of_real_add RS sym,
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1047
                        left_distrib,real_of_nat_Suc] 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1048
                 delsimps [complex_of_real_add]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1049
by (case_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1050
by (auto_tac (claset(), 
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1051
              simpset() addsimps [mult_assoc, add_commute]));
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1052
by (auto_tac (claset(),simpset() addsimps [mult_commute]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1053
qed "CDERIV_pow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1054
Addsimps [CDERIV_pow,simplify (simpset()) CDERIV_pow];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1055
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1056
(* NS version *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1057
Goal "NSCDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1058
by (simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1059
qed "NSCDERIV_pow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1060
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
  1061
Goal "[|CDERIV f x :> D; D = E|] ==> CDERIV f x :> E";
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1062
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1063
qed "lemma_CDERIV_subst";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1064
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1065
(*used once, in NSCDERIV_inverse*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1066
Goal "[| h: CInfinitesimal; x ~= 0 |] ==> hcomplex_of_complex x + h ~= 0";
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
  1067
by (Clarify_tac 1);
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
  1068
by (dtac (thm"equals_zero_I") 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1069
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1070
qed "CInfinitesimal_add_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1071
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1072
(*Can't get rid of x ~= 0 because it isn't continuous at zero*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1073
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1074
Goalw [nscderiv_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1075
     "x ~= 0 ==> NSCDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1076
by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1077
by (forward_tac [CInfinitesimal_add_not_zero] 1);
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 14320
diff changeset
  1078
by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_commute,numeral_2_eq_2]) 2); 
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1079
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1080
     simpset() addsimps [starfunC_inverse_inverse,hcomplex_diff_def] 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1081
               delsimps [minus_mult_left RS sym, minus_mult_right RS sym]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1082
by (asm_simp_tac
14320
fb7a114826be tidying up hcomplex arithmetic
paulson
parents: 14318
diff changeset
  1083
     (simpset() addsimps [inverse_add,
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1084
          inverse_mult_distrib RS sym, inverse_minus_eq RS sym] 
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1085
          @ add_ac @ mult_ac 
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14323
diff changeset
  1086
       delsimps [inverse_minus_eq,
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1087
		 inverse_mult_distrib, minus_mult_right RS sym, minus_mult_left RS sym] ) 1);
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1088
by (asm_simp_tac (HOL_ss addsimps [mult_assoc RS sym, right_distrib]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1089
by (res_inst_tac [("y"," inverse(- hcomplex_of_complex x * hcomplex_of_complex x)")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1090
                 capprox_trans 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1091
by (rtac inverse_add_CInfinitesimal_capprox2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1092
by (auto_tac (claset() addSDs [hcomplex_of_complex_CFinite_diff_CInfinitesimal] addIs [CFinite_mult], 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1093
         simpset() addsimps [inverse_minus_eq RS sym]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1094
by (rtac CInfinitesimal_CFinite_mult2 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1095
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1096
qed "NSCDERIV_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1097
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1098
Goal "x ~= 0 ==> CDERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1099
by (asm_simp_tac (simpset() addsimps [NSCDERIV_inverse,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1100
         NSCDERIV_CDERIV_iff RS sym] delsimps [complexpow_Suc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1101
qed "CDERIV_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1102
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1103
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1104
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1105
(* Derivative of inverse                                                              *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1106
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1107
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1108
Goal "[| CDERIV f x :> d; f(x) ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1109
\     ==> CDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1110
by (rtac (mult_commute RS subst) 1);
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1111
by (asm_simp_tac (simpset() addsimps [minus_mult_left,
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1112
    power_inverse] delsimps [complexpow_Suc, minus_mult_left RS sym]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1113
by (fold_goals_tac [o_def]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1114
by (blast_tac (claset() addSIs [CDERIV_chain,CDERIV_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1115
qed "CDERIV_inverse_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1116
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1117
Goal "[| NSCDERIV f x :> d; f(x) ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1118
\     ==> NSCDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1119
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1120
            CDERIV_inverse_fun] delsimps [complexpow_Suc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1121
qed "NSCDERIV_inverse_fun";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1122
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1123
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1124
(* Derivative of quotient                                                             *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1125
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1126
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1127
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1128
Goal "x ~= (0::complex) \\<Longrightarrow> (x * inverse(x) ^ 2) = inverse x";
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 14320
diff changeset
  1129
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2]));
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1130
qed "lemma_complex_mult_inverse_squared";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1131
Addsimps [lemma_complex_mult_inverse_squared];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1132
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1133
Goalw [complex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1134
     "[| CDERIV f x :> d; CDERIV g x :> e; g(x) ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1135
\      ==> CDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1136
by (dres_inst_tac [("f","g")] CDERIV_inverse_fun 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1137
by (dtac CDERIV_mult 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1138
by (REPEAT(assume_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1139
by (asm_full_simp_tac
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1140
    (simpset() addsimps [complex_divide_def, right_distrib,
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1141
                         power_inverse,minus_mult_left] @ mult_ac 
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1142
       delsimps [complexpow_Suc, minus_mult_right RS sym, minus_mult_left RS sym]) 1);
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1143
qed "CDERIV_quotient";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1144
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1145
Goal "[| NSCDERIV f x :> d; NSCDERIV g x :> e; g(x) ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1146
\      ==> NSCDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1147
by (asm_full_simp_tac (simpset() addsimps [NSCDERIV_CDERIV_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1148
            CDERIV_quotient] delsimps [complexpow_Suc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1149
qed "NSCDERIV_quotient";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1150
 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1151
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1152
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1153
(* Caratheodory formulation of derivative at a point: standard proof                  *)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
  1154
(*-----------------------------------------------------------------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1155
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1156
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1157
Goalw [CLIM_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1158
      "[| ALL x. x ~= a --> (f x = g x) |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1159
\           ==> (f -- a --C> l) = (g -- a --C> l)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1160
by (auto_tac (claset(), simpset() addsimps [complex_add_minus_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1161
qed "CLIM_equal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1162
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1163
Goal "[| (%x. f(x) + -g(x)) -- a --C> 0; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1164
\        g -- a --C> l |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1165
\      ==> f -- a --C> l";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1166
by (dtac CLIM_add 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1167
by (auto_tac (claset(), simpset() addsimps [complex_add_assoc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1168
qed "CLIM_trans";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1169
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1170
Goal "(CDERIV f x :> l) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1171
\     (EX g. (ALL z. f z - f x = g z * (z - x)) & isContc g x & g x = l)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1172
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1173
by (res_inst_tac 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1174
    [("x","%z. if  z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14365
diff changeset
  1175
by (auto_tac (claset(),simpset() addsimps [mult_assoc,
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1176
    CLAIM "z ~= x ==> z - x ~= (0::complex)"]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1177
by (auto_tac (claset(),simpset() addsimps [isContc_iff,CDERIV_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1178
by (ALLGOALS(rtac (CLIM_equal RS iffD1)));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1179
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1180
qed "CARAT_CDERIV";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1181
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1182
Goal "NSCDERIV f x :> l ==> \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1183
\     EX g. (ALL z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1184
by (auto_tac (claset(),simpset() addsimps [NSCDERIV_CDERIV_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1185
    isNSContc_isContc_iff,CARAT_CDERIV]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1186
qed "CARAT_NSCDERIV";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1187
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1188
(* How about a NS proof? *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1189
Goal "(ALL z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1190
\     ==> NSCDERIV f x :> l";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1191
by (auto_tac (claset(), 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14373
diff changeset
  1192
              simpset() delsimprocs field_cancel_factor
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1193
                        addsimps [NSCDERIV_iff2]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1194
by (asm_full_simp_tac (simpset() addsimps [isNSContc_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1195
qed "CARAT_CDERIVD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1196