src/HOL/Nonstandard_Analysis/NSComplex.thy
author paulson <lp15@cam.ac.uk>
Wed, 18 Sep 2019 14:41:37 +0100
changeset 70723 4e39d87c9737
parent 69597 ff784d5a5bfb
child 73926 5f71c16f0b37
permissions -rw-r--r--
imported new material mostly due to Sébastien Gouëzel
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61981
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/NSComplex.thy
41959
b460124855b8 tuned headers;
wenzelm
parents: 41413
diff changeset
     2
    Author:     Jacques D. Fleuriot, University of Edinburgh
b460124855b8 tuned headers;
wenzelm
parents: 41413
diff changeset
     3
    Author:     Lawrence C Paulson
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
     6
section \<open>Nonstandard Complex Numbers\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     7
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
theory NSComplex
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
     9
  imports NSA
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41959
diff changeset
    12
type_synonym hcomplex = "complex star"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    14
abbreviation hcomplex_of_complex :: "complex \<Rightarrow> complex star"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    15
  where "hcomplex_of_complex \<equiv> star_of"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    17
abbreviation hcmod :: "complex star \<Rightarrow> real star"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    18
  where "hcmod \<equiv> hnorm"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    21
subsubsection \<open>Real and Imaginary parts\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    22
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    23
definition hRe :: "hcomplex \<Rightarrow> hypreal"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    24
  where "hRe = *f* Re"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    25
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    26
definition hIm :: "hcomplex \<Rightarrow> hypreal"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    27
  where "hIm = *f* Im"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    28
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    30
subsubsection \<open>Imaginary unit\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    31
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    32
definition iii :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    33
  where "iii = star_of \<i>"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    35
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    36
subsubsection \<open>Complex conjugate\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    37
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    38
definition hcnj :: "hcomplex \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    39
  where "hcnj = *f* cnj"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    42
subsubsection \<open>Argand\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    44
definition hsgn :: "hcomplex \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    45
  where "hsgn = *f* sgn"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    47
definition harg :: "hcomplex \<Rightarrow> hypreal"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    48
  where "harg = *f* arg"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    50
definition  \<comment> \<open>abbreviation for \<open>cos a + i sin a\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    51
  hcis :: "hypreal \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    52
  where "hcis = *f* cis"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    54
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    55
subsubsection \<open>Injection from hyperreals\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    56
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    57
abbreviation hcomplex_of_hypreal :: "hypreal \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    58
  where "hcomplex_of_hypreal \<equiv> of_hypreal"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    60
definition  \<comment> \<open>abbreviation for \<open>r * (cos a + i sin a)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    61
  hrcis :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    62
  where "hrcis = *f2* rcis"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    63
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    64
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    65
subsubsection \<open>\<open>e ^ (x + iy)\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    66
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    67
definition hExp :: "hcomplex \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    68
  where "hExp = *f* exp"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    69
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    70
definition HComplex :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hcomplex"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    71
  where "HComplex = *f2* Complex"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    72
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    73
lemmas hcomplex_defs [transfer_unfold] =
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
  hRe_def hIm_def iii_def hcnj_def hsgn_def harg_def hcis_def
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58878
diff changeset
    75
  hrcis_def hExp_def HComplex_def
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
lemma Standard_hRe [simp]: "x \<in> Standard \<Longrightarrow> hRe x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    78
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    79
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    80
lemma Standard_hIm [simp]: "x \<in> Standard \<Longrightarrow> hIm x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    81
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    82
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
lemma Standard_iii [simp]: "iii \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    84
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    85
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    86
lemma Standard_hcnj [simp]: "x \<in> Standard \<Longrightarrow> hcnj x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    87
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    89
lemma Standard_hsgn [simp]: "x \<in> Standard \<Longrightarrow> hsgn x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    90
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    92
lemma Standard_harg [simp]: "x \<in> Standard \<Longrightarrow> harg x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    93
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    94
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
lemma Standard_hcis [simp]: "r \<in> Standard \<Longrightarrow> hcis r \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    96
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    97
59658
0cc388370041 sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents: 58878
diff changeset
    98
lemma Standard_hExp [simp]: "x \<in> Standard \<Longrightarrow> hExp x \<in> Standard"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
    99
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   101
lemma Standard_hrcis [simp]: "r \<in> Standard \<Longrightarrow> s \<in> Standard \<Longrightarrow> hrcis r s \<in> Standard"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   102
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   103
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   104
lemma Standard_HComplex [simp]: "r \<in> Standard \<Longrightarrow> s \<in> Standard \<Longrightarrow> HComplex r s \<in> Standard"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   105
  by (simp add: hcomplex_defs)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   106
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   107
lemma hcmod_def: "hcmod = *f* cmod"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   108
  by (rule hnorm_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   109
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   110
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   111
subsection \<open>Properties of Nonstandard Real and Imaginary Parts\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   113
lemma hcomplex_hRe_hIm_cancel_iff: "\<And>w z. w = z \<longleftrightarrow> hRe w = hRe z \<and> hIm w = hIm z"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 65274
diff changeset
   114
  by transfer (rule complex_eq_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   115
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   116
lemma hcomplex_equality [intro?]: "\<And>z w. hRe z = hRe w \<Longrightarrow> hIm z = hIm w \<Longrightarrow> z = w"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 65274
diff changeset
   117
  by transfer (rule complex_eqI)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   119
lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   120
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   121
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   122
lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   123
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
lemma hcomplex_hRe_one [simp]: "hRe 1 = 1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   126
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   127
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
lemma hcomplex_hIm_one [simp]: "hIm 1 = 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   129
  by transfer simp
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   130
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   131
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   132
subsection \<open>Addition for Nonstandard Complex Numbers\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   133
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   134
lemma hRe_add: "\<And>x y. hRe (x + y) = hRe x + hRe y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   135
  by transfer simp
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   136
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   137
lemma hIm_add: "\<And>x y. hIm (x + y) = hIm x + hIm y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   138
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   140
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   141
subsection \<open>More Minus Laws\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   142
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   143
lemma hRe_minus: "\<And>z. hRe (- z) = - hRe z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   144
  by transfer (rule uminus_complex.sel)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   145
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   146
lemma hIm_minus: "\<And>z. hIm (- z) = - hIm z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   147
  by transfer (rule uminus_complex.sel)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   148
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   149
lemma hcomplex_add_minus_eq_minus: "x + y = 0 \<Longrightarrow> x = - y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   150
  for x y :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   151
  apply (drule minus_unique)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   152
  apply (simp add: minus_equation_iff [of x y])
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   153
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   154
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   155
lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   156
  by transfer (rule i_squared)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   157
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   158
lemma hcomplex_i_mult_left [simp]: "\<And>z. iii * (iii * z) = - z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   159
  by transfer (rule complex_i_mult_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   160
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   161
lemma hcomplex_i_not_zero [simp]: "iii \<noteq> 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   162
  by transfer (rule complex_i_not_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   163
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   164
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   165
subsection \<open>More Multiplication Laws\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   166
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   167
lemma hcomplex_mult_minus_one: "- 1 * z = - z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   168
  for z :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   169
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   170
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   171
lemma hcomplex_mult_minus_one_right: "z * - 1 = - z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   172
  for z :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   173
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   174
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   175
lemma hcomplex_mult_left_cancel: "c \<noteq> 0 \<Longrightarrow> c * a = c * b \<longleftrightarrow> a = b"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   176
  for a b c :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   177
  by simp
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   178
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   179
lemma hcomplex_mult_right_cancel: "c \<noteq> 0 \<Longrightarrow> a * c = b * c \<longleftrightarrow> a = b"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   180
  for a b c :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   181
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   182
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   183
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   184
subsection \<open>Subtraction and Division\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   185
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   186
(* TODO: delete *)
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   187
lemma hcomplex_diff_eq_eq [simp]: "x - y = z \<longleftrightarrow> x = z + y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   188
  for x y z :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   189
  by (rule diff_eq_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   190
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   191
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68499
diff changeset
   192
subsection \<open>Embedding Properties for \<^term>\<open>hcomplex_of_hypreal\<close> Map\<close>
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   193
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   194
lemma hRe_hcomplex_of_hypreal [simp]: "\<And>z. hRe (hcomplex_of_hypreal z) = z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   195
  by transfer (rule Re_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   197
lemma hIm_hcomplex_of_hypreal [simp]: "\<And>z. hIm (hcomplex_of_hypreal z) = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   198
  by transfer (rule Im_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
70723
4e39d87c9737 imported new material mostly due to Sébastien Gouëzel
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   200
lemma hcomplex_of_epsilon_not_zero [simp]: "hcomplex_of_hypreal \<epsilon> \<noteq> 0"
4e39d87c9737 imported new material mostly due to Sébastien Gouëzel
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   201
  by (simp add: epsilon_not_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   202
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   203
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   204
subsection \<open>\<open>HComplex\<close> theorems\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   205
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   206
lemma hRe_HComplex [simp]: "\<And>x y. hRe (HComplex x y) = x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   207
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   208
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   209
lemma hIm_HComplex [simp]: "\<And>x y. hIm (HComplex x y) = y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   210
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   211
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   212
lemma hcomplex_surj [simp]: "\<And>z. HComplex (hRe z) (hIm z) = z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   213
  by transfer (rule complex_surj)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   214
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   215
lemma hcomplex_induct [case_names rect(*, induct type: hcomplex*)]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   216
  "(\<And>x y. P (HComplex x y)) \<Longrightarrow> P z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   217
  by (rule hcomplex_surj [THEN subst]) blast
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   218
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   219
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   220
subsection \<open>Modulus (Absolute Value) of Nonstandard Complex Number\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   221
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   222
lemma hcomplex_of_hypreal_abs:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   223
  "hcomplex_of_hypreal \<bar>x\<bar> = hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   224
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   225
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   226
lemma HComplex_inject [simp]: "\<And>x y x' y'. HComplex x y = HComplex x' y' \<longleftrightarrow> x = x' \<and> y = y'"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   227
  by transfer (rule complex.inject)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   228
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   229
lemma HComplex_add [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   230
  "\<And>x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1 + x2) (y1 + y2)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   231
  by transfer (rule complex_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   232
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   233
lemma HComplex_minus [simp]: "\<And>x y. - HComplex x y = HComplex (- x) (- y)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   234
  by transfer (rule complex_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   235
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   236
lemma HComplex_diff [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   237
  "\<And>x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1 - x2) (y1 - y2)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   238
  by transfer (rule complex_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   239
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   240
lemma HComplex_mult [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   241
  "\<And>x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 = HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   242
  by transfer (rule complex_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   243
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   244
text \<open>\<open>HComplex_inverse\<close> is proved below.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   245
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   246
lemma hcomplex_of_hypreal_eq: "\<And>r. hcomplex_of_hypreal r = HComplex r 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   247
  by transfer (rule complex_of_real_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   249
lemma HComplex_add_hcomplex_of_hypreal [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   250
  "\<And>x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x + r) y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   251
  by transfer (rule Complex_add_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   252
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   253
lemma hcomplex_of_hypreal_add_HComplex [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   254
  "\<And>r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r + x) y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   255
  by transfer (rule complex_of_real_add_Complex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   256
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   257
lemma HComplex_mult_hcomplex_of_hypreal:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   258
  "\<And>x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x * r) (y * r)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   259
  by transfer (rule Complex_mult_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   260
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   261
lemma hcomplex_of_hypreal_mult_HComplex:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   262
  "\<And>r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r * x) (r * y)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   263
  by transfer (rule complex_of_real_mult_Complex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   264
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   265
lemma i_hcomplex_of_hypreal [simp]: "\<And>r. iii * hcomplex_of_hypreal r = HComplex 0 r"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   266
  by transfer (rule i_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   267
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   268
lemma hcomplex_of_hypreal_i [simp]: "\<And>r. hcomplex_of_hypreal r * iii = HComplex 0 r"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   269
  by transfer (rule complex_of_real_i)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   270
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   271
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   272
subsection \<open>Conjugation\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   273
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   274
lemma hcomplex_hcnj_cancel_iff [iff]: "\<And>x y. hcnj x = hcnj y \<longleftrightarrow> x = y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   275
  by transfer (rule complex_cnj_cancel_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   276
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   277
lemma hcomplex_hcnj_hcnj [simp]: "\<And>z. hcnj (hcnj z) = z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   278
  by transfer (rule complex_cnj_cnj)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   279
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   280
lemma hcomplex_hcnj_hcomplex_of_hypreal [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   281
  "\<And>x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   282
  by transfer (rule complex_cnj_complex_of_real)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   283
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   284
lemma hcomplex_hmod_hcnj [simp]: "\<And>z. hcmod (hcnj z) = hcmod z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   285
  by transfer (rule complex_mod_cnj)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   286
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   287
lemma hcomplex_hcnj_minus: "\<And>z. hcnj (- z) = - hcnj z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   288
  by transfer (rule complex_cnj_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   289
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   290
lemma hcomplex_hcnj_inverse: "\<And>z. hcnj (inverse z) = inverse (hcnj z)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   291
  by transfer (rule complex_cnj_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   292
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   293
lemma hcomplex_hcnj_add: "\<And>w z. hcnj (w + z) = hcnj w + hcnj z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   294
  by transfer (rule complex_cnj_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   295
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   296
lemma hcomplex_hcnj_diff: "\<And>w z. hcnj (w - z) = hcnj w - hcnj z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   297
  by transfer (rule complex_cnj_diff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   298
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   299
lemma hcomplex_hcnj_mult: "\<And>w z. hcnj (w * z) = hcnj w * hcnj z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   300
  by transfer (rule complex_cnj_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   301
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   302
lemma hcomplex_hcnj_divide: "\<And>w z. hcnj (w / z) = hcnj w / hcnj z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   303
  by transfer (rule complex_cnj_divide)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   304
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   305
lemma hcnj_one [simp]: "hcnj 1 = 1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   306
  by transfer (rule complex_cnj_one)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   307
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   308
lemma hcomplex_hcnj_zero [simp]: "hcnj 0 = 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   309
  by transfer (rule complex_cnj_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   310
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   311
lemma hcomplex_hcnj_zero_iff [iff]: "\<And>z. hcnj z = 0 \<longleftrightarrow> z = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   312
  by transfer (rule complex_cnj_zero_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   313
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   314
lemma hcomplex_mult_hcnj: "\<And>z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   315
  by transfer (rule complex_mult_cnj)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   316
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   317
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68499
diff changeset
   318
subsection \<open>More Theorems about the Function \<^term>\<open>hcmod\<close>\<close>
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   319
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   320
lemma hcmod_hcomplex_of_hypreal_of_nat [simp]:
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   321
  "hcmod (hcomplex_of_hypreal (hypreal_of_nat n)) = hypreal_of_nat n"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   322
  by simp
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   323
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   324
lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]:
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   325
  "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   326
  by simp
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   327
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   328
lemma hcmod_mult_hcnj: "\<And>z. hcmod (z * hcnj z) = (hcmod z)\<^sup>2"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   329
  by transfer (rule complex_mod_mult_cnj)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   330
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   331
lemma hcmod_triangle_ineq2 [simp]: "\<And>a b. hcmod (b + a) - hcmod b \<le> hcmod a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   332
  by transfer (rule complex_mod_triangle_ineq2)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   333
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   334
lemma hcmod_diff_ineq [simp]: "\<And>a b. hcmod a - hcmod b \<le> hcmod (a + b)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   335
  by transfer (rule norm_diff_ineq)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   336
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   337
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   338
subsection \<open>Exponentiation\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   339
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   340
lemma hcomplexpow_0 [simp]: "z ^ 0 = 1"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   341
  for z :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   342
  by (rule power_0)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   343
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   344
lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = z * (z ^ n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   345
  for z :: hcomplex
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   346
  by (rule power_Suc)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   347
53077
a1b3784f8129 more symbols;
wenzelm
parents: 51525
diff changeset
   348
lemma hcomplexpow_i_squared [simp]: "iii\<^sup>2 = -1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   349
  by transfer (rule power2_i)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   350
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   351
lemma hcomplex_of_hypreal_pow: "\<And>x. hcomplex_of_hypreal (x ^ n) = hcomplex_of_hypreal x ^ n"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   352
  by transfer (rule of_real_power)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   353
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   354
lemma hcomplex_hcnj_pow: "\<And>z. hcnj (z ^ n) = hcnj z ^ n"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   355
  by transfer (rule complex_cnj_power)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   356
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   357
lemma hcmod_hcomplexpow: "\<And>x. hcmod (x ^ n) = hcmod x ^ n"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   358
  by transfer (rule norm_power)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   359
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   360
lemma hcpow_minus:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   361
  "\<And>x n. (- x :: hcomplex) pow n = (if ( *p* even) n then (x pow n) else - (x pow n))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   362
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   363
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   364
lemma hcpow_mult: "(r * s) pow n = (r pow n) * (s pow n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   365
  for r s :: hcomplex
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 56889
diff changeset
   366
  by (fact hyperpow_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   367
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   368
lemma hcpow_zero2 [simp]: "\<And>n. 0 pow (hSuc n) = (0::'a::semiring_1 star)"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60017
diff changeset
   369
  by transfer (rule power_0_Suc)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   370
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   371
lemma hcpow_not_zero [simp,intro]: "\<And>r n. r \<noteq> 0 \<Longrightarrow> r pow n \<noteq> (0::hcomplex)"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60017
diff changeset
   372
  by (fact hyperpow_not_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   373
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   374
lemma hcpow_zero_zero: "r pow n = 0 \<Longrightarrow> r = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   375
  for r :: hcomplex
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60017
diff changeset
   376
  by (blast intro: ccontr dest: hcpow_not_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   377
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   378
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68499
diff changeset
   379
subsection \<open>The Function \<^term>\<open>hsgn\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   380
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   381
lemma hsgn_zero [simp]: "hsgn 0 = 0"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   382
  by transfer (rule sgn_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   383
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   384
lemma hsgn_one [simp]: "hsgn 1 = 1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   385
  by transfer (rule sgn_one)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   386
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   387
lemma hsgn_minus: "\<And>z. hsgn (- z) = - hsgn z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   388
  by transfer (rule sgn_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   389
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   390
lemma hsgn_eq: "\<And>z. hsgn z = z / hcomplex_of_hypreal (hcmod z)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   391
  by transfer (rule sgn_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   392
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   393
lemma hcmod_i: "\<And>x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   394
  by transfer (rule complex_norm)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   395
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   396
lemma hcomplex_eq_cancel_iff1 [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   397
  "hcomplex_of_hypreal xa = HComplex x y \<longleftrightarrow> xa = x \<and> y = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   398
  by (simp add: hcomplex_of_hypreal_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   399
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   400
lemma hcomplex_eq_cancel_iff2 [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   401
  "HComplex x y = hcomplex_of_hypreal xa \<longleftrightarrow> x = xa \<and> y = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   402
  by (simp add: hcomplex_of_hypreal_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   403
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   404
lemma HComplex_eq_0 [simp]: "\<And>x y. HComplex x y = 0 \<longleftrightarrow> x = 0 \<and> y = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   405
  by transfer (rule Complex_eq_0)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   406
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   407
lemma HComplex_eq_1 [simp]: "\<And>x y. HComplex x y = 1 \<longleftrightarrow> x = 1 \<and> y = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   408
  by transfer (rule Complex_eq_1)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   409
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   410
lemma i_eq_HComplex_0_1: "iii = HComplex 0 1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   411
  by transfer (simp add: complex_eq_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   412
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   413
lemma HComplex_eq_i [simp]: "\<And>x y. HComplex x y = iii \<longleftrightarrow> x = 0 \<and> y = 1"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   414
  by transfer (rule Complex_eq_i)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   415
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   416
lemma hRe_hsgn [simp]: "\<And>z. hRe (hsgn z) = hRe z / hcmod z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   417
  by transfer (rule Re_sgn)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   418
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   419
lemma hIm_hsgn [simp]: "\<And>z. hIm (hsgn z) = hIm z / hcmod z"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   420
  by transfer (rule Im_sgn)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   421
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   422
lemma HComplex_inverse: "\<And>x y. inverse (HComplex x y) = HComplex (x / (x\<^sup>2 + y\<^sup>2)) (- y / (x\<^sup>2 + y\<^sup>2))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   423
  by transfer (rule complex_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   424
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   425
lemma hRe_mult_i_eq[simp]: "\<And>y. hRe (iii * hcomplex_of_hypreal y) = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   426
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   427
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   428
lemma hIm_mult_i_eq [simp]: "\<And>y. hIm (iii * hcomplex_of_hypreal y) = y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   429
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   430
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   431
lemma hcmod_mult_i [simp]: "\<And>y. hcmod (iii * hcomplex_of_hypreal y) = \<bar>y\<bar>"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   432
  by transfer (simp add: norm_complex_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   433
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   434
lemma hcmod_mult_i2 [simp]: "\<And>y. hcmod (hcomplex_of_hypreal y * iii) = \<bar>y\<bar>"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   435
  by transfer (simp add: norm_complex_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   436
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   437
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   438
subsubsection \<open>\<open>harg\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   439
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   440
lemma cos_harg_i_mult_zero [simp]: "\<And>y. y \<noteq> 0 \<Longrightarrow> ( *f* cos) (harg (HComplex 0 y)) = 0"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 64435
diff changeset
   441
  by transfer (simp add: Complex_eq)
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   442
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   443
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   444
subsection \<open>Polar Form for Nonstandard Complex Numbers\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   445
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   446
lemma complex_split_polar2: "\<forall>n. \<exists>r a. (z n) = complex_of_real r * Complex (cos a) (sin a)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 64435
diff changeset
   447
  unfolding Complex_eq by (auto intro: complex_split_polar)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   448
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   449
lemma hcomplex_split_polar:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   450
  "\<And>z. \<exists>r a. z = hcomplex_of_hypreal r * (HComplex (( *f* cos) a) (( *f* sin) a))"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 64435
diff changeset
   451
  by transfer (simp add: Complex_eq complex_split_polar)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   452
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   453
lemma hcis_eq:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   454
  "\<And>a. hcis a = hcomplex_of_hypreal (( *f* cos) a) + iii * hcomplex_of_hypreal (( *f* sin) a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   455
  by transfer (simp add: complex_eq_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   456
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   457
lemma hrcis_Ex: "\<And>z. \<exists>r a. z = hrcis r a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   458
  by transfer (rule rcis_Ex)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   459
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   460
lemma hRe_hcomplex_polar [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   461
  "\<And>r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* cos) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   462
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   463
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   464
lemma hRe_hrcis [simp]: "\<And>r a. hRe (hrcis r a) = r * ( *f* cos) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   465
  by transfer (rule Re_rcis)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   466
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   467
lemma hIm_hcomplex_polar [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   468
  "\<And>r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* sin) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   469
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   470
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   471
lemma hIm_hrcis [simp]: "\<And>r a. hIm (hrcis r a) = r * ( *f* sin) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   472
  by transfer (rule Im_rcis)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   473
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   474
lemma hcmod_unit_one [simp]: "\<And>a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   475
  by transfer (simp add: cmod_unit_one)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   476
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   477
lemma hcmod_complex_polar [simp]:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   478
  "\<And>r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \<bar>r\<bar>"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 64435
diff changeset
   479
  by transfer (simp add: Complex_eq cmod_complex_polar)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   480
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   481
lemma hcmod_hrcis [simp]: "\<And>r a. hcmod(hrcis r a) = \<bar>r\<bar>"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   482
  by transfer (rule complex_mod_rcis)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   483
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   484
text \<open>\<open>(r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   485
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   486
lemma hcis_hrcis_eq: "\<And>a. hcis a = hrcis 1 a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   487
  by transfer (rule cis_rcis_eq)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   488
declare hcis_hrcis_eq [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   489
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   490
lemma hrcis_mult: "\<And>a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1 * r2) (a + b)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   491
  by transfer (rule rcis_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   492
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   493
lemma hcis_mult: "\<And>a b. hcis a * hcis b = hcis (a + b)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   494
  by transfer (rule cis_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   495
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   496
lemma hcis_zero [simp]: "hcis 0 = 1"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   497
  by transfer (rule cis_zero)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   498
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   499
lemma hrcis_zero_mod [simp]: "\<And>a. hrcis 0 a = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   500
  by transfer (rule rcis_zero_mod)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   501
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   502
lemma hrcis_zero_arg [simp]: "\<And>r. hrcis r 0 = hcomplex_of_hypreal r"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   503
  by transfer (rule rcis_zero_arg)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   504
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   505
lemma hcomplex_i_mult_minus [simp]: "\<And>x. iii * (iii * x) = - x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   506
  by transfer (rule complex_i_mult_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   507
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   508
lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   509
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   510
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   511
lemma hcis_hypreal_of_nat_Suc_mult:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   512
  "\<And>a. hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   513
  by transfer (simp add: distrib_right cis_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   514
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   515
lemma NSDeMoivre: "\<And>a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   516
  by transfer (rule DeMoivre)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   517
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   518
lemma hcis_hypreal_of_hypnat_Suc_mult:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   519
  "\<And>a n. hcis (hypreal_of_hypnat (n + 1) * a) = hcis a * hcis (hypreal_of_hypnat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   520
  by transfer (simp add: distrib_right cis_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   521
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   522
lemma NSDeMoivre_ext: "\<And>a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   523
  by transfer (rule DeMoivre)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   524
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   525
lemma NSDeMoivre2: "\<And>a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   526
  by transfer (rule DeMoivre2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   527
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   528
lemma DeMoivre2_ext: "\<And>a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   529
  by transfer (rule DeMoivre2)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   530
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   531
lemma hcis_inverse [simp]: "\<And>a. inverse (hcis a) = hcis (- a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   532
  by transfer (rule cis_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   533
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   534
lemma hrcis_inverse: "\<And>a r. inverse (hrcis r a) = hrcis (inverse r) (- a)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   535
  by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   536
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   537
lemma hRe_hcis [simp]: "\<And>a. hRe (hcis a) = ( *f* cos) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   538
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   539
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   540
lemma hIm_hcis [simp]: "\<And>a. hIm (hcis a) = ( *f* sin) a"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   541
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   542
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   543
lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe (hcis a ^ n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   544
  by (simp add: NSDeMoivre)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   545
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   546
lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm (hcis a ^ n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   547
  by (simp add: NSDeMoivre)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   548
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   549
lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe (hcis a pow n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   550
  by (simp add: NSDeMoivre_ext)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   551
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   552
lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm (hcis a pow n)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   553
  by (simp add: NSDeMoivre_ext)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   554
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   555
lemma hExp_add: "\<And>a b. hExp (a + b) = hExp a * hExp b"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   556
  by transfer (rule exp_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   557
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   558
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68499
diff changeset
   559
subsection \<open>\<^term>\<open>hcomplex_of_complex\<close>: the Injection from type \<^typ>\<open>complex\<close> to to \<^typ>\<open>hcomplex\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   560
63589
58aab4745e85 more symbols;
wenzelm
parents: 62479
diff changeset
   561
lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex \<i>"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   562
  by (rule iii_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   563
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   564
lemma hRe_hcomplex_of_complex: "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   565
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   566
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   567
lemma hIm_hcomplex_of_complex: "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   568
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   569
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   570
lemma hcmod_hcomplex_of_complex: "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   571
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   572
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   573
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   574
subsection \<open>Numerals and Arithmetic\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   575
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   576
lemma hcomplex_of_hypreal_eq_hcomplex_of_complex:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   577
  "hcomplex_of_hypreal (hypreal_of_real x) = hcomplex_of_complex (complex_of_real x)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   578
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   579
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44846
diff changeset
   580
lemma hcomplex_hypreal_numeral:
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44846
diff changeset
   581
  "hcomplex_of_complex (numeral w) = hcomplex_of_hypreal(numeral w)"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   582
  by transfer (rule of_real_numeral [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   583
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44846
diff changeset
   584
lemma hcomplex_hypreal_neg_numeral:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 53077
diff changeset
   585
  "hcomplex_of_complex (- numeral w) = hcomplex_of_hypreal(- numeral w)"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   586
  by transfer (rule of_real_neg_numeral [symmetric])
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44846
diff changeset
   587
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   588
lemma hcomplex_numeral_hcnj [simp]: "hcnj (numeral v :: hcomplex) = numeral v"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   589
  by transfer (rule complex_cnj_numeral)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   590
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   591
lemma hcomplex_numeral_hcmod [simp]: "hcmod (numeral v :: hcomplex) = (numeral v :: hypreal)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   592
  by transfer (rule norm_numeral)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 44846
diff changeset
   593
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   594
lemma hcomplex_neg_numeral_hcmod [simp]: "hcmod (- numeral v :: hcomplex) = (numeral v :: hypreal)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   595
  by transfer (rule norm_neg_numeral)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   596
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   597
lemma hcomplex_numeral_hRe [simp]: "hRe (numeral v :: hcomplex) = numeral v"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   598
  by transfer (rule complex_Re_numeral)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   599
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   600
lemma hcomplex_numeral_hIm [simp]: "hIm (numeral v :: hcomplex) = 0"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 63589
diff changeset
   601
  by transfer (rule complex_Im_numeral)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   602
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   603
end