src/HOL/Complex/Complex.thy
author nipkow
Wed, 18 Aug 2004 11:09:40 +0200
changeset 15140 322485b816ac
parent 15131 c69542757a4d
child 15234 ec91a90c604e
permissions -rw-r--r--
import -> imports
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:       Complex.thy
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
     2
    ID:      $Id$
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     3
    Author:      Jacques D. Fleuriot
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
     4
    Copyright:   2001 University of Edinburgh
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
     5
    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
13957
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
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
     8
header {* Complex Numbers: Rectangular and Polar Representations *}
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    10
theory Complex
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    11
imports "../Hyperreal/HLog"
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    12
begin
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    13
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    14
datatype complex = Complex real real
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    15
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14443
diff changeset
    16
instance complex :: "{zero, one, plus, times, minus, inverse, power}" ..
13957
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
consts
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    19
  "ii"    :: complex    ("\<i>")
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    20
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    21
consts Re :: "complex => real"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    22
primrec "Re (Complex x y) = x"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    23
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    24
consts Im :: "complex => real"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    25
primrec "Im (Complex x y) = y"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    26
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    27
lemma complex_surj [simp]: "Complex (Re z) (Im z) = z"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    28
  by (induct z) simp
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    29
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    30
constdefs
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
  (*----------- modulus ------------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    33
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    34
  cmod :: "complex => real"
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    35
  "cmod z == sqrt(Re(z) ^ 2 + Im(z) ^ 2)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    36
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    37
  (*----- injection from reals -----*)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    38
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    39
  complex_of_real :: "real => complex"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    40
  "complex_of_real r == Complex r 0"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    41
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    42
  (*------- complex conjugate ------*)
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    43
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    44
  cnj :: "complex => complex"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    45
  "cnj z == Complex (Re z) (-Im z)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    46
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    47
  (*------------ Argand -------------*)
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    48
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    49
  sgn :: "complex => complex"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    50
  "sgn z == z / complex_of_real(cmod z)"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    51
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    52
  arg :: "complex => real"
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
    53
  "arg z == @a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \<le> pi"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    54
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    55
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    56
defs (overloaded)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    57
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    58
  complex_zero_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    59
  "0 == Complex 0 0"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    60
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    61
  complex_one_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    62
  "1 == Complex 1 0"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    63
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    64
  i_def: "ii == Complex 0 1"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    65
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    66
  complex_minus_def: "- z == Complex (- Re z) (- Im z)"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    67
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    68
  complex_inverse_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    69
   "inverse z ==
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    70
    Complex (Re z / ((Re z)\<twosuperior> + (Im z)\<twosuperior>)) (- Im z / ((Re z)\<twosuperior> + (Im z)\<twosuperior>))"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    71
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    72
  complex_add_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    73
    "z + w == Complex (Re z + Re w) (Im z + Im w)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    74
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    75
  complex_diff_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    76
    "z - w == z + - (w::complex)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    77
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
    78
  complex_mult_def:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    79
    "z * w == Complex (Re z * Re w - Im z * Im w) (Re z * Im w + Im z * Re w)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    80
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    81
  complex_divide_def: "w / (z::complex) == w * inverse z"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    82
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    83
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    84
constdefs
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    85
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    86
  (* abbreviation for (cos a + i sin a) *)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    87
  cis :: "real => complex"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
    88
  "cis a == Complex (cos a) (sin a)"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    89
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    90
  (* abbreviation for r*(cos a + i sin a) *)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    91
  rcis :: "[real, real] => complex"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    92
  "rcis r a == complex_of_real r * cis a"
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    93
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    94
  (* e ^ (x + iy) *)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    95
  expi :: "complex => complex"
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
    96
  "expi z == complex_of_real(exp (Re z)) * cis (Im z)"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    97
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
    98
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
    99
lemma complex_equality [intro?]: "Re z = Re w ==> Im z = Im w ==> z = w"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   100
  by (induct z, induct w) simp
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   101
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   102
lemma Re [simp]: "Re(Complex x y) = x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   103
by simp
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   104
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   105
lemma Im [simp]: "Im(Complex x y) = y"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   106
by simp
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   107
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   108
lemma complex_Re_Im_cancel_iff: "(w=z) = (Re(w) = Re(z) & Im(w) = Im(z))"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   109
by (induct w, induct z, simp)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   110
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   111
lemma complex_Re_zero [simp]: "Re 0 = 0"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   112
by (simp add: complex_zero_def)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   113
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   114
lemma complex_Im_zero [simp]: "Im 0 = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   115
by (simp add: complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   116
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   117
lemma complex_Re_one [simp]: "Re 1 = 1"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   118
by (simp add: complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   119
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   120
lemma complex_Im_one [simp]: "Im 1 = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   121
by (simp add: complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   122
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   123
lemma complex_Re_i [simp]: "Re(ii) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   124
by (simp add: i_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   125
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   126
lemma complex_Im_i [simp]: "Im(ii) = 1"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   127
by (simp add: i_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   128
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   129
lemma Re_complex_of_real [simp]: "Re(complex_of_real z) = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   130
by (simp add: complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   131
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   132
lemma Im_complex_of_real [simp]: "Im(complex_of_real z) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   133
by (simp add: complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   134
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   135
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   136
subsection{*Unary Minus*}
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   137
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   138
lemma complex_minus [simp]: "- (Complex x y) = Complex (-x) (-y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   139
by (simp add: complex_minus_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   140
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   141
lemma complex_Re_minus [simp]: "Re (-z) = - Re z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   142
by (simp add: complex_minus_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   143
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   144
lemma complex_Im_minus [simp]: "Im (-z) = - Im z"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   145
by (simp add: complex_minus_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   146
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   147
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   148
subsection{*Addition*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   149
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   150
lemma complex_add [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   151
     "Complex x1 y1 + Complex x2 y2 = Complex (x1+x2) (y1+y2)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   152
by (simp add: complex_add_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   153
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   154
lemma complex_Re_add [simp]: "Re(x + y) = Re(x) + Re(y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   155
by (simp add: complex_add_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   156
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   157
lemma complex_Im_add [simp]: "Im(x + y) = Im(x) + Im(y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   158
by (simp add: complex_add_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   159
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   160
lemma complex_add_commute: "(u::complex) + v = v + u"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   161
by (simp add: complex_add_def add_commute)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   162
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   163
lemma complex_add_assoc: "((u::complex) + v) + w = u + (v + w)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   164
by (simp add: complex_add_def add_assoc)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   165
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   166
lemma complex_add_zero_left: "(0::complex) + z = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   167
by (simp add: complex_add_def complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   168
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   169
lemma complex_add_zero_right: "z + (0::complex) = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   170
by (simp add: complex_add_def complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   171
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   172
lemma complex_add_minus_left: "-z + z = (0::complex)"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   173
by (simp add: complex_add_def complex_minus_def complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   174
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   175
lemma complex_diff:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   176
      "Complex x1 y1 - Complex x2 y2 = Complex (x1-x2) (y1-y2)"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   177
by (simp add: complex_add_def complex_minus_def complex_diff_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   178
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   179
lemma complex_Re_diff [simp]: "Re(x - y) = Re(x) - Re(y)"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   180
by (simp add: complex_diff_def)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   181
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   182
lemma complex_Im_diff [simp]: "Im(x - y) = Im(x) - Im(y)"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   183
by (simp add: complex_diff_def)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   184
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   185
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   186
subsection{*Multiplication*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   187
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   188
lemma complex_mult [simp]:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   189
     "Complex x1 y1 * Complex x2 y2 = Complex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   190
by (simp add: complex_mult_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   191
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   192
lemma complex_mult_commute: "(w::complex) * z = z * w"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   193
by (simp add: complex_mult_def mult_commute add_commute)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   194
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   195
lemma complex_mult_assoc: "((u::complex) * v) * w = u * (v * w)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   196
by (simp add: complex_mult_def mult_ac add_ac
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   197
              right_diff_distrib right_distrib left_diff_distrib left_distrib)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   198
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   199
lemma complex_mult_one_left: "(1::complex) * z = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   200
by (simp add: complex_mult_def complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   201
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   202
lemma complex_mult_one_right: "z * (1::complex) = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   203
by (simp add: complex_mult_def complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   204
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   205
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   206
subsection{*Inverse*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   207
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   208
lemma complex_inverse [simp]:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   209
     "inverse (Complex x y) = Complex (x/(x ^ 2 + y ^ 2)) (-y/(x ^ 2 + y ^ 2))"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   210
by (simp add: complex_inverse_def)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   211
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   212
lemma complex_mult_inv_left: "z \<noteq> (0::complex) ==> inverse(z) * z = 1"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   213
apply (induct z)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   214
apply (rename_tac x y)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   215
apply (auto simp add: complex_mult complex_inverse complex_one_def
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   216
      complex_zero_def add_divide_distrib [symmetric] power2_eq_square mult_ac)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   217
apply (drule_tac y = y in real_sum_squares_not_zero)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   218
apply (drule_tac [2] x = x in real_sum_squares_not_zero2, auto)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   219
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   220
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   221
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   222
subsection {* The field of complex numbers *}
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   223
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   224
instance complex :: field
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   225
proof
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   226
  fix z u v w :: complex
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   227
  show "(u + v) + w = u + (v + w)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   228
    by (rule complex_add_assoc)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   229
  show "z + w = w + z"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   230
    by (rule complex_add_commute)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   231
  show "0 + z = z"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   232
    by (rule complex_add_zero_left)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   233
  show "-z + z = 0"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   234
    by (rule complex_add_minus_left)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   235
  show "z - w = z + -w"
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   236
    by (simp add: complex_diff_def)
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   237
  show "(u * v) * w = u * (v * w)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   238
    by (rule complex_mult_assoc)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   239
  show "z * w = w * z"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   240
    by (rule complex_mult_commute)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   241
  show "1 * z = z"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   242
    by (rule complex_mult_one_left)
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14335
diff changeset
   243
  show "0 \<noteq> (1::complex)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   244
    by (simp add: complex_zero_def complex_one_def)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   245
  show "(u + v) * w = u * w + v * w"
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14387
diff changeset
   246
    by (simp add: complex_mult_def complex_add_def left_distrib 
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14387
diff changeset
   247
                  diff_minus add_ac)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   248
  show "z / w = z * inverse w"
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   249
    by (simp add: complex_divide_def)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   250
  assume "w \<noteq> 0"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   251
  thus "inverse w * w = 1"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   252
    by (simp add: complex_mult_inv_left)
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   253
qed
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   254
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   255
instance complex :: division_by_zero
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   256
proof
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   257
  show "inverse 0 = (0::complex)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   258
    by (simp add: complex_inverse_def complex_zero_def)
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   259
qed
14335
9c0b5e081037 conversion of Real/PReal to Isar script;
paulson
parents: 14334
diff changeset
   260
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   261
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   262
subsection{*Embedding Properties for @{term complex_of_real} Map*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   263
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   264
lemma Complex_add_complex_of_real [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   265
     "Complex x y + complex_of_real r = Complex (x+r) y"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   266
by (simp add: complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   267
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   268
lemma complex_of_real_add_Complex [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   269
     "complex_of_real r + Complex x y = Complex (r+x) y"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   270
by (simp add: i_def complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   271
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   272
lemma Complex_mult_complex_of_real:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   273
     "Complex x y * complex_of_real r = Complex (x*r) (y*r)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   274
by (simp add: complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   275
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   276
lemma complex_of_real_mult_Complex:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   277
     "complex_of_real r * Complex x y = Complex (r*x) (r*y)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   278
by (simp add: i_def complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   279
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   280
lemma i_complex_of_real [simp]: "ii * complex_of_real r = Complex 0 r"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   281
by (simp add: i_def complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   282
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   283
lemma complex_of_real_i [simp]: "complex_of_real r * ii = Complex 0 r"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   284
by (simp add: i_def complex_of_real_def)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   285
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   286
lemma complex_of_real_one [simp]: "complex_of_real 1 = 1"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   287
by (simp add: complex_one_def complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   288
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   289
lemma complex_of_real_zero [simp]: "complex_of_real 0 = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   290
by (simp add: complex_zero_def complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   291
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   292
lemma complex_of_real_eq_iff [iff]:
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   293
     "(complex_of_real x = complex_of_real y) = (x = y)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   294
by (simp add: complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   295
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   296
lemma complex_of_real_minus [simp]: "complex_of_real(-x) = - complex_of_real x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   297
by (simp add: complex_of_real_def complex_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   298
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   299
lemma complex_of_real_inverse [simp]:
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   300
     "complex_of_real(inverse x) = inverse(complex_of_real x)"
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   301
apply (case_tac "x=0", simp)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   302
apply (simp add: complex_of_real_def divide_inverse power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   303
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   304
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   305
lemma complex_of_real_add [simp]:
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   306
     "complex_of_real (x + y) = complex_of_real x + complex_of_real y"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   307
by (simp add: complex_add complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   308
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   309
lemma complex_of_real_diff [simp]:
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   310
     "complex_of_real (x - y) = complex_of_real x - complex_of_real y"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   311
by (simp add: complex_of_real_minus diff_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   312
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   313
lemma complex_of_real_mult [simp]:
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   314
     "complex_of_real (x * y) = complex_of_real x * complex_of_real y"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   315
by (simp add: complex_mult complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   316
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   317
lemma complex_of_real_divide [simp]:
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   318
      "complex_of_real(x/y) = complex_of_real x / complex_of_real y"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   319
apply (simp add: complex_divide_def)
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   320
apply (case_tac "y=0", simp)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   321
apply (simp add: complex_of_real_mult complex_of_real_inverse 
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   322
                 divide_inverse)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   323
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   324
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   325
lemma complex_mod [simp]: "cmod (Complex x y) = sqrt(x ^ 2 + y ^ 2)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   326
by (simp add: cmod_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   327
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   328
lemma complex_mod_zero [simp]: "cmod(0) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   329
by (simp add: cmod_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   330
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   331
lemma complex_mod_one [simp]: "cmod(1) = 1"
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   332
by (simp add: cmod_def power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   333
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   334
lemma complex_mod_complex_of_real [simp]: "cmod(complex_of_real x) = abs x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   335
by (simp add: complex_of_real_def power2_eq_square complex_mod)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   336
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   337
lemma complex_of_real_abs:
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   338
     "complex_of_real (abs x) = complex_of_real(cmod(complex_of_real x))"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   339
by simp
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   340
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   341
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   342
subsection{*The Functions @{term Re} and @{term Im}*}
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   343
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   344
lemma complex_Re_mult_eq: "Re (w * z) = Re w * Re z - Im w * Im z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   345
by (induct z, induct w, simp add: complex_mult)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   346
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   347
lemma complex_Im_mult_eq: "Im (w * z) = Re w * Im z + Im w * Re z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   348
by (induct z, induct w, simp add: complex_mult)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   349
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   350
lemma Re_i_times [simp]: "Re(ii * z) = - Im z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   351
by (simp add: complex_Re_mult_eq) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   352
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   353
lemma Re_times_i [simp]: "Re(z * ii) = - Im z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   354
by (simp add: complex_Re_mult_eq) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   355
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   356
lemma Im_i_times [simp]: "Im(ii * z) = Re z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   357
by (simp add: complex_Im_mult_eq) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   358
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   359
lemma Im_times_i [simp]: "Im(z * ii) = Re z"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   360
by (simp add: complex_Im_mult_eq) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   361
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   362
lemma complex_Re_mult: "[| Im w = 0; Im z = 0 |] ==> Re(w * z) = Re(w) * Re(z)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   363
by (simp add: complex_Re_mult_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   364
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   365
lemma complex_Re_mult_complex_of_real [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   366
     "Re (z * complex_of_real c) = Re(z) * c"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   367
by (simp add: complex_Re_mult_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   368
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   369
lemma complex_Im_mult_complex_of_real [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   370
     "Im (z * complex_of_real c) = Im(z) * c"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   371
by (simp add: complex_Im_mult_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   372
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   373
lemma complex_Re_mult_complex_of_real2 [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   374
     "Re (complex_of_real c * z) = c * Re(z)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   375
by (simp add: complex_Re_mult_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   376
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   377
lemma complex_Im_mult_complex_of_real2 [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   378
     "Im (complex_of_real c * z) = c * Im(z)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   379
by (simp add: complex_Im_mult_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   380
 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   381
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   382
subsection{*Conjugation is an Automorphism*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   383
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   384
lemma complex_cnj: "cnj (Complex x y) = Complex x (-y)"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   385
by (simp add: cnj_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   386
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   387
lemma complex_cnj_cancel_iff [simp]: "(cnj x = cnj y) = (x = y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   388
by (simp add: cnj_def complex_Re_Im_cancel_iff)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   389
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   390
lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   391
by (simp add: cnj_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   392
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   393
lemma complex_cnj_complex_of_real [simp]:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   394
     "cnj (complex_of_real x) = complex_of_real x"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   395
by (simp add: complex_of_real_def complex_cnj)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   396
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   397
lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   398
by (induct z, simp add: complex_cnj complex_mod power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   399
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   400
lemma complex_cnj_minus: "cnj (-z) = - cnj z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   401
by (simp add: cnj_def complex_minus complex_Re_minus complex_Im_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   402
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   403
lemma complex_cnj_inverse: "cnj(inverse z) = inverse(cnj z)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   404
by (induct z, simp add: complex_cnj complex_inverse power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   405
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   406
lemma complex_cnj_add: "cnj(w + z) = cnj(w) + cnj(z)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   407
by (induct w, induct z, simp add: complex_cnj complex_add)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   408
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   409
lemma complex_cnj_diff: "cnj(w - z) = cnj(w) - cnj(z)"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   410
by (simp add: diff_minus complex_cnj_add complex_cnj_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   411
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   412
lemma complex_cnj_mult: "cnj(w * z) = cnj(w) * cnj(z)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   413
by (induct w, induct z, simp add: complex_cnj complex_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   414
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   415
lemma complex_cnj_divide: "cnj(w / z) = (cnj w)/(cnj z)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   416
by (simp add: complex_divide_def complex_cnj_mult complex_cnj_inverse)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   417
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   418
lemma complex_cnj_one [simp]: "cnj 1 = 1"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   419
by (simp add: cnj_def complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   420
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   421
lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re(z))"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   422
by (induct z, simp add: complex_add complex_cnj complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   423
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   424
lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im(z)) * ii"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   425
apply (induct z)
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   426
apply (simp add: complex_add complex_cnj complex_of_real_def diff_minus
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   427
                 complex_minus i_def complex_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   428
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   429
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   430
lemma complex_cnj_zero [simp]: "cnj 0 = 0"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   431
by (simp add: cnj_def complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   432
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   433
lemma complex_cnj_zero_iff [iff]: "(cnj z = 0) = (z = 0)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   434
by (induct z, simp add: complex_zero_def complex_cnj)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   435
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   436
lemma complex_mult_cnj: "z * cnj z = complex_of_real (Re(z) ^ 2 + Im(z) ^ 2)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   437
by (induct z,
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   438
    simp add: complex_cnj complex_mult complex_of_real_def power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   439
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   440
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   441
subsection{*Modulus*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   442
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   443
lemma complex_mod_eq_zero_cancel [simp]: "(cmod x = 0) = (x = 0)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   444
apply (induct x)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15013
diff changeset
   445
apply (auto iff: real_0_le_add_iff 
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15013
diff changeset
   446
            intro: real_sum_squares_cancel real_sum_squares_cancel2
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   447
            simp add: complex_mod complex_zero_def power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   448
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   449
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   450
lemma complex_mod_complex_of_real_of_nat [simp]:
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   451
     "cmod (complex_of_real(real (n::nat))) = real n"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   452
by simp
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   453
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   454
lemma complex_mod_minus [simp]: "cmod (-x) = cmod(x)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   455
by (induct x, simp add: complex_mod complex_minus power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   456
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   457
lemma complex_mod_mult_cnj: "cmod(z * cnj(z)) = cmod(z) ^ 2"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   458
apply (induct z, simp add: complex_mod complex_cnj complex_mult)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15013
diff changeset
   459
apply (simp add: power2_eq_square abs_if linorder_not_less real_0_le_add_iff)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   460
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   461
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   462
lemma complex_mod_squared: "cmod(Complex x y) ^ 2 = x ^ 2 + y ^ 2"
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   463
by (simp add: cmod_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   464
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   465
lemma complex_mod_ge_zero [simp]: "0 \<le> cmod x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   466
by (simp add: cmod_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   467
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   468
lemma abs_cmod_cancel [simp]: "abs(cmod x) = cmod x"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   469
by (simp add: abs_if linorder_not_less)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   470
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   471
lemma complex_mod_mult: "cmod(x*y) = cmod(x) * cmod(y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   472
apply (induct x, induct y)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   473
apply (auto simp add: complex_mult complex_mod real_sqrt_mult_distrib2[symmetric])
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   474
apply (rule_tac n = 1 in power_inject_base)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   475
apply (auto simp add: power2_eq_square [symmetric] simp del: realpow_Suc)
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   476
apply (auto simp add: real_diff_def power2_eq_square right_distrib left_distrib 
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   477
                      add_ac mult_ac)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   478
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   479
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   480
lemma cmod_unit_one [simp]: "cmod (Complex (cos a) (sin a)) = 1"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   481
by (simp add: cmod_def) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   482
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   483
lemma cmod_complex_polar [simp]:
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   484
     "cmod (complex_of_real r * Complex (cos a) (sin a)) = abs r"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   485
by (simp only: cmod_unit_one complex_mod_mult, simp) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   486
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   487
lemma complex_mod_add_squared_eq:
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   488
     "cmod(x + y) ^ 2 = cmod(x) ^ 2 + cmod(y) ^ 2 + 2 * Re(x * cnj y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   489
apply (induct x, induct y)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   490
apply (auto simp add: complex_add complex_mod_squared complex_mult complex_cnj real_diff_def simp del: realpow_Suc)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   491
apply (auto simp add: right_distrib left_distrib power2_eq_square mult_ac add_ac)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   492
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   493
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   494
lemma complex_Re_mult_cnj_le_cmod [simp]: "Re(x * cnj y) \<le> cmod(x * cnj y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   495
apply (induct x, induct y)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   496
apply (auto simp add: complex_mod complex_mult complex_cnj real_diff_def simp del: realpow_Suc)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   497
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   498
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   499
lemma complex_Re_mult_cnj_le_cmod2 [simp]: "Re(x * cnj y) \<le> cmod(x * y)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   500
by (insert complex_Re_mult_cnj_le_cmod [of x y], simp add: complex_mod_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   501
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   502
lemma real_sum_squared_expand:
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   503
     "((x::real) + y) ^ 2 = x ^ 2 + y ^ 2 + 2 * x * y"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   504
by (simp add: left_distrib right_distrib power2_eq_square)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   505
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   506
lemma complex_mod_triangle_squared [simp]:
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   507
     "cmod (x + y) ^ 2 \<le> (cmod(x) + cmod(y)) ^ 2"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   508
by (simp add: real_sum_squared_expand complex_mod_add_squared_eq real_mult_assoc complex_mod_mult [symmetric])
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   509
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   510
lemma complex_mod_minus_le_complex_mod [simp]: "- cmod x \<le> cmod x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   511
by (rule order_trans [OF _ complex_mod_ge_zero], simp)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   512
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   513
lemma complex_mod_triangle_ineq [simp]: "cmod (x + y) \<le> cmod(x) + cmod(y)"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   514
apply (rule_tac n = 1 in realpow_increasing)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   515
apply (auto intro:  order_trans [OF _ complex_mod_ge_zero]
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15013
diff changeset
   516
            simp add: add_increasing power2_eq_square [symmetric])
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   517
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   518
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   519
lemma complex_mod_triangle_ineq2 [simp]: "cmod(b + a) - cmod b \<le> cmod a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   520
by (insert complex_mod_triangle_ineq [THEN add_right_mono, of b a"-cmod b"], simp)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   521
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   522
lemma complex_mod_diff_commute: "cmod (x - y) = cmod (y - x)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   523
apply (induct x, induct y)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   524
apply (auto simp add: complex_diff complex_mod right_diff_distrib power2_eq_square left_diff_distrib add_ac mult_ac)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   525
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   526
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   527
lemma complex_mod_add_less:
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   528
     "[| cmod x < r; cmod y < s |] ==> cmod (x + y) < r + s"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   529
by (auto intro: order_le_less_trans complex_mod_triangle_ineq)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   530
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   531
lemma complex_mod_mult_less:
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   532
     "[| cmod x < r; cmod y < s |] ==> cmod (x * y) < r * s"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   533
by (auto intro: real_mult_less_mono' simp add: complex_mod_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   534
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   535
lemma complex_mod_diff_ineq [simp]: "cmod(a) - cmod(b) \<le> cmod(a + b)"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   536
apply (rule linorder_cases [of "cmod(a)" "cmod (b)"])
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   537
apply auto
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   538
apply (rule order_trans [of _ 0], rule order_less_imp_le)
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   539
apply (simp add: compare_rls, simp)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   540
apply (simp add: compare_rls)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   541
apply (rule complex_mod_minus [THEN subst])
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   542
apply (rule order_trans)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   543
apply (rule_tac [2] complex_mod_triangle_ineq)
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   544
apply (auto simp add: add_ac)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   545
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   546
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   547
lemma complex_Re_le_cmod [simp]: "Re z \<le> cmod z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   548
by (induct z, simp add: complex_mod del: realpow_Suc)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   549
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   550
lemma complex_mod_gt_zero: "z \<noteq> 0 ==> 0 < cmod z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   551
apply (insert complex_mod_ge_zero [of z])
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   552
apply (drule order_le_imp_less_or_eq, auto)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   553
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   554
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   555
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   556
subsection{*A Few More Theorems*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   557
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   558
lemma complex_mod_inverse: "cmod(inverse x) = inverse(cmod x)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   559
apply (case_tac "x=0", simp)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   560
apply (rule_tac c1 = "cmod x" in real_mult_left_cancel [THEN iffD1])
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   561
apply (auto simp add: complex_mod_mult [symmetric])
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   562
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   563
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   564
lemma complex_mod_divide: "cmod(x/y) = cmod(x)/(cmod y)"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   565
by (simp add: complex_divide_def divide_inverse complex_mod_mult complex_mod_inverse)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   566
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   567
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   568
subsection{*Exponentiation*}
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   569
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   570
primrec
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   571
     complexpow_0:   "z ^ 0       = 1"
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   572
     complexpow_Suc: "z ^ (Suc n) = (z::complex) * (z ^ n)"
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   573
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   574
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14691
diff changeset
   575
instance complex :: recpower
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   576
proof
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   577
  fix z :: complex
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   578
  fix n :: nat
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   579
  show "z^0 = 1" by simp
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   580
  show "z^(Suc n) = z * (z^n)" by simp
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   581
qed
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   582
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   583
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   584
lemma complex_of_real_pow: "complex_of_real (x ^ n) = (complex_of_real x) ^ n"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   585
apply (induct_tac "n")
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   586
apply (auto simp add: complex_of_real_mult [symmetric])
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   587
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   588
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   589
lemma complex_cnj_pow: "cnj(z ^ n) = cnj(z) ^ n"
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   590
apply (induct_tac "n")
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   591
apply (auto simp add: complex_cnj_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   592
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   593
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   594
lemma complex_mod_complexpow: "cmod(x ^ n) = cmod(x) ^ n"
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   595
apply (induct_tac "n")
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   596
apply (auto simp add: complex_mod_mult)
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   597
done
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   598
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   599
lemma complexpow_i_squared [simp]: "ii ^ 2 = -(1::complex)"
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   600
by (simp add: i_def complex_mult complex_one_def complex_minus numeral_2_eq_2)
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   601
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   602
lemma complex_i_not_zero [simp]: "ii \<noteq> 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   603
by (simp add: i_def complex_zero_def)
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   604
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   605
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   606
subsection{*The Function @{term sgn}*}
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   607
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   608
lemma sgn_zero [simp]: "sgn 0 = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   609
by (simp add: sgn_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   610
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   611
lemma sgn_one [simp]: "sgn 1 = 1"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   612
by (simp add: sgn_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   613
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   614
lemma sgn_minus: "sgn (-z) = - sgn(z)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   615
by (simp add: sgn_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   616
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   617
lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   618
by (simp add: sgn_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   619
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   620
lemma i_mult_eq: "ii * ii = complex_of_real (-1)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   621
by (simp add: i_def complex_of_real_def complex_mult complex_add)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   622
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   623
lemma i_mult_eq2 [simp]: "ii * ii = -(1::complex)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   624
by (simp add: i_def complex_one_def complex_mult complex_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   625
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   626
lemma complex_eq_cancel_iff2 [simp]:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   627
     "(Complex x y = complex_of_real xa) = (x = xa & y = 0)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   628
by (simp add: complex_of_real_def) 
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   629
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   630
lemma complex_eq_cancel_iff2a [simp]:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   631
     "(Complex x y = complex_of_real xa) = (x = xa & y = 0)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   632
by (simp add: complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   633
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   634
lemma Complex_eq_0 [simp]: "(Complex x y = 0) = (x = 0 & y = 0)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   635
by (simp add: complex_zero_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   636
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   637
lemma Complex_eq_1 [simp]: "(Complex x y = 1) = (x = 1 & y = 0)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   638
by (simp add: complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   639
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   640
lemma Complex_eq_i [simp]: "(Complex x y = ii) = (x = 0 & y = 1)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   641
by (simp add: i_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   642
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   643
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   644
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   645
lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   646
proof (induct z)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   647
  case (Complex x y)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   648
    have "sqrt (x\<twosuperior> + y\<twosuperior>) * inverse (x\<twosuperior> + y\<twosuperior>) = inverse (sqrt (x\<twosuperior> + y\<twosuperior>))"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   649
      by (simp add: divide_inverse [symmetric] sqrt_divide_self_eq)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   650
    thus "Re (sgn (Complex x y)) = Re (Complex x y) /cmod (Complex x y)"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   651
       by (simp add: sgn_def complex_of_real_def divide_inverse)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   652
qed
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   653
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   654
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   655
lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   656
proof (induct z)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   657
  case (Complex x y)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   658
    have "sqrt (x\<twosuperior> + y\<twosuperior>) * inverse (x\<twosuperior> + y\<twosuperior>) = inverse (sqrt (x\<twosuperior> + y\<twosuperior>))"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   659
      by (simp add: divide_inverse [symmetric] sqrt_divide_self_eq)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   660
    thus "Im (sgn (Complex x y)) = Im (Complex x y) /cmod (Complex x y)"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   661
       by (simp add: sgn_def complex_of_real_def divide_inverse)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   662
qed
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   663
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   664
lemma complex_inverse_complex_split:
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   665
     "inverse(complex_of_real x + ii * complex_of_real y) =
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   666
      complex_of_real(x/(x ^ 2 + y ^ 2)) -
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   667
      ii * complex_of_real(y/(x ^ 2 + y ^ 2))"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   668
by (simp add: complex_of_real_def i_def complex_mult complex_add
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   669
         diff_minus complex_minus complex_inverse divide_inverse)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   670
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   671
(*----------------------------------------------------------------------------*)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   672
(* Many of the theorems below need to be moved elsewhere e.g. Transc. Also *)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   673
(* many of the theorems are not used - so should they be kept?                *)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   674
(*----------------------------------------------------------------------------*)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   675
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   676
lemma complex_of_real_zero_iff [simp]: "(complex_of_real y = 0) = (y = 0)"
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   677
by (auto simp add: complex_zero_def complex_of_real_def)
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   678
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   679
lemma cos_arg_i_mult_zero_pos:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   680
   "0 < y ==> cos (arg(Complex 0 y)) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   681
apply (simp add: arg_def abs_if)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   682
apply (rule_tac a = "pi/2" in someI2, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   683
apply (rule order_less_trans [of _ 0], auto)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   684
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   685
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   686
lemma cos_arg_i_mult_zero_neg:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   687
   "y < 0 ==> cos (arg(Complex 0 y)) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   688
apply (simp add: arg_def abs_if)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   689
apply (rule_tac a = "- pi/2" in someI2, auto)
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   690
apply (rule order_trans [of _ 0], auto)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   691
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   692
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   693
lemma cos_arg_i_mult_zero [simp]:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   694
     "y \<noteq> 0 ==> cos (arg(Complex 0 y)) = 0"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   695
by (auto simp add: linorder_neq_iff cos_arg_i_mult_zero_pos cos_arg_i_mult_zero_neg)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   696
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   697
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   698
subsection{*Finally! Polar Form for Complex Numbers*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   699
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   700
lemma complex_split_polar:
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   701
     "\<exists>r a. z = complex_of_real r * (Complex (cos a) (sin a))"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   702
apply (induct z) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   703
apply (auto simp add: polar_Ex complex_of_real_mult_Complex)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   704
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   705
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   706
lemma rcis_Ex: "\<exists>r a. z = rcis r a"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   707
apply (induct z) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   708
apply (simp add: rcis_def cis_def polar_Ex complex_of_real_mult_Complex)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   709
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   710
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   711
lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   712
by (simp add: rcis_def cis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   713
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   714
lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   715
by (simp add: rcis_def cis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   716
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   717
lemma sin_cos_squared_add2_mult: "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior>"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   718
proof -
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   719
  have "(r * cos a)\<twosuperior> + (r * sin a)\<twosuperior> = r\<twosuperior> * ((cos a)\<twosuperior> + (sin a)\<twosuperior>)"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   720
    by (simp only: power_mult_distrib right_distrib) 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   721
  thus ?thesis by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   722
qed
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   723
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   724
lemma complex_mod_rcis [simp]: "cmod(rcis r a) = abs r"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   725
by (simp add: rcis_def cis_def sin_cos_squared_add2_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   726
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   727
lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   728
apply (simp add: cmod_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   729
apply (rule real_sqrt_eq_iff [THEN iffD2])
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   730
apply (auto simp add: complex_mult_cnj)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   731
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   732
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   733
lemma complex_Re_cnj [simp]: "Re(cnj z) = Re z"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   734
by (induct z, simp add: complex_cnj)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   735
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   736
lemma complex_Im_cnj [simp]: "Im(cnj z) = - Im z"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   737
by (induct z, simp add: complex_cnj)
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   738
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   739
lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   740
by (induct z, simp add: complex_cnj complex_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   741
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   742
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   743
(*---------------------------------------------------------------------------*)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   744
(*  (r1 * cis a) * (r2 * cis b) = r1 * r2 * cis (a + b)                      *)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   745
(*---------------------------------------------------------------------------*)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   746
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   747
lemma cis_rcis_eq: "cis a = rcis 1 a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   748
by (simp add: rcis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   749
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   750
lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1*r2) (a + b)"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   751
by (simp add: rcis_def cis_def cos_add sin_add right_distrib right_diff_distrib
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   752
              complex_of_real_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   753
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   754
lemma cis_mult: "cis a * cis b = cis (a + b)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   755
by (simp add: cis_rcis_eq rcis_mult)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   756
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   757
lemma cis_zero [simp]: "cis 0 = 1"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   758
by (simp add: cis_def complex_one_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   759
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   760
lemma rcis_zero_mod [simp]: "rcis 0 a = 0"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   761
by (simp add: rcis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   762
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   763
lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   764
by (simp add: rcis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   765
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   766
lemma complex_of_real_minus_one:
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   767
   "complex_of_real (-(1::real)) = -(1::complex)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   768
by (simp add: complex_of_real_def complex_one_def complex_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   769
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   770
lemma complex_i_mult_minus [simp]: "ii * (ii * x) = - x"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   771
by (simp add: complex_mult_assoc [symmetric])
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   772
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   773
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   774
lemma cis_real_of_nat_Suc_mult:
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   775
   "cis (real (Suc n) * a) = cis a * cis (real n * a)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14374
diff changeset
   776
by (simp add: cis_def real_of_nat_Suc left_distrib cos_add sin_add right_distrib)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   777
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   778
lemma DeMoivre: "(cis a) ^ n = cis (real n * a)"
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   779
apply (induct_tac "n")
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   780
apply (auto simp add: cis_real_of_nat_Suc_mult)
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   781
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   782
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   783
lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   784
by (simp add: rcis_def power_mult_distrib DeMoivre complex_of_real_pow)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   785
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   786
lemma cis_inverse [simp]: "inverse(cis a) = cis (-a)"
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   787
by (simp add: cis_def complex_inverse_complex_split complex_of_real_minus 
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   788
              diff_minus)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   789
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   790
lemma rcis_inverse: "inverse(rcis r a) = rcis (1/r) (-a)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   791
by (simp add: divide_inverse rcis_def complex_of_real_inverse)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   792
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   793
lemma cis_divide: "cis a / cis b = cis (a - b)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   794
by (simp add: complex_divide_def cis_mult real_diff_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   795
14354
988aa4648597 types complex and hcomplex are now instances of class ringpower:
paulson
parents: 14353
diff changeset
   796
lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1/r2) (a - b)"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   797
apply (simp add: complex_divide_def)
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   798
apply (case_tac "r2=0", simp)
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   799
apply (simp add: rcis_inverse rcis_mult real_diff_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   800
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   801
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   802
lemma Re_cis [simp]: "Re(cis a) = cos a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   803
by (simp add: cis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   804
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   805
lemma Im_cis [simp]: "Im(cis a) = sin a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   806
by (simp add: cis_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   807
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   808
lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re(cis a ^ n)"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   809
by (auto simp add: DeMoivre)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   810
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   811
lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im(cis a ^ n)"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   812
by (auto simp add: DeMoivre)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   813
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   814
lemma expi_add: "expi(a + b) = expi(a) * expi(b)"
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   815
by (simp add: expi_def complex_Re_add exp_add complex_Im_add 
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   816
              cis_mult [symmetric] complex_of_real_mult mult_ac)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   817
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   818
lemma expi_zero [simp]: "expi (0::complex) = 1"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   819
by (simp add: expi_def)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   820
14374
61de62096768 further tidying of the complex numbers
paulson
parents: 14373
diff changeset
   821
lemma complex_expi_Ex: "\<exists>a r. z = complex_of_real r * expi a"
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14354
diff changeset
   822
apply (insert rcis_Ex [of z])
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   823
apply (auto simp add: expi_def rcis_def complex_mult_assoc [symmetric] complex_of_real_mult)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14323
diff changeset
   824
apply (rule_tac x = "ii * complex_of_real a" in exI, auto)
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   825
done
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   826
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   827
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   828
subsection{*Numerals and Arithmetic*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   829
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   830
instance complex :: number ..
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   831
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   832
defs (overloaded)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   833
  complex_number_of_def: "(number_of w :: complex) == of_int (Rep_Bin w)"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   834
    --{*the type constraint is essential!*}
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   835
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   836
instance complex :: number_ring
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   837
by (intro_classes, simp add: complex_number_of_def) 
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   838
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   839
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   840
lemma complex_of_real_of_nat [simp]: "complex_of_real (of_nat n) = of_nat n"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   841
by (induct n, simp_all) 
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   842
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   843
lemma complex_of_real_of_int [simp]: "complex_of_real (of_int z) = of_int z"
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   844
proof (cases z)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   845
  case (1 n)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   846
    thus ?thesis by simp
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   847
next
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   848
  case (2 n)
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   849
    thus ?thesis 
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   850
      by (simp only: of_int_minus complex_of_real_minus, simp)
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   851
qed
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   852
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   853
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   854
text{*Collapse applications of @{term complex_of_real} to @{term number_of}*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   855
lemma complex_number_of [simp]: "complex_of_real (number_of w) = number_of w"
15013
34264f5e4691 new treatment of binary numerals
paulson
parents: 15003
diff changeset
   856
by (simp add: complex_number_of_def real_number_of_def) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   857
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   858
text{*This theorem is necessary because theorems such as
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   859
   @{text iszero_number_of_0} only hold for ordered rings. They cannot
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   860
   be generalized to fields in general because they fail for finite fields.
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   861
   They work for type complex because the reals can be embedded in them.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   862
lemma iszero_complex_number_of [simp]:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   863
     "iszero (number_of w :: complex) = iszero (number_of w :: real)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   864
by (simp only: complex_of_real_zero_iff complex_number_of [symmetric] 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   865
               iszero_def)  
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   866
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   867
lemma complex_number_of_cnj [simp]: "cnj(number_of v :: complex) = number_of v"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   868
apply (subst complex_number_of [symmetric])
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   869
apply (rule complex_cnj_complex_of_real)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   870
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   871
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   872
lemma complex_number_of_cmod: 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   873
      "cmod(number_of v :: complex) = abs (number_of v :: real)"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   874
by (simp only: complex_number_of [symmetric] complex_mod_complex_of_real)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   875
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   876
lemma complex_number_of_Re [simp]: "Re(number_of v :: complex) = number_of v"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   877
by (simp only: complex_number_of [symmetric] Re_complex_of_real)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   878
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   879
lemma complex_number_of_Im [simp]: "Im(number_of v :: complex) = 0"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   880
by (simp only: complex_number_of [symmetric] Im_complex_of_real)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   881
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   882
lemma expi_two_pi_i [simp]: "expi((2::complex) * complex_of_real pi * ii) = 1"
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   883
by (simp add: expi_def complex_Re_mult_eq complex_Im_mult_eq cis_def)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   884
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   885
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   886
(*examples:
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   887
print_depth 22
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   888
set timing;
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   889
set trace_simp;
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   890
fun test s = (Goal s, by (Simp_tac 1)); 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   891
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   892
test "23 * ii + 45 * ii= (x::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   893
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   894
test "5 * ii + 12 - 45 * ii= (x::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   895
test "5 * ii + 40 - 12 * ii + 9 = (x::complex) + 89 * ii";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   896
test "5 * ii + 40 - 12 * ii + 9 - 78 = (x::complex) + 89 * ii";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   897
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   898
test "l + 10 * ii + 90 + 3*l +  9 + 45 * ii= (x::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   899
test "87 + 10 * ii + 90 + 3*7 +  9 + 45 * ii= (x::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   900
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   901
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   902
fun test s = (Goal s; by (Asm_simp_tac 1)); 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   903
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   904
test "x*k = k*(y::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   905
test "k = k*(y::complex)"; 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   906
test "a*(b*c) = (b::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   907
test "a*(b*c) = d*(b::complex)*(x*a)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   908
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   909
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   910
test "(x*k) / (k*(y::complex)) = (uu::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   911
test "(k) / (k*(y::complex)) = (uu::complex)"; 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   912
test "(a*(b*c)) / ((b::complex)) = (uu::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   913
test "(a*(b*c)) / (d*(b::complex)*(x*a)) = (uu::complex)";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   914
15003
6145dd7538d7 replaced monomorphic abs definitions by abs_if
paulson
parents: 14691
diff changeset
   915
FIXME: what do we do about this?
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   916
test "a*(b*c)/(y*z) = d*(b::complex)*(x*a)/z";
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   917
*)
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   918
14323
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   919
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   920
ML
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   921
{*
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   922
val complex_zero_def = thm"complex_zero_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   923
val complex_one_def = thm"complex_one_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   924
val complex_minus_def = thm"complex_minus_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   925
val complex_divide_def = thm"complex_divide_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   926
val complex_mult_def = thm"complex_mult_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   927
val complex_add_def = thm"complex_add_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   928
val complex_of_real_def = thm"complex_of_real_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   929
val i_def = thm"i_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   930
val expi_def = thm"expi_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   931
val cis_def = thm"cis_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   932
val rcis_def = thm"rcis_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   933
val cmod_def = thm"cmod_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   934
val cnj_def = thm"cnj_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   935
val sgn_def = thm"sgn_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   936
val arg_def = thm"arg_def";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   937
val complexpow_0 = thm"complexpow_0";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   938
val complexpow_Suc = thm"complexpow_Suc";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   939
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   940
val Re = thm"Re";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   941
val Im = thm"Im";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   942
val complex_Re_Im_cancel_iff = thm"complex_Re_Im_cancel_iff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   943
val complex_Re_zero = thm"complex_Re_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   944
val complex_Im_zero = thm"complex_Im_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   945
val complex_Re_one = thm"complex_Re_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   946
val complex_Im_one = thm"complex_Im_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   947
val complex_Re_i = thm"complex_Re_i";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   948
val complex_Im_i = thm"complex_Im_i";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   949
val Re_complex_of_real = thm"Re_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   950
val Im_complex_of_real = thm"Im_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   951
val complex_minus = thm"complex_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   952
val complex_Re_minus = thm"complex_Re_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   953
val complex_Im_minus = thm"complex_Im_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   954
val complex_add = thm"complex_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   955
val complex_Re_add = thm"complex_Re_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   956
val complex_Im_add = thm"complex_Im_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   957
val complex_add_commute = thm"complex_add_commute";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   958
val complex_add_assoc = thm"complex_add_assoc";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   959
val complex_add_zero_left = thm"complex_add_zero_left";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   960
val complex_add_zero_right = thm"complex_add_zero_right";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   961
val complex_diff = thm"complex_diff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   962
val complex_mult = thm"complex_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   963
val complex_mult_one_left = thm"complex_mult_one_left";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   964
val complex_mult_one_right = thm"complex_mult_one_right";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   965
val complex_inverse = thm"complex_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   966
val complex_of_real_one = thm"complex_of_real_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   967
val complex_of_real_zero = thm"complex_of_real_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   968
val complex_of_real_eq_iff = thm"complex_of_real_eq_iff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   969
val complex_of_real_minus = thm"complex_of_real_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   970
val complex_of_real_inverse = thm"complex_of_real_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   971
val complex_of_real_add = thm"complex_of_real_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   972
val complex_of_real_diff = thm"complex_of_real_diff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   973
val complex_of_real_mult = thm"complex_of_real_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   974
val complex_of_real_divide = thm"complex_of_real_divide";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   975
val complex_of_real_pow = thm"complex_of_real_pow";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   976
val complex_mod = thm"complex_mod";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   977
val complex_mod_zero = thm"complex_mod_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   978
val complex_mod_one = thm"complex_mod_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   979
val complex_mod_complex_of_real = thm"complex_mod_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   980
val complex_of_real_abs = thm"complex_of_real_abs";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   981
val complex_cnj = thm"complex_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   982
val complex_cnj_cancel_iff = thm"complex_cnj_cancel_iff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   983
val complex_cnj_cnj = thm"complex_cnj_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   984
val complex_cnj_complex_of_real = thm"complex_cnj_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   985
val complex_mod_cnj = thm"complex_mod_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   986
val complex_cnj_minus = thm"complex_cnj_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   987
val complex_cnj_inverse = thm"complex_cnj_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   988
val complex_cnj_add = thm"complex_cnj_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   989
val complex_cnj_diff = thm"complex_cnj_diff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   990
val complex_cnj_mult = thm"complex_cnj_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   991
val complex_cnj_divide = thm"complex_cnj_divide";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   992
val complex_cnj_one = thm"complex_cnj_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   993
val complex_cnj_pow = thm"complex_cnj_pow";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   994
val complex_add_cnj = thm"complex_add_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   995
val complex_diff_cnj = thm"complex_diff_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   996
val complex_cnj_zero = thm"complex_cnj_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   997
val complex_cnj_zero_iff = thm"complex_cnj_zero_iff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   998
val complex_mult_cnj = thm"complex_mult_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
   999
val complex_mod_eq_zero_cancel = thm"complex_mod_eq_zero_cancel";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1000
val complex_mod_complex_of_real_of_nat = thm"complex_mod_complex_of_real_of_nat";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1001
val complex_mod_minus = thm"complex_mod_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1002
val complex_mod_mult_cnj = thm"complex_mod_mult_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1003
val complex_mod_squared = thm"complex_mod_squared";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1004
val complex_mod_ge_zero = thm"complex_mod_ge_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1005
val abs_cmod_cancel = thm"abs_cmod_cancel";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1006
val complex_mod_mult = thm"complex_mod_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1007
val complex_mod_add_squared_eq = thm"complex_mod_add_squared_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1008
val complex_Re_mult_cnj_le_cmod = thm"complex_Re_mult_cnj_le_cmod";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1009
val complex_Re_mult_cnj_le_cmod2 = thm"complex_Re_mult_cnj_le_cmod2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1010
val real_sum_squared_expand = thm"real_sum_squared_expand";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1011
val complex_mod_triangle_squared = thm"complex_mod_triangle_squared";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1012
val complex_mod_minus_le_complex_mod = thm"complex_mod_minus_le_complex_mod";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1013
val complex_mod_triangle_ineq = thm"complex_mod_triangle_ineq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1014
val complex_mod_triangle_ineq2 = thm"complex_mod_triangle_ineq2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1015
val complex_mod_diff_commute = thm"complex_mod_diff_commute";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1016
val complex_mod_add_less = thm"complex_mod_add_less";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1017
val complex_mod_mult_less = thm"complex_mod_mult_less";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1018
val complex_mod_diff_ineq = thm"complex_mod_diff_ineq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1019
val complex_Re_le_cmod = thm"complex_Re_le_cmod";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1020
val complex_mod_gt_zero = thm"complex_mod_gt_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1021
val complex_mod_complexpow = thm"complex_mod_complexpow";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1022
val complex_mod_inverse = thm"complex_mod_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1023
val complex_mod_divide = thm"complex_mod_divide";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1024
val complexpow_i_squared = thm"complexpow_i_squared";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1025
val complex_i_not_zero = thm"complex_i_not_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1026
val sgn_zero = thm"sgn_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1027
val sgn_one = thm"sgn_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1028
val sgn_minus = thm"sgn_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1029
val sgn_eq = thm"sgn_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1030
val i_mult_eq = thm"i_mult_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1031
val i_mult_eq2 = thm"i_mult_eq2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1032
val Re_sgn = thm"Re_sgn";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1033
val Im_sgn = thm"Im_sgn";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1034
val complex_inverse_complex_split = thm"complex_inverse_complex_split";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1035
val cos_arg_i_mult_zero = thm"cos_arg_i_mult_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1036
val complex_of_real_zero_iff = thm"complex_of_real_zero_iff";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1037
val rcis_Ex = thm"rcis_Ex";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1038
val Re_rcis = thm"Re_rcis";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1039
val Im_rcis = thm"Im_rcis";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1040
val complex_mod_rcis = thm"complex_mod_rcis";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1041
val complex_mod_sqrt_Re_mult_cnj = thm"complex_mod_sqrt_Re_mult_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1042
val complex_Re_cnj = thm"complex_Re_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1043
val complex_Im_cnj = thm"complex_Im_cnj";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1044
val complex_In_mult_cnj_zero = thm"complex_In_mult_cnj_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1045
val complex_Re_mult = thm"complex_Re_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1046
val complex_Re_mult_complex_of_real = thm"complex_Re_mult_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1047
val complex_Im_mult_complex_of_real = thm"complex_Im_mult_complex_of_real";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1048
val complex_Re_mult_complex_of_real2 = thm"complex_Re_mult_complex_of_real2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1049
val complex_Im_mult_complex_of_real2 = thm"complex_Im_mult_complex_of_real2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1050
val cis_rcis_eq = thm"cis_rcis_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1051
val rcis_mult = thm"rcis_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1052
val cis_mult = thm"cis_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1053
val cis_zero = thm"cis_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1054
val rcis_zero_mod = thm"rcis_zero_mod";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1055
val rcis_zero_arg = thm"rcis_zero_arg";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1056
val complex_of_real_minus_one = thm"complex_of_real_minus_one";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1057
val complex_i_mult_minus = thm"complex_i_mult_minus";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1058
val cis_real_of_nat_Suc_mult = thm"cis_real_of_nat_Suc_mult";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1059
val DeMoivre = thm"DeMoivre";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1060
val DeMoivre2 = thm"DeMoivre2";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1061
val cis_inverse = thm"cis_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1062
val rcis_inverse = thm"rcis_inverse";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1063
val cis_divide = thm"cis_divide";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1064
val rcis_divide = thm"rcis_divide";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1065
val Re_cis = thm"Re_cis";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1066
val Im_cis = thm"Im_cis";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1067
val cos_n_Re_cis_pow_n = thm"cos_n_Re_cis_pow_n";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1068
val sin_n_Im_cis_pow_n = thm"sin_n_Im_cis_pow_n";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1069
val expi_add = thm"expi_add";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1070
val expi_zero = thm"expi_zero";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1071
val complex_Re_mult_eq = thm"complex_Re_mult_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1072
val complex_Im_mult_eq = thm"complex_Im_mult_eq";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1073
val complex_expi_Ex = thm"complex_expi_Ex";
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1074
*}
27724f528f82 converting Complex/Complex.ML to Isar
paulson
parents: 13957
diff changeset
  1075
13957
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1076
end
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1077
10dbf16be15f new session Complex for the complex numbers
paulson
parents:
diff changeset
  1078