| author | wenzelm | 
| Mon, 11 Dec 2023 13:40:02 +0100 | |
| changeset 79243 | 0e15387c0300 | 
| parent 70217 | 1f04832cbfcf | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/NSCA.thy | 
| 2 | Author: Jacques D. Fleuriot | |
| 3 | Copyright: 2001, 2002 University of Edinburgh | |
| 27468 | 4 | *) | 
| 5 | ||
| 61975 | 6 | section\<open>Non-Standard Complex Analysis\<close> | 
| 27468 | 7 | |
| 8 | theory NSCA | |
| 28952 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 haftmann parents: 
28562diff
changeset | 9 | imports NSComplex HTranscendental | 
| 27468 | 10 | begin | 
| 11 | ||
| 12 | abbreviation | |
| 13 | (* standard complex numbers reagarded as an embedded subset of NS complex *) | |
| 14 | SComplex :: "hcomplex set" where | |
| 15 | "SComplex \<equiv> Standard" | |
| 16 | ||
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
67091diff
changeset | 17 | definition \<comment> \<open>standard part map\<close> | 
| 27468 | 18 | stc :: "hcomplex => hcomplex" where | 
| 67091 | 19 | "stc x = (SOME r. x \<in> HFinite \<and> r\<in>SComplex \<and> r \<approx> x)" | 
| 27468 | 20 | |
| 21 | ||
| 61975 | 22 | subsection\<open>Closure Laws for SComplex, the Standard Complex Numbers\<close> | 
| 27468 | 23 | |
| 24 | lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)" | |
| 70216 | 25 | using Standard_minus by fastforce | 
| 27468 | 26 | |
| 27 | lemma SComplex_add_cancel: | |
| 70216 | 28 | "\<lbrakk>x + y \<in> SComplex; y \<in> SComplex\<rbrakk> \<Longrightarrow> x \<in> SComplex" | 
| 29 | using Standard_diff by fastforce | |
| 27468 | 30 | |
| 31 | lemma SReal_hcmod_hcomplex_of_complex [simp]: | |
| 70216 | 32 | "hcmod (hcomplex_of_complex r) \<in> \<real>" | 
| 33 | by (simp add: Reals_eq_Standard) | |
| 27468 | 34 | |
| 70216 | 35 | lemma SReal_hcmod_numeral: "hcmod (numeral w ::hcomplex) \<in> \<real>" | 
| 36 | by simp | |
| 27468 | 37 | |
| 70216 | 38 | lemma SReal_hcmod_SComplex: "x \<in> SComplex \<Longrightarrow> hcmod x \<in> \<real>" | 
| 39 | by (simp add: Reals_eq_Standard) | |
| 27468 | 40 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
37887diff
changeset | 41 | lemma SComplex_divide_numeral: | 
| 70216 | 42 | "r \<in> SComplex \<Longrightarrow> r/(numeral w::hcomplex) \<in> SComplex" | 
| 43 | by simp | |
| 27468 | 44 | |
| 45 | lemma SComplex_UNIV_complex: | |
| 70216 | 46 |   "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
 | 
| 47 | by simp | |
| 27468 | 48 | |
| 49 | lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)" | |
| 70216 | 50 | by (simp add: Standard_def image_def) | 
| 27468 | 51 | |
| 52 | lemma hcomplex_of_complex_image: | |
| 70216 | 53 | "range hcomplex_of_complex = SComplex" | 
| 54 | by (simp add: Standard_def) | |
| 27468 | 55 | |
| 56 | lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV" | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 57 | by (auto simp add: Standard_def image_def) (metis inj_star_of inv_f_f) | 
| 27468 | 58 | |
| 59 | lemma SComplex_hcomplex_of_complex_image: | |
| 67613 | 60 | "\<lbrakk>\<exists>x. x \<in> P; P \<le> SComplex\<rbrakk> \<Longrightarrow> \<exists>Q. P = hcomplex_of_complex ` Q" | 
| 70216 | 61 | by (metis Standard_def subset_imageE) | 
| 27468 | 62 | |
| 63 | lemma SComplex_SReal_dense: | |
| 67613 | 64 | "\<lbrakk>x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y | 
| 65 | \<rbrakk> \<Longrightarrow> \<exists>r \<in> Reals. hcmod x< r \<and> r < hcmod y" | |
| 70216 | 66 | by (simp add: SReal_dense SReal_hcmod_SComplex) | 
| 27468 | 67 | |
| 68 | ||
| 61975 | 69 | subsection\<open>The Finite Elements form a Subring\<close> | 
| 27468 | 70 | |
| 71 | lemma HFinite_hcmod_hcomplex_of_complex [simp]: | |
| 70216 | 72 | "hcmod (hcomplex_of_complex r) \<in> HFinite" | 
| 73 | by (auto intro!: SReal_subset_HFinite [THEN subsetD]) | |
| 27468 | 74 | |
| 70217 | 75 | lemma HFinite_hcmod_iff [simp]: "hcmod x \<in> HFinite \<longleftrightarrow> x \<in> HFinite" | 
| 70216 | 76 | by (simp add: HFinite_def) | 
| 27468 | 77 | |
| 78 | lemma HFinite_bounded_hcmod: | |
| 67613 | 79 | "\<lbrakk>x \<in> HFinite; y \<le> hcmod x; 0 \<le> y\<rbrakk> \<Longrightarrow> y \<in> HFinite" | 
| 70217 | 80 | using HFinite_bounded HFinite_hcmod_iff by blast | 
| 27468 | 81 | |
| 82 | ||
| 61975 | 83 | subsection\<open>The Complex Infinitesimals form a Subring\<close> | 
| 27468 | 84 | |
| 85 | lemma Infinitesimal_hcmod_iff: | |
| 70216 | 86 | "(z \<in> Infinitesimal) = (hcmod z \<in> Infinitesimal)" | 
| 87 | by (simp add: Infinitesimal_def) | |
| 27468 | 88 | |
| 89 | lemma HInfinite_hcmod_iff: "(z \<in> HInfinite) = (hcmod z \<in> HInfinite)" | |
| 70216 | 90 | by (simp add: HInfinite_def) | 
| 27468 | 91 | |
| 92 | lemma HFinite_diff_Infinitesimal_hcmod: | |
| 70216 | 93 | "x \<in> HFinite - Infinitesimal \<Longrightarrow> hcmod x \<in> HFinite - Infinitesimal" | 
| 70217 | 94 | by (simp add: Infinitesimal_hcmod_iff) | 
| 27468 | 95 | |
| 96 | lemma hcmod_less_Infinitesimal: | |
| 70216 | 97 | "\<lbrakk>e \<in> Infinitesimal; hcmod x < hcmod e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal" | 
| 98 | by (auto elim: hrabs_less_Infinitesimal simp add: Infinitesimal_hcmod_iff) | |
| 27468 | 99 | |
| 100 | lemma hcmod_le_Infinitesimal: | |
| 70216 | 101 | "\<lbrakk>e \<in> Infinitesimal; hcmod x \<le> hcmod e\<rbrakk> \<Longrightarrow> x \<in> Infinitesimal" | 
| 102 | by (auto elim: hrabs_le_Infinitesimal simp add: Infinitesimal_hcmod_iff) | |
| 27468 | 103 | |
| 104 | ||
| 61975 | 105 | subsection\<open>The ``Infinitely Close'' Relation\<close> | 
| 27468 | 106 | |
| 70216 | 107 | lemma approx_SComplex_mult_cancel_zero: | 
| 108 | "\<lbrakk>a \<in> SComplex; a \<noteq> 0; a*x \<approx> 0\<rbrakk> \<Longrightarrow> x \<approx> 0" | |
| 109 | by (metis Infinitesimal_mult_disj SComplex_iff mem_infmal_iff star_of_Infinitesimal_iff_0 star_zero_def) | |
| 27468 | 110 | |
| 70216 | 111 | lemma approx_mult_SComplex1: "\<lbrakk>a \<in> SComplex; x \<approx> 0\<rbrakk> \<Longrightarrow> x*a \<approx> 0" | 
| 112 | using SComplex_iff approx_mult_subst_star_of by fastforce | |
| 27468 | 113 | |
| 70216 | 114 | lemma approx_mult_SComplex2: "\<lbrakk>a \<in> SComplex; x \<approx> 0\<rbrakk> \<Longrightarrow> a*x \<approx> 0" | 
| 115 | by (metis approx_mult_SComplex1 mult.commute) | |
| 27468 | 116 | |
| 117 | lemma approx_mult_SComplex_zero_cancel_iff [simp]: | |
| 70216 | 118 | "\<lbrakk>a \<in> SComplex; a \<noteq> 0\<rbrakk> \<Longrightarrow> (a*x \<approx> 0) = (x \<approx> 0)" | 
| 119 | using approx_SComplex_mult_cancel_zero approx_mult_SComplex2 by blast | |
| 27468 | 120 | |
| 121 | lemma approx_SComplex_mult_cancel: | |
| 70216 | 122 | "\<lbrakk>a \<in> SComplex; a \<noteq> 0; a*w \<approx> a*z\<rbrakk> \<Longrightarrow> w \<approx> z" | 
| 123 | by (metis approx_SComplex_mult_cancel_zero approx_minus_iff right_diff_distrib) | |
| 27468 | 124 | |
| 125 | lemma approx_SComplex_mult_cancel_iff1 [simp]: | |
| 70216 | 126 | "\<lbrakk>a \<in> SComplex; a \<noteq> 0\<rbrakk> \<Longrightarrow> (a*w \<approx> a*z) = (w \<approx> z)" | 
| 127 | by (metis HFinite_star_of SComplex_iff approx_SComplex_mult_cancel approx_mult2) | |
| 27468 | 128 | |
| 129 | (* TODO: generalize following theorems: hcmod -> hnorm *) | |
| 130 | ||
| 61982 | 131 | lemma approx_hcmod_approx_zero: "(x \<approx> y) = (hcmod (y - x) \<approx> 0)" | 
| 70216 | 132 | by (simp add: Infinitesimal_hcmod_iff approx_def hnorm_minus_commute) | 
| 27468 | 133 | |
| 61982 | 134 | lemma approx_approx_zero_iff: "(x \<approx> 0) = (hcmod x \<approx> 0)" | 
| 27468 | 135 | by (simp add: approx_hcmod_approx_zero) | 
| 136 | ||
| 61982 | 137 | lemma approx_minus_zero_cancel_iff [simp]: "(-x \<approx> 0) = (x \<approx> 0)" | 
| 27468 | 138 | by (simp add: approx_def) | 
| 139 | ||
| 140 | lemma Infinitesimal_hcmod_add_diff: | |
| 70216 | 141 | "u \<approx> 0 \<Longrightarrow> hcmod(x + u) - hcmod x \<in> Infinitesimal" | 
| 142 | by (metis add.commute add.left_neutral approx_add_right_iff approx_def approx_hnorm) | |
| 27468 | 143 | |
| 70216 | 144 | lemma approx_hcmod_add_hcmod: "u \<approx> 0 \<Longrightarrow> hcmod(x + u) \<approx> hcmod x" | 
| 145 | using Infinitesimal_hcmod_add_diff approx_def by blast | |
| 27468 | 146 | |
| 147 | ||
| 61975 | 148 | subsection\<open>Zero is the Only Infinitesimal Complex Number\<close> | 
| 27468 | 149 | |
| 150 | lemma Infinitesimal_less_SComplex: | |
| 70216 | 151 | "\<lbrakk>x \<in> SComplex; y \<in> Infinitesimal; 0 < hcmod x\<rbrakk> \<Longrightarrow> hcmod y < hcmod x" | 
| 152 | by (auto intro: Infinitesimal_less_SReal SReal_hcmod_SComplex simp add: Infinitesimal_hcmod_iff) | |
| 27468 | 153 | |
| 154 | lemma SComplex_Int_Infinitesimal_zero: "SComplex Int Infinitesimal = {0}"
 | |
| 70216 | 155 | by (auto simp add: Standard_def Infinitesimal_hcmod_iff) | 
| 27468 | 156 | |
| 157 | lemma SComplex_Infinitesimal_zero: | |
| 70216 | 158 | "\<lbrakk>x \<in> SComplex; x \<in> Infinitesimal\<rbrakk> \<Longrightarrow> x = 0" | 
| 159 | using SComplex_iff by auto | |
| 27468 | 160 | |
| 161 | lemma SComplex_HFinite_diff_Infinitesimal: | |
| 70216 | 162 | "\<lbrakk>x \<in> SComplex; x \<noteq> 0\<rbrakk> \<Longrightarrow> x \<in> HFinite - Infinitesimal" | 
| 163 | using SComplex_iff by auto | |
| 27468 | 164 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
37887diff
changeset | 165 | lemma numeral_not_Infinitesimal [simp]: | 
| 70216 | 166 | "numeral w \<noteq> (0::hcomplex) \<Longrightarrow> (numeral w::hcomplex) \<notin> Infinitesimal" | 
| 167 | by (fast dest: Standard_numeral [THEN SComplex_Infinitesimal_zero]) | |
| 27468 | 168 | |
| 169 | lemma approx_SComplex_not_zero: | |
| 70216 | 170 | "\<lbrakk>y \<in> SComplex; x \<approx> y; y\<noteq> 0\<rbrakk> \<Longrightarrow> x \<noteq> 0" | 
| 171 | by (auto dest: SComplex_Infinitesimal_zero approx_sym [THEN mem_infmal_iff [THEN iffD2]]) | |
| 27468 | 172 | |
| 173 | lemma SComplex_approx_iff: | |
| 70216 | 174 | "\<lbrakk>x \<in> SComplex; y \<in> SComplex\<rbrakk> \<Longrightarrow> (x \<approx> y) = (x = y)" | 
| 175 | by (auto simp add: Standard_def) | |
| 27468 | 176 | |
| 177 | lemma approx_unique_complex: | |
| 70216 | 178 | "\<lbrakk>r \<in> SComplex; s \<in> SComplex; r \<approx> x; s \<approx> x\<rbrakk> \<Longrightarrow> r = s" | 
| 179 | by (blast intro: SComplex_approx_iff [THEN iffD1] approx_trans2) | |
| 27468 | 180 | |
| 69597 | 181 | subsection \<open>Properties of \<^term>\<open>hRe\<close>, \<^term>\<open>hIm\<close> and \<^term>\<open>HComplex\<close>\<close> | 
| 27468 | 182 | |
| 183 | lemma abs_hRe_le_hcmod: "\<And>x. \<bar>hRe x\<bar> \<le> hcmod x" | |
| 70216 | 184 | by transfer (rule abs_Re_le_cmod) | 
| 27468 | 185 | |
| 186 | lemma abs_hIm_le_hcmod: "\<And>x. \<bar>hIm x\<bar> \<le> hcmod x" | |
| 70216 | 187 | by transfer (rule abs_Im_le_cmod) | 
| 27468 | 188 | |
| 189 | lemma Infinitesimal_hRe: "x \<in> Infinitesimal \<Longrightarrow> hRe x \<in> Infinitesimal" | |
| 70216 | 190 | using Infinitesimal_hcmod_iff abs_hRe_le_hcmod hrabs_le_Infinitesimal by blast | 
| 27468 | 191 | |
| 192 | lemma Infinitesimal_hIm: "x \<in> Infinitesimal \<Longrightarrow> hIm x \<in> Infinitesimal" | |
| 70216 | 193 | using Infinitesimal_hcmod_iff abs_hIm_le_hcmod hrabs_le_Infinitesimal by blast | 
| 27468 | 194 | |
| 195 | lemma Infinitesimal_HComplex: | |
| 70217 | 196 | assumes x: "x \<in> Infinitesimal" and y: "y \<in> Infinitesimal" | 
| 197 | shows "HComplex x y \<in> Infinitesimal" | |
| 198 | proof - | |
| 199 | have "hcmod (HComplex 0 y) \<in> Infinitesimal" | |
| 200 | by (simp add: hcmod_i y) | |
| 201 | moreover have "hcmod (hcomplex_of_hypreal x) \<in> Infinitesimal" | |
| 202 | using Infinitesimal_hcmod_iff Infinitesimal_of_hypreal_iff x by blast | |
| 203 | ultimately have "hcmod (HComplex x y) \<in> Infinitesimal" | |
| 204 | by (metis Infinitesimal_add Infinitesimal_hcmod_iff add.right_neutral hcomplex_of_hypreal_add_HComplex) | |
| 205 | then show ?thesis | |
| 206 | by (simp add: Infinitesimal_hnorm_iff) | |
| 207 | qed | |
| 27468 | 208 | |
| 209 | lemma hcomplex_Infinitesimal_iff: | |
| 70217 | 210 | "(x \<in> Infinitesimal) \<longleftrightarrow> (hRe x \<in> Infinitesimal \<and> hIm x \<in> Infinitesimal)" | 
| 70216 | 211 | using Infinitesimal_HComplex Infinitesimal_hIm Infinitesimal_hRe by fastforce | 
| 27468 | 212 | |
| 213 | lemma hRe_diff [simp]: "\<And>x y. hRe (x - y) = hRe x - hRe y" | |
| 70216 | 214 | by transfer simp | 
| 27468 | 215 | |
| 216 | lemma hIm_diff [simp]: "\<And>x y. hIm (x - y) = hIm x - hIm y" | |
| 70216 | 217 | by transfer simp | 
| 27468 | 218 | |
| 219 | lemma approx_hRe: "x \<approx> y \<Longrightarrow> hRe x \<approx> hRe y" | |
| 70216 | 220 | unfolding approx_def by (drule Infinitesimal_hRe) simp | 
| 27468 | 221 | |
| 222 | lemma approx_hIm: "x \<approx> y \<Longrightarrow> hIm x \<approx> hIm y" | |
| 70216 | 223 | unfolding approx_def by (drule Infinitesimal_hIm) simp | 
| 27468 | 224 | |
| 225 | lemma approx_HComplex: | |
| 226 | "\<lbrakk>a \<approx> b; c \<approx> d\<rbrakk> \<Longrightarrow> HComplex a c \<approx> HComplex b d" | |
| 70216 | 227 | unfolding approx_def by (simp add: Infinitesimal_HComplex) | 
| 27468 | 228 | |
| 229 | lemma hcomplex_approx_iff: | |
| 230 | "(x \<approx> y) = (hRe x \<approx> hRe y \<and> hIm x \<approx> hIm y)" | |
| 70216 | 231 | unfolding approx_def by (simp add: hcomplex_Infinitesimal_iff) | 
| 27468 | 232 | |
| 233 | lemma HFinite_hRe: "x \<in> HFinite \<Longrightarrow> hRe x \<in> HFinite" | |
| 70216 | 234 | using HFinite_bounded_hcmod abs_ge_zero abs_hRe_le_hcmod by blast | 
| 27468 | 235 | |
| 236 | lemma HFinite_hIm: "x \<in> HFinite \<Longrightarrow> hIm x \<in> HFinite" | |
| 70216 | 237 | using HFinite_bounded_hcmod abs_ge_zero abs_hIm_le_hcmod by blast | 
| 27468 | 238 | |
| 239 | lemma HFinite_HComplex: | |
| 70217 | 240 | assumes "x \<in> HFinite" "y \<in> HFinite" | 
| 241 | shows "HComplex x y \<in> HFinite" | |
| 242 | proof - | |
| 243 | have "HComplex x 0 \<in> HFinite" "HComplex 0 y \<in> HFinite" | |
| 244 | using HFinite_hcmod_iff assms hcmod_i by fastforce+ | |
| 245 | then have "HComplex x 0 + HComplex 0 y \<in> HFinite" | |
| 246 | using HFinite_add by blast | |
| 247 | then show ?thesis | |
| 248 | by simp | |
| 249 | qed | |
| 27468 | 250 | |
| 251 | lemma hcomplex_HFinite_iff: | |
| 252 | "(x \<in> HFinite) = (hRe x \<in> HFinite \<and> hIm x \<in> HFinite)" | |
| 70216 | 253 | using HFinite_HComplex HFinite_hIm HFinite_hRe by fastforce | 
| 27468 | 254 | |
| 255 | lemma hcomplex_HInfinite_iff: | |
| 256 | "(x \<in> HInfinite) = (hRe x \<in> HInfinite \<or> hIm x \<in> HInfinite)" | |
| 70216 | 257 | by (simp add: HInfinite_HFinite_iff hcomplex_HFinite_iff) | 
| 27468 | 258 | |
| 259 | lemma hcomplex_of_hypreal_approx_iff [simp]: | |
| 70216 | 260 | "(hcomplex_of_hypreal x \<approx> hcomplex_of_hypreal z) = (x \<approx> z)" | 
| 261 | by (simp add: hcomplex_approx_iff) | |
| 27468 | 262 | |
| 263 | (* Here we go - easy proof now!! *) | |
| 70217 | 264 | lemma stc_part_Ex: | 
| 265 | assumes "x \<in> HFinite" | |
| 266 | shows "\<exists>t \<in> SComplex. x \<approx> t" | |
| 267 | proof - | |
| 268 | let ?t = "HComplex (st (hRe x)) (st (hIm x))" | |
| 269 | have "?t \<in> SComplex" | |
| 270 | using HFinite_hIm HFinite_hRe Reals_eq_Standard assms st_SReal by auto | |
| 271 | moreover have "x \<approx> ?t" | |
| 272 | by (simp add: HFinite_hIm HFinite_hRe assms hcomplex_approx_iff st_HFinite st_eq_approx) | |
| 273 | ultimately show ?thesis .. | |
| 274 | qed | |
| 27468 | 275 | |
| 67613 | 276 | lemma stc_part_Ex1: "x \<in> HFinite \<Longrightarrow> \<exists>!t. t \<in> SComplex \<and> x \<approx> t" | 
| 70216 | 277 | using approx_sym approx_unique_complex stc_part_Ex by blast | 
| 27468 | 278 | |
| 279 | ||
| 61975 | 280 | subsection\<open>Theorems About Monads\<close> | 
| 27468 | 281 | |
| 67613 | 282 | lemma monad_zero_hcmod_iff: "(x \<in> monad 0) = (hcmod x \<in> monad 0)" | 
| 70216 | 283 | by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff) | 
| 27468 | 284 | |
| 285 | ||
| 61975 | 286 | subsection\<open>Theorems About Standard Part\<close> | 
| 27468 | 287 | |
| 70216 | 288 | lemma stc_approx_self: "x \<in> HFinite \<Longrightarrow> stc x \<approx> x" | 
| 289 | unfolding stc_def | |
| 290 | by (metis (no_types, lifting) approx_reorient someI_ex stc_part_Ex1) | |
| 27468 | 291 | |
| 70216 | 292 | lemma stc_SComplex: "x \<in> HFinite \<Longrightarrow> stc x \<in> SComplex" | 
| 293 | unfolding stc_def | |
| 294 | by (metis (no_types, lifting) SComplex_iff approx_sym someI_ex stc_part_Ex) | |
| 27468 | 295 | |
| 70216 | 296 | lemma stc_HFinite: "x \<in> HFinite \<Longrightarrow> stc x \<in> HFinite" | 
| 297 | by (erule stc_SComplex [THEN Standard_subset_HFinite [THEN subsetD]]) | |
| 27468 | 298 | |
| 299 | lemma stc_unique: "\<lbrakk>y \<in> SComplex; y \<approx> x\<rbrakk> \<Longrightarrow> stc x = y" | |
| 70216 | 300 | by (metis SComplex_approx_iff SComplex_iff approx_monad_iff approx_star_of_HFinite stc_SComplex stc_approx_self) | 
| 27468 | 301 | |
| 70216 | 302 | lemma stc_SComplex_eq [simp]: "x \<in> SComplex \<Longrightarrow> stc x = x" | 
| 303 | by (simp add: stc_unique) | |
| 27468 | 304 | |
| 305 | lemma stc_eq_approx: | |
| 70216 | 306 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite; stc x = stc y\<rbrakk> \<Longrightarrow> x \<approx> y" | 
| 307 | by (auto dest!: stc_approx_self elim!: approx_trans3) | |
| 27468 | 308 | |
| 309 | lemma approx_stc_eq: | |
| 70216 | 310 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite; x \<approx> y\<rbrakk> \<Longrightarrow> stc x = stc y" | 
| 311 | by (metis approx_sym approx_trans3 stc_part_Ex1 stc_unique) | |
| 27468 | 312 | |
| 313 | lemma stc_eq_approx_iff: | |
| 70216 | 314 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> (x \<approx> y) = (stc x = stc y)" | 
| 315 | by (blast intro: approx_stc_eq stc_eq_approx) | |
| 27468 | 316 | |
| 317 | lemma stc_Infinitesimal_add_SComplex: | |
| 70216 | 318 | "\<lbrakk>x \<in> SComplex; e \<in> Infinitesimal\<rbrakk> \<Longrightarrow> stc(x + e) = x" | 
| 319 | using Infinitesimal_add_approx_self stc_unique by blast | |
| 27468 | 320 | |
| 321 | lemma stc_Infinitesimal_add_SComplex2: | |
| 70216 | 322 | "\<lbrakk>x \<in> SComplex; e \<in> Infinitesimal\<rbrakk> \<Longrightarrow> stc(e + x) = x" | 
| 323 | using Infinitesimal_add_approx_self2 stc_unique by blast | |
| 27468 | 324 | |
| 325 | lemma HFinite_stc_Infinitesimal_add: | |
| 70216 | 326 | "x \<in> HFinite \<Longrightarrow> \<exists>e \<in> Infinitesimal. x = stc(x) + e" | 
| 327 | by (blast dest!: stc_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) | |
| 27468 | 328 | |
| 329 | lemma stc_add: | |
| 70216 | 330 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> stc (x + y) = stc(x) + stc(y)" | 
| 331 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_add) | |
| 27468 | 332 | |
| 70216 | 333 | lemma stc_zero: "stc 0 = 0" | 
| 334 | by simp | |
| 27468 | 335 | |
| 70216 | 336 | lemma stc_one: "stc 1 = 1" | 
| 337 | by simp | |
| 27468 | 338 | |
| 70216 | 339 | lemma stc_minus: "y \<in> HFinite \<Longrightarrow> stc(-y) = -stc(y)" | 
| 340 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_minus) | |
| 27468 | 341 | |
| 342 | lemma stc_diff: | |
| 70216 | 343 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> stc (x-y) = stc(x) - stc(y)" | 
| 344 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_diff) | |
| 27468 | 345 | |
| 346 | lemma stc_mult: | |
| 70216 | 347 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> | 
| 348 | \<Longrightarrow> stc (x * y) = stc(x) * stc(y)" | |
| 349 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_mult_HFinite) | |
| 27468 | 350 | |
| 70216 | 351 | lemma stc_Infinitesimal: "x \<in> Infinitesimal \<Longrightarrow> stc x = 0" | 
| 352 | by (simp add: stc_unique mem_infmal_iff) | |
| 27468 | 353 | |
| 70216 | 354 | lemma stc_not_Infinitesimal: "stc(x) \<noteq> 0 \<Longrightarrow> x \<notin> Infinitesimal" | 
| 355 | by (fast intro: stc_Infinitesimal) | |
| 27468 | 356 | |
| 357 | lemma stc_inverse: | |
| 70216 | 358 | "\<lbrakk>x \<in> HFinite; stc x \<noteq> 0\<rbrakk> \<Longrightarrow> stc(inverse x) = inverse (stc x)" | 
| 70217 | 359 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_inverse stc_not_Infinitesimal) | 
| 27468 | 360 | |
| 361 | lemma stc_divide [simp]: | |
| 70216 | 362 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite; stc y \<noteq> 0\<rbrakk> | 
| 363 | \<Longrightarrow> stc(x/y) = (stc x) / (stc y)" | |
| 364 | by (simp add: divide_inverse stc_mult stc_not_Infinitesimal HFinite_inverse stc_inverse) | |
| 27468 | 365 | |
| 70216 | 366 | lemma stc_idempotent [simp]: "x \<in> HFinite \<Longrightarrow> stc(stc(x)) = stc(x)" | 
| 367 | by (blast intro: stc_HFinite stc_approx_self approx_stc_eq) | |
| 27468 | 368 | |
| 369 | lemma HFinite_HFinite_hcomplex_of_hypreal: | |
| 70216 | 370 | "z \<in> HFinite \<Longrightarrow> hcomplex_of_hypreal z \<in> HFinite" | 
| 371 | by (simp add: hcomplex_HFinite_iff) | |
| 27468 | 372 | |
| 373 | lemma SComplex_SReal_hcomplex_of_hypreal: | |
| 70216 | 374 | "x \<in> \<real> \<Longrightarrow> hcomplex_of_hypreal x \<in> SComplex" | 
| 375 | by (simp add: Reals_eq_Standard) | |
| 27468 | 376 | |
| 377 | lemma stc_hcomplex_of_hypreal: | |
| 70216 | 378 | "z \<in> HFinite \<Longrightarrow> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)" | 
| 70217 | 379 | by (simp add: SComplex_SReal_hcomplex_of_hypreal st_SReal st_approx_self stc_unique) | 
| 27468 | 380 | |
| 70217 | 381 | lemma hmod_stc_eq: | 
| 382 | assumes "x \<in> HFinite" | |
| 383 | shows "hcmod(stc x) = st(hcmod x)" | |
| 384 | by (metis SReal_hcmod_SComplex approx_HFinite approx_hnorm assms st_unique stc_SComplex_eq stc_eq_approx_iff stc_part_Ex) | |
| 27468 | 385 | |
| 386 | lemma Infinitesimal_hcnj_iff [simp]: | |
| 70217 | 387 | "(hcnj z \<in> Infinitesimal) \<longleftrightarrow> (z \<in> Infinitesimal)" | 
| 388 | by (simp add: Infinitesimal_hcmod_iff) | |
| 27468 | 389 | |
| 390 | end |