| author | wenzelm | 
| Thu, 08 Nov 2007 20:52:27 +0100 | |
| changeset 25356 | 059c03630d6e | 
| parent 22956 | 617140080e6a | 
| child 26119 | cb9bdde1b444 | 
| permissions | -rw-r--r-- | 
| 13957 | 1 | (* Title : NSCA.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 2001,2002 University of Edinburgh | |
| 4 | *) | |
| 5 | ||
| 14408 | 6 | header{*Non-Standard Complex Analysis*}
 | 
| 13957 | 7 | |
| 15131 | 8 | theory NSCA | 
| 22655 | 9 | imports NSComplex "../Hyperreal/HTranscendental" | 
| 15131 | 10 | begin | 
| 13957 | 11 | |
| 21830 | 12 | abbreviation | 
| 13957 | 13 | (* standard complex numbers reagarded as an embedded subset of NS complex *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20724diff
changeset | 14 | SComplex :: "hcomplex set" where | 
| 21830 | 15 | "SComplex \<equiv> Standard" | 
| 13957 | 16 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20724diff
changeset | 17 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20724diff
changeset | 18 | stc :: "hcomplex => hcomplex" where | 
| 14408 | 19 |     --{* standard part map*}
 | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 20 | "stc x = (SOME r. x \<in> HFinite & r:SComplex & r @= x)" | 
| 14408 | 21 | |
| 22 | ||
| 23 | subsection{*Closure Laws for SComplex, the Standard Complex Numbers*}
 | |
| 24 | ||
| 25 | lemma SComplex_minus_iff [simp]: "(-x \<in> SComplex) = (x \<in> SComplex)" | |
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 26 | by (auto, drule Standard_minus, auto) | 
| 14408 | 27 | |
| 28 | lemma SComplex_add_cancel: | |
| 29 | "[| x + y \<in> SComplex; y \<in> SComplex |] ==> x \<in> SComplex" | |
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 30 | by (drule (1) Standard_diff, simp) | 
| 14408 | 31 | |
| 32 | lemma SReal_hcmod_hcomplex_of_complex [simp]: | |
| 33 | "hcmod (hcomplex_of_complex r) \<in> Reals" | |
| 21830 | 34 | by (simp add: Reals_eq_Standard) | 
| 14408 | 35 | |
| 36 | lemma SReal_hcmod_number_of [simp]: "hcmod (number_of w ::hcomplex) \<in> Reals" | |
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 37 | by (simp add: Reals_eq_Standard) | 
| 14408 | 38 | |
| 39 | lemma SReal_hcmod_SComplex: "x \<in> SComplex ==> hcmod x \<in> Reals" | |
| 21830 | 40 | by (simp add: Reals_eq_Standard) | 
| 14408 | 41 | |
| 42 | lemma SComplex_divide_number_of: | |
| 43 | "r \<in> SComplex ==> r/(number_of w::hcomplex) \<in> SComplex" | |
| 21830 | 44 | by simp | 
| 14408 | 45 | |
| 46 | lemma SComplex_UNIV_complex: | |
| 47 |      "{x. hcomplex_of_complex x \<in> SComplex} = (UNIV::complex set)"
 | |
| 21830 | 48 | by simp | 
| 14408 | 49 | |
| 50 | lemma SComplex_iff: "(x \<in> SComplex) = (\<exists>y. x = hcomplex_of_complex y)" | |
| 21830 | 51 | by (simp add: Standard_def image_def) | 
| 14408 | 52 | |
| 53 | lemma hcomplex_of_complex_image: | |
| 54 | "hcomplex_of_complex `(UNIV::complex set) = SComplex" | |
| 21830 | 55 | by (simp add: Standard_def) | 
| 14408 | 56 | |
| 57 | lemma inv_hcomplex_of_complex_image: "inv hcomplex_of_complex `SComplex = UNIV" | |
| 21830 | 58 | apply (auto simp add: Standard_def image_def) | 
| 14408 | 59 | apply (rule inj_hcomplex_of_complex [THEN inv_f_f, THEN subst], blast) | 
| 60 | done | |
| 61 | ||
| 62 | lemma SComplex_hcomplex_of_complex_image: | |
| 63 | "[| \<exists>x. x: P; P \<le> SComplex |] ==> \<exists>Q. P = hcomplex_of_complex ` Q" | |
| 21830 | 64 | apply (simp add: Standard_def, blast) | 
| 14408 | 65 | done | 
| 66 | ||
| 67 | lemma SComplex_SReal_dense: | |
| 68 | "[| x \<in> SComplex; y \<in> SComplex; hcmod x < hcmod y | |
| 69 | |] ==> \<exists>r \<in> Reals. hcmod x< r & r < hcmod y" | |
| 70 | apply (auto intro: SReal_dense simp add: SReal_hcmod_SComplex) | |
| 71 | done | |
| 72 | ||
| 73 | lemma SComplex_hcmod_SReal: | |
| 74 | "z \<in> SComplex ==> hcmod z \<in> Reals" | |
| 21830 | 75 | by (simp add: Reals_eq_Standard) | 
| 14408 | 76 | |
| 77 | (* | |
| 78 | Goalw [SComplex_def,SReal_def] "hcmod z \<in> Reals ==> z \<in> SComplex" | |
| 79 | by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1);
 | |
| 80 | by (auto_tac (claset(),simpset() addsimps [hcmod,hypreal_of_real_def, | |
| 81 | hcomplex_of_complex_def,cmod_def])); | |
| 82 | *) | |
| 83 | ||
| 84 | ||
| 85 | subsection{*The Finite Elements form a Subring*}
 | |
| 86 | ||
| 87 | lemma HFinite_hcmod_hcomplex_of_complex [simp]: | |
| 88 | "hcmod (hcomplex_of_complex r) \<in> HFinite" | |
| 89 | by (auto intro!: SReal_subset_HFinite [THEN subsetD]) | |
| 90 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 91 | lemma HFinite_hcmod_iff: "(x \<in> HFinite) = (hcmod x \<in> HFinite)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 92 | by (simp add: HFinite_def) | 
| 14408 | 93 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 94 | lemma HFinite_bounded_hcmod: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 95 | "[|x \<in> HFinite; y \<le> hcmod x; 0 \<le> y |] ==> y: HFinite" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 96 | by (auto intro: HFinite_bounded simp add: HFinite_hcmod_iff) | 
| 14408 | 97 | |
| 98 | ||
| 99 | subsection{*The Complex Infinitesimals form a Subring*}
 | |
| 100 | ||
| 101 | lemma hcomplex_sum_of_halves: "x/(2::hcomplex) + x/(2::hcomplex) = x" | |
| 102 | by auto | |
| 103 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 104 | lemma Infinitesimal_hcmod_iff: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 105 | "(z \<in> Infinitesimal) = (hcmod z \<in> Infinitesimal)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 106 | by (simp add: Infinitesimal_def) | 
| 14408 | 107 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 108 | lemma HInfinite_hcmod_iff: "(z \<in> HInfinite) = (hcmod z \<in> HInfinite)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 109 | by (simp add: HInfinite_def) | 
| 14408 | 110 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 111 | lemma HFinite_diff_Infinitesimal_hcmod: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 112 | "x \<in> HFinite - Infinitesimal ==> hcmod x \<in> HFinite - Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 113 | by (simp add: HFinite_hcmod_iff Infinitesimal_hcmod_iff) | 
| 14408 | 114 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 115 | lemma hcmod_less_Infinitesimal: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 116 | "[| e \<in> Infinitesimal; hcmod x < hcmod e |] ==> x \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 117 | by (auto elim: hrabs_less_Infinitesimal simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 118 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 119 | lemma hcmod_le_Infinitesimal: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 120 | "[| e \<in> Infinitesimal; hcmod x \<le> hcmod e |] ==> x \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 121 | by (auto elim: hrabs_le_Infinitesimal simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 122 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 123 | lemma Infinitesimal_interval_hcmod: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 124 | "[| e \<in> Infinitesimal; | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 125 | e' \<in> Infinitesimal; | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 126 | hcmod e' < hcmod x ; hcmod x < hcmod e | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 127 | |] ==> x \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 128 | by (auto intro: Infinitesimal_interval simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 129 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 130 | lemma Infinitesimal_interval2_hcmod: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 131 | "[| e \<in> Infinitesimal; | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 132 | e' \<in> Infinitesimal; | 
| 14408 | 133 | hcmod e' \<le> hcmod x ; hcmod x \<le> hcmod e | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 134 | |] ==> x \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 135 | by (auto intro: Infinitesimal_interval2 simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 136 | |
| 137 | ||
| 138 | subsection{*The ``Infinitely Close'' Relation*}
 | |
| 139 | ||
| 140 | (* | |
| 141 | Goalw [capprox_def,approx_def] "(z @c= w) = (hcmod z @= hcmod w)" | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 142 | by (auto_tac (claset(),simpset() addsimps [Infinitesimal_hcmod_iff])); | 
| 14408 | 143 | *) | 
| 144 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 145 | lemma approx_SComplex_mult_cancel_zero: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 146 | "[| a \<in> SComplex; a \<noteq> 0; a*x @= 0 |] ==> x @= 0" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 147 | apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]]) | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 148 | apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) | 
| 14408 | 149 | done | 
| 150 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 151 | lemma approx_mult_SComplex1: "[| a \<in> SComplex; x @= 0 |] ==> x*a @= 0" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 152 | by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult1) | 
| 14408 | 153 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 154 | lemma approx_mult_SComplex2: "[| a \<in> SComplex; x @= 0 |] ==> a*x @= 0" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 155 | by (auto dest: Standard_subset_HFinite [THEN subsetD] approx_mult2) | 
| 14408 | 156 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 157 | lemma approx_mult_SComplex_zero_cancel_iff [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 158 | "[|a \<in> SComplex; a \<noteq> 0 |] ==> (a*x @= 0) = (x @= 0)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 159 | by (blast intro: approx_SComplex_mult_cancel_zero approx_mult_SComplex2) | 
| 14408 | 160 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 161 | lemma approx_SComplex_mult_cancel: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 162 | "[| a \<in> SComplex; a \<noteq> 0; a* w @= a*z |] ==> w @= z" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 163 | apply (drule Standard_inverse [THEN Standard_subset_HFinite [THEN subsetD]]) | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 164 | apply (auto dest: approx_mult2 simp add: mult_assoc [symmetric]) | 
| 14408 | 165 | done | 
| 166 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 167 | lemma approx_SComplex_mult_cancel_iff1 [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 168 | "[| a \<in> SComplex; a \<noteq> 0|] ==> (a* w @= a*z) = (w @= z)" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 169 | by (auto intro!: approx_mult2 Standard_subset_HFinite [THEN subsetD] | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 170 | intro: approx_SComplex_mult_cancel) | 
| 14408 | 171 | |
| 21810 
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
 huffman parents: 
21404diff
changeset | 172 | (* TODO: generalize following theorems: hcmod -> hnorm *) | 
| 
b2d23672b003
generalized some lemmas; removed redundant lemmas; cleaned up some proofs
 huffman parents: 
21404diff
changeset | 173 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 174 | lemma approx_hcmod_approx_zero: "(x @= y) = (hcmod (y - x) @= 0)" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 175 | apply (subst hnorm_minus_commute) | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 176 | apply (simp add: approx_def Infinitesimal_hcmod_iff diff_minus) | 
| 14408 | 177 | done | 
| 178 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 179 | lemma approx_approx_zero_iff: "(x @= 0) = (hcmod x @= 0)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 180 | by (simp add: approx_hcmod_approx_zero) | 
| 14408 | 181 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 182 | lemma approx_minus_zero_cancel_iff [simp]: "(-x @= 0) = (x @= 0)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 183 | by (simp add: approx_def) | 
| 14408 | 184 | |
| 185 | lemma Infinitesimal_hcmod_add_diff: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 186 | "u @= 0 ==> hcmod(x + u) - hcmod x \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 187 | apply (drule approx_approx_zero_iff [THEN iffD1]) | 
| 14408 | 188 | apply (rule_tac e = "hcmod u" and e' = "- hcmod u" in Infinitesimal_interval2) | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 189 | apply (auto simp add: mem_infmal_iff [symmetric] diff_def) | 
| 14408 | 190 | apply (rule_tac c1 = "hcmod x" in add_le_cancel_left [THEN iffD1]) | 
| 191 | apply (auto simp add: diff_minus [symmetric]) | |
| 192 | done | |
| 193 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 194 | lemma approx_hcmod_add_hcmod: "u @= 0 ==> hcmod(x + u) @= hcmod x" | 
| 14408 | 195 | apply (rule approx_minus_iff [THEN iffD2]) | 
| 196 | apply (auto intro: Infinitesimal_hcmod_add_diff simp add: mem_infmal_iff [symmetric] diff_minus [symmetric]) | |
| 197 | done | |
| 198 | ||
| 13957 | 199 | |
| 14408 | 200 | subsection{*Zero is the Only Infinitesimal Complex Number*}
 | 
| 201 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 202 | lemma Infinitesimal_less_SComplex: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 203 | "[| x \<in> SComplex; y \<in> Infinitesimal; 0 < hcmod x |] ==> hcmod y < hcmod x" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 204 | by (auto intro: Infinitesimal_less_SReal SComplex_hcmod_SReal simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 205 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 206 | lemma SComplex_Int_Infinitesimal_zero: "SComplex Int Infinitesimal = {0}"
 | 
| 21830 | 207 | by (auto simp add: Standard_def Infinitesimal_hcmod_iff) | 
| 14408 | 208 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 209 | lemma SComplex_Infinitesimal_zero: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 210 | "[| x \<in> SComplex; x \<in> Infinitesimal|] ==> x = 0" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 211 | by (cut_tac SComplex_Int_Infinitesimal_zero, blast) | 
| 14408 | 212 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 213 | lemma SComplex_HFinite_diff_Infinitesimal: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 214 | "[| x \<in> SComplex; x \<noteq> 0 |] ==> x \<in> HFinite - Infinitesimal" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 215 | by (auto dest: SComplex_Infinitesimal_zero Standard_subset_HFinite [THEN subsetD]) | 
| 14408 | 216 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 217 | lemma hcomplex_of_complex_HFinite_diff_Infinitesimal: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 218 | "hcomplex_of_complex x \<noteq> 0 | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 219 | ==> hcomplex_of_complex x \<in> HFinite - Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 220 | by (rule SComplex_HFinite_diff_Infinitesimal, auto) | 
| 14408 | 221 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 222 | lemma number_of_not_Infinitesimal [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 223 | "number_of w \<noteq> (0::hcomplex) ==> (number_of w::hcomplex) \<notin> Infinitesimal" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 224 | by (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero]) | 
| 14408 | 225 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 226 | lemma approx_SComplex_not_zero: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 227 | "[| y \<in> SComplex; x @= y; y\<noteq> 0 |] ==> x \<noteq> 0" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 228 | by (auto dest: SComplex_Infinitesimal_zero approx_sym [THEN mem_infmal_iff [THEN iffD2]]) | 
| 14408 | 229 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 230 | lemma SComplex_approx_iff: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 231 | "[|x \<in> SComplex; y \<in> SComplex|] ==> (x @= y) = (x = y)" | 
| 21830 | 232 | by (auto simp add: Standard_def) | 
| 14408 | 233 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 234 | lemma number_of_Infinitesimal_iff [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 235 | "((number_of w :: hcomplex) \<in> Infinitesimal) = | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 236 | (number_of w = (0::hcomplex))" | 
| 14408 | 237 | apply (rule iffI) | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 238 | apply (fast dest: Standard_number_of [THEN SComplex_Infinitesimal_zero]) | 
| 14408 | 239 | apply (simp (no_asm_simp)) | 
| 240 | done | |
| 241 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 242 | lemma approx_unique_complex: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 243 | "[| r \<in> SComplex; s \<in> SComplex; r @= x; s @= x|] ==> r = s" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 244 | by (blast intro: SComplex_approx_iff [THEN iffD1] approx_trans2) | 
| 14408 | 245 | |
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 246 | subsection {* Properties of @{term hRe}, @{term hIm} and @{term HComplex} *}
 | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 247 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 248 | lemma abs_Re_le_cmod: "\<bar>Re x\<bar> \<le> cmod x" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 249 | by (induct x) simp | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 250 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 251 | lemma abs_Im_le_cmod: "\<bar>Im x\<bar> \<le> cmod x" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 252 | by (induct x) simp | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 253 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 254 | lemma abs_hRe_le_hcmod: "\<And>x. \<bar>hRe x\<bar> \<le> hcmod x" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 255 | by transfer (rule abs_Re_le_cmod) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 256 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 257 | lemma abs_hIm_le_hcmod: "\<And>x. \<bar>hIm x\<bar> \<le> hcmod x" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 258 | by transfer (rule abs_Im_le_cmod) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 259 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 260 | lemma Infinitesimal_hRe: "x \<in> Infinitesimal \<Longrightarrow> hRe x \<in> Infinitesimal" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 261 | apply (rule InfinitesimalI2, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 262 | apply (rule order_le_less_trans [OF abs_hRe_le_hcmod]) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 263 | apply (erule (1) InfinitesimalD2) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 264 | done | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 265 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 266 | lemma Infinitesimal_hIm: "x \<in> Infinitesimal \<Longrightarrow> hIm x \<in> Infinitesimal" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 267 | apply (rule InfinitesimalI2, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 268 | apply (rule order_le_less_trans [OF abs_hIm_le_hcmod]) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 269 | apply (erule (1) InfinitesimalD2) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 270 | done | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 271 | |
| 22956 
617140080e6a
define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
 huffman parents: 
22883diff
changeset | 272 | lemma real_sqrt_lessI: "\<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> sqrt x < u" | 
| 
617140080e6a
define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
 huffman parents: 
22883diff
changeset | 273 | (* TODO: this belongs somewhere else *) | 
| 
617140080e6a
define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
 huffman parents: 
22883diff
changeset | 274 | by (frule real_sqrt_less_mono) simp | 
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 275 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 276 | lemma hypreal_sqrt_lessI: | 
| 22956 
617140080e6a
define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
 huffman parents: 
22883diff
changeset | 277 | "\<And>x u. \<lbrakk>0 < u; x < u\<twosuperior>\<rbrakk> \<Longrightarrow> ( *f* sqrt) x < u" | 
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 278 | by transfer (rule real_sqrt_lessI) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 279 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 280 | lemma hypreal_sqrt_ge_zero: "\<And>x. 0 \<le> x \<Longrightarrow> 0 \<le> ( *f* sqrt) x" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 281 | by transfer (rule real_sqrt_ge_zero) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 282 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 283 | lemma Infinitesimal_sqrt: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 284 | "\<lbrakk>x \<in> Infinitesimal; 0 \<le> x\<rbrakk> \<Longrightarrow> ( *f* sqrt) x \<in> Infinitesimal" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 285 | apply (rule InfinitesimalI2) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 286 | apply (drule_tac r="r\<twosuperior>" in InfinitesimalD2, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 287 | apply (simp add: hypreal_sqrt_ge_zero) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 288 | apply (rule hypreal_sqrt_lessI, simp_all) | 
| 14408 | 289 | done | 
| 290 | ||
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 291 | lemma Infinitesimal_HComplex: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 292 | "\<lbrakk>x \<in> Infinitesimal; y \<in> Infinitesimal\<rbrakk> \<Longrightarrow> HComplex x y \<in> Infinitesimal" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 293 | apply (rule Infinitesimal_hcmod_iff [THEN iffD2]) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 294 | apply (simp add: hcmod_i) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 295 | apply (rule Infinitesimal_sqrt) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 296 | apply (rule Infinitesimal_add) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 297 | apply (erule Infinitesimal_hrealpow, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 298 | apply (erule Infinitesimal_hrealpow, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 299 | apply (rule add_nonneg_nonneg) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 300 | apply (rule zero_le_power2) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 301 | apply (rule zero_le_power2) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 302 | done | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 303 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 304 | lemma hcomplex_Infinitesimal_iff: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 305 | "(x \<in> Infinitesimal) = (hRe x \<in> Infinitesimal \<and> hIm x \<in> Infinitesimal)" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 306 | apply (safe intro!: Infinitesimal_hRe Infinitesimal_hIm) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 307 | apply (drule (1) Infinitesimal_HComplex, simp) | 
| 14408 | 308 | done | 
| 309 | ||
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 310 | lemma hRe_diff [simp]: "\<And>x y. hRe (x - y) = hRe x - hRe y" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 311 | by transfer (rule complex_Re_diff) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 312 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 313 | lemma hIm_diff [simp]: "\<And>x y. hIm (x - y) = hIm x - hIm y" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 314 | by transfer (rule complex_Im_diff) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 315 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 316 | lemma approx_hRe: "x \<approx> y \<Longrightarrow> hRe x \<approx> hRe y" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 317 | unfolding approx_def by (drule Infinitesimal_hRe) simp | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 318 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 319 | lemma approx_hIm: "x \<approx> y \<Longrightarrow> hIm x \<approx> hIm y" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 320 | unfolding approx_def by (drule Infinitesimal_hIm) simp | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 321 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 322 | lemma approx_HComplex: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 323 | "\<lbrakk>a \<approx> b; c \<approx> d\<rbrakk> \<Longrightarrow> HComplex a c \<approx> HComplex b d" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 324 | unfolding approx_def by (simp add: Infinitesimal_HComplex) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 325 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 326 | lemma hcomplex_approx_iff: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 327 | "(x \<approx> y) = (hRe x \<approx> hRe y \<and> hIm x \<approx> hIm y)" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 328 | unfolding approx_def by (simp add: hcomplex_Infinitesimal_iff) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 329 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 330 | lemma HFinite_hRe: "x \<in> HFinite \<Longrightarrow> hRe x \<in> HFinite" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 331 | apply (auto simp add: HFinite_def SReal_def) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 332 | apply (rule_tac x="star_of r" in exI, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 333 | apply (erule order_le_less_trans [OF abs_hRe_le_hcmod]) | 
| 14408 | 334 | done | 
| 335 | ||
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 336 | lemma HFinite_hIm: "x \<in> HFinite \<Longrightarrow> hIm x \<in> HFinite" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 337 | apply (auto simp add: HFinite_def SReal_def) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 338 | apply (rule_tac x="star_of r" in exI, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 339 | apply (erule order_le_less_trans [OF abs_hIm_le_hcmod]) | 
| 14408 | 340 | done | 
| 341 | ||
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 342 | lemma HFinite_HComplex: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 343 | "\<lbrakk>x \<in> HFinite; y \<in> HFinite\<rbrakk> \<Longrightarrow> HComplex x y \<in> HFinite" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 344 | apply (subgoal_tac "HComplex x 0 + HComplex 0 y \<in> HFinite", simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 345 | apply (rule HFinite_add) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 346 | apply (simp add: HFinite_hcmod_iff hcmod_i) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 347 | apply (simp add: HFinite_hcmod_iff hcmod_i) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 348 | done | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 349 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 350 | lemma hcomplex_HFinite_iff: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 351 | "(x \<in> HFinite) = (hRe x \<in> HFinite \<and> hIm x \<in> HFinite)" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 352 | apply (safe intro!: HFinite_hRe HFinite_hIm) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 353 | apply (drule (1) HFinite_HComplex, simp) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 354 | done | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 355 | |
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 356 | lemma hcomplex_HInfinite_iff: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 357 | "(x \<in> HInfinite) = (hRe x \<in> HInfinite \<or> hIm x \<in> HInfinite)" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 358 | by (simp add: HInfinite_HFinite_iff hcomplex_HFinite_iff) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 359 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 360 | lemma hcomplex_of_hypreal_approx_iff [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 361 | "(hcomplex_of_hypreal x @= hcomplex_of_hypreal z) = (x @= z)" | 
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 362 | by (simp add: hcomplex_approx_iff) | 
| 14408 | 363 | |
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 364 | lemma Standard_HComplex: | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 365 | "\<lbrakk>x \<in> Standard; y \<in> Standard\<rbrakk> \<Longrightarrow> HComplex x y \<in> Standard" | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 366 | by (simp add: HComplex_def) | 
| 14408 | 367 | |
| 368 | (* Here we go - easy proof now!! *) | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 369 | lemma stc_part_Ex: "x:HFinite ==> \<exists>t \<in> SComplex. x @= t" | 
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 370 | apply (simp add: hcomplex_HFinite_iff hcomplex_approx_iff) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 371 | apply (rule_tac x="HComplex (st (hRe x)) (st (hIm x))" in bexI) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 372 | apply (simp add: st_approx_self [THEN approx_sym]) | 
| 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 373 | apply (simp add: Standard_HComplex st_SReal [unfolded Reals_eq_Standard]) | 
| 14408 | 374 | done | 
| 375 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 376 | lemma stc_part_Ex1: "x:HFinite ==> EX! t. t \<in> SComplex & x @= t" | 
| 14408 | 377 | apply (drule stc_part_Ex, safe) | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 378 | apply (drule_tac [2] approx_sym, drule_tac [2] approx_sym, drule_tac [2] approx_sym) | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 379 | apply (auto intro!: approx_unique_complex) | 
| 14408 | 380 | done | 
| 381 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 382 | lemmas hcomplex_of_complex_approx_inverse = | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 383 | hcomplex_of_complex_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] | 
| 14408 | 384 | |
| 385 | ||
| 386 | subsection{*Theorems About Monads*}
 | |
| 387 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 388 | lemma monad_zero_hcmod_iff: "(x \<in> monad 0) = (hcmod x:monad 0)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 389 | by (simp add: Infinitesimal_monad_zero_iff [symmetric] Infinitesimal_hcmod_iff) | 
| 14408 | 390 | |
| 391 | ||
| 392 | subsection{*Theorems About Standard Part*}
 | |
| 393 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 394 | lemma stc_approx_self: "x \<in> HFinite ==> stc x @= x" | 
| 14408 | 395 | apply (simp add: stc_def) | 
| 396 | apply (frule stc_part_Ex, safe) | |
| 397 | apply (rule someI2) | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 398 | apply (auto intro: approx_sym) | 
| 14408 | 399 | done | 
| 400 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 401 | lemma stc_SComplex: "x \<in> HFinite ==> stc x \<in> SComplex" | 
| 14408 | 402 | apply (simp add: stc_def) | 
| 403 | apply (frule stc_part_Ex, safe) | |
| 404 | apply (rule someI2) | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 405 | apply (auto intro: approx_sym) | 
| 14408 | 406 | done | 
| 407 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 408 | lemma stc_HFinite: "x \<in> HFinite ==> stc x \<in> HFinite" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 409 | by (erule stc_SComplex [THEN Standard_subset_HFinite [THEN subsetD]]) | 
| 14408 | 410 | |
| 20724 | 411 | lemma stc_unique: "\<lbrakk>y \<in> SComplex; y \<approx> x\<rbrakk> \<Longrightarrow> stc x = y" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 412 | apply (frule Standard_subset_HFinite [THEN subsetD]) | 
| 20724 | 413 | apply (drule (1) approx_HFinite) | 
| 414 | apply (unfold stc_def) | |
| 14408 | 415 | apply (rule some_equality) | 
| 20724 | 416 | apply (auto intro: approx_unique_complex) | 
| 417 | done | |
| 418 | ||
| 419 | lemma stc_SComplex_eq [simp]: "x \<in> SComplex ==> stc x = x" | |
| 420 | apply (erule stc_unique) | |
| 421 | apply (rule approx_refl) | |
| 14408 | 422 | done | 
| 423 | ||
| 424 | lemma stc_hcomplex_of_complex: | |
| 425 | "stc (hcomplex_of_complex x) = hcomplex_of_complex x" | |
| 426 | by auto | |
| 427 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 428 | lemma stc_eq_approx: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 429 | "[| x \<in> HFinite; y \<in> HFinite; stc x = stc y |] ==> x @= y" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 430 | by (auto dest!: stc_approx_self elim!: approx_trans3) | 
| 14408 | 431 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 432 | lemma approx_stc_eq: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 433 | "[| x \<in> HFinite; y \<in> HFinite; x @= y |] ==> stc x = stc y" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 434 | by (blast intro: approx_trans approx_trans2 SComplex_approx_iff [THEN iffD1] | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 435 | dest: stc_approx_self stc_SComplex) | 
| 13957 | 436 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 437 | lemma stc_eq_approx_iff: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 438 | "[| x \<in> HFinite; y \<in> HFinite|] ==> (x @= y) = (stc x = stc y)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 439 | by (blast intro: approx_stc_eq stc_eq_approx) | 
| 14408 | 440 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 441 | lemma stc_Infinitesimal_add_SComplex: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 442 | "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(x + e) = x" | 
| 20724 | 443 | apply (erule stc_unique) | 
| 444 | apply (erule Infinitesimal_add_approx_self) | |
| 14408 | 445 | done | 
| 446 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 447 | lemma stc_Infinitesimal_add_SComplex2: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 448 | "[| x \<in> SComplex; e \<in> Infinitesimal |] ==> stc(e + x) = x" | 
| 20724 | 449 | apply (erule stc_unique) | 
| 450 | apply (erule Infinitesimal_add_approx_self2) | |
| 14408 | 451 | done | 
| 452 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 453 | lemma HFinite_stc_Infinitesimal_add: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 454 | "x \<in> HFinite ==> \<exists>e \<in> Infinitesimal. x = stc(x) + e" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 455 | by (blast dest!: stc_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) | 
| 14408 | 456 | |
| 457 | lemma stc_add: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 458 | "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x + y) = stc(x) + stc(y)" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 459 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_add) | 
| 14408 | 460 | |
| 461 | lemma stc_number_of [simp]: "stc (number_of w) = number_of w" | |
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 462 | by (rule Standard_number_of [THEN stc_SComplex_eq]) | 
| 14408 | 463 | |
| 464 | lemma stc_zero [simp]: "stc 0 = 0" | |
| 465 | by simp | |
| 466 | ||
| 467 | lemma stc_one [simp]: "stc 1 = 1" | |
| 468 | by simp | |
| 469 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 470 | lemma stc_minus: "y \<in> HFinite ==> stc(-y) = -stc(y)" | 
| 20724 | 471 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_minus) | 
| 14408 | 472 | |
| 473 | lemma stc_diff: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 474 | "[| x \<in> HFinite; y \<in> HFinite |] ==> stc (x-y) = stc(x) - stc(y)" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 475 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_diff) | 
| 14408 | 476 | |
| 477 | lemma stc_mult: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 478 | "[| x \<in> HFinite; y \<in> HFinite |] | 
| 14408 | 479 | ==> stc (x * y) = stc(x) * stc(y)" | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 480 | by (simp add: stc_unique stc_SComplex stc_approx_self approx_mult_HFinite) | 
| 14408 | 481 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 482 | lemma stc_Infinitesimal: "x \<in> Infinitesimal ==> stc x = 0" | 
| 20724 | 483 | by (simp add: stc_unique mem_infmal_iff) | 
| 14408 | 484 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 485 | lemma stc_not_Infinitesimal: "stc(x) \<noteq> 0 ==> x \<notin> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 486 | by (fast intro: stc_Infinitesimal) | 
| 14408 | 487 | |
| 488 | lemma stc_inverse: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 489 | "[| x \<in> HFinite; stc x \<noteq> 0 |] | 
| 14408 | 490 | ==> stc(inverse x) = inverse (stc x)" | 
| 20724 | 491 | apply (drule stc_not_Infinitesimal) | 
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 492 | apply (simp add: stc_unique stc_SComplex stc_approx_self approx_inverse) | 
| 14408 | 493 | done | 
| 494 | ||
| 495 | lemma stc_divide [simp]: | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 496 | "[| x \<in> HFinite; y \<in> HFinite; stc y \<noteq> 0 |] | 
| 14408 | 497 | ==> stc(x/y) = (stc x) / (stc y)" | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 498 | by (simp add: divide_inverse stc_mult stc_not_Infinitesimal HFinite_inverse stc_inverse) | 
| 14408 | 499 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 500 | lemma stc_idempotent [simp]: "x \<in> HFinite ==> stc(stc(x)) = stc(x)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 501 | by (blast intro: stc_HFinite stc_approx_self approx_stc_eq) | 
| 14408 | 502 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 503 | lemma HFinite_HFinite_hcomplex_of_hypreal: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 504 | "z \<in> HFinite ==> hcomplex_of_hypreal z \<in> HFinite" | 
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21830diff
changeset | 505 | by (simp add: hcomplex_HFinite_iff) | 
| 14408 | 506 | |
| 507 | lemma SComplex_SReal_hcomplex_of_hypreal: | |
| 508 | "x \<in> Reals ==> hcomplex_of_hypreal x \<in> SComplex" | |
| 22883 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 509 | apply (rule Standard_of_hypreal) | 
| 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 510 | apply (simp add: Reals_eq_Standard) | 
| 
005be8dafce0
hcomplex_of_hypreal abbreviates of_hypreal; removed redundant lemmas
 huffman parents: 
22655diff
changeset | 511 | done | 
| 14408 | 512 | |
| 513 | lemma stc_hcomplex_of_hypreal: | |
| 514 | "z \<in> HFinite ==> stc(hcomplex_of_hypreal z) = hcomplex_of_hypreal (st z)" | |
| 20724 | 515 | apply (rule stc_unique) | 
| 516 | apply (rule SComplex_SReal_hcomplex_of_hypreal) | |
| 517 | apply (erule st_SReal) | |
| 518 | apply (simp add: hcomplex_of_hypreal_approx_iff st_approx_self) | |
| 14408 | 519 | done | 
| 520 | ||
| 521 | (* | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 522 | Goal "x \<in> HFinite ==> hcmod(stc x) = st(hcmod x)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 523 | by (dtac stc_approx_self 1) | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 524 | by (auto_tac (claset(),simpset() addsimps [bex_Infinitesimal_iff2 RS sym])); | 
| 14408 | 525 | |
| 526 | ||
| 527 | approx_hcmod_add_hcmod | |
| 528 | *) | |
| 529 | ||
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 530 | lemma Infinitesimal_hcnj_iff [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 531 | "(hcnj z \<in> Infinitesimal) = (z \<in> Infinitesimal)" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 532 | by (simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 533 | |
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 534 | lemma Infinitesimal_hcomplex_of_hypreal_epsilon [simp]: | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 535 | "hcomplex_of_hypreal epsilon \<in> Infinitesimal" | 
| 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20557diff
changeset | 536 | by (simp add: Infinitesimal_hcmod_iff) | 
| 14408 | 537 | |
| 13957 | 538 | end |