src/HOL/Complex/NSCA.ML
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 13957 10dbf16be15f
child 14299 0b5c0b0a3eba
permissions -rw-r--r--
HOL-Real -> HOL-Complex
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       : NSCA.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,2002 University of Edinburgh
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Description : Infinite, infinitesimal complex number etc! 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     5
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     6
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     7
(*--------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     8
(* Closure laws for members of (embedded) set standard complex SComplex                 *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     9
(* -------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
Goalw [SComplex_def] "[| (x::hcomplex): SComplex; y: SComplex |] ==> x + y: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
by (res_inst_tac [("x","r + ra")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
qed "SComplex_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
Goalw [SComplex_def] "[| (x::hcomplex): SComplex; y: SComplex |] ==> x * y: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
by (res_inst_tac [("x","r * ra")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
qed "SComplex_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
Goalw [SComplex_def] "x: SComplex ==> inverse x : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
by (blast_tac (claset() addIs [hcomplex_of_complex_inverse RS sym]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
qed "SComplex_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
Goal "[| x: SComplex;  y: SComplex |] ==> x/y: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
by (asm_simp_tac (simpset() addsimps [SComplex_mult,SComplex_inverse,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
                                      hcomplex_divide_def]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
qed "SComplex_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    31
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    32
Goalw [SComplex_def] "x: SComplex ==> -x : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
by (blast_tac (claset() addIs [hcomplex_of_complex_minus RS sym]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
qed "SComplex_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
Goal "(-x : SComplex) = (x: SComplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
by (etac SComplex_minus 2); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
by (dtac SComplex_minus 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
qed "SComplex_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
Addsimps [SComplex_minus_iff]; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
Goal "[| x + y : SComplex; y: SComplex |] ==> x: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
by (dres_inst_tac [("x","y")] SComplex_minus 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
by (dtac SComplex_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
by (assume_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
qed "SComplex_add_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
Goalw [hcomplex_of_complex_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
     "hcmod (hcomplex_of_complex r) : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
by (simp_tac (simpset() addsimps [hcmod,SReal_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
    hypreal_of_real_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
qed "SReal_hcmod_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
Addsimps [SReal_hcmod_hcomplex_of_complex];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
Goalw [hcomplex_number_of_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
    "hcmod (number_of w ::hcomplex) : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
by (rtac SReal_hcmod_hcomplex_of_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
qed "SReal_hcmod_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
Addsimps [SReal_hcmod_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
Goalw [SComplex_def] "x: SComplex ==> hcmod x : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
qed "SReal_hcmod_SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
Goalw [SComplex_def] "hcomplex_of_complex x: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
qed "SComplex_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
Addsimps [SComplex_hcomplex_of_complex];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    72
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    73
Goalw [hcomplex_number_of_def] "(number_of w ::hcomplex) : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
by (rtac SComplex_hcomplex_of_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
qed "SComplex_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
Addsimps [SComplex_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
Goalw [hcomplex_divide_def] "r : SComplex ==> r/(number_of w::hcomplex) : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
by (blast_tac (claset() addSIs [SComplex_number_of, SComplex_mult, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
                                SComplex_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
qed "SComplex_divide_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
Goalw [SComplex_def] "{x. hcomplex_of_complex x : SComplex} = (UNIV::complex set)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
qed "SComplex_UNIV_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    87
Goalw [SComplex_def] "(x: SComplex) = (EX y. x = hcomplex_of_complex y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
qed "SComplex_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
Goalw [SComplex_def] "hcomplex_of_complex `(UNIV::complex set) = SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
qed "hcomplex_of_complex_image";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
Goalw [SComplex_def] "inv hcomplex_of_complex `SComplex = (UNIV::complex set)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
by (rtac (inj_hcomplex_of_complex RS inv_f_f RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    98
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    99
qed "inv_hcomplex_of_complex_image";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
Goalw [SComplex_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
      "[| EX x. x: P; P <= SComplex |] ==> EX Q. P = hcomplex_of_complex ` Q";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
by (Best_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
qed "SComplex_hcomplex_of_complex_image";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
Goal "[| (x::hcomplex): SComplex; y: SComplex; hcmod x < hcmod y \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
\     |] ==> EX r: Reals. hcmod x< r & r < hcmod y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
by (auto_tac (claset() addIs [SReal_dense], simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
    addsimps [SReal_hcmod_SComplex]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
qed "SComplex_SReal_dense";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
Goalw [SComplex_def,SReal_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
      "z : SComplex ==> hcmod z : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
    hcomplex_of_complex_def,cmod_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
by (res_inst_tac [("x","cmod r")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
by (ultra_tac (claset(),simpset() addsimps [cmod_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   119
qed "SComplex_hcmod_SReal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   120
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
Goal "0 : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
by (auto_tac (claset(),simpset() addsimps [SComplex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   123
qed "SComplex_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
Addsimps [SComplex_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   125
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
Goal "1 : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
by (auto_tac (claset(),simpset() addsimps [SComplex_def,hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
    hcomplex_one_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   129
qed "SComplex_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   130
Addsimps [SComplex_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   132
(*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   133
Goalw [SComplex_def,SReal_def] "hcmod z : Reals ==> z : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   134
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   135
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
    hcomplex_of_complex_def,cmod_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   137
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   138
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   139
(*--------------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
(*        Set of finite elements is a subring of the extended complex numbers                 *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
(* -------------------------------------------------------------------------------------------*)
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
Goalw [CFinite_def] "[|x : CFinite; y : CFinite|] ==> (x+y) : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
by (blast_tac (claset() addSIs [SReal_add,hcmod_add_less]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
qed "CFinite_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   147
Goalw [CFinite_def] "[|x : CFinite; y : CFinite|] ==> x*y : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   148
by (blast_tac (claset() addSIs [SReal_mult,hcmod_mult_less]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   149
qed "CFinite_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   150
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   151
Goalw [CFinite_def] "(-x : CFinite) = (x : CFinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   152
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   153
qed "CFinite_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   154
Addsimps [CFinite_minus_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   155
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   156
Goalw [SComplex_def,CFinite_def] "SComplex <= CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   157
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   158
by (res_inst_tac [("x","1 + hcmod(hcomplex_of_complex r)")] bexI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   159
by (auto_tac (claset() addIs [SReal_add],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   160
qed "SComplex_subset_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   161
Addsimps [ SComplex_subset_CFinite];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   162
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   163
Goal "hcmod (hcomplex_of_complex r) : HFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   164
by (auto_tac (claset() addSIs [ SReal_subset_HFinite RS subsetD],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   165
qed "HFinite_hcmod_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   166
Addsimps [HFinite_hcmod_hcomplex_of_complex];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   167
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   168
Goal "hcomplex_of_complex x: CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   169
by (auto_tac (claset() addSIs [ SComplex_subset_CFinite RS subsetD],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   170
qed "CFinite_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   171
Addsimps [CFinite_hcomplex_of_complex];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   172
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   173
Goalw [CFinite_def] "x : CFinite ==> EX t: Reals. hcmod x < t";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   174
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   175
qed "CFiniteD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   176
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   177
Goalw [CFinite_def] "(x : CFinite) = (hcmod x : HFinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   178
by (auto_tac (claset(), simpset() addsimps [HFinite_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   179
qed "CFinite_hcmod_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   180
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   181
Goal "number_of w : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   182
by (rtac (SComplex_number_of RS (SComplex_subset_CFinite RS subsetD)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   183
qed "CFinite_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   184
Addsimps [CFinite_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   185
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   186
Goal "[|x : CFinite; y <= hcmod x; 0 <= y |] ==> y: HFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   187
by (auto_tac (claset() addIs [HFinite_bounded],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   188
    [CFinite_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   189
qed "CFinite_bounded";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   190
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   191
(*--------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   192
(* Set of complex infinitesimals is a subring of the nonstandard complex numbers        *) 
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
	 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   195
Goalw [CInfinitesimal_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   196
      "x : CInfinitesimal ==> ALL r: Reals. 0 < r --> hcmod x < r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   197
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   198
qed "CInfinitesimalD";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   200
Goalw [CInfinitesimal_def] "0 : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   201
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   202
qed "CInfinitesimal_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   203
AddIffs [CInfinitesimal_zero];
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 "x/(2::hcomplex) + x/(2::hcomplex) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   206
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   207
qed "hcomplex_sum_of_halves";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   208
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   209
Goalw [CInfinitesimal_def,Infinitesimal_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   210
   "(z : CInfinitesimal) = (hcmod z : Infinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   211
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   212
qed "CInfinitesimal_hcmod_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   213
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   214
Goal "1 ~: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   215
by (simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   216
qed "one_not_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   217
Addsimps [one_not_CInfinitesimal];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   218
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   219
Goal "[| x : CInfinitesimal; y : CInfinitesimal |] ==> (x+y) : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   220
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   221
by (rtac hrabs_le_Infinitesimal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   222
by (res_inst_tac [("y","hcmod y")] Infinitesimal_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   223
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   224
qed "CInfinitesimal_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   225
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   226
Goalw [CInfinitesimal_def] "(-x:CInfinitesimal) = (x:CInfinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   227
by (Full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   228
qed "CInfinitesimal_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   229
Addsimps [CInfinitesimal_minus_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   230
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   231
Goal "[| x : CInfinitesimal;  y : CInfinitesimal |] ==> x-y : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   232
by (asm_simp_tac
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   233
    (simpset() addsimps [hcomplex_diff_def, CInfinitesimal_add]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   234
qed "CInfinitesimal_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   235
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   236
Goal "[| x : CInfinitesimal; y : CInfinitesimal |] ==> (x * y) : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   237
by (auto_tac (claset() addIs [Infinitesimal_mult],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   238
    [CInfinitesimal_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   239
qed "CInfinitesimal_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 "[| x : CInfinitesimal; y : CFinite |] ==> (x * y) : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   242
by (auto_tac (claset() addIs [Infinitesimal_HFinite_mult],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   243
    addsimps [CInfinitesimal_hcmod_iff,CFinite_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   244
qed "CInfinitesimal_CFinite_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
Goal "[| x : CInfinitesimal; y : CFinite |] ==> (y * x) : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   247
by (auto_tac (claset() addDs [CInfinitesimal_CFinite_mult],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   248
              simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   249
qed "CInfinitesimal_CFinite_mult2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   250
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   251
Goalw [CInfinite_def,HInfinite_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   252
   "(z : CInfinite) = (hcmod z : HInfinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   253
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   254
qed "CInfinite_hcmod_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   255
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   256
Goal "x: CInfinite ==> inverse x: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   257
by (auto_tac (claset() addIs [HInfinite_inverse_Infinitesimal],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   258
    simpset() addsimps [CInfinitesimal_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   259
    CInfinite_hcmod_iff,hcmod_hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   260
qed "CInfinite_inverse_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   261
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   262
Goal "[|x: CInfinite; y: CInfinite|] ==> (x*y): CInfinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   263
by (auto_tac (claset() addIs [HInfinite_mult],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   264
    [CInfinite_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   265
qed "CInfinite_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   266
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   267
Goalw [CInfinite_def] "(-x : CInfinite) = (x : CInfinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   268
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   269
qed "CInfinite_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   270
Addsimps [CInfinite_minus_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   271
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   272
Goal "[|a: CFinite; b: CFinite; c: CFinite|] \ 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   273
\     ==> a*a + b*b + c*c : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   274
by (auto_tac (claset() addIs [CFinite_mult,CFinite_add], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   275
qed "CFinite_sum_squares";
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 "x ~: CInfinitesimal ==> x ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   278
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   279
qed "not_CInfinitesimal_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   280
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   281
Goal "x: CFinite - CInfinitesimal ==> x ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   282
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   283
qed "not_CInfinitesimal_not_zero2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   284
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   285
Goal "x : CFinite - CInfinitesimal ==> hcmod x : HFinite - Infinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   286
by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   287
qed "CFinite_diff_CInfinitesimal_hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   288
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   289
Goal "[| e : CInfinitesimal; hcmod x < hcmod e |] ==> x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   290
by (auto_tac (claset() addIs [hrabs_less_Infinitesimal],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   291
    addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   292
qed "hcmod_less_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   293
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   294
Goal "[| e : CInfinitesimal; hcmod x <= hcmod e |] ==> x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   295
by (auto_tac (claset() addIs [hrabs_le_Infinitesimal],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   296
    addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   297
qed "hcmod_le_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   298
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   299
Goal  "[| e : CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   300
\         e' : CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   301
\         hcmod e' < hcmod x ; hcmod x < hcmod e \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   302
\      |] ==> x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   303
by (auto_tac (claset() addIs [Infinitesimal_interval],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   304
    addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   305
qed "CInfinitesimal_interval";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   306
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   307
Goal "[| e : CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   308
\        e' : CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   309
\        hcmod e' <= hcmod x ; hcmod x <= hcmod e \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   310
\     |] ==> x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   311
by (auto_tac (claset() addIs [Infinitesimal_interval2],simpset() 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   312
    addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   313
qed "CInfinitesimal_interval2";
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 ~: CInfinitesimal;  y ~: CInfinitesimal|] ==> (x*y) ~: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   316
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   317
by (dtac not_Infinitesimal_mult 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   318
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   319
qed "not_CInfinitesimal_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   320
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   321
Goal "x*y : CInfinitesimal ==> x : CInfinitesimal | y : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   322
by (auto_tac (claset() addDs [Infinitesimal_mult_disj],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   323
    [CInfinitesimal_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   324
qed "CInfinitesimal_mult_disj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   325
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   326
Goal "[| x : CFinite - CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   327
\                  y : CFinite - CInfinitesimal \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   328
\               |] ==> (x*y) : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   329
by (Clarify_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   330
by (blast_tac (claset() addDs [CFinite_mult,not_CInfinitesimal_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   331
qed "CFinite_CInfinitesimal_diff_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   332
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   333
Goal "CInfinitesimal <= CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   334
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   335
    simpset() addsimps [CInfinitesimal_hcmod_iff,CFinite_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   336
qed "CInfinitesimal_subset_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   337
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   338
Goal "x: CInfinitesimal ==> x * hcomplex_of_complex r : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   339
by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   340
    simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   341
qed "CInfinitesimal_hcomplex_of_complex_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   342
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   343
Goal "x: CInfinitesimal ==> hcomplex_of_complex r * x: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   344
by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   345
    simpset() addsimps [CInfinitesimal_hcmod_iff,hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   346
qed "CInfinitesimal_hcomplex_of_complex_mult2";
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   349
(*--------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   350
(* Infinitely close relation @c=                                                        *)
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
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
Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   355
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   356
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   357
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   358
Goal "x:CInfinitesimal = (x @c= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   359
by (simp_tac (simpset() addsimps [CInfinitesimal_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   360
    capprox_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   361
qed "mem_cinfmal_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   362
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   363
Goalw [capprox_def,hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   364
      " (x @c= y) = (x + -y @c= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   365
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   366
qed "capprox_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   367
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   368
Goalw [capprox_def,hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   369
      " (x @c= y) = (-y + x @c= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   370
by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   371
qed "capprox_minus_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   372
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   373
Goalw [capprox_def] "x @c= x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   374
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   375
qed "capprox_refl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   376
Addsimps [capprox_refl];
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 [capprox_def,CInfinitesimal_def]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   379
     "x @c= y ==> y @c= x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   380
by (auto_tac (claset() addSDs [bspec],simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   381
    [hcmod_diff_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   382
qed "capprox_sym";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   383
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   384
Goalw [capprox_def]  "[| x @c= y; y @c= z |] ==> x @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   385
by (dtac CInfinitesimal_add 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   386
by (assume_tac 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   387
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   388
qed "capprox_trans";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   389
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   390
Goal "[| r @c= x; s @c= x |] ==> r @c= s";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   391
by (blast_tac (claset() addIs [capprox_sym, capprox_trans]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   392
qed "capprox_trans2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   393
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   394
Goal "[| x @c= r; x @c= s|] ==> r @c= s";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   395
by (blast_tac (claset() addIs [capprox_sym, capprox_trans]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   396
qed "capprox_trans3";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   397
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   398
Goal "(number_of w @c= x) = (x @c= number_of w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   399
by (blast_tac (claset() addIs [capprox_sym]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   400
qed "number_of_capprox_reorient";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   401
Addsimps [number_of_capprox_reorient];
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
Goal "(x-y : CInfinitesimal) = (x @c= y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   404
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   405
              simpset() addsimps [hcomplex_diff_def, capprox_minus_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   406
                                  mem_cinfmal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   407
qed "CInfinitesimal_capprox_minus";
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
Goalw [cmonad_def] "(x @c= y) = (cmonad(x)=cmonad(y))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   410
by (auto_tac (claset() addDs [capprox_sym] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   411
                       addSEs [capprox_trans,equalityCE],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   412
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   413
qed "capprox_monad_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   414
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   415
Goal "[| x: CInfinitesimal; y: CInfinitesimal |] ==> x @c= y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   416
by (asm_full_simp_tac (simpset() addsimps [mem_cinfmal_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   417
by (blast_tac (claset() addIs [capprox_trans, capprox_sym]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   418
qed "Infinitesimal_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   419
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   420
val prem1::prem2::rest = 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   421
goalw thy [capprox_def,hcomplex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   422
     "[| a @c= b; c @c= d |] ==> a+c @c= b+d";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   423
by (rtac (hcomplex_minus_add_distrib RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   424
by (rtac (hcomplex_add_assoc RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   425
by (res_inst_tac [("y1","c")] (hcomplex_add_left_commute RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   426
by (rtac (hcomplex_add_assoc RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   427
by (rtac ([prem1,prem2] MRS CInfinitesimal_add) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   428
qed "capprox_add";
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
Goal "a @c= b ==> -a @c= -b";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
by (rtac ((capprox_minus_iff RS iffD2) RS capprox_sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   432
by (dtac (capprox_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   433
by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   434
qed "capprox_minus";
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
Goal "-a @c= -b ==> a @c= b";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   437
by (auto_tac (claset() addDs [capprox_minus], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   438
qed "capprox_minus2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   439
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   440
Goal "(-a @c= -b) = (a @c= b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   441
by (blast_tac (claset() addIs [capprox_minus,capprox_minus2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   442
qed "capprox_minus_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   443
Addsimps [capprox_minus_cancel];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   444
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   445
Goal "[| a @c= b; c @c= d |] ==> a + -c @c= b + -d";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   446
by (blast_tac (claset() addSIs [capprox_add,capprox_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   447
qed "capprox_add_minus";
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
Goalw [capprox_def,hcomplex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   450
      "[| a @c= b; c: CFinite|] ==> a*c @c= b*c"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   451
by (asm_full_simp_tac (simpset() addsimps [CInfinitesimal_CFinite_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   452
    hcomplex_minus_mult_eq1,hcomplex_add_mult_distrib RS sym] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   453
    delsimps [hcomplex_minus_mult_eq1 RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   454
qed "capprox_mult1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   455
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   456
Goal "[|a @c= b; c: CFinite|] ==> c*a @c= c*b"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   457
by (asm_simp_tac (simpset() addsimps [capprox_mult1,hcomplex_mult_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   458
qed "capprox_mult2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   459
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   460
Goal "[|u @c= v*x; x @c= y; v: CFinite|] ==> u @c= v*y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   461
by (fast_tac (claset() addIs [capprox_mult2,capprox_trans]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   462
qed "capprox_mult_subst";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   463
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   464
Goal "[| u @c= x*v; x @c= y; v: CFinite |] ==> u @c= y*v";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   465
by (fast_tac (claset() addIs [capprox_mult1,capprox_trans]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   466
qed "capprox_mult_subst2";
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 "[| u @c= x*hcomplex_of_complex v; x @c= y |] ==> u @c= y*hcomplex_of_complex v";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   469
by (auto_tac (claset() addIs [capprox_mult_subst2], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   470
qed "capprox_mult_subst_SComplex";
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
Goalw [capprox_def]  "a = b ==> a @c= b";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   473
by (Asm_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   474
qed "capprox_eq_imp";
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
Goal "x: CInfinitesimal ==> -x @c= x"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   477
by (fast_tac (HOL_cs addIs [CInfinitesimal_minus_iff RS iffD2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   478
    mem_cinfmal_iff RS iffD1,capprox_trans2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   479
qed "CInfinitesimal_minus_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   480
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   481
Goalw [capprox_def]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   482
     "(EX y: CInfinitesimal. x - z = y) = (x @c= z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   483
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   484
qed "bex_CInfinitesimal_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   485
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   486
Goal "(EX y: CInfinitesimal. x = z + y) = (x @c= z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   487
by (asm_full_simp_tac (simpset() addsimps [bex_CInfinitesimal_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   488
by (Force_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   489
qed "bex_CInfinitesimal_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   490
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   491
Goal "[| y: CInfinitesimal; x + y = z |] ==> x @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   492
by (rtac (bex_CInfinitesimal_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   493
by (dtac (CInfinitesimal_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   494
by (auto_tac (claset(), simpset() addsimps [hcomplex_add_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   495
qed "CInfinitesimal_add_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   496
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   497
Goal "y: CInfinitesimal ==> x @c= x + y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   498
by (rtac (bex_CInfinitesimal_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   499
by (dtac (CInfinitesimal_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   500
by (auto_tac (claset(), simpset() addsimps [hcomplex_add_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   501
qed "CInfinitesimal_add_capprox_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   502
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   503
Goal "y: CInfinitesimal ==> x @c= y + x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   504
by (auto_tac (claset() addDs [CInfinitesimal_add_capprox_self],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   505
    simpset() addsimps [hcomplex_add_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   506
qed "CInfinitesimal_add_capprox_self2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   507
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   508
Goal "y: CInfinitesimal ==> x @c= x + -y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   509
by (blast_tac (claset() addSIs [CInfinitesimal_add_capprox_self,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   510
                                CInfinitesimal_minus_iff RS iffD2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   511
qed "CInfinitesimal_add_minus_capprox_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   512
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   513
Goal "[| y: CInfinitesimal; x+y @c= z|] ==> x @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   514
by (dres_inst_tac [("x","x")] (CInfinitesimal_add_capprox_self RS capprox_sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   515
by (etac (capprox_trans3 RS capprox_sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   516
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   517
qed "CInfinitesimal_add_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   518
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   519
Goal "[| y: CInfinitesimal; x @c= z + y|] ==> x @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   520
by (dres_inst_tac [("x","z")] (CInfinitesimal_add_capprox_self2  RS capprox_sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   521
by (etac (capprox_trans3 RS capprox_sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   522
by (asm_full_simp_tac (simpset() addsimps [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   523
by (etac capprox_sym 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   524
qed "CInfinitesimal_add_right_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   525
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   526
Goal "d + b  @c= d + c ==> b @c= c";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   527
by (dtac (capprox_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   528
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   529
    [hcomplex_minus_add_distrib,capprox_minus_iff RS sym] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   530
    @ hcomplex_add_ac) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   531
qed "capprox_add_left_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   532
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   533
Goal "b + d @c= c + d ==> b @c= c";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   534
by (rtac capprox_add_left_cancel 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   535
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   536
    [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   537
qed "capprox_add_right_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   538
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   539
Goal "b @c= c ==> d + b @c= d + c";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   540
by (rtac (capprox_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   541
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   542
    [capprox_minus_iff RS sym] @ hcomplex_add_ac) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   543
qed "capprox_add_mono1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   544
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   545
Goal "b @c= c ==> b + a @c= c + a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   546
by (asm_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   547
    [hcomplex_add_commute,capprox_add_mono1]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   548
qed "capprox_add_mono2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   549
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   550
Goal "(a + b @c= a + c) = (b @c= c)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   551
by (fast_tac (claset() addEs [capprox_add_left_cancel,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   552
    capprox_add_mono1]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   553
qed "capprox_add_left_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   554
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   555
AddIffs [capprox_add_left_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   556
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
Goal "(b + a @c= c + a) = (b @c= c)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   559
by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   560
qed "capprox_add_right_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
AddIffs [capprox_add_right_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   563
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   564
Goal "[| x: CFinite; x @c= y |] ==> y: CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   565
by (dtac (bex_CInfinitesimal_iff2 RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   566
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   567
by (dtac (CInfinitesimal_subset_CFinite RS subsetD 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   568
          RS (CFinite_minus_iff RS iffD2)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   569
by (dtac CFinite_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   570
by (assume_tac 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   571
qed "capprox_CFinite";
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
Goal "x @c= hcomplex_of_complex D ==> x: CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   574
by (rtac (capprox_sym RSN (2,capprox_CFinite)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   575
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   576
qed "capprox_hcomplex_of_complex_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   577
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   578
Goal "[|a @c= b; c @c= d; b: CFinite; d: CFinite|] ==> a*c @c= b*d";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   579
by (rtac capprox_trans 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   580
by (rtac capprox_mult2 2); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   581
by (rtac capprox_mult1 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   582
by (blast_tac (claset() addIs [capprox_CFinite, capprox_sym]) 2);  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   583
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   584
qed "capprox_mult_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   585
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   586
Goal "[|a @c= hcomplex_of_complex b; c @c= hcomplex_of_complex d |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   587
\     ==> a*c @c= hcomplex_of_complex b * hcomplex_of_complex d";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   588
by (blast_tac (claset() addSIs [capprox_mult_CFinite,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   589
            capprox_hcomplex_of_complex_CFinite,CFinite_hcomplex_of_complex]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   590
qed "capprox_mult_hcomplex_of_complex";
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 "[| a: SComplex; a ~= 0; a*x @c= 0 |] ==> x @c= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   593
by (dtac (SComplex_inverse RS (SComplex_subset_CFinite RS subsetD)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   594
by (auto_tac (claset() addDs [capprox_mult2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   595
    simpset() addsimps [hcomplex_mult_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   596
qed "capprox_SComplex_mult_cancel_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   597
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   598
Goal "[| a: SComplex; x @c= 0 |] ==> x*a @c= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   599
by (auto_tac (claset() addDs [(SComplex_subset_CFinite RS subsetD),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   600
              capprox_mult1], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   601
qed "capprox_mult_SComplex1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   602
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   603
Goal "[| a: SComplex; x @c= 0 |] ==> a*x @c= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   604
by (auto_tac (claset() addDs [(SComplex_subset_CFinite RS subsetD),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   605
              capprox_mult2], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   606
qed "capprox_mult_SComplex2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   607
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   608
Goal "[|a : SComplex; a ~= 0 |] ==> (a*x @c= 0) = (x @c= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   609
by (blast_tac (claset() addIs [capprox_SComplex_mult_cancel_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   610
    capprox_mult_SComplex2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   611
qed "capprox_mult_SComplex_zero_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   612
Addsimps [capprox_mult_SComplex_zero_cancel_iff];
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
Goal "[| a: SComplex; a ~= 0; a* w @c= a*z |] ==> w @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   615
by (dtac (SComplex_inverse RS (SComplex_subset_CFinite RS subsetD)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   616
by (auto_tac (claset() addDs [capprox_mult2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   617
    simpset() addsimps [hcomplex_mult_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   618
qed "capprox_SComplex_mult_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   619
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   620
Goal "[| a: SComplex; a ~= 0|] ==> (a* w @c= a*z) = (w @c= z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   621
by (auto_tac (claset() addSIs [capprox_mult2,SComplex_subset_CFinite RS subsetD] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   622
    addIs [capprox_SComplex_mult_cancel], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   623
qed "capprox_SComplex_mult_cancel_iff1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   624
Addsimps [capprox_SComplex_mult_cancel_iff1];
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
Goal "(x @c= y) = (hcmod (y - x) @= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   627
by (rtac (capprox_minus_iff RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   628
by (auto_tac (claset(),simpset() addsimps [capprox_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   629
    CInfinitesimal_hcmod_iff,mem_infmal_iff,symmetric hcomplex_diff_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   630
    hcmod_diff_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   631
qed "capprox_hcmod_approx_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   632
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   633
Goal "(x @c= 0) = (hcmod x @= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   634
by (auto_tac (claset(),simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   635
    [capprox_hcmod_approx_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   636
qed "capprox_approx_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   637
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   638
Goal "(-x @c= 0) = (x @c= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   639
by (auto_tac (claset(),simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   640
    [capprox_hcmod_approx_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   641
qed "capprox_minus_zero_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   642
Addsimps [capprox_minus_zero_cancel_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   643
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   644
Goal "u @c= 0 ==> hcmod(x + u) - hcmod x : Infinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   645
by (res_inst_tac [("e","hcmod u"),("e'","- hcmod u")] Infinitesimal_interval2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   646
by (auto_tac (claset() addDs [capprox_approx_zero_iff RS iffD1], 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   647
    simpset() addsimps [mem_infmal_iff RS sym,hypreal_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   648
by (res_inst_tac [("C","hcmod x")] hypreal_le_add_left_cancel 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   649
by (auto_tac (claset(),simpset() addsimps [symmetric hypreal_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   650
qed "Infinitesimal_hcmod_add_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   651
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   652
Goal "u @c= 0 ==> hcmod(x + u) @= hcmod x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   653
by (rtac (approx_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   654
by (auto_tac (claset() addIs [Infinitesimal_hcmod_add_diff],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   655
    simpset() addsimps [mem_infmal_iff RS sym,symmetric hypreal_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   656
qed "approx_hcmod_add_hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   657
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   658
Goal "x @c= y ==> hcmod x @= hcmod y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   659
by (auto_tac (claset() addIs [approx_hcmod_add_hcmod] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   660
    addSDs [bex_CInfinitesimal_iff2 RS iffD2],simpset() addsimps [mem_cinfmal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   661
qed "capprox_hcmod_approx";
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
(*--------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   664
(* zero is the only complex number that is also infinitesimal                           *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   665
(*--------------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   666
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   667
Goal "[| x: SComplex; y: CInfinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   668
by (auto_tac (claset() addSIs [Infinitesimal_less_SReal,SComplex_hcmod_SReal],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   669
    simpset() addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   670
qed "CInfinitesimal_less_SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   671
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   672
Goal "y: CInfinitesimal ==> ALL r: SComplex. 0 < hcmod r --> hcmod y < hcmod r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   673
by (blast_tac (claset() addIs [CInfinitesimal_less_SComplex]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   674
qed "CInfinitesimal_less_SComplex2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   675
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   676
Goal "SComplex Int CInfinitesimal = {0}";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   677
by (auto_tac (claset(),simpset() addsimps [SComplex_def,CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   678
by (cut_inst_tac [("r","r")] SReal_hcmod_hcomplex_of_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   679
by (dres_inst_tac [("A","Reals")] IntI 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   680
by (subgoal_tac "hcmod (hcomplex_of_complex r) = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   681
by (Asm_full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   682
by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   683
by (rotate_tac 2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   684
by (Asm_full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   685
qed "SComplex_Int_CInfinitesimal_zero";
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
Goal "[| x: SComplex; x: CInfinitesimal|] ==> x = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   688
by (cut_facts_tac [SComplex_Int_CInfinitesimal_zero] 1);
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 "SComplex_CInfinitesimal_zero";
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
Goal "[| x : SComplex; x ~= 0 |] ==> x : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   693
by (auto_tac (claset() addDs [SComplex_CInfinitesimal_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   694
                              SComplex_subset_CFinite RS subsetD], 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   695
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   696
qed "SComplex_CFinite_diff_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   697
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   698
Goal "hcomplex_of_complex x ~= 0 ==> hcomplex_of_complex x : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   699
by (rtac SComplex_CFinite_diff_CInfinitesimal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   700
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   701
qed "hcomplex_of_complex_CFinite_diff_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   702
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   703
Goal "(hcomplex_of_complex x : CInfinitesimal) = (x=0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   704
by (auto_tac (claset(), simpset() addsimps [hcomplex_of_complex_zero]));  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   705
by (rtac ccontr 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   706
by (rtac (hcomplex_of_complex_CFinite_diff_CInfinitesimal RS DiffD2) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   707
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   708
qed "hcomplex_of_complex_CInfinitesimal_iff_0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   709
AddIffs [hcomplex_of_complex_CInfinitesimal_iff_0];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   710
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   711
Goal "number_of w ~= (0::hcomplex) ==> number_of w ~: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   712
by (fast_tac (claset() addDs [SComplex_number_of RS SComplex_CInfinitesimal_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   713
qed "number_of_not_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   714
Addsimps [number_of_not_CInfinitesimal];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   715
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   716
Goal "[| y: SComplex; x @c= y; y~= 0 |] ==> x ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   717
by (auto_tac (claset() addDs [SComplex_CInfinitesimal_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   718
    capprox_sym RS (mem_cinfmal_iff RS iffD2)],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   719
qed "capprox_SComplex_not_zero";
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
Goal "[| x @c= y; y : CFinite - CInfinitesimal |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   722
\     ==> x : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   723
by (auto_tac (claset() addIs [capprox_sym RSN (2,capprox_CFinite)],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   724
              simpset() addsimps [mem_cinfmal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   725
by (dtac capprox_trans3 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   726
by (blast_tac (claset() addDs [capprox_sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   727
qed "CFinite_diff_CInfinitesimal_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   728
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   729
Goal "[| y ~= 0;  y: CInfinitesimal;  x/y : CFinite |] ==> x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   730
by (dtac CInfinitesimal_CFinite_mult2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   731
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   732
by (asm_full_simp_tac 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   733
    (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   734
qed "CInfinitesimal_ratio";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   735
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   736
Goal "[|x: SComplex; y: SComplex|] ==> (x @c= y) = (x = y)"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   737
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   738
by (rewrite_goals_tac [capprox_def]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   739
by (dres_inst_tac [("x","y")] SComplex_minus 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   740
by (dtac SComplex_add 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   741
by (rtac (CLAIM "x - y = 0 ==> x = (y::hcomplex)") 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   742
by (rtac SComplex_CInfinitesimal_zero 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   743
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   744
qed "SComplex_capprox_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   745
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   746
Goal "(number_of v @c= number_of w) = (number_of v = (number_of w :: hcomplex))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   747
by (rtac SComplex_capprox_iff 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   748
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   749
qed "number_of_capprox_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   750
Addsimps [number_of_capprox_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   751
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   752
Goal "(number_of w : CInfinitesimal) = (number_of w = (0::hcomplex))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   753
by (rtac iffI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   754
by (fast_tac (claset() addDs [SComplex_number_of RS SComplex_CInfinitesimal_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   755
by (Asm_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   756
qed "number_of_CInfinitesimal_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   757
Addsimps [number_of_CInfinitesimal_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
Goal "(hcomplex_of_complex k @c= hcomplex_of_complex m) = (k = m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   760
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   761
by (rtac (inj_hcomplex_of_complex RS injD) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   762
by (rtac (SComplex_capprox_iff RS iffD1) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   763
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   764
qed "hcomplex_of_complex_approx_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   765
Addsimps [hcomplex_of_complex_approx_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   766
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   767
Goal "(hcomplex_of_complex k @c= number_of w) = (k = number_of w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   768
by (stac (hcomplex_of_complex_approx_iff RS sym) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   769
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   770
qed "hcomplex_of_complex_capprox_number_of_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   771
Addsimps [hcomplex_of_complex_capprox_number_of_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   772
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   773
Goal "[| r: SComplex; s: SComplex; r @c= x; s @c= x|] ==> r = s";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   774
by (blast_tac (claset() addIs [(SComplex_capprox_iff RS iffD1),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   775
               capprox_trans2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   776
qed "capprox_unique_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   777
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   778
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n}) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   779
\     ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) @= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   780
\         Abs_hypreal(hyprel `` {%n. Re(Y n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   781
by (auto_tac (claset(),simpset() addsimps [approx_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   782
by (dtac (capprox_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   783
by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   784
    mem_cinfmal_iff RS sym,CInfinitesimal_hcmod_iff,hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   785
    Infinitesimal_FreeUltrafilterNat_iff2]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   786
by (dres_inst_tac [("x","m")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   787
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   788
by (res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   789
by (res_inst_tac [("z","Y x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   790
by (auto_tac (claset(),simpset() addsimps [complex_minus,complex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   791
    complex_mod] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   792
by (rtac order_le_less_trans 1 THEN assume_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   793
by (dres_inst_tac [("t","Ya x")] sym 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   794
by (auto_tac (claset(),simpset() addsimps [abs_eqI1] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   795
qed "hcomplex_capproxD1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   796
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   797
(* same proof *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   798
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n}) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   799
\     ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) @= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   800
\         Abs_hypreal(hyprel `` {%n. Im(Y n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   801
by (auto_tac (claset(),simpset() addsimps [approx_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   802
by (dtac (capprox_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   803
by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   804
    mem_cinfmal_iff RS sym,CInfinitesimal_hcmod_iff,hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   805
    Infinitesimal_FreeUltrafilterNat_iff2]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   806
by (dres_inst_tac [("x","m")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   807
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   808
by (res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   809
by (res_inst_tac [("z","Y x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   810
by (auto_tac (claset(),simpset() addsimps [complex_minus,complex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   811
    complex_mod] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   812
by (rtac order_le_less_trans 1 THEN assume_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   813
by (dres_inst_tac [("t","Ya x")] sym 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   814
by (auto_tac (claset(),simpset() addsimps [abs_eqI1] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   815
qed "hcomplex_capproxD2";
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
Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) @= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   818
\        Abs_hypreal(hyprel `` {%n. Re(Y n)}); \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   819
\        Abs_hypreal(hyprel `` {%n. Im(X n)}) @= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   820
\        Abs_hypreal(hyprel `` {%n. Im(Y n)}) \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   821
\     |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   822
by (dtac (approx_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   823
by (dtac (approx_minus_iff RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   824
by (rtac (capprox_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   825
by (auto_tac (claset(),simpset() addsimps [mem_cinfmal_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   826
    mem_infmal_iff RS sym,hypreal_minus,hypreal_add,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   827
    hcomplex_add,CInfinitesimal_hcmod_iff,hcmod,Infinitesimal_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   828
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   829
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   830
by (dres_inst_tac [("x","u/2")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   831
by (dres_inst_tac [("x","u/2")] spec 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   832
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   833
by (TRYALL(Force_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   834
by (ultra_tac (claset(),HOL_ss) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   835
by (dtac sym 1 THEN dtac sym 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   836
by (res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   837
by (res_inst_tac [("z","Y x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   838
by (auto_tac (claset(),HOL_ss addsimps [complex_minus,complex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   839
    complex_mod,pair_mem_complex RS Abs_complex_inverse,snd_conv,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   840
    fst_conv,two_eq_Suc_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   841
by (rtac (realpow_two_abs RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   842
by (res_inst_tac [("x1","xa + - xb")] (realpow_two_abs RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   843
by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   844
by (rtac lemma_sqrt_hcomplex_capprox 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   845
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   846
qed "hcomplex_capproxI";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   847
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   848
Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) @c= Abs_hcomplex(hcomplexrel``{%n. Y n})) =\
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   849
\      (Abs_hypreal(hyprel `` {%n. Re(X n)}) @= Abs_hypreal(hyprel `` {%n. Re(Y n)}) & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   850
\       Abs_hypreal(hyprel `` {%n. Im(X n)}) @= Abs_hypreal(hyprel `` {%n. Im(Y n)}))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   851
by (blast_tac (claset() addIs [hcomplex_capproxI,hcomplex_capproxD1,hcomplex_capproxD2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   852
qed "capprox_approx_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   853
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   854
Goal "(hcomplex_of_hypreal x @c= hcomplex_of_hypreal z) = (x @= z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   855
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   856
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   857
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   858
    capprox_approx_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   859
qed "hcomplex_of_hypreal_capprox_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   860
Addsimps [hcomplex_of_hypreal_capprox_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   861
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   862
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   863
\     ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   864
by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   865
    hcmod,HFinite_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   866
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   867
by (res_inst_tac [("x","u")] exI 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   868
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   869
by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   870
by (auto_tac (claset(),simpset() addsimps [complex_mod,two_eq_Suc_Suc] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   871
by (rtac ccontr 1 THEN dtac real_leI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   872
by (dtac order_less_le_trans 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   873
by (dtac (real_sqrt_ge_abs1 RSN (2,order_less_le_trans)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   874
by (auto_tac (claset(),simpset() addsimps [two_eq_Suc_Suc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   875
qed "CFinite_HFinite_Re";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   876
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   877
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   878
\     ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   879
by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   880
    hcmod,HFinite_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   881
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   882
by (res_inst_tac [("x","u")] exI 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   883
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   884
by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   885
by (auto_tac (claset(),simpset() addsimps [complex_mod] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   886
by (rtac ccontr 1 THEN dtac real_leI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   887
by (dtac order_less_le_trans 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   888
by (dtac (real_sqrt_ge_abs2 RSN (2,order_less_le_trans)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   889
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   890
qed "CFinite_HFinite_Im";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   891
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   892
Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   893
\        Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   894
\     |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   895
by (auto_tac (claset(),simpset() addsimps [CFinite_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   896
    hcmod,HFinite_FreeUltrafilterNat_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   897
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   898
by (res_inst_tac [("x","2*(u + ua)")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   899
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   900
by (dtac sym 1 THEN res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   901
by (auto_tac (claset(),simpset() addsimps [complex_mod,two_eq_Suc_Suc] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   902
by (subgoal_tac "0 < u" 1 THEN arith_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   903
by (subgoal_tac "0 < ua" 1 THEN arith_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   904
by (rtac (realpow_two_abs RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   905
by (res_inst_tac [("x1","Y x")] (realpow_two_abs RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   906
by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   907
by (rtac lemma_sqrt_hcomplex_capprox 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   908
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   909
qed "HFinite_Re_Im_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   910
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   911
Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CFinite) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   912
\     (Abs_hypreal(hyprel `` {%n. Re(X n)}) : HFinite & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   913
\      Abs_hypreal(hyprel `` {%n. Im(X n)}) : HFinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   914
by (blast_tac (claset() addIs [HFinite_Re_Im_CFinite,CFinite_HFinite_Im,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   915
    CFinite_HFinite_Re]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   916
qed "CFinite_HFinite_iff";
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
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   919
\     ==> Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   920
by (auto_tac (claset(),simpset() addsimps [SComplex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   921
    hcomplex_of_complex_def,SReal_def,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   922
by (res_inst_tac [("x","Re r")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   923
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   924
qed "SComplex_Re_SReal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   925
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   926
Goal "Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   927
\     ==> Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   928
by (auto_tac (claset(),simpset() addsimps [SComplex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   929
    hcomplex_of_complex_def,SReal_def,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   930
by (res_inst_tac [("x","Im r")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   931
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   932
qed "SComplex_Im_SReal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   933
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   934
Goal "[| Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   935
\        Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   936
\     |] ==> Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   937
by (auto_tac (claset(),simpset() addsimps [SComplex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   938
    hcomplex_of_complex_def,SReal_def,hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   939
by (res_inst_tac [("x","complex_of_real r + ii  * complex_of_real ra")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   940
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   941
by (res_inst_tac [("z","X x")] eq_Abs_complex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   942
by (auto_tac (claset(),simpset() addsimps [complex_of_real_def,i_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   943
    complex_add,complex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   944
qed "Reals_Re_Im_SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   945
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   946
Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : SComplex) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   947
\     (Abs_hypreal(hyprel `` {%n. Re(X n)}) : Reals & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   948
\      Abs_hypreal(hyprel `` {%n. Im(X n)}) : Reals)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   949
by (blast_tac (claset() addIs [SComplex_Re_SReal,SComplex_Im_SReal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   950
    Reals_Re_Im_SComplex]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   951
qed "SComplex_SReal_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   952
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   953
Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CInfinitesimal) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   954
\     (Abs_hypreal(hyprel `` {%n. Re(X n)}) : Infinitesimal & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   955
\      Abs_hypreal(hyprel `` {%n. Im(X n)}) : Infinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   956
by (auto_tac (claset(),simpset() addsimps [mem_cinfmal_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   957
    mem_infmal_iff,hcomplex_zero_num,hypreal_zero_num,capprox_approx_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   958
qed "CInfinitesimal_Infinitesimal_iff";
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
(*** more lemmas ****)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   961
Goal "(EX t. P t) = (EX X. P (Abs_hcomplex(hcomplexrel `` {X})))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   962
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   963
by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   964
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   965
qed "eq_Abs_hcomplex_EX";
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 "(EX t : A. P t) = (EX X. (Abs_hcomplex(hcomplexrel `` {X})) : A & \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   968
\                        P (Abs_hcomplex(hcomplexrel `` {X})))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   969
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   970
by (res_inst_tac [("z","t")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   971
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   972
qed "eq_Abs_hcomplex_Bex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   973
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   974
(* Here we go - easy proof now!! *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   975
Goal "x:CFinite ==> EX t: SComplex. x @c= t";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   976
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   977
by (auto_tac (claset(),simpset() addsimps [CFinite_HFinite_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   978
    eq_Abs_hcomplex_Bex,SComplex_SReal_iff,capprox_approx_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   979
by (REPEAT(dtac st_part_Ex 1 THEN Step_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   980
by (res_inst_tac [("z","t")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   981
by (res_inst_tac [("z","ta")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   982
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   983
by (res_inst_tac [("x","%n. complex_of_real (xa n) + ii * complex_of_real (xb n)")]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   984
    exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   985
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   986
qed "stc_part_Ex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   987
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   988
Goal "x:CFinite ==> EX! t. t : SComplex &  x @c= t";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   989
by (dtac stc_part_Ex 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   990
by (dtac capprox_sym 2 THEN dtac capprox_sym 2 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   991
    THEN dtac capprox_sym 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   992
by (auto_tac (claset() addSIs [capprox_unique_complex], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   993
qed "stc_part_Ex1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   994
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   995
Goalw [CFinite_def,CInfinite_def] "CFinite Int CInfinite = {}";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   996
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   997
qed "CFinite_Int_CInfinite_empty";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   998
Addsimps [CFinite_Int_CInfinite_empty];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   999
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1000
Goal "x: CFinite ==> x ~: CInfinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1001
by (EVERY1[Step_tac, dtac IntI, assume_tac]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1002
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1003
qed "CFinite_not_CInfinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1004
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1005
Goal "x~: CFinite ==> x: CInfinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1006
by (auto_tac (claset() addIs [not_HFinite_HInfinite],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1007
    simpset() addsimps [CFinite_hcmod_iff,CInfinite_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1008
qed "not_CFinite_CInfinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1009
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1010
Goal "x : CInfinite | x : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1011
by (blast_tac (claset() addIs [not_CFinite_CInfinite]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1012
qed "CInfinite_CFinite_disj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1013
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1014
Goal "(x : CInfinite) = (x ~: CFinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1015
by (blast_tac (claset() addDs [CFinite_not_CInfinite,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1016
               not_CFinite_CInfinite]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1017
qed "CInfinite_CFinite_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1018
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1019
Goal "(x : CFinite) = (x ~: CInfinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1020
by (simp_tac (simpset() addsimps [CInfinite_CFinite_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1021
qed "CFinite_CInfinite_iff";
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 "x ~: CInfinitesimal ==> x : CInfinite | x : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1024
by (fast_tac (claset() addIs [not_CFinite_CInfinite]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1025
qed "CInfinite_diff_CFinite_CInfinitesimal_disj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1026
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1027
Goal "[| x : CFinite; x ~: CInfinitesimal |] ==> inverse x : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1028
by (cut_inst_tac [("x","inverse x")] CInfinite_CFinite_disj 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1029
by (auto_tac (claset() addSDs [CInfinite_inverse_CInfinitesimal], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1030
qed "CFinite_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1031
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1032
Goal "x : CFinite - CInfinitesimal ==> inverse x : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1033
by (blast_tac (claset() addIs [CFinite_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1034
qed "CFinite_inverse2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1035
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1036
Goal "x ~: CInfinitesimal ==> inverse(x) : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1037
by (dtac CInfinite_diff_CFinite_CInfinitesimal_disj 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1038
by (blast_tac (claset() addIs [CFinite_inverse,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1039
                 CInfinite_inverse_CInfinitesimal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1040
                 CInfinitesimal_subset_CFinite RS subsetD]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1041
qed "CInfinitesimal_inverse_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1042
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1043
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1044
Goal "x : CFinite - CInfinitesimal ==> inverse x : CFinite - CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1045
by (auto_tac (claset() addIs [CInfinitesimal_inverse_CFinite], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1046
by (dtac CInfinitesimal_CFinite_mult2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1047
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1048
by (asm_full_simp_tac (simpset() addsimps [not_CInfinitesimal_not_zero]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1049
qed "CFinite_not_CInfinitesimal_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1050
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1051
Goal "[| x @c= y; y :  CFinite - CInfinitesimal |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1052
\     ==> inverse x @c= inverse y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1053
by (forward_tac [CFinite_diff_CInfinitesimal_capprox] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1054
by (assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1055
by (forward_tac [not_CInfinitesimal_not_zero2] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1056
by (forw_inst_tac [("x","x")] not_CInfinitesimal_not_zero2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1057
by (REPEAT(dtac CFinite_inverse2 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1058
by (dtac capprox_mult2 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1059
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1060
by (dres_inst_tac [("c","inverse x")] capprox_mult1 1 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1061
    THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1062
by (auto_tac (claset() addIs [capprox_sym],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1063
    simpset() addsimps [hcomplex_mult_assoc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1064
qed "capprox_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1065
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1066
bind_thm ("hcomplex_of_complex_capprox_inverse",
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1067
       hcomplex_of_complex_CFinite_diff_CInfinitesimal RSN (2, capprox_inverse));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1068
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1069
Goal "[| x: CFinite - CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1070
\        h : CInfinitesimal |] ==> inverse(x + h) @c= inverse x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1071
by (auto_tac (claset() addIs [capprox_inverse, capprox_sym, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1072
                              CInfinitesimal_add_capprox_self], 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1073
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1074
qed "inverse_add_CInfinitesimal_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1075
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1076
Goal "[| x: CFinite - CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1077
\        h : CInfinitesimal |] ==> inverse(h + x) @c= inverse x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1078
by (rtac (hcomplex_add_commute RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1079
by (blast_tac (claset() addIs [inverse_add_CInfinitesimal_capprox]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1080
qed "inverse_add_CInfinitesimal_capprox2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1081
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1082
Goal "[| x : CFinite - CInfinitesimal; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1083
\        h : CInfinitesimal |] ==> inverse(x + h) - inverse x @c= h"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1084
by (rtac capprox_trans2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1085
by (auto_tac (claset() addIs [inverse_add_CInfinitesimal_capprox],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1086
              simpset() addsimps [mem_cinfmal_iff,hcomplex_diff_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1087
                                  capprox_minus_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1088
qed "inverse_add_CInfinitesimal_approx_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1089
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1090
Goal "(x*x : CInfinitesimal) = (x : CInfinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1091
by (auto_tac (claset(), simpset() addsimps [CInfinitesimal_hcmod_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1092
    hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1093
qed "CInfinitesimal_square_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1094
AddIffs [CInfinitesimal_square_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1095
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1096
Goal "[| a: CFinite-CInfinitesimal; a*w @c= a*z |] ==> w @c= z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1097
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1098
by (ftac CFinite_inverse 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1099
by (dtac not_CInfinitesimal_not_zero 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1100
by (auto_tac (claset() addDs [capprox_mult2],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1101
    simpset() addsimps [hcomplex_mult_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1102
qed "capprox_CFinite_mult_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1103
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1104
Goal "a: CFinite-CInfinitesimal ==> (a * w @c= a * z) = (w @c= z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1105
by (auto_tac (claset() addIs [capprox_mult2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1106
    capprox_CFinite_mult_cancel], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1107
qed "capprox_CFinite_mult_cancel_iff1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1108
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1109
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1110
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1111
(* Theorems about monads                                                     *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1112
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1113
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1114
Goalw [cmonad_def] "(x @c= y) = (cmonad(x)=cmonad(y))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1115
by (auto_tac (claset() addDs [capprox_sym] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1116
                       addSEs [capprox_trans,equalityCE],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1117
              simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1118
qed "capprox_cmonad_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1119
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1120
Goal "e : CInfinitesimal ==> cmonad (x+e) = cmonad x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1121
by (fast_tac (claset() addSIs [CInfinitesimal_add_capprox_self RS capprox_sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1122
    capprox_cmonad_iff RS iffD1]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1123
qed "CInfinitesimal_cmonad_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1124
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1125
Goalw [cmonad_def] "(u:cmonad x) = (-u:cmonad (-x))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1126
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1127
qed "mem_cmonad_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1128
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1129
Goalw [cmonad_def] "(x:CInfinitesimal) = (x:cmonad 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1130
by (auto_tac (claset() addIs [capprox_sym],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1131
    simpset() addsimps [mem_cinfmal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1132
qed "CInfinitesimal_cmonad_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1133
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1134
Goal "(x:cmonad 0) = (-x:cmonad 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1135
by (simp_tac (simpset() addsimps [CInfinitesimal_cmonad_zero_iff RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1136
qed "cmonad_zero_minus_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1137
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1138
Goal "(x:cmonad 0) = (hcmod x:monad 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1139
by (auto_tac (claset(), simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1140
    [CInfinitesimal_cmonad_zero_iff RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1141
     CInfinitesimal_hcmod_iff,Infinitesimal_monad_zero_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1142
qed "cmonad_zero_hcmod_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1143
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1144
Goalw [cmonad_def] "x:cmonad x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1145
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1146
qed "mem_cmonad_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1147
Addsimps [mem_cmonad_self];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1148
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1149
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1150
(* Theorems about standard part                                              *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1151
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1152
Goalw [stc_def] "x: CFinite ==> stc x @c= x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1153
by (forward_tac [stc_part_Ex] 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1154
by (rtac someI2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1155
by (auto_tac (claset() addIs [capprox_sym], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1156
qed "stc_capprox_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1157
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1158
Goalw [stc_def] "x: CFinite ==> stc x: SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1159
by (forward_tac [stc_part_Ex] 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1160
by (rtac someI2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1161
by (auto_tac (claset() addIs [capprox_sym], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1162
qed "stc_SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1163
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1164
Goal "x: CFinite ==> stc x: CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1165
by (etac (stc_SComplex RS (SComplex_subset_CFinite RS subsetD)) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1166
qed "stc_CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1167
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1168
Goalw [stc_def] "x: SComplex ==> stc x = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1169
by (rtac some_equality 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1170
by (auto_tac (claset() addIs [(SComplex_subset_CFinite RS subsetD)],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1171
by (blast_tac (claset() addDs [SComplex_capprox_iff RS iffD1]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1172
qed "stc_SComplex_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1173
Addsimps [stc_SComplex_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1174
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1175
Goal "stc (hcomplex_of_complex x) = hcomplex_of_complex x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1176
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1177
qed "stc_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1178
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1179
Goal "[| x: CFinite; y: CFinite; stc x = stc y |] ==> x @c= y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1180
by (auto_tac (claset() addSDs [stc_capprox_self] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1181
              addSEs [capprox_trans3], simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1182
qed "stc_eq_capprox";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1183
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1184
Goal "[| x: CFinite; y: CFinite; x @c= y |] ==> stc x = stc y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1185
by (EVERY1 [forward_tac [stc_capprox_self],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1186
    forw_inst_tac [("x","y")] stc_capprox_self,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1187
    dtac stc_SComplex,dtac stc_SComplex]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1188
by (fast_tac (claset() addEs [capprox_trans,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1189
    capprox_trans2,SComplex_capprox_iff RS iffD1]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1190
qed "capprox_stc_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1191
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1192
Goal "[| x: CFinite; y: CFinite|] ==> (x @c= y) = (stc x = stc y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1193
by (blast_tac (claset() addIs [capprox_stc_eq,stc_eq_capprox]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1194
qed "stc_eq_capprox_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1195
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1196
Goal "[| x: SComplex; e: CInfinitesimal |] ==> stc(x + e) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1197
by (forward_tac [stc_SComplex_eq RS subst] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1198
by (assume_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1199
by (forward_tac [SComplex_subset_CFinite RS subsetD] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1200
by (forward_tac [CInfinitesimal_subset_CFinite RS subsetD] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1201
by (dtac stc_SComplex_eq 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1202
by (rtac capprox_stc_eq 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1203
by (auto_tac (claset() addIs  [CFinite_add],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1204
    simpset() addsimps [CInfinitesimal_add_capprox_self 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1205
    RS capprox_sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1206
qed "stc_CInfinitesimal_add_SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1207
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1208
Goal "[| x: SComplex; e: CInfinitesimal |] ==> stc(e + x) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1209
by (rtac (hcomplex_add_commute RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1210
by (blast_tac (claset() addSIs [stc_CInfinitesimal_add_SComplex]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1211
qed "stc_CInfinitesimal_add_SComplex2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1212
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1213
Goal "x: CFinite ==> EX e: CInfinitesimal. x = stc(x) + e";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1214
by (blast_tac (claset() addSDs [(stc_capprox_self RS 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1215
    capprox_sym),bex_CInfinitesimal_iff2 RS iffD2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1216
qed "CFinite_stc_CInfinitesimal_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1217
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1218
Goal "[| x: CFinite; y: CFinite |] ==> stc (x + y) = stc(x) + stc(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1219
by (forward_tac [CFinite_stc_CInfinitesimal_add] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1220
by (forw_inst_tac [("x","y")] CFinite_stc_CInfinitesimal_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1221
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1222
by (subgoal_tac "stc (x + y) = stc ((stc x + e) + (stc y + ea))" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1223
by (dtac sym 2 THEN dtac sym 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1224
by (Asm_full_simp_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1225
by (asm_simp_tac (simpset() addsimps hcomplex_add_ac) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1226
by (REPEAT(dtac stc_SComplex 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1227
by (dtac SComplex_add 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1228
by (dtac CInfinitesimal_add 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1229
by (rtac (hcomplex_add_assoc RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1230
by (blast_tac (claset() addSIs [stc_CInfinitesimal_add_SComplex2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1231
qed "stc_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1232
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1233
Goal "stc (number_of w) = number_of w";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1234
by (rtac (SComplex_number_of RS stc_SComplex_eq) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1235
qed "stc_number_of";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1236
Addsimps [stc_number_of];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1237
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1238
Goal "stc 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1239
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1240
qed "stc_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1241
Addsimps [stc_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1242
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1243
Goal "stc 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1244
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1245
qed "stc_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1246
Addsimps [stc_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1247
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1248
Goal "y: CFinite ==> stc(-y) = -stc(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1249
by (forward_tac [CFinite_minus_iff RS iffD2] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1250
by (rtac hcomplex_add_minus_eq_minus 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1251
by (dtac (stc_add RS sym) 1 THEN assume_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1252
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1253
qed "stc_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1254
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1255
Goalw [hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1256
     "[| x: CFinite; y: CFinite |] ==> stc (x-y) = stc(x) - stc(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1257
by (forw_inst_tac [("y1","y")] (stc_minus RS sym) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1258
by (dres_inst_tac [("x1","y")] (CFinite_minus_iff RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1259
by (auto_tac (claset() addIs [stc_add],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1260
qed "stc_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1261
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1262
Goal "[| x: CFinite; y: CFinite; \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1263
\        e: CInfinitesimal;       \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1264
\        ea: CInfinitesimal |]   \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1265
\      ==> e*y + x*ea + e*ea: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1266
by (forw_inst_tac [("x","e"),("y","y")] CInfinitesimal_CFinite_mult 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1267
by (forw_inst_tac [("x","ea"),("y","x")] CInfinitesimal_CFinite_mult 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1268
by (dtac CInfinitesimal_mult 3);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1269
by (auto_tac (claset() addIs [CInfinitesimal_add],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1270
              simpset() addsimps hcomplex_add_ac @ hcomplex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1271
qed "lemma_stc_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1272
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1273
Goal "[| x: CFinite; y: CFinite |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1274
\              ==> stc (x * y) = stc(x) * stc(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1275
by (forward_tac [CFinite_stc_CInfinitesimal_add] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1276
by (forw_inst_tac [("x","y")] CFinite_stc_CInfinitesimal_add 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1277
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1278
by (subgoal_tac "stc (x * y) = stc ((stc x + e) * (stc y + ea))" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1279
by (dtac sym 2 THEN dtac sym 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1280
by (Asm_full_simp_tac 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1281
by (thin_tac "x = stc x + e" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1282
by (thin_tac "y = stc y + ea" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1283
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1284
    [hcomplex_add_mult_distrib,hcomplex_add_mult_distrib2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1285
by (REPEAT(dtac stc_SComplex 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1286
by (full_simp_tac (simpset() addsimps [hcomplex_add_assoc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1287
by (rtac stc_CInfinitesimal_add_SComplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1288
by (blast_tac (claset() addSIs [SComplex_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1289
by (REPEAT(dtac (SComplex_subset_CFinite RS subsetD) 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1290
by (rtac (hcomplex_add_assoc RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1291
by (blast_tac (claset() addSIs [lemma_stc_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1292
qed "stc_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1293
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1294
Goal "x: CInfinitesimal ==> stc x = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1295
by (rtac (stc_zero RS subst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1296
by (rtac capprox_stc_eq 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1297
by (auto_tac (claset() addIs [CInfinitesimal_subset_CFinite RS subsetD],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1298
              simpset() addsimps [mem_cinfmal_iff RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1299
qed "stc_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1300
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1301
Goal "stc(x) ~= 0 ==> x ~: CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1302
by (fast_tac (claset() addIs [stc_CInfinitesimal]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1303
qed "stc_not_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1304
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1305
Goal "[| x: CFinite; stc x ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1306
\     ==> stc(inverse x) = inverse (stc x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1307
by (res_inst_tac [("c1","stc x")] (hcomplex_mult_left_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1308
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1309
       simpset() addsimps [stc_mult RS sym, stc_not_CInfinitesimal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1310
                           CFinite_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1311
by (stac hcomplex_mult_inv_right 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1312
by Auto_tac;  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1313
qed "stc_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1314
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1315
Goal "[| x: CFinite; y: CFinite; stc y ~= 0 |] \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1316
\     ==> stc(x/y) = (stc x) / (stc y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1317
by (auto_tac (claset(),
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1318
      simpset() addsimps [hcomplex_divide_def, stc_mult, stc_not_CInfinitesimal, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1319
                          CFinite_inverse, stc_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1320
qed "stc_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1321
Addsimps [stc_divide];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1322
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1323
Goal "x: CFinite ==> stc(stc(x)) = stc(x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1324
by (blast_tac (claset() addIs [stc_CFinite, stc_capprox_self,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1325
                               capprox_stc_eq]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1326
qed "stc_idempotent";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1327
Addsimps [stc_idempotent];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1328
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1329
Goal "z : HFinite ==> hcomplex_of_hypreal z : CFinite";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1330
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1331
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1332
    CFinite_HFinite_iff,symmetric hypreal_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1333
qed "CFinite_HFinite_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1334
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1335
Goal "x : Reals ==>  hcomplex_of_hypreal x : SComplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1336
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1337
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1338
    SComplex_SReal_iff,symmetric hypreal_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1339
qed "SComplex_SReal_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1340
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1341
Goalw [st_def,stc_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1342
 "z : HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1343
by (ftac st_part_Ex 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1344
by (rtac someI2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1345
by (auto_tac (claset() addIs [approx_sym],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1346
by (dtac CFinite_HFinite_hcomplex_of_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1347
by (ftac stc_part_Ex 1 THEN Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1348
by (rtac someI2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1349
by (auto_tac (claset() addIs [capprox_sym] addSIs [capprox_unique_complex] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1350
    addDs [SComplex_SReal_hcomplex_of_hypreal],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1351
qed "stc_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1352
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1353
(*
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1354
Goal "x: CFinite ==> hcmod(stc x) = st(hcmod x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1355
by (dtac stc_capprox_self 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1356
by (auto_tac (claset(),simpset() addsimps [bex_CInfinitesimal_iff2 RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1357
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1358
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1359
approx_hcmod_add_hcmod
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1360
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1361
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1362
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1363
(* More nonstandard complex specific theorems                                *)        
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1364
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1365
Goal "(hcnj z : CInfinitesimal) = (z : CInfinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1366
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1367
qed "CInfinitesimal_hcnj_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1368
Addsimps [CInfinitesimal_hcnj_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1369
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1370
Goal "(Abs_hcomplex(hcomplexrel ``{%n. X n}) : CInfinite) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1371
\     (Abs_hypreal(hyprel `` {%n. Re(X n)}) : HInfinite | \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1372
\      Abs_hypreal(hyprel `` {%n. Im(X n)}) : HInfinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1373
by (auto_tac (claset(),simpset() addsimps [CInfinite_CFinite_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1374
    HInfinite_HFinite_iff,CFinite_HFinite_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1375
qed "CInfinite_HInfinite_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1376
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1377
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CInfinitesimal) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1378
\     (x : Infinitesimal & y : Infinitesimal)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1379
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1380
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1381
by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1382
    hcomplex_of_hypreal,CInfinitesimal_Infinitesimal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1383
qed "hcomplex_split_CInfinitesimal_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1384
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1385
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CFinite) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1386
\     (x : HFinite & y : HFinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1387
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1388
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1389
by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1390
    hcomplex_of_hypreal,CFinite_HFinite_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1391
qed "hcomplex_split_CFinite_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1392
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1393
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : SComplex) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1394
\     (x : Reals & y : Reals)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1395
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1396
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1397
by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1398
    hcomplex_of_hypreal,SComplex_SReal_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1399
qed "hcomplex_split_SComplex_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1400
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1401
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y : CInfinite) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1402
\     (x : HInfinite | y : HInfinite)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1403
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1404
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1405
by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1406
    hcomplex_of_hypreal,CInfinite_HInfinite_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1407
qed "hcomplex_split_CInfinite_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1408
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1409
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y @c= \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1410
\      hcomplex_of_hypreal x' + iii * hcomplex_of_hypreal y') = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1411
\     (x @= x' & y @= y')";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1412
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1413
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1414
by (res_inst_tac [("z","x'")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1415
by (res_inst_tac [("z","y'")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1416
by (auto_tac (claset(),simpset() addsimps [iii_def,hcomplex_add,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1417
    hcomplex_of_hypreal,capprox_approx_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1418
qed "hcomplex_split_capprox_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1419
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1420
(*** More theorems ***)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1421
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1422
Goal "ALL n. cmod (X n - x) < inverse (real (Suc n)) ==> \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1423
\     Abs_hcomplex(hcomplexrel``{X}) - hcomplex_of_complex x : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1424
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1425
    CInfinitesimal_hcmod_iff,hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1426
    Infinitesimal_FreeUltrafilterNat_iff,hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1427
by (rtac bexI 1 THEN Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1428
by (auto_tac (claset() addDs [FreeUltrafilterNat_inverse_real_of_posnat,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1429
                  FreeUltrafilterNat_all,FreeUltrafilterNat_Int] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1430
           addIs [order_less_trans, FreeUltrafilterNat_subset],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1431
      simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1432
qed "complex_seq_to_hcomplex_CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1433
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1434
Goal "hcomplex_of_hypreal epsilon : CInfinitesimal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1435
by (auto_tac (claset(),simpset() addsimps [CInfinitesimal_hcmod_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1436
qed "CInfinitesimal_hcomplex_of_hypreal_epsilon";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1437
Addsimps [CInfinitesimal_hcomplex_of_hypreal_epsilon];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1438
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1439
Goal "(hcomplex_of_complex z @c= 0) = (z = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1440
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_zero RS sym]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1441
    delsimps [hcomplex_of_complex_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1442
qed "hcomplex_of_complex_approx_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1443
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1444
Goal "(0 @c= hcomplex_of_complex z) = (z = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1445
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_zero RS sym]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1446
    delsimps [hcomplex_of_complex_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1447
qed "hcomplex_of_complex_approx_zero_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1448
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1449
Addsimps [hcomplex_of_complex_approx_zero_iff,hcomplex_of_complex_approx_zero_iff2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1450
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1451