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