src/HOL/Complex/NSComplex.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:       NSComplex.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
    Copyhright:  2001  University of Edinburgh
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Description: Nonstandard Complex numbers
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
Goalw [hcomplexrel_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     8
   "((X,Y): hcomplexrel) = ({n. X n = Y n}: FreeUltrafilterNat)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     9
by (Fast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    10
qed "hcomplexrel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    11
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    12
Goalw [hcomplexrel_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
     "!!X. {n. X n = Y n}: FreeUltrafilterNat \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    14
\      ==> (X,Y): hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
by (Fast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    16
qed "hcomplexrelI";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    17
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    18
Goalw [hcomplexrel_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    19
  "p: hcomplexrel --> (EX X Y. \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    20
\                 p = (X,Y) & {n. X n = Y n} : FreeUltrafilterNat)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    21
by (Fast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    22
qed "hcomplexrelE_lemma";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    23
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    24
val [major,minor] = goal thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    25
  "[| p: hcomplexrel;  \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    26
\     !!X Y. [| p = (X,Y); {n. X n = Y n}: FreeUltrafilterNat\
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    27
\                    |] ==> Q |] ==> Q";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    28
by (cut_facts_tac [major RS (hcomplexrelE_lemma RS mp)] 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
qed "hcomplexrelE";
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
AddSIs [hcomplexrelI];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
AddSEs [hcomplexrelE];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    34
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    35
Goalw [hcomplexrel_def] "(x,x): hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    37
qed "hcomplexrel_refl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    38
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    39
Goalw [hcomplexrel_def] "(x,y): hcomplexrel --> (y,x):hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    40
by (auto_tac (claset() addIs [lemma_perm RS subst],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    41
qed_spec_mp "hcomplexrel_sym";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
Goalw [hcomplexrel_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    44
      "(x,y): hcomplexrel --> (y,z):hcomplexrel --> (x,z):hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    45
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    47
qed_spec_mp "hcomplexrel_trans";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    49
Goalw [equiv_def, refl_def, sym_def, trans_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
    "equiv {x::nat=>complex. True} hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
by (auto_tac (claset() addSIs [hcomplexrel_refl] addSEs 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    52
    [hcomplexrel_sym,hcomplexrel_trans] delrules [hcomplexrelI,hcomplexrelE],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    53
    simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    54
qed "equiv_hcomplexrel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    56
val equiv_hcomplexrel_iff =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    57
    [TrueI, TrueI] MRS 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    58
    ([CollectI, CollectI] MRS 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    59
    (equiv_hcomplexrel RS eq_equiv_class_iff));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    61
Goalw  [hcomplex_def,hcomplexrel_def,quotient_def] "hcomplexrel``{x}:hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    62
by (Blast_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    63
qed "hcomplexrel_in_hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    64
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    65
Goal "inj_on Abs_hcomplex hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    66
by (rtac inj_on_inverseI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    67
by (etac Abs_hcomplex_inverse 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    68
qed "inj_on_Abs_hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    69
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    70
Addsimps [equiv_hcomplexrel_iff,inj_on_Abs_hcomplex RS inj_on_iff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
          hcomplexrel_iff, hcomplexrel_in_hcomplex, Abs_hcomplex_inverse];
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
Addsimps [equiv_hcomplexrel RS eq_equiv_class_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
val eq_hcomplexrelD = equiv_hcomplexrel RSN (2,eq_equiv_class);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    75
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    76
Goal "inj(Rep_hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
by (rtac inj_inverseI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    78
by (rtac Rep_hcomplex_inverse 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    79
qed "inj_Rep_hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    81
Goalw [hcomplexrel_def] "x: hcomplexrel `` {x}";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    82
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
qed "lemma_hcomplexrel_refl";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
Addsimps [lemma_hcomplexrel_refl];
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 [hcomplex_def] "{} ~: hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    88
by (auto_tac (claset() addSEs [quotientE],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
qed "hcomplex_empty_not_mem";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
Addsimps [hcomplex_empty_not_mem];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    91
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
Goal "Rep_hcomplex x ~= {}";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
by (cut_inst_tac [("x","x")] Rep_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    95
qed "Rep_hcomplex_nonempty";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
Addsimps [Rep_hcomplex_nonempty];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    97
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    98
val [prem] = goal thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    99
    "(!!x. z = Abs_hcomplex(hcomplexrel `` {x}) ==> P) ==> P";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   100
by (res_inst_tac [("x1","z")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   101
    (rewrite_rule [hcomplex_def] Rep_hcomplex RS quotientE) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   102
by (dres_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   103
by (res_inst_tac [("x","x")] prem 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   104
by (asm_full_simp_tac (simpset() addsimps [Rep_hcomplex_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   105
qed "eq_Abs_hcomplex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   106
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   107
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   108
(* Properties of nonstandard real and imaginary parts                    *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   109
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   110
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   111
Goalw [hRe_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   112
     "hRe(Abs_hcomplex (hcomplexrel `` {X})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   113
\     Abs_hypreal(hyprel `` {%n. Re(X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   114
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   115
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   116
qed "hRe";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   117
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   118
Goalw [hIm_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   119
     "hIm(Abs_hcomplex (hcomplexrel `` {X})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   120
\     Abs_hypreal(hyprel `` {%n. Im(X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   121
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   122
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   123
qed "hIm";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   124
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   125
Goal "(w=z) = (hRe(w) = hRe(z) & hIm(w) = hIm(z))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   126
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   127
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   128
by (auto_tac (claset(),simpset() addsimps [hRe,hIm,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   129
    complex_Re_Im_cancel_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   130
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   131
qed "hcomplex_hRe_hIm_cancel_iff";
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 [hcomplex_zero_def] "hRe 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   134
by (simp_tac (simpset() addsimps [hRe,hypreal_zero_num]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   135
qed "hcomplex_hRe_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   136
Addsimps [hcomplex_hRe_zero];
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
Goalw [hcomplex_zero_def] "hIm 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   139
by (simp_tac (simpset() addsimps [hIm,hypreal_zero_num]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   140
qed "hcomplex_hIm_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   141
Addsimps [hcomplex_hIm_zero];
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 [hcomplex_one_def] "hRe 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   144
by (simp_tac (simpset() addsimps [hRe,hypreal_one_num]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   145
qed "hcomplex_hRe_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   146
Addsimps [hcomplex_hRe_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   147
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   148
Goalw [hcomplex_one_def] "hIm 1 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   149
by (simp_tac (simpset() addsimps [hIm,hypreal_one_def,hypreal_zero_num]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   150
qed "hcomplex_hIm_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   151
Addsimps [hcomplex_hIm_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   152
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   153
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   154
(*   hcomplex_of_complex: the injection from complex to hcomplex         *)
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   157
Goal "inj(hcomplex_of_complex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   158
by (rtac injI 1 THEN rtac ccontr 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   159
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_complex_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   160
qed "inj_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   161
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   162
Goalw [iii_def,hcomplex_of_complex_def] "iii = hcomplex_of_complex ii";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   163
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   164
qed "hcomplex_of_complex_i";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   165
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   166
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   167
(*   Addition for nonstandard complex numbers: hcomplex_add              *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   168
(* ----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   169
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   170
Goalw [congruent2_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   171
    "congruent2 hcomplexrel (%X Y. hcomplexrel `` {%n. X n + Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   172
by (safe_tac (claset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   173
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   174
qed "hcomplex_add_congruent2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   175
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   176
Goalw [hcomplex_add_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   177
  "Abs_hcomplex(hcomplexrel``{%n. X n}) + Abs_hcomplex(hcomplexrel``{%n. Y n}) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   178
\  Abs_hcomplex(hcomplexrel``{%n. X n + Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   179
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   180
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   181
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   182
qed "hcomplex_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   183
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   184
Goal "(z::hcomplex) + w = w + z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   185
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   186
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   187
by (asm_simp_tac (simpset() addsimps (complex_add_ac @ [hcomplex_add])) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   188
qed "hcomplex_add_commute";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   189
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   190
Goal "((z1::hcomplex) + z2) + z3 = z1 + (z2 + z3)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   191
by (res_inst_tac [("z","z1")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   192
by (res_inst_tac [("z","z2")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   193
by (res_inst_tac [("z","z3")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   194
by (asm_simp_tac (simpset() addsimps [hcomplex_add,complex_add_assoc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   195
qed "hcomplex_add_assoc";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   196
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   197
(*For AC rewriting*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   198
Goal "(x::hcomplex)+(y+z)=y+(x+z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   199
by (rtac (hcomplex_add_commute RS trans) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   200
by (rtac (hcomplex_add_assoc RS trans) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   201
by (rtac (hcomplex_add_commute RS arg_cong) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   202
qed "hcomplex_add_left_commute";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   203
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   204
(* hcomplex addition is an AC operator *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   205
val hcomplex_add_ac = [hcomplex_add_assoc,hcomplex_add_commute,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   206
                      hcomplex_add_left_commute];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   207
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   208
Goalw [hcomplex_zero_def] "(0::hcomplex) + z = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   209
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   210
by (asm_full_simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   211
    [hcomplex_add]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   212
qed "hcomplex_add_zero_left";
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 "z + (0::hcomplex) = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   215
by (simp_tac (simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   216
    [hcomplex_add_zero_left,hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   217
qed "hcomplex_add_zero_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   218
Addsimps [hcomplex_add_zero_left,hcomplex_add_zero_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   219
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   220
Goal "hRe(x + y) = hRe(x) + hRe(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   221
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   222
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   223
by (auto_tac (claset(),simpset() addsimps [hRe,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   224
    hypreal_add,complex_Re_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   225
qed "hRe_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   226
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   227
Goal "hIm(x + y) = hIm(x) + hIm(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   228
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   229
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   230
by (auto_tac (claset(),simpset() addsimps [hIm,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   231
    hypreal_add,complex_Im_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   232
qed "hIm_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   233
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   234
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   235
(* hypreal_minus: additive inverse on nonstandard complex                *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   236
(* ----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   237
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   238
Goalw [congruent_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   239
  "congruent hcomplexrel (%X. hcomplexrel `` {%n. - (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   240
by (safe_tac (claset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   241
by (ALLGOALS(Ultra_tac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   242
qed "hcomplex_minus_congruent";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   243
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   244
Goalw [hcomplex_minus_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   245
  "- (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   246
\     Abs_hcomplex(hcomplexrel `` {%n. -(X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   247
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   248
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   249
qed "hcomplex_minus";
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
Goal "- (- z) = (z::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   252
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   253
by (asm_simp_tac (simpset() addsimps [hcomplex_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   254
qed "hcomplex_minus_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   255
Addsimps [hcomplex_minus_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   256
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   257
Goal "inj(%z::hcomplex. -z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   258
by (rtac injI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   259
by (dres_inst_tac [("f","uminus")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   260
by (Asm_full_simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   261
qed "inj_hcomplex_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   262
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   263
Goalw [hcomplex_zero_def] "- 0 = (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   264
by (simp_tac (simpset() addsimps [hcomplex_minus]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   265
qed "hcomplex_minus_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   266
Addsimps [hcomplex_minus_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   267
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   268
Goal "(-x = 0) = (x = (0::hcomplex))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   269
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   270
by (auto_tac (claset(),simpset() addsimps [hcomplex_zero_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   271
    hcomplex_minus] @ complex_add_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   272
qed "hcomplex_minus_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   273
Addsimps [hcomplex_minus_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   274
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   275
Goal "(0 = -x) = (x = (0::hcomplex))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   276
by (auto_tac (claset() addDs [sym],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   277
qed "hcomplex_minus_zero_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   278
Addsimps [hcomplex_minus_zero_iff2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   279
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   280
Goal "(-x ~= 0) = (x ~= (0::hcomplex))"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   281
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   282
qed "hcomplex_minus_not_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   283
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   284
Goal "z + - z = (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   285
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   286
by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   287
    hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   288
qed "hcomplex_add_minus_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   289
Addsimps [hcomplex_add_minus_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   290
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   291
Goal "-z + z = (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   292
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   293
by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   294
    hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   295
qed "hcomplex_add_minus_left";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   296
Addsimps [hcomplex_add_minus_left];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   297
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   298
Goal "z + (- z + w) = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   299
by (simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   300
qed "hcomplex_add_minus_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   301
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   302
Goal "(-z) + (z + w) = (w::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   303
by (simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   304
qed "hcomplex_minus_add_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   305
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   306
Goal "hRe(-z) = - hRe(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   307
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   308
by (auto_tac (claset(),simpset() addsimps [hRe,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   309
    hypreal_minus,complex_Re_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   310
qed "hRe_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   311
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   312
Goal "hIm(-z) = - hIm(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   313
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   314
by (auto_tac (claset(),simpset() addsimps [hIm,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   315
    hypreal_minus,complex_Im_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   316
qed "hIm_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   317
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   318
Goalw [hcomplex_zero_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   319
      "x + y = (0::hcomplex) ==> x = -y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   320
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   321
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   322
by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   323
by (ultra_tac (claset() addIs [complex_add_minus_eq_minus],simpset()) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   324
qed "hcomplex_add_minus_eq_minus";
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 + y) = -x + -(y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   327
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   328
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   329
by (auto_tac (claset(),simpset() addsimps [hcomplex_add,hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   330
qed "hcomplex_minus_add_distrib";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   331
Addsimps [hcomplex_minus_add_distrib];
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 "((x::hcomplex) + y = x + z) = (y = z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   334
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   335
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   336
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   337
by (auto_tac (claset(),simpset() addsimps [hcomplex_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   338
qed "hcomplex_add_left_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   339
AddIffs [hcomplex_add_left_cancel];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   340
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   341
Goal "(y + (x::hcomplex)= z + x) = (y = z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   342
by (simp_tac (simpset() addsimps [hcomplex_add_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   343
qed "hcomplex_add_right_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   344
AddIffs [hcomplex_add_right_cancel];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   345
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   346
Goal "((x::hcomplex) = y) = ((0::hcomplex) = x + - y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   347
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   348
by (res_inst_tac [("x1","-y")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   349
      (hcomplex_add_right_cancel RS iffD1) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   350
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   351
qed "hcomplex_eq_minus_iff"; 
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
Goal "((x::hcomplex) = y) = (x + - y = (0::hcomplex))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   354
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   355
by (res_inst_tac [("x1","-y")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   356
      (hcomplex_add_right_cancel RS iffD1) 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   357
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   358
qed "hcomplex_eq_minus_iff2"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   359
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   360
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   361
(* Subraction for nonstandard complex numbers: hcomplex_diff             *)
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   364
Goalw [hcomplex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   365
  "Abs_hcomplex(hcomplexrel``{%n. X n}) - Abs_hcomplex(hcomplexrel``{%n. Y n}) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   366
\  Abs_hcomplex(hcomplexrel``{%n. X n - Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   367
by (auto_tac (claset(),simpset() addsimps [hcomplex_minus,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   368
    complex_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   369
qed "hcomplex_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   370
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   371
Goalw [hcomplex_diff_def] "(z::hcomplex) - z = (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   372
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   373
qed "hcomplex_diff_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   374
Addsimps [hcomplex_diff_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   375
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   376
Goal "(0::hcomplex) - x = -x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   377
by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   378
qed "hcomplex_diff_0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   379
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   380
Goal "x - (0::hcomplex) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   381
by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   382
qed "hcomplex_diff_0_right";
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
Goal "x - x = (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   385
by (simp_tac (simpset() addsimps [hcomplex_diff_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   386
qed "hcomplex_diff_self";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   387
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   388
Addsimps [hcomplex_diff_0, hcomplex_diff_0_right, hcomplex_diff_self];
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 "((x::hcomplex) - y = z) = (x = z + y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   391
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def,hcomplex_add_assoc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   392
qed "hcomplex_diff_eq_eq";
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
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   395
(* Multiplication for nonstandard complex numbers: hcomplex_mult         *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   396
(* ----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   397
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   398
Goalw [hcomplex_mult_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   399
  "Abs_hcomplex(hcomplexrel``{%n. X n}) * Abs_hcomplex(hcomplexrel``{%n. Y n}) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   400
\  Abs_hcomplex(hcomplexrel``{%n. X n * Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   401
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   402
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   403
qed "hcomplex_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   404
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   405
Goal "(w::hcomplex) * z = z * w";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   406
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   407
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   408
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   409
    complex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   410
qed "hcomplex_mult_commute";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   411
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   412
Goal "((u::hcomplex) * v) * w = u * (v * w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   413
by (res_inst_tac [("z","u")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   414
by (res_inst_tac [("z","v")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   415
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   416
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   417
    complex_mult_assoc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   418
qed "hcomplex_mult_assoc";
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
Goal "(x::hcomplex) * (y * z) = y * (x * z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   421
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   422
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   423
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   424
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   425
    complex_mult_left_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   426
qed "hcomplex_mult_left_commute";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   427
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   428
val hcomplex_mult_ac = [hcomplex_mult_assoc,hcomplex_mult_commute,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   429
                        hcomplex_mult_left_commute];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   430
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   431
Goalw [hcomplex_one_def] "(1::hcomplex) * z = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   432
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   433
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   434
qed "hcomplex_mult_one_left";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   435
Addsimps [hcomplex_mult_one_left];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   436
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   437
Goal "z * (1::hcomplex) = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   438
by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   439
qed "hcomplex_mult_one_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   440
Addsimps [hcomplex_mult_one_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   441
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   442
Goalw [hcomplex_zero_def] "(0::hcomplex) * z = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   443
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   444
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   445
qed "hcomplex_mult_zero_left";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   446
Addsimps [hcomplex_mult_zero_left];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   447
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   448
Goal "z * (0::hcomplex) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   449
by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   450
qed "hcomplex_mult_zero_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   451
Addsimps [hcomplex_mult_zero_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   452
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   453
Goal "-(x * y) = -x * (y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   454
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   455
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   456
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   457
    hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   458
qed "hcomplex_minus_mult_eq1";
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 "-(x * y) = x * -(y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   461
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   462
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   463
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   464
    hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   465
qed "hcomplex_minus_mult_eq2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   466
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   467
Addsimps [hcomplex_minus_mult_eq1 RS sym,hcomplex_minus_mult_eq2 RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   468
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   469
Goal "- 1 * (z::hcomplex) = -z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   470
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   471
qed "hcomplex_mult_minus_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   472
Addsimps [hcomplex_mult_minus_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   473
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   474
Goal "(z::hcomplex) * - 1 = -z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   475
by (stac hcomplex_mult_commute 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   476
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   477
qed "hcomplex_mult_minus_one_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   478
Addsimps [hcomplex_mult_minus_one_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   479
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   480
Goal "-x * -y = x * (y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   481
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   482
qed "hcomplex_minus_mult_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   483
Addsimps [hcomplex_minus_mult_cancel];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   484
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   485
Goal "-x * y = x * -(y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   486
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   487
qed "hcomplex_minus_mult_commute";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   488
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   489
qed_goal "hcomplex_add_assoc_cong" thy
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   490
    "!!z. (z::hcomplex) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   491
 (fn _ => [(asm_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1)]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   492
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   493
qed_goal "hcomplex_add_assoc_swap" thy "(z::hcomplex) + (v + w) = v + (z + w)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   494
 (fn _ => [(REPEAT (ares_tac [hcomplex_add_commute RS hcomplex_add_assoc_cong] 1))]);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   495
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   496
Goal "((z1::hcomplex) + z2) * w = (z1 * w) + (z2 * w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   497
by (res_inst_tac [("z","z1")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   498
by (res_inst_tac [("z","z2")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   499
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   500
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   501
    complex_add_mult_distrib]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   502
qed "hcomplex_add_mult_distrib";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   503
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   504
Goal "(w::hcomplex) * (z1 + z2) = (w * z1) + (w * z2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   505
by (res_inst_tac [("z1","z1 + z2")] (hcomplex_mult_commute RS ssubst) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   506
by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   507
by (simp_tac (simpset() addsimps [hcomplex_mult_commute]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   508
qed "hcomplex_add_mult_distrib2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   509
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   510
Goalw [hcomplex_zero_def,hcomplex_one_def] "(0::hcomplex) ~= (1::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   511
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   512
qed "hcomplex_zero_not_eq_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   513
Addsimps [hcomplex_zero_not_eq_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   514
Addsimps [hcomplex_zero_not_eq_one RS not_sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   515
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   516
(*-----------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   517
(* Inverse of nonstandard complex number                                 *)
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   520
Goalw [hcinv_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   521
  "inverse (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   522
\     Abs_hcomplex(hcomplexrel `` {%n. inverse (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   523
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   524
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   525
qed "hcomplex_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   526
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   527
Goalw [hcomplex_zero_def] "inverse (0::hcomplex) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   528
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   529
qed "HCOMPLEX_INVERSE_ZERO";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   530
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   531
Goal "a / (0::hcomplex) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   532
by (simp_tac (simpset() addsimps [hcomplex_divide_def, HCOMPLEX_INVERSE_ZERO]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   533
qed "HCOMPLEX_DIVISION_BY_ZERO";  (*NOT for adding to default simpset*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   534
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   535
fun hcomplex_div_undefined_case_tac s i =
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   536
  case_tac s i THEN 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   537
  asm_simp_tac (simpset() addsimps [HCOMPLEX_DIVISION_BY_ZERO, HCOMPLEX_INVERSE_ZERO]) i;
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
Goalw [hcomplex_zero_def,hcomplex_one_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   540
      "z ~= (0::hcomplex) ==> inverse(z) * z = (1::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   541
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   542
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   543
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   544
by (rtac ccontr 1 THEN dtac (complex_mult_inv_left) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   545
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   546
qed "hcomplex_mult_inv_left";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   547
Addsimps [hcomplex_mult_inv_left];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   548
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   549
Goal "z ~= (0::hcomplex) ==> z * inverse(z) = (1::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   550
by (auto_tac (claset() addIs [hcomplex_mult_commute RS subst],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   551
qed "hcomplex_mult_inv_right";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   552
Addsimps [hcomplex_mult_inv_right];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   553
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   554
Goal "(c::hcomplex) ~= (0::hcomplex) ==> (c*a=c*b) = (a=b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   555
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   556
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   557
by (asm_full_simp_tac (simpset() addsimps hcomplex_mult_ac)  1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   558
qed "hcomplex_mult_left_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   559
        
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   560
Goal "(c::hcomplex) ~= (0::hcomplex) ==> (a*c=b*c) = (a=b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   561
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   562
by (dres_inst_tac [("f","%x. x*inverse c")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   563
by (asm_full_simp_tac (simpset() addsimps hcomplex_mult_ac)  1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   564
qed "hcomplex_mult_right_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   565
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   566
Goal "z ~= (0::hcomplex) ==> inverse(z) ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   567
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   568
by (ftac (hcomplex_mult_right_cancel RS iffD2) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   569
by (thin_tac "inverse z = 0" 2);
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 "hcomplex_inverse_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   572
Addsimps [hcomplex_inverse_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   573
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   574
Goal "[| x ~= (0::hcomplex); y ~= 0 |] ==> x * y ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   575
by (Step_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   576
by (dres_inst_tac [("f","%z. inverse x*z")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   577
by (asm_full_simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   578
qed "hcomplex_mult_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   579
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   580
bind_thm ("hcomplex_mult_not_zeroE",hcomplex_mult_not_zero RS notE);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   581
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   582
Goal "inverse(inverse x) = (x::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   583
by (hcomplex_div_undefined_case_tac "x = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   584
by (res_inst_tac [("c1","inverse x")] (hcomplex_mult_right_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   585
by (etac hcomplex_inverse_not_zero 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   586
by (auto_tac (claset() addDs [hcomplex_inverse_not_zero],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   587
qed "hcomplex_inverse_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   588
Addsimps [hcomplex_inverse_inverse];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   589
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   590
Goalw [hcomplex_one_def] "inverse((1::hcomplex)) = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   591
by (simp_tac (simpset() addsimps [hcomplex_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   592
qed "hcomplex_inverse_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   593
Addsimps [hcomplex_inverse_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   594
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   595
Goal "inverse(-x) = -inverse(x::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   596
by (hcomplex_div_undefined_case_tac "x = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   597
by (res_inst_tac [("c1","-x")] (hcomplex_mult_right_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   598
by (stac hcomplex_mult_inv_left 2);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   599
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   600
qed "hcomplex_minus_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   601
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   602
Goal "inverse(x*y) = inverse x * inverse (y::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   603
by (hcomplex_div_undefined_case_tac "x = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   604
by (hcomplex_div_undefined_case_tac "y = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   605
by (res_inst_tac [("c1","x*y")] (hcomplex_mult_left_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   606
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_not_zero]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   607
    @ hcomplex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   608
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_not_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   609
    hcomplex_mult_assoc RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   610
qed "hcomplex_inverse_distrib";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   611
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   612
(*** division ***)
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
(* adding some of these theorems to simpset as for reals: not 100% convinced for some*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   615
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   616
Goal "(x::hcomplex) * (y/z) = (x*y)/z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   617
by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_mult_assoc]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   618
qed "hcomplex_times_divide1_eq";
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 "(y/z) * (x::hcomplex) = (y*x)/z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   621
by (simp_tac (simpset() addsimps [hcomplex_divide_def] @ hcomplex_mult_ac) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   622
qed "hcomplex_times_divide2_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   623
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   624
Addsimps [hcomplex_times_divide1_eq, hcomplex_times_divide2_eq];
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::hcomplex) / (y/z) = (x*z)/y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   627
by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_inverse_distrib]@
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   628
                                  hcomplex_mult_ac) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   629
qed "hcomplex_divide_divide1_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   630
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   631
Goal "((x::hcomplex) / y) / z = x/(y*z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   632
by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_inverse_distrib, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   633
                                  hcomplex_mult_assoc]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   634
qed "hcomplex_divide_divide2_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   635
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   636
Addsimps [hcomplex_divide_divide1_eq, hcomplex_divide_divide2_eq];
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
(** As with multiplication, pull minus signs OUT of the / operator **)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   639
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   640
Goal "(-x) / (y::hcomplex) = - (x/y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   641
by (simp_tac (simpset() addsimps [hcomplex_divide_def]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   642
qed "hcomplex_minus_divide_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   643
Addsimps [hcomplex_minus_divide_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   644
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   645
Goal "(x / -(y::hcomplex)) = - (x/y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   646
by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_minus_inverse]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   647
qed "hcomplex_divide_minus_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   648
Addsimps [hcomplex_divide_minus_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   649
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   650
Goal "(x+y)/(z::hcomplex) = x/z + y/z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   651
by (simp_tac (simpset() addsimps [hcomplex_divide_def, hcomplex_add_mult_distrib]) 1); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   652
qed "hcomplex_add_divide_distrib";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   653
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   654
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   655
(* Embedding properties for hcomplex_of_hypreal map                          *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   656
(*---------------------------------------------------------------------------*)
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
Goalw [hcomplex_of_hypreal_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   659
  "hcomplex_of_hypreal (Abs_hypreal(hyprel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   660
\     Abs_hcomplex(hcomplexrel `` {%n. complex_of_real (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   661
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   662
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   663
qed "hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   664
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   665
Goal "inj hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   666
by (rtac injI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   667
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   668
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   669
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   670
qed "inj_hcomplex_of_hypreal";
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 "(hcomplex_of_hypreal x = hcomplex_of_hypreal y) = (x = y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   673
by (auto_tac (claset() addDs [inj_hcomplex_of_hypreal RS injD],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   674
qed "hcomplex_of_hypreal_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   675
AddIffs [hcomplex_of_hypreal_cancel_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   676
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   677
Goal "hcomplex_of_hypreal(-x) = - hcomplex_of_hypreal x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   678
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   679
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   680
    hcomplex_minus,hypreal_minus,complex_of_real_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   681
qed "hcomplex_of_hypreal_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   682
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   683
Goal "hcomplex_of_hypreal(inverse x) = inverse(hcomplex_of_hypreal x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   684
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   685
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   686
    hypreal_inverse,hcomplex_inverse,complex_of_real_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   687
qed "hcomplex_of_hypreal_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   688
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   689
Goal "hcomplex_of_hypreal x + hcomplex_of_hypreal y = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   690
\     hcomplex_of_hypreal (x + y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   691
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   692
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   693
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   694
    hypreal_add,hcomplex_add,complex_of_real_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   695
qed "hcomplex_of_hypreal_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   696
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   697
Goalw [hcomplex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   698
     "hcomplex_of_hypreal x - hcomplex_of_hypreal y = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   699
\     hcomplex_of_hypreal (x - y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   700
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_minus 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   701
    RS sym,hcomplex_of_hypreal_add,hypreal_diff_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   702
qed "hcomplex_of_hypreal_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   703
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   704
Goal "hcomplex_of_hypreal x * hcomplex_of_hypreal y = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   705
\     hcomplex_of_hypreal (x * y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   706
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   707
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   708
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   709
    hypreal_mult,hcomplex_mult,complex_of_real_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   710
qed "hcomplex_of_hypreal_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   711
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   712
Goalw [hcomplex_divide_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   713
  "hcomplex_of_hypreal x / hcomplex_of_hypreal y = hcomplex_of_hypreal(x/y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   714
by (hypreal_div_undefined_case_tac "y=0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   715
by (simp_tac (simpset() addsimps [rename_numerals HYPREAL_DIVISION_BY_ZERO,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   716
    HCOMPLEX_INVERSE_ZERO]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   717
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   718
    hcomplex_of_hypreal_inverse RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   719
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   720
qed "hcomplex_of_hypreal_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   721
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   722
Goalw [hcomplex_one_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   723
      "hcomplex_of_hypreal 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   724
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hypreal_one_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   725
qed "hcomplex_of_hypreal_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   726
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   727
Goalw [hcomplex_zero_def,hypreal_zero_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   728
      "hcomplex_of_hypreal 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   729
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   730
qed "hcomplex_of_hypreal_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   731
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   732
Addsimps [hcomplex_of_hypreal_one,hcomplex_of_hypreal_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   733
          rename_numerals hcomplex_of_hypreal_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   734
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   735
Goal "hcomplex_of_hypreal (x ^ n) = (hcomplex_of_hypreal x) ^ n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   736
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   737
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal_mult RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   738
qed "hcomplex_of_hypreal_pow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   739
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   740
Goal "hRe(hcomplex_of_hypreal z) = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   741
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   742
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hRe]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   743
qed "hRe_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   744
Addsimps [hRe_hcomplex_of_hypreal];
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 "hIm(hcomplex_of_hypreal z) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   747
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   748
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hIm,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   749
    hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   750
qed "hIm_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   751
Addsimps [hIm_hcomplex_of_hypreal];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   752
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   753
Goal "hcomplex_of_hypreal epsilon ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   754
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   755
    epsilon_def,hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   756
qed "hcomplex_of_hypreal_epsilon_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   757
Addsimps [hcomplex_of_hypreal_epsilon_not_zero];
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
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   760
(*  Modulus (absolute value) of nonstandard complex number                   *) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   761
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   762
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   763
Goalw [hcmod_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   764
  "hcmod (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   765
\     Abs_hypreal(hyprel `` {%n. cmod (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   766
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   767
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   768
qed "hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   769
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   770
Goalw [hcomplex_zero_def,hypreal_zero_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   771
      "hcmod(0) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   772
by (auto_tac (claset(),simpset() addsimps [hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   773
qed "hcmod_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   774
Addsimps [hcmod_zero,rename_numerals hcmod_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   775
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   776
Goalw [hcomplex_one_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   777
      "hcmod(1) = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   778
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_one_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   779
qed "hcmod_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   780
Addsimps [hcmod_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   781
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   782
Goal "hcmod(hcomplex_of_hypreal x) = abs x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   783
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   784
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   785
    hypreal_hrabs]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   786
qed "hcmod_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   787
Addsimps [hcmod_hcomplex_of_hypreal];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   788
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   789
Goal "hcomplex_of_hypreal (abs x) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   790
\     hcomplex_of_hypreal(hcmod(hcomplex_of_hypreal x))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   791
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   792
qed "hcomplex_of_hypreal_abs";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   793
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   794
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   795
(*                   conjugation                                             *)
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   798
Goalw [hcnj_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   799
  "hcnj (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   800
\     Abs_hcomplex(hcomplexrel `` {%n. cnj(X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   801
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   802
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   803
qed "hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   804
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   805
Goal "inj hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   806
by (rtac injI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   807
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   808
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   809
by (auto_tac (claset(),simpset() addsimps [hcnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   810
qed "inj_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   811
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   812
Goal "(hcnj x = hcnj y) = (x = y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   813
by (auto_tac (claset() addDs [inj_hcnj RS injD],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   814
qed "hcomplex_hcnj_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   815
Addsimps [hcomplex_hcnj_cancel_iff];
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 "hcnj (hcnj z) = z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   818
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   819
by (auto_tac (claset(),simpset() addsimps [hcnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   820
qed "hcomplex_hcnj_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   821
Addsimps [hcomplex_hcnj_hcnj];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   822
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   823
Goal "hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   824
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   825
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   826
qed "hcomplex_hcnj_hcomplex_of_hypreal";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   827
Addsimps [hcomplex_hcnj_hcomplex_of_hypreal];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   828
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   829
Goal "hcmod (hcnj z) = hcmod z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   830
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   831
by (auto_tac (claset(),simpset() addsimps [hcnj,hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   832
qed "hcomplex_hmod_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   833
Addsimps [hcomplex_hmod_hcnj];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   834
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   835
Goal "hcnj (-z) = - hcnj z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   836
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   837
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   838
    complex_cnj_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   839
qed "hcomplex_hcnj_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   840
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   841
Goal "hcnj(inverse z) = inverse(hcnj z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   842
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   843
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_inverse,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   844
    complex_cnj_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   845
qed "hcomplex_hcnj_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   846
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   847
Goal "hcnj(w + z) = hcnj(w) + hcnj(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   848
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   849
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   850
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   851
    complex_cnj_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   852
qed "hcomplex_hcnj_add";
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 "hcnj(w - z) = hcnj(w) - hcnj(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   855
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   856
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   857
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   858
    complex_cnj_diff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   859
qed "hcomplex_hcnj_diff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   860
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   861
Goal "hcnj(w * z) = hcnj(w) * hcnj(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   862
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   863
by (res_inst_tac [("z","w")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   864
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   865
    complex_cnj_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   866
qed "hcomplex_hcnj_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   867
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   868
Goalw [hcomplex_divide_def] "hcnj(w / z) = (hcnj w)/(hcnj z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   869
by (simp_tac (simpset() addsimps [hcomplex_hcnj_mult,hcomplex_hcnj_inverse]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   870
qed "hcomplex_hcnj_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   871
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   872
Goalw [hcomplex_one_def] "hcnj 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   873
by (simp_tac (simpset() addsimps [hcnj]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   874
qed "hcnj_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   875
Addsimps [hcnj_one];
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 "hcnj(z ^ n) = hcnj(z) ^ n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   878
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   879
by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   880
qed "hcomplex_hcnj_pow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   881
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   882
(* MOVE to NSComplexBin
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   883
Goal "z + hcnj z = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   884
\     hcomplex_of_hypreal (2 * hRe(z))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   885
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   886
by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   887
    hypreal_mult,hcomplex_of_hypreal,complex_add_cnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   888
qed "hcomplex_add_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   889
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   890
Goal "z - hcnj z = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   891
\     hcomplex_of_hypreal (hypreal_of_real 2 * hIm(z)) * iii";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   892
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   893
by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   894
    hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   895
    complex_diff_cnj,iii_def,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   896
qed "hcomplex_diff_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   897
*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   898
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   899
Goalw [hcomplex_zero_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   900
      "hcnj 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   901
by (auto_tac (claset(),simpset() addsimps [hcnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   902
qed "hcomplex_hcnj_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   903
Addsimps [hcomplex_hcnj_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   904
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   905
Goal "(hcnj z = 0) = (z = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   906
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   907
by (auto_tac (claset(),simpset() addsimps [hcomplex_zero_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   908
    hcnj]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   909
qed "hcomplex_hcnj_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   910
AddIffs [hcomplex_hcnj_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   911
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   912
Goal "z * hcnj z = hcomplex_of_hypreal (hRe(z) ^ 2 + hIm(z) ^ 2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   913
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   914
by (auto_tac (claset(),simpset() addsimps [hcnj,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   915
    hcomplex_of_hypreal,hRe,hIm,hypreal_add,hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   916
    complex_mult_cnj,two_eq_Suc_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   917
qed "hcomplex_mult_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   918
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   919
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   920
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   921
(*  some algebra etc.                                                        *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   922
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   923
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   924
Goal "(x*y = (0::hcomplex)) = (x = 0 | y = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   925
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   926
by (auto_tac (claset() addIs [ccontr] addDs 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   927
    [hcomplex_mult_not_zero],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   928
qed "hcomplex_mult_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   929
Addsimps [hcomplex_mult_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   930
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   931
Goal "(x + y = x) = (y = (0::hcomplex))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   932
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   933
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   934
by (auto_tac (claset(),simpset() addsimps [hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   935
    hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   936
qed "hcomplex_add_left_cancel_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   937
Addsimps [hcomplex_add_left_cancel_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   938
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   939
Goalw [hcomplex_diff_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   940
      "((z1::hcomplex) - z2) * w = (z1 * w) - (z2 * w)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   941
by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   942
qed "hcomplex_diff_mult_distrib";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   943
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   944
Goalw [hcomplex_diff_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   945
      "(w::hcomplex) * (z1 - z2) = (w * z1) - (w * z2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   946
by (simp_tac (simpset() addsimps [hcomplex_add_mult_distrib2]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   947
qed "hcomplex_diff_mult_distrib2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   948
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   949
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   950
(*  More theorems about hcmod                                                *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   951
(*---------------------------------------------------------------------------*)
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 "(hcmod x = 0) = (x = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   954
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   955
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_zero_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   956
    hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   957
qed "hcomplex_hcmod_eq_zero_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   958
Addsimps [hcomplex_hcmod_eq_zero_cancel];
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
(* not proved already? strange! *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   961
Goalw [hypreal_le_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   962
      "(hypreal_of_nat n <= hypreal_of_nat m) = (n <= m)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   963
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   964
qed "hypreal_of_nat_le_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   965
Addsimps [hypreal_of_nat_le_iff];
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 "0 <= hypreal_of_nat n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   968
by (simp_tac (simpset() addsimps [hypreal_of_nat_zero RS sym]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   969
    delsimps [hypreal_of_nat_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   970
qed "hypreal_of_nat_ge_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   971
Addsimps [hypreal_of_nat_ge_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   972
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   973
Addsimps [hypreal_of_nat_ge_zero RS hrabs_eqI1];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   974
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   975
Goal "0 <= hypreal_of_hypnat n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   976
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   977
by (asm_simp_tac (simpset() addsimps [hypreal_of_hypnat,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   978
    hypreal_zero_num,hypreal_le]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   979
qed "hypreal_of_hypnat_ge_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   980
Addsimps [hypreal_of_hypnat_ge_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   981
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   982
Addsimps [hypreal_of_hypnat_ge_zero RS hrabs_eqI1];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   983
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   984
Goal "hcmod (hcomplex_of_hypreal(hypreal_of_nat n)) = hypreal_of_nat n";
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 "hcmod_hcomplex_of_hypreal_of_nat";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   987
Addsimps [hcmod_hcomplex_of_hypreal_of_nat];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   988
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   989
Goal "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   990
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   991
qed "hcmod_hcomplex_of_hypreal_of_hypnat";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   992
Addsimps [hcmod_hcomplex_of_hypreal_of_hypnat];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   993
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   994
Goal "hcmod (-x) = hcmod(x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   995
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   996
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   997
qed "hcmod_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
   998
Addsimps [hcmod_minus];
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 "hcmod(z * hcnj(z)) = hcmod(z) ^ 2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1001
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1002
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1003
    hcnj,hypreal_mult,complex_mod_mult_cnj,two_eq_Suc_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1004
qed "hcmod_mult_hcnj";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1005
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1006
Goal "(0::hypreal) <= hcmod x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1007
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1008
by (auto_tac (claset(),simpset() addsimps [hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1009
    hypreal_zero_num,hypreal_le]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1010
qed "hcmod_ge_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1011
Addsimps [hcmod_ge_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1012
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1013
Goal "abs(hcmod x) = hcmod x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1014
by (auto_tac (claset() addIs [hrabs_eqI1],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1015
qed "hrabs_hcmod_cancel";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1016
Addsimps [hrabs_hcmod_cancel];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1017
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1018
Goal "hcmod(x*y) = hcmod(x) * hcmod(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1019
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1020
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1021
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1022
    hypreal_mult,complex_mod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1023
qed "hcmod_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1024
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1025
Goal "hcmod(x + y) ^ 2 = hcmod(x) ^ 2 + hcmod(y) ^ 2 + 2 * hRe(x * hcnj y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1026
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1027
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1028
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1029
    hypreal_mult,hRe,hcnj,hcomplex_mult,two_eq_Suc_Suc,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1030
    realpow_two RS sym] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1031
by (auto_tac (claset(),simpset() addsimps [two_eq_Suc_Suc RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1032
    complex_mod_add_squared_eq,hypreal_add RS sym,hypreal_mult RS sym,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1033
    symmetric hypreal_of_real_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1034
qed "hcmod_add_squared_eq";
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 "hRe(x * hcnj y) <= hcmod(x * hcnj y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1037
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1038
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1039
by (auto_tac (claset(),simpset() addsimps [hcmod,hcnj,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1040
    hcomplex_mult,hRe,hypreal_le]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1041
qed "hcomplex_hRe_mult_hcnj_le_hcmod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1042
Addsimps [hcomplex_hRe_mult_hcnj_le_hcmod];
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 "hRe(x * hcnj y) <= hcmod(x * y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1045
by (cut_inst_tac [("x","x"),("y","y")] hcomplex_hRe_mult_hcnj_le_hcmod 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1046
by (asm_full_simp_tac (simpset() addsimps [hcmod_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1047
qed "hcomplex_hRe_mult_hcnj_le_hcmod2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1048
Addsimps [hcomplex_hRe_mult_hcnj_le_hcmod2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1049
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1050
Goal "hcmod (x + y) ^ 2 <= (hcmod(x) + hcmod(y)) ^ 2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1051
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1052
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1053
by (auto_tac (claset(),simpset() addsimps [hcmod,hcnj,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1054
    hcomplex_add,hypreal_mult,hypreal_add,hypreal_le,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1055
    realpow_two RS sym,two_eq_Suc_Suc] delsimps [realpow_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1056
by (simp_tac (simpset() addsimps [two_eq_Suc_Suc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1057
qed "hcmod_triangle_squared";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1058
Addsimps [hcmod_triangle_squared];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1059
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1060
Goal "hcmod (x + y) <= hcmod(x) + hcmod(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1061
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1062
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1063
by (auto_tac (claset(),simpset() addsimps [hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1064
    hcomplex_add,hypreal_add,hypreal_le]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1065
qed "hcmod_triangle_ineq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1066
Addsimps [hcmod_triangle_ineq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1067
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1068
Goal "hcmod(b + a) - hcmod b <= hcmod a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1069
by (cut_inst_tac [("x1","b"),("y1","a"),("x","-hcmod b")]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1070
   (hcmod_triangle_ineq RS hypreal_add_le_mono1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1071
by (asm_full_simp_tac (simpset() addsimps hypreal_add_ac) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1072
qed "hcmod_triangle_ineq2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1073
Addsimps [hcmod_triangle_ineq2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1074
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1075
Goal "hcmod (x - y) = hcmod (y - x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1076
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1077
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1078
by (auto_tac (claset(),simpset() addsimps [hcmod,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1079
    hcomplex_diff,complex_mod_diff_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1080
qed "hcmod_diff_commute";
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 "[| hcmod x < r; hcmod y < s |] ==> hcmod (x + y) < r + s";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1083
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1084
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1085
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1086
by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1087
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1088
    hypreal_add,hypreal_less]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1089
by (ultra_tac (claset() addIs [complex_mod_add_less],simpset()) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1090
qed "hcmod_add_less";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1091
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1092
Goal "[| hcmod x < r; hcmod y < s |] ==> hcmod (x * y) < r * s";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1093
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1094
by (res_inst_tac [("z","y")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1095
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1096
by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1097
by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1098
    hypreal_less,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1099
by (ultra_tac (claset() addIs [complex_mod_mult_less],simpset()) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1100
qed "hcmod_mult_less";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1101
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1102
goal NSComplex.thy "hcmod(a) - hcmod(b) <= hcmod(a + b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1103
by (res_inst_tac [("z","a")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1104
by (res_inst_tac [("z","b")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1105
by (auto_tac (claset(),simpset() addsimps [hcmod,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1106
    hypreal_diff,hypreal_le]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1107
qed "hcmod_diff_ineq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1108
Addsimps [hcmod_diff_ineq];
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
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1112
(*                       a few nonlinear theorems                            *)
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
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1115
Goalw [hcpow_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1116
  "Abs_hcomplex(hcomplexrel``{%n. X n}) hcpow \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1117
\  Abs_hypnat(hypnatrel``{%n. Y n}) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1118
\  Abs_hcomplex(hcomplexrel``{%n. X n ^ Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1119
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1120
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1121
qed "hcpow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1122
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1123
Goal "hcomplex_of_hypreal (x pow n) = (hcomplex_of_hypreal x) hcpow n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1124
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1125
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1126
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1127
    hyperpow,hcpow,complex_of_real_pow]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1128
qed "hcomplex_of_hypreal_hyperpow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1129
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1130
Goal "hcmod(x ^ n) = hcmod(x) ^ n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1131
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1132
by (auto_tac (claset(),simpset() addsimps [hcmod_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1133
qed "hcmod_hcomplexpow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1134
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1135
Goal "hcmod(x hcpow n) = hcmod(x) pow n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1136
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1137
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1138
by (auto_tac (claset(),simpset() addsimps [hcpow,hyperpow,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1139
    hcmod,complex_mod_complexpow]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1140
qed "hcmod_hcpow";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1141
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1142
Goal "(-x::hcomplex) ^ n = (if even n then (x ^ n) else -(x ^ n))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1143
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1144
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1145
qed "hcomplexpow_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1146
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1147
Goal "(-x::hcomplex) hcpow n = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1148
\     (if ( *pNat* even) n then (x hcpow n) else -(x hcpow n))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1149
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1150
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1151
by (auto_tac (claset(),simpset() addsimps [hcpow,hyperpow,starPNat,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1152
    hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1153
by (ALLGOALS(ultra_tac (claset(),simpset() addsimps [complexpow_minus])));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1154
qed "hcpow_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1155
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1156
Goal "inverse(-x) = - inverse (x::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1157
by (res_inst_tac [("z","x")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1158
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1159
    complex_inverse_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1160
qed "hccomplex_inverse_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1161
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1162
Goalw [hcomplex_divide_def] "x / (1::hcomplex) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1163
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1164
qed "hcomplex_div_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1165
Addsimps [hcomplex_div_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1166
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1167
Goal "hcmod(inverse x) = inverse(hcmod x)"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1168
by (hcomplex_div_undefined_case_tac "x = 0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1169
by (res_inst_tac [("c1","hcmod x")] (hypreal_mult_left_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1170
by (auto_tac (claset(),simpset() addsimps [hcmod_mult RS sym]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1171
qed "hcmod_hcomplex_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1172
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1173
Goalw [hcomplex_divide_def,hypreal_divide_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1174
      "hcmod(x/y) = hcmod(x)/(hcmod y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1175
by (auto_tac (claset(),simpset() addsimps [hcmod_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1176
    hcmod_hcomplex_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1177
qed "hcmod_divide";
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
Goalw [hcomplex_divide_def]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1180
      "inverse(x/y) = y/(x::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1181
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse_distrib,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1182
    hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1183
qed "hcomplex_inverse_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1184
Addsimps [hcomplex_inverse_divide];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1185
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1186
Goal "((r::hcomplex) * s) ^ n = (r ^ n) * (s ^ n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1187
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1188
by (auto_tac (claset(),simpset() addsimps hcomplex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1189
qed "hcomplexpow_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1190
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1191
Goal "((r::hcomplex) * s) hcpow n = (r hcpow n) * (s hcpow n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1192
by (res_inst_tac [("z","r")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1193
by (res_inst_tac [("z","s")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1194
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1195
by (auto_tac (claset(),simpset() addsimps [hcpow,hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1196
    hcomplex_mult,complexpow_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1197
qed "hcpow_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1198
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1199
Goal "(0::hcomplex) ^ (Suc n) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1200
by (Auto_tac);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1201
qed "hcomplexpow_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1202
Addsimps [hcomplexpow_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1203
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1204
Goalw [hcomplex_zero_def,hypnat_one_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1205
   "0 hcpow (n + 1) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1206
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1207
by (auto_tac (claset(),simpset() addsimps [hcpow,hypnat_add]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1208
qed "hcpow_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1209
Addsimps [hcpow_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1210
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1211
Goalw [hSuc_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1212
   "0 hcpow (hSuc n) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1213
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1214
qed "hcpow_zero2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1215
Addsimps [hcpow_zero2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1216
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1217
Goal "r ~= (0::hcomplex) --> r ^ n ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1218
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1219
by (auto_tac (claset(),simpset() addsimps 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1220
    [hcomplex_mult_not_zero]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1221
qed_spec_mp "hcomplexpow_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1222
Addsimps [hcomplexpow_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1223
AddIs [hcomplexpow_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1224
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1225
Goal "r ~= 0 ==> r hcpow n ~= (0::hcomplex)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1226
by (res_inst_tac [("z","r")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1227
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1228
by (auto_tac (claset(),simpset() addsimps [hcpow,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1229
    hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1230
by (ultra_tac (claset() addDs [complexpow_zero_zero],simpset()) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1231
qed "hcpow_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1232
Addsimps [hcpow_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1233
AddIs [hcpow_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1234
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1235
Goal "r ^ n = (0::hcomplex) ==> r = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1236
by (blast_tac (claset() addIs [ccontr] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1237
    addDs [hcomplexpow_not_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1238
qed "hcomplexpow_zero_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1239
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1240
Goal "r hcpow n = (0::hcomplex) ==> r = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1241
by (blast_tac (claset() addIs [ccontr] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1242
    addDs [hcpow_not_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1243
qed "hcpow_zero_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1244
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1245
Goalw [iii_def] "iii * iii = - 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1246
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1247
    hcomplex_one_def,hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1248
qed "hcomplex_i_mult_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1249
Addsimps [hcomplex_i_mult_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1250
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1251
Goal "iii ^ 2 = - 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1252
by (simp_tac (simpset() addsimps [two_eq_Suc_Suc]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1253
qed "hcomplexpow_i_squared";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1254
Addsimps [hcomplexpow_i_squared];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1255
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1256
Goalw [iii_def,hcomplex_zero_def] "iii ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1257
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1258
qed "hcomplex_i_not_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1259
Addsimps [hcomplex_i_not_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1260
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1261
Goal "x * y ~= (0::hcomplex) ==> x ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1262
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1263
qed "hcomplex_mult_eq_zero_cancel1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1264
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1265
Goal "x * y ~= (0::hcomplex) ==> y ~= 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1266
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1267
qed "hcomplex_mult_eq_zero_cancel2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1268
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1269
Goal "(x * y ~= (0::hcomplex)) = (x ~= 0 & y ~= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1270
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1271
qed "hcomplex_mult_not_eq_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1272
AddIffs [hcomplex_mult_not_eq_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1273
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1274
Goalw [hcomplex_divide_def,complex_divide_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1275
  "Abs_hcomplex(hcomplexrel``{%n. X n}) / Abs_hcomplex(hcomplexrel``{%n. Y n}) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1276
\  Abs_hcomplex(hcomplexrel``{%n. X n / Y n})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1277
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1278
qed "hcomplex_divide";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1279
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1280
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1281
(*                             hsgn                                          *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1282
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1283
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1284
Goalw [hsgn_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1285
  "hsgn (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1286
\     Abs_hcomplex(hcomplexrel `` {%n. sgn (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1287
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1288
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1289
qed "hsgn";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1290
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1291
Addsimps [rename_numerals sgn_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1292
Goalw [hcomplex_zero_def] "hsgn 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1293
by (simp_tac (simpset() addsimps [hsgn]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1294
qed "hsgn_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1295
Addsimps[hsgn_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1296
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1297
Addsimps [rename_numerals sgn_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1298
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1299
Goalw [hcomplex_one_def] "hsgn 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1300
by (simp_tac (simpset() addsimps [hsgn]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1301
qed "hsgn_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1302
Addsimps[hsgn_one];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1303
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1304
Goal "hsgn (-z) = - hsgn(z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1305
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1306
by (auto_tac (claset(),simpset() addsimps [hsgn,hcomplex_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1307
    sgn_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1308
qed "hsgn_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1309
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1310
Goal "hsgn z = z / hcomplex_of_hypreal (hcmod z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1311
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1312
by (auto_tac (claset(),simpset() addsimps [hsgn,hcomplex_divide,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1313
    hcomplex_of_hypreal,hcmod,sgn_eq]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1314
qed "hsgn_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1315
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1316
Goal "(EX (x::hypreal) y. P x y) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1317
\     (EX f g. P (Abs_hypreal(hyprel `` {f})) (Abs_hypreal(hyprel `` {g})))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1318
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1319
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1320
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1321
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1322
qed "lemma_hypreal_P_EX2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1323
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1324
Goal "ALL (n::nat). EX x y. (z n) = complex_of_real(x) + ii * complex_of_real(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1325
by (blast_tac (claset() addIs [complex_split]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1326
qed "complex_split2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1327
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1328
(* Interesting proof! *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1329
Goal "EX x y. z = hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1330
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1331
by (auto_tac (claset(),simpset() addsimps [lemma_hypreal_P_EX2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1332
    hcomplex_of_hypreal,iii_def,hcomplex_add,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1333
by (cut_inst_tac [("z","x")] complex_split2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1334
by (REPEAT(dtac choice 1 THEN Step_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1335
by (res_inst_tac [("x","f")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1336
by (res_inst_tac [("x","fa")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1337
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1338
qed "hcomplex_split";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1339
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1340
Goal "hRe(hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1341
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1342
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1343
by (auto_tac (claset(),simpset() addsimps [hRe,iii_def,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1344
    hcomplex_mult,hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1345
qed "hRe_hcomplex_i";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1346
Addsimps [hRe_hcomplex_i];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1347
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1348
Goal "hIm(hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1349
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1350
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1351
by (auto_tac (claset(),simpset() addsimps [hIm,iii_def,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1352
    hcomplex_mult,hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1353
qed "hIm_hcomplex_i";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1354
Addsimps [hIm_hcomplex_i];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1355
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1356
Goal "hcmod (hcomplex_of_hypreal(x) + iii * hcomplex_of_hypreal(y)) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1357
\     ( *f* sqrt) (x ^ 2 + y ^ 2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1358
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1359
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1360
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1361
    iii_def,hcomplex_add,hcomplex_mult,starfun,hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1362
    hypreal_add,hcmod,cmod_i,two_eq_Suc_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1363
qed "hcmod_i";
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
Goalw [iii_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1366
     "hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1367
\     hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1368
\      ==> xa = xb";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1369
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1370
by (res_inst_tac [("z","ya")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1371
by (res_inst_tac [("z","xb")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1372
by (res_inst_tac [("z","yb")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1373
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1374
    hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1375
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1376
qed "hcomplex_eq_hRe_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1377
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1378
Goalw [iii_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1379
     "hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1380
\     hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1381
\      ==> ya = yb";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1382
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1383
by (res_inst_tac [("z","ya")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1384
by (res_inst_tac [("z","xb")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1385
by (res_inst_tac [("z","yb")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1386
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1387
    hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1388
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1389
qed "hcomplex_eq_hIm_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1390
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1391
Goal "(hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1392
\      hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1393
\     ((xa = xb) & (ya = yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1394
by (auto_tac (claset() addIs [hcomplex_eq_hIm_eq,hcomplex_eq_hRe_eq],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1395
qed "hcomplex_eq_cancel_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1396
Addsimps [hcomplex_eq_cancel_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1397
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1398
Goal "(hcomplex_of_hypreal xa + hcomplex_of_hypreal ya * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1399
\      hcomplex_of_hypreal xb + hcomplex_of_hypreal yb * iii ) = ((xa = xb) & (ya = yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1400
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1401
qed "hcomplex_eq_cancel_iffA";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1402
AddIffs [hcomplex_eq_cancel_iffA];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1403
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1404
Goal "(hcomplex_of_hypreal xa + hcomplex_of_hypreal ya * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1405
\      hcomplex_of_hypreal xb + iii * hcomplex_of_hypreal yb) = ((xa = xb) & (ya = yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1406
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1407
qed "hcomplex_eq_cancel_iffB";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1408
AddIffs [hcomplex_eq_cancel_iffB];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1409
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1410
Goal "(hcomplex_of_hypreal xa + iii * hcomplex_of_hypreal ya  = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1411
\      hcomplex_of_hypreal xb + hcomplex_of_hypreal yb * iii) = ((xa = xb) & (ya = yb))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1412
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1413
qed "hcomplex_eq_cancel_iffC";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1414
AddIffs [hcomplex_eq_cancel_iffC];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1415
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1416
Goal"(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1417
\     hcomplex_of_hypreal xa) = (x = xa & y = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1418
by (cut_inst_tac [("xa","x"),("ya","y"),("xb","xa"),("yb","0")]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1419
    hcomplex_eq_cancel_iff 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1420
by (asm_full_simp_tac (simpset() delsimps [hcomplex_eq_cancel_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1421
qed "hcomplex_eq_cancel_iff2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1422
Addsimps [hcomplex_eq_cancel_iff2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1423
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1424
Goal"(hcomplex_of_hypreal x + hcomplex_of_hypreal y * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1425
\     hcomplex_of_hypreal xa) = (x = xa & y = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1426
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1427
qed "hcomplex_eq_cancel_iff2a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1428
Addsimps [hcomplex_eq_cancel_iff2a];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1429
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1430
Goal "(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1431
\     iii * hcomplex_of_hypreal ya) = (x = 0 & y = ya)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1432
by (cut_inst_tac [("xa","x"),("ya","y"),("xb","0"),("yb","ya")]  
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1433
    hcomplex_eq_cancel_iff 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1434
by (asm_full_simp_tac (simpset() delsimps [hcomplex_eq_cancel_iff]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1435
qed "hcomplex_eq_cancel_iff3";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1436
Addsimps [hcomplex_eq_cancel_iff3];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1437
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1438
Goal "(hcomplex_of_hypreal x + hcomplex_of_hypreal y * iii = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1439
\     iii * hcomplex_of_hypreal ya) = (x = 0 & y = ya)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1440
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1441
qed "hcomplex_eq_cancel_iff3a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1442
Addsimps [hcomplex_eq_cancel_iff3a];
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
Goalw [iii_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1445
     "hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = 0 \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1446
\     ==> x = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1447
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1448
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1449
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1450
    hcomplex_add,hcomplex_mult,hcomplex_zero_def,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1451
by (ultra_tac (claset(),simpset() addsimps [complex_split_Re_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1452
qed "hcomplex_split_hRe_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1453
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1454
Goalw [iii_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1455
     "hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y = 0 \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1456
\     ==> y = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1457
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1458
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1459
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1460
    hcomplex_add,hcomplex_mult,hcomplex_zero_def,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1461
by (ultra_tac (claset(),simpset() addsimps [complex_split_Im_zero]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1462
qed "hcomplex_split_hIm_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1463
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1464
Goal "hRe(hsgn z) = hRe(z)/hcmod z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1465
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1466
by (auto_tac (claset(),simpset() addsimps [hsgn,hcmod,hRe,hypreal_divide]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1467
qed "hRe_hsgn";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1468
Addsimps [hRe_hsgn];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1469
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1470
Goal "hIm(hsgn z) = hIm(z)/hcmod z";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1471
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1472
by (auto_tac (claset(),simpset() addsimps [hsgn,hcmod,hIm,hypreal_divide]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1473
qed "hIm_hsgn";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1474
Addsimps [hIm_hsgn];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1475
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1476
Goal "(x*x + y*y = 0) = ((x::real) = 0 & y = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1477
by (auto_tac (claset() addIs [real_sum_squares_cancel],simpset()));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1478
qed "real_two_squares_add_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1479
Addsimps [real_two_squares_add_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1480
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1481
Goal "inverse(hcomplex_of_hypreal x + iii * hcomplex_of_hypreal y) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1482
\     hcomplex_of_hypreal(x/(x ^ 2 + y ^ 2)) - \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1483
\     iii * hcomplex_of_hypreal(y/(x ^ 2 + y ^ 2))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1484
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1485
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1486
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1487
    hcomplex_mult,hcomplex_add,iii_def,starfun,hypreal_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1488
    hypreal_add,hcomplex_inverse,hypreal_divide,hcomplex_diff,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1489
    complex_inverse_complex_split,two_eq_Suc_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1490
qed "hcomplex_inverse_complex_split";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1491
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1492
Goalw [iii_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1493
    "hRe (iii * hcomplex_of_hypreal y) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1494
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1495
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1496
    hcomplex_mult,hRe,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1497
qed "hRe_mult_i_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1498
Addsimps [hRe_mult_i_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1499
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1500
Goalw [iii_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1501
    "hIm (iii * hcomplex_of_hypreal y) = y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1502
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1503
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1504
    hcomplex_mult,hIm,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1505
qed "hIm_mult_i_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1506
Addsimps [hIm_mult_i_eq];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1507
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1508
Goal "hcmod (iii * hcomplex_of_hypreal y) = abs y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1509
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1510
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1511
    hcmod,hypreal_hrabs,iii_def,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1512
qed "hcmod_mult_i";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1513
Addsimps [hcmod_mult_i];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1514
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1515
Goal "hcmod (hcomplex_of_hypreal y * iii) = abs y";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1516
by (auto_tac (claset(),simpset() addsimps [hcomplex_mult_commute]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1517
qed "hcmod_mult_i2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1518
Addsimps [hcmod_mult_i2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1519
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1520
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1521
(*  harg                                                                     *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1522
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1523
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1524
Goalw [harg_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1525
  "harg (Abs_hcomplex(hcomplexrel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1526
\     Abs_hypreal(hyprel `` {%n. arg (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1527
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1528
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1529
qed "harg";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1530
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1531
Goal "0 < y ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1532
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1533
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1534
    iii_def,hcomplex_mult,hypreal_zero_num,hypreal_less,starfun,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1535
    harg]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1536
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1537
qed "cos_harg_i_mult_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1538
Addsimps [cos_harg_i_mult_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1539
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1540
Goal "y < 0 ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1541
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1542
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1543
    iii_def,hcomplex_mult,hypreal_zero_num,hypreal_less,starfun,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1544
    harg]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1545
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1546
qed "cos_harg_i_mult_zero2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1547
Addsimps [cos_harg_i_mult_zero2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1548
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1549
Goal "(hcomplex_of_hypreal y ~= 0) = (y ~= 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1550
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1551
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1552
    hypreal_zero_num,hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1553
qed "hcomplex_of_hypreal_not_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1554
Addsimps [hcomplex_of_hypreal_not_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1555
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1556
Goal "(hcomplex_of_hypreal y = 0) = (y = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1557
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1558
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1559
    hypreal_zero_num,hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1560
qed "hcomplex_of_hypreal_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1561
Addsimps [hcomplex_of_hypreal_zero_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1562
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1563
Goal "y ~= 0 ==> ( *f* cos) (harg(iii * hcomplex_of_hypreal y)) = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1564
by (cut_inst_tac [("x","y"),("y","0")] hypreal_linear 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1565
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1566
qed "cos_harg_i_mult_zero3";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1567
Addsimps [cos_harg_i_mult_zero3];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1568
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1569
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1570
(* Polar form for nonstandard complex numbers                                 *) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1571
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1572
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1573
Goal "ALL n. EX r a. (z n) = complex_of_real r * \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1574
\     (complex_of_real(cos a) + ii * complex_of_real(sin a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1575
by (blast_tac (claset() addIs [complex_split_polar]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1576
qed "complex_split_polar2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1577
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1578
Goal 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1579
  "EX r a. z = hcomplex_of_hypreal r * \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1580
\  (hcomplex_of_hypreal(( *f* cos) a) + iii * hcomplex_of_hypreal(( *f* sin) a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1581
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1582
by (auto_tac (claset(),simpset() addsimps [lemma_hypreal_P_EX2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1583
    hcomplex_of_hypreal,iii_def,starfun,hcomplex_add,hcomplex_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1584
by (cut_inst_tac [("z","x")] complex_split_polar2 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1585
by (REPEAT(dtac choice 1 THEN Step_tac 1));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1586
by (res_inst_tac [("x","f")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1587
by (res_inst_tac [("x","fa")] exI 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1588
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1589
qed "hcomplex_split_polar";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1590
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1591
Goalw [hcis_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1592
  "hcis (Abs_hypreal(hyprel `` {%n. X n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1593
\     Abs_hcomplex(hcomplexrel `` {%n. cis (X n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1594
by (res_inst_tac [("f","Abs_hcomplex")] arg_cong 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1595
by (Auto_tac THEN Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1596
qed "hcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1597
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1598
Goal 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1599
   "hcis a = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1600
\   (hcomplex_of_hypreal(( *f* cos) a) + \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1601
\   iii * hcomplex_of_hypreal(( *f* sin) a))";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1602
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1603
by (auto_tac (claset(),simpset() addsimps [starfun, hcis,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1604
    hcomplex_of_hypreal,iii_def,hcomplex_mult,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1605
    cis_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1606
qed "hcis_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1607
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1608
Goalw [hrcis_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1609
  "hrcis (Abs_hypreal(hyprel `` {%n. X n})) (Abs_hypreal(hyprel `` {%n. Y n})) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1610
\     Abs_hcomplex(hcomplexrel `` {%n. rcis (X n) (Y n)})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1611
by (auto_tac (claset(),simpset() addsimps [hcomplex_of_hypreal,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1612
    hcis,rcis_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1613
qed "hrcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1614
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1615
Goal "EX r a. z = hrcis r a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1616
by (simp_tac (simpset() addsimps [hrcis_def,hcis_eq]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1617
by (rtac hcomplex_split_polar 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1618
qed "hrcis_Ex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1619
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1620
Goal "hRe(hcomplex_of_hypreal r * \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1621
\     (hcomplex_of_hypreal(( *f* cos) a) + \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1622
\      iii * hcomplex_of_hypreal(( *f* sin) a))) = r * ( *f* cos) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1623
by (auto_tac (claset(),simpset() addsimps [hcomplex_add_mult_distrib2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1624
    hcomplex_of_hypreal_mult] @ hcomplex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1625
qed "hRe_hcomplex_polar";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1626
Addsimps [hRe_hcomplex_polar];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1627
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1628
Goalw [hrcis_def] "hRe(hrcis r a) = r * ( *f* cos) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1629
by (auto_tac (claset(),simpset() addsimps [hcis_eq]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1630
qed "hRe_hrcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1631
Addsimps [hRe_hrcis];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1632
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1633
Goal "hIm(hcomplex_of_hypreal r * \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1634
\     (hcomplex_of_hypreal(( *f* cos) a) + \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1635
\      iii * hcomplex_of_hypreal(( *f* sin) a))) = r * ( *f* sin) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1636
by (auto_tac (claset(),simpset() addsimps [hcomplex_add_mult_distrib2,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1637
    hcomplex_of_hypreal_mult] @ hcomplex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1638
qed "hIm_hcomplex_polar";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1639
Addsimps [hIm_hcomplex_polar];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1640
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1641
Goalw [hrcis_def] "hIm(hrcis r a) = r * ( *f* sin) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1642
by (auto_tac (claset(),simpset() addsimps [hcis_eq]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1643
qed "hIm_hrcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1644
Addsimps [hIm_hrcis];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1645
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1646
Goal "hcmod (hcomplex_of_hypreal r * \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1647
\     (hcomplex_of_hypreal(( *f* cos) a) + \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1648
\      iii * hcomplex_of_hypreal(( *f* sin) a))) = abs r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1649
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1650
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1651
by (auto_tac (claset(),simpset() addsimps [iii_def,starfun,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1652
    hcomplex_of_hypreal,hcomplex_mult,hcmod,hcomplex_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1653
    hypreal_hrabs]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1654
qed "hcmod_complex_polar";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1655
Addsimps [hcmod_complex_polar];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1656
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1657
Goalw [hrcis_def] "hcmod(hrcis r a) = abs r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1658
by (auto_tac (claset(),simpset() addsimps [hcis_eq]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1659
qed "hcmod_hrcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1660
Addsimps [hcmod_hrcis];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1661
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1662
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1663
(*  (r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)                *) 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1664
(*---------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1665
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1666
Goalw [hrcis_def] "hcis a = hrcis 1 a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1667
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1668
qed "hcis_hrcis_eq";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1669
Addsimps [hcis_hrcis_eq RS sym];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1670
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1671
Goalw [hrcis_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1672
  "hrcis r1 a * hrcis r2 b = hrcis (r1*r2) (a + b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1673
by (res_inst_tac [("z","r1")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1674
by (res_inst_tac [("z","r2")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1675
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1676
by (res_inst_tac [("z","b")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1677
by (auto_tac (claset(),simpset() addsimps [hrcis,hcis,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1678
    hypreal_add,hypreal_mult,hcomplex_of_hypreal,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1679
    hcomplex_mult,cis_mult RS sym,complex_of_real_mult
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1680
    RS sym] addsimps complex_mult_ac));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1681
qed "hrcis_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1682
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1683
Goal "hcis a * hcis b = hcis (a + b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1684
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1685
by (res_inst_tac [("z","b")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1686
by (auto_tac (claset(),simpset() addsimps [hcis,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1687
    hypreal_add,cis_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1688
qed "hcis_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1689
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1690
Goalw [hcomplex_one_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1691
  "hcis 0 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1692
by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1693
qed "hcis_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1694
Addsimps [hcis_zero];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1695
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1696
Goalw [hcomplex_zero_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1697
  "hrcis 0 a = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1698
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1699
by (auto_tac (claset(),simpset() addsimps [hrcis,hypreal_zero_num]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1700
qed "hrcis_zero_mod";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1701
Addsimps [hrcis_zero_mod];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1702
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1703
Goal "hrcis r 0 = hcomplex_of_hypreal r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1704
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1705
by (auto_tac (claset(),simpset() addsimps [hrcis,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1706
    hypreal_zero_num,hcomplex_of_hypreal]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1707
qed "hrcis_zero_arg";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1708
Addsimps [hrcis_zero_arg];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1709
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1710
Goal "iii * (iii * x) = - x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1711
by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1712
qed "hcomplex_i_mult_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1713
Addsimps [hcomplex_i_mult_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1714
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1715
Goal "iii * iii * x = - x";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1716
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1717
qed "hcomplex_i_mult_minus2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1718
Addsimps [hcomplex_i_mult_minus2];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1719
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1720
(* Move to one of the hyperreal theories *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1721
Goal "hypreal_of_nat m = Abs_hypreal(hyprel `` {%n. real m})";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1722
by (induct_tac "m" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1723
by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1724
    hypreal_of_nat_Suc,hypreal_zero_num,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1725
    hypreal_one_num,hypreal_add,real_of_nat_Suc]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1726
qed "hypreal_of_nat";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1727
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1728
Goal
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1729
   "hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1730
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1731
by (auto_tac (claset(),simpset() addsimps [hypreal_of_nat,hcis,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1732
    hypreal_mult,hcomplex_mult,cis_real_of_nat_Suc_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1733
qed "hcis_hypreal_of_nat_Suc_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1734
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1735
Goal "(hcis a) ^ n = hcis (hypreal_of_nat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1736
by (induct_tac "n" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1737
by (auto_tac (claset(),simpset() addsimps [hcis_hypreal_of_nat_Suc_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1738
qed "NSDeMoivre";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1739
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1740
Goal "hcis (hypreal_of_hypnat (n + 1) * a) = \
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1741
\     hcis a * hcis (hypreal_of_hypnat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1742
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1743
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1744
by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_of_hypnat,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1745
    hypnat_add,hypnat_one_def,hypreal_mult,hcomplex_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1746
    cis_real_of_nat_Suc_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1747
qed "hcis_hypreal_of_hypnat_Suc_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1748
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1749
Goal "(hcis a) hcpow n = hcis (hypreal_of_hypnat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1750
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1751
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1752
by (auto_tac (claset(),simpset() addsimps [hcis,hypreal_of_hypnat,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1753
    hypreal_mult,hcpow,DeMoivre]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1754
qed "NSDeMoivre_ext";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1755
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1756
Goalw [hrcis_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1757
  "(hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1758
by (auto_tac (claset(),simpset() addsimps [hcomplexpow_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1759
    NSDeMoivre,hcomplex_of_hypreal_pow]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1760
qed "DeMoivre2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1761
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1762
Goalw [hrcis_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1763
  "(hrcis r a) hcpow n = hrcis (r pow n) (hypreal_of_hypnat n * a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1764
by (auto_tac (claset(),simpset() addsimps [hcpow_mult,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1765
    NSDeMoivre_ext,hcomplex_of_hypreal_hyperpow]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1766
qed "DeMoivre2_ext";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1767
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1768
Goal "inverse(hcis a) = hcis (-a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1769
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1770
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hcis,hypreal_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1771
qed "hcis_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1772
Addsimps [hcis_inverse];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1773
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1774
Goal "inverse(hrcis r a) = hrcis (inverse r) (-a)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1775
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1776
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1777
by (auto_tac (claset(),simpset() addsimps [hcomplex_inverse,hrcis,hypreal_minus,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1778
    hypreal_inverse,rcis_inverse]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1779
by (Ultra_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1780
by (rewtac real_divide_def);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1781
by (auto_tac (claset(),simpset() addsimps [INVERSE_ZERO]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1782
qed "hrcis_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1783
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1784
Goal "hRe(hcis a) = ( *f* cos) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1785
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1786
by (auto_tac (claset(),simpset() addsimps [hcis,starfun,hRe]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1787
qed "hRe_hcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1788
Addsimps [hRe_hcis];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1789
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1790
Goal "hIm(hcis a) = ( *f* sin) a";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1791
by (res_inst_tac [("z","a")] eq_Abs_hypreal 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1792
by (auto_tac (claset(),simpset() addsimps [hcis,starfun,hIm]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1793
qed "hIm_hcis";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1794
Addsimps [hIm_hcis];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1795
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1796
Goal "( *f* cos) (hypreal_of_nat n * a) = hRe(hcis a ^ n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1797
by (auto_tac (claset(),simpset() addsimps [NSDeMoivre]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1798
qed "cos_n_hRe_hcis_pow_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1799
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1800
Goal "( *f* sin) (hypreal_of_nat n * a) = hIm(hcis a ^ n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1801
by (auto_tac (claset(),simpset() addsimps [NSDeMoivre]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1802
qed "sin_n_hIm_hcis_pow_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1803
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1804
Goal "( *f* cos) (hypreal_of_hypnat n * a) = hRe(hcis a hcpow n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1805
by (auto_tac (claset(),simpset() addsimps [NSDeMoivre_ext]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1806
qed "cos_n_hRe_hcis_hcpow_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1807
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1808
Goal "( *f* sin) (hypreal_of_hypnat n * a) = hIm(hcis a hcpow n)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1809
by (auto_tac (claset(),simpset() addsimps [NSDeMoivre_ext]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1810
qed "sin_n_hIm_hcis_hcpow_n";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1811
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1812
Goalw [hexpi_def] "hexpi(a + b) = hexpi(a) * hexpi(b)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1813
by (res_inst_tac [("z","a")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1814
by (res_inst_tac [("z","b")] eq_Abs_hcomplex 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1815
by (auto_tac (claset(),simpset() addsimps [hcis,hRe,hIm,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1816
    hcomplex_add,hcomplex_mult,hypreal_mult,starfun,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1817
    hcomplex_of_hypreal,cis_mult RS sym,complex_Im_add,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1818
    complex_Re_add,exp_add,complex_of_real_mult]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1819
qed "hexpi_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1820
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1821
(*----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1822
(* hcomplex_of_complex  preserves field and order properties                        *)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1823
(*----------------------------------------------------------------------------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1824
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1825
Goalw [hcomplex_of_complex_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1826
     "hcomplex_of_complex (z1 + z2) = hcomplex_of_complex z1 + hcomplex_of_complex z2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1827
by (simp_tac (simpset() addsimps [hcomplex_add]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1828
qed "hcomplex_of_complex_add";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1829
Addsimps [hcomplex_of_complex_add];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1830
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1831
Goalw [hcomplex_of_complex_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1832
     "hcomplex_of_complex (z1 * z2) = hcomplex_of_complex z1 * hcomplex_of_complex z2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1833
by (simp_tac (simpset() addsimps [hcomplex_mult]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1834
qed "hcomplex_of_complex_mult";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1835
Addsimps [hcomplex_of_complex_mult];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1836
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1837
Goalw [hcomplex_of_complex_def]
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1838
 "(hcomplex_of_complex z1 = hcomplex_of_complex z2) = (z1 = z2)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1839
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1840
qed "hcomplex_of_complex_eq_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1841
Addsimps [hcomplex_of_complex_eq_iff];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1842
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1843
Goalw [hcomplex_of_complex_def] "hcomplex_of_complex (-r) = - hcomplex_of_complex  r";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1844
by (auto_tac (claset(),simpset() addsimps [hcomplex_minus]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1845
qed "hcomplex_of_complex_minus";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1846
Addsimps [hcomplex_of_complex_minus];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1847
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1848
Goalw [hcomplex_of_complex_def,hcomplex_one_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1849
      "hcomplex_of_complex 1 = 1";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1850
by Auto_tac;
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1851
qed "hcomplex_of_complex_one";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1852
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1853
Goalw [hcomplex_of_complex_def,hcomplex_zero_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1854
      "hcomplex_of_complex 0 = 0";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1855
by (Simp_tac 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1856
qed "hcomplex_of_complex_zero";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1857
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1858
Goal "(hcomplex_of_complex r = 0) = (r = 0)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1859
by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1860
    simpset() addsimps [hcomplex_of_complex_def,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1861
                        hcomplex_zero_def]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1862
qed "hcomplex_of_complex_zero_iff";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1863
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1864
Goal "hcomplex_of_complex (inverse r) = inverse (hcomplex_of_complex r)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1865
by (case_tac "r=0" 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1866
by (asm_simp_tac (simpset() addsimps [COMPLEX_INVERSE_ZERO, 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1867
                              HCOMPLEX_INVERSE_ZERO, hcomplex_of_complex_zero,
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1868
                              COMPLEX_DIVIDE_ZERO]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1869
by (res_inst_tac [("c1","hcomplex_of_complex r")] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1870
    (hcomplex_mult_left_cancel RS iffD1) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1871
by (stac (hcomplex_of_complex_mult RS sym) 2); 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1872
by (auto_tac (claset(), 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1873
         simpset() addsimps [hcomplex_of_complex_one, hcomplex_of_complex_zero_iff]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1874
qed "hcomplex_of_complex_inverse";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1875
Addsimps [hcomplex_of_complex_inverse];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1876
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1877
Goal "hcomplex_of_complex (z1 / z2) = hcomplex_of_complex z1 / hcomplex_of_complex z2";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1878
by (simp_tac (simpset() addsimps [hcomplex_divide_def, complex_divide_def]) 1);
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1879
qed "hcomplex_of_complex_divide"; 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1880
Addsimps [hcomplex_of_complex_divide];
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1881
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1882
Goalw [hcomplex_of_complex_def,hypreal_of_real_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1883
   "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1884
by (auto_tac (claset(),simpset() addsimps [hRe]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1885
qed "hRe_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1886
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1887
Goalw [hcomplex_of_complex_def,hypreal_of_real_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1888
   "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1889
by (auto_tac (claset(),simpset() addsimps [hIm]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1890
qed "hIm_hcomplex_of_complex";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1891
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1892
Goalw [hypreal_of_real_def,hcomplex_of_complex_def] 
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1893
     "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)";
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1894
by (auto_tac (claset(),simpset() addsimps [hcmod]));
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1895
qed "hcmod_hcomplex_of_complex";