| author | nipkow | 
| Mon, 22 Nov 2004 11:53:56 +0100 | |
| changeset 15305 | 0bd9eedaa301 | 
| parent 15234 | ec91a90c604e | 
| child 17292 | 5c613b64bf0a | 
| permissions | -rw-r--r-- | 
| 13957 | 1 | (* Title : CStar.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 2001 University of Edinburgh | |
| 4 | *) | |
| 5 | ||
| 14407 | 6 | header{*Star-transforms in NSA, Extending Sets of Complex Numbers
 | 
| 7 | and Complex Functions*} | |
| 8 | ||
| 15131 | 9 | theory CStar | 
| 15140 | 10 | imports NSCA | 
| 15131 | 11 | begin | 
| 13957 | 12 | |
| 13 | constdefs | |
| 14 | ||
| 15 | (* nonstandard extension of sets *) | |
| 14407 | 16 |     starsetC :: "complex set => hcomplex set"          ("*sc* _" [80] 80)
 | 
| 17 |     "*sc* A  == {x. \<forall>X \<in> Rep_hcomplex(x). {n. X n \<in> A} \<in> FreeUltrafilterNat}"
 | |
| 13957 | 18 | |
| 19 | (* internal sets *) | |
| 14407 | 20 |     starsetC_n :: "(nat => complex set) => hcomplex set"  ("*scn* _" [80] 80)
 | 
| 21 |     "*scn* As  == {x. \<forall>X \<in> Rep_hcomplex(x). 
 | |
| 22 |                       {n. X n \<in> (As n)} \<in> FreeUltrafilterNat}"
 | |
| 23 | ||
| 13957 | 24 | InternalCSets :: "hcomplex set set" | 
| 14407 | 25 |     "InternalCSets == {X. \<exists>As. X = *scn* As}"
 | 
| 13957 | 26 | |
| 27 | (* star transform of functions f: Complex --> Complex *) | |
| 28 | ||
| 14407 | 29 | starfunC :: "(complex => complex) => hcomplex => hcomplex" | 
| 30 |                 ("*fc* _" [80] 80)
 | |
| 31 | "*fc* f == | |
| 32 |         (%x. Abs_hcomplex(\<Union>X \<in> Rep_hcomplex(x). hcomplexrel``{%n. f (X n)}))"
 | |
| 13957 | 33 | |
| 14407 | 34 | starfunC_n :: "(nat => (complex => complex)) => hcomplex => hcomplex" | 
| 35 |                   ("*fcn* _" [80] 80)
 | |
| 36 | "*fcn* F == | |
| 37 |       (%x. Abs_hcomplex(\<Union>X \<in> Rep_hcomplex(x). hcomplexrel``{%n. (F n)(X n)}))"
 | |
| 13957 | 38 | |
| 14407 | 39 | InternalCFuns :: "(hcomplex => hcomplex) set" | 
| 40 |     "InternalCFuns == {X. \<exists>F. X = *fcn* F}"
 | |
| 13957 | 41 | |
| 42 | ||
| 43 | (* star transform of functions f: Real --> Complex *) | |
| 44 | ||
| 14407 | 45 | starfunRC :: "(real => complex) => hypreal => hcomplex" | 
| 46 |                  ("*fRc* _" [80] 80)
 | |
| 47 |     "*fRc* f  == (%x. Abs_hcomplex(\<Union>X \<in> Rep_hypreal(x). hcomplexrel``{%n. f (X n)}))"
 | |
| 13957 | 48 | |
| 14407 | 49 | starfunRC_n :: "(nat => (real => complex)) => hypreal => hcomplex" | 
| 50 |                    ("*fRcn* _" [80] 80)
 | |
| 51 |     "*fRcn* F  == (%x. Abs_hcomplex(\<Union>X \<in> Rep_hypreal(x). hcomplexrel``{%n. (F n)(X n)}))"
 | |
| 13957 | 52 | |
| 14407 | 53 | InternalRCFuns :: "(hypreal => hcomplex) set" | 
| 54 |     "InternalRCFuns == {X. \<exists>F. X = *fRcn* F}"
 | |
| 13957 | 55 | |
| 56 | (* star transform of functions f: Complex --> Real; needed for Re and Im parts *) | |
| 57 | ||
| 14407 | 58 | starfunCR :: "(complex => real) => hcomplex => hypreal" | 
| 59 |                  ("*fcR* _" [80] 80)
 | |
| 60 |     "*fcR* f  == (%x. Abs_hypreal(\<Union>X \<in> Rep_hcomplex(x). hyprel``{%n. f (X n)}))"
 | |
| 61 | ||
| 62 | starfunCR_n :: "(nat => (complex => real)) => hcomplex => hypreal" | |
| 63 |                    ("*fcRn* _" [80] 80)
 | |
| 64 |     "*fcRn* F  == (%x. Abs_hypreal(\<Union>X \<in> Rep_hcomplex(x). hyprel``{%n. (F n)(X n)}))"
 | |
| 65 | ||
| 66 | InternalCRFuns :: "(hcomplex => hypreal) set" | |
| 67 |     "InternalCRFuns == {X. \<exists>F. X = *fcRn* F}"
 | |
| 68 | ||
| 69 | ||
| 70 | subsection{*Properties of the *-Transform Applied to Sets of Reals*}
 | |
| 71 | ||
| 72 | lemma STARC_complex_set [simp]: "*sc*(UNIV::complex set) = (UNIV)" | |
| 73 | by (simp add: starsetC_def) | |
| 74 | declare STARC_complex_set | |
| 75 | ||
| 76 | lemma STARC_empty_set: "*sc* {} = {}"
 | |
| 77 | by (simp add: starsetC_def) | |
| 78 | declare STARC_empty_set [simp] | |
| 79 | ||
| 80 | lemma STARC_Un: "*sc* (A Un B) = *sc* A Un *sc* B" | |
| 81 | apply (auto simp add: starsetC_def) | |
| 82 | apply (drule bspec, assumption) | |
| 83 | apply (rule_tac z = x in eq_Abs_hcomplex, simp, ultra) | |
| 84 | apply (blast intro: FreeUltrafilterNat_subset)+ | |
| 85 | done | |
| 86 | ||
| 87 | lemma starsetC_n_Un: "*scn* (%n. (A n) Un (B n)) = *scn* A Un *scn* B" | |
| 88 | apply (auto simp add: starsetC_n_def) | |
| 89 | apply (drule_tac x = Xa in bspec) | |
| 90 | apply (rule_tac [2] z = x in eq_Abs_hcomplex) | |
| 91 | apply (auto dest!: bspec, ultra+) | |
| 92 | done | |
| 93 | ||
| 94 | lemma InternalCSets_Un: | |
| 95 | "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X Un Y) \<in> InternalCSets" | |
| 96 | by (auto simp add: InternalCSets_def starsetC_n_Un [symmetric]) | |
| 97 | ||
| 98 | lemma STARC_Int: "*sc* (A Int B) = *sc* A Int *sc* B" | |
| 99 | apply (auto simp add: starsetC_def) | |
| 100 | prefer 3 apply (blast intro: FreeUltrafilterNat_Int FreeUltrafilterNat_subset) | |
| 101 | apply (blast intro: FreeUltrafilterNat_subset)+ | |
| 102 | done | |
| 103 | ||
| 104 | lemma starsetC_n_Int: "*scn* (%n. (A n) Int (B n)) = *scn* A Int *scn* B" | |
| 105 | apply (auto simp add: starsetC_n_def) | |
| 106 | apply (auto dest!: bspec, ultra+) | |
| 107 | done | |
| 108 | ||
| 109 | lemma InternalCSets_Int: | |
| 110 | "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X Int Y) \<in> InternalCSets" | |
| 111 | by (auto simp add: InternalCSets_def starsetC_n_Int [symmetric]) | |
| 112 | ||
| 113 | lemma STARC_Compl: "*sc* -A = -( *sc* A)" | |
| 114 | apply (auto simp add: starsetC_def) | |
| 115 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 116 | apply (rule_tac [2] z = x in eq_Abs_hcomplex) | |
| 117 | apply (auto dest!: bspec, ultra+) | |
| 118 | done | |
| 119 | ||
| 120 | lemma starsetC_n_Compl: "*scn* ((%n. - A n)) = -( *scn* A)" | |
| 121 | apply (auto simp add: starsetC_n_def) | |
| 122 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 123 | apply (rule_tac [2] z = x in eq_Abs_hcomplex) | |
| 124 | apply (auto dest!: bspec, ultra+) | |
| 125 | done | |
| 126 | ||
| 127 | lemma InternalCSets_Compl: "X :InternalCSets ==> -X \<in> InternalCSets" | |
| 128 | by (auto simp add: InternalCSets_def starsetC_n_Compl [symmetric]) | |
| 129 | ||
| 130 | lemma STARC_mem_Compl: "x \<notin> *sc* F ==> x \<in> *sc* (- F)" | |
| 131 | by (simp add: STARC_Compl) | |
| 132 | ||
| 133 | lemma STARC_diff: "*sc* (A - B) = *sc* A - *sc* B" | |
| 134 | by (simp add: Diff_eq STARC_Int STARC_Compl) | |
| 135 | ||
| 136 | lemma starsetC_n_diff: | |
| 137 | "*scn* (%n. (A n) - (B n)) = *scn* A - *scn* B" | |
| 138 | apply (auto simp add: starsetC_n_def) | |
| 139 | apply (rule_tac [2] z = x in eq_Abs_hcomplex) | |
| 140 | apply (rule_tac [3] z = x in eq_Abs_hcomplex) | |
| 141 | apply (auto dest!: bspec, ultra+) | |
| 142 | done | |
| 143 | ||
| 144 | lemma InternalCSets_diff: | |
| 145 | "[| X \<in> InternalCSets; Y \<in> InternalCSets |] ==> (X - Y) \<in> InternalCSets" | |
| 146 | by (auto simp add: InternalCSets_def starsetC_n_diff [symmetric]) | |
| 147 | ||
| 148 | lemma STARC_subset: "A \<le> B ==> *sc* A \<le> *sc* B" | |
| 149 | apply (simp add: starsetC_def) | |
| 150 | apply (blast intro: FreeUltrafilterNat_subset)+ | |
| 151 | done | |
| 152 | ||
| 153 | lemma STARC_mem: "a \<in> A ==> hcomplex_of_complex a \<in> *sc* A" | |
| 154 | apply (simp add: starsetC_def hcomplex_of_complex_def) | |
| 155 | apply (auto intro: FreeUltrafilterNat_subset) | |
| 156 | done | |
| 157 | ||
| 158 | lemma STARC_hcomplex_of_complex_image_subset: | |
| 159 | "hcomplex_of_complex ` A \<le> *sc* A" | |
| 160 | apply (auto simp add: starsetC_def hcomplex_of_complex_def) | |
| 161 | apply (blast intro: FreeUltrafilterNat_subset) | |
| 162 | done | |
| 163 | ||
| 164 | lemma STARC_SComplex_subset: "SComplex \<le> *sc* (UNIV:: complex set)" | |
| 165 | by auto | |
| 166 | ||
| 167 | lemma STARC_hcomplex_of_complex_Int: | |
| 168 | "*sc* X Int SComplex = hcomplex_of_complex ` X" | |
| 169 | apply (auto simp add: starsetC_def hcomplex_of_complex_def SComplex_def) | |
| 170 | apply (fold hcomplex_of_complex_def) | |
| 171 | apply (rule imageI, rule ccontr) | |
| 172 | apply (drule bspec) | |
| 173 | apply (rule lemma_hcomplexrel_refl) | |
| 174 | prefer 2 apply (blast intro: FreeUltrafilterNat_subset, auto) | |
| 175 | done | |
| 176 | ||
| 177 | lemma lemma_not_hcomplexA: | |
| 178 | "x \<notin> hcomplex_of_complex ` A ==> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y" | |
| 179 | by auto | |
| 180 | ||
| 181 | lemma starsetC_starsetC_n_eq: "*sc* X = *scn* (%n. X)" | |
| 182 | by (simp add: starsetC_n_def starsetC_def) | |
| 183 | ||
| 184 | lemma InternalCSets_starsetC_n [simp]: "( *sc* X) \<in> InternalCSets" | |
| 185 | by (auto simp add: InternalCSets_def starsetC_starsetC_n_eq) | |
| 186 | ||
| 187 | lemma InternalCSets_UNIV_diff: | |
| 188 | "X \<in> InternalCSets ==> UNIV - X \<in> InternalCSets" | |
| 189 | by (auto intro: InternalCSets_Compl) | |
| 190 | ||
| 191 | text{*Nonstandard extension of a set (defined using a constant sequence) as a special case of an internal set*}
 | |
| 192 | ||
| 193 | lemma starsetC_n_starsetC: "\<forall>n. (As n = A) ==> *scn* As = *sc* A" | |
| 194 | by (simp add:starsetC_n_def starsetC_def) | |
| 195 | ||
| 196 | ||
| 197 | subsection{*Theorems about Nonstandard Extensions of Functions*}
 | |
| 198 | ||
| 199 | lemma starfunC_n_starfunC: "\<forall>n. (F n = f) ==> *fcn* F = *fc* f" | |
| 200 | by (simp add: starfunC_n_def starfunC_def) | |
| 201 | ||
| 202 | lemma starfunRC_n_starfunRC: "\<forall>n. (F n = f) ==> *fRcn* F = *fRc* f" | |
| 203 | by (simp add: starfunRC_n_def starfunRC_def) | |
| 204 | ||
| 205 | lemma starfunCR_n_starfunCR: "\<forall>n. (F n = f) ==> *fcRn* F = *fcR* f" | |
| 206 | by (simp add: starfunCR_n_def starfunCR_def) | |
| 207 | ||
| 208 | lemma starfunC_congruent: | |
| 15169 | 209 |       "(%X. hcomplexrel``{%n. f (X n)}) respects hcomplexrel"
 | 
| 210 | by (auto simp add: hcomplexrel_iff congruent_def, ultra) | |
| 14407 | 211 | |
| 212 | (* f::complex => complex *) | |
| 213 | lemma starfunC: | |
| 214 |       "( *fc* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) =
 | |
| 215 |        Abs_hcomplex(hcomplexrel `` {%n. f (X n)})"
 | |
| 216 | apply (simp add: starfunC_def) | |
| 217 | apply (rule arg_cong [where f = Abs_hcomplex]) | |
| 218 | apply (auto iff add: hcomplexrel_iff, ultra) | |
| 219 | done | |
| 220 | ||
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 221 | lemma cstarfun_if_eq: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 222 | "w \<noteq> hcomplex_of_complex x | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 223 | ==> ( *fc* (\<lambda>z. if z = x then a else g z)) w = ( *fc* g) w" | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 224 | apply (cases w) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 225 | apply (simp add: hcomplex_of_complex_def starfunC, ultra) | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 226 | done | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 227 | |
| 14407 | 228 | lemma starfunRC: | 
| 229 |       "( *fRc* f) (Abs_hypreal(hyprel``{%n. X n})) =
 | |
| 230 |        Abs_hcomplex(hcomplexrel `` {%n. f (X n)})"
 | |
| 231 | apply (simp add: starfunRC_def) | |
| 232 | apply (rule arg_cong [where f = Abs_hcomplex], auto, ultra) | |
| 233 | done | |
| 234 | ||
| 235 | lemma starfunCR: | |
| 236 |       "( *fcR* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) =
 | |
| 237 |        Abs_hypreal(hyprel `` {%n. f (X n)})"
 | |
| 238 | apply (simp add: starfunCR_def) | |
| 239 | apply (rule arg_cong [where f = Abs_hypreal]) | |
| 240 | apply (auto iff add: hcomplexrel_iff, ultra) | |
| 241 | done | |
| 242 | ||
| 243 | (** multiplication: ( *f) x ( *g) = *(f x g) **) | |
| 244 | ||
| 245 | lemma starfunC_mult: "( *fc* f) z * ( *fc* g) z = ( *fc* (%x. f x * g x)) z" | |
| 246 | apply (rule_tac z = z in eq_Abs_hcomplex) | |
| 247 | apply (auto simp add: starfunC hcomplex_mult) | |
| 248 | done | |
| 249 | declare starfunC_mult [symmetric, simp] | |
| 250 | ||
| 251 | lemma starfunRC_mult: | |
| 252 | "( *fRc* f) z * ( *fRc* g) z = ( *fRc* (%x. f x * g x)) z" | |
| 14469 | 253 | apply (cases z) | 
| 14407 | 254 | apply (simp add: starfunRC hcomplex_mult) | 
| 255 | done | |
| 256 | declare starfunRC_mult [symmetric, simp] | |
| 257 | ||
| 258 | lemma starfunCR_mult: | |
| 259 | "( *fcR* f) z * ( *fcR* g) z = ( *fcR* (%x. f x * g x)) z" | |
| 260 | apply (rule_tac z = z in eq_Abs_hcomplex) | |
| 261 | apply (simp add: starfunCR hypreal_mult) | |
| 262 | done | |
| 263 | declare starfunCR_mult [symmetric, simp] | |
| 264 | ||
| 265 | (** addition: ( *f) + ( *g) = *(f + g) **) | |
| 266 | ||
| 267 | lemma starfunC_add: "( *fc* f) z + ( *fc* g) z = ( *fc* (%x. f x + g x)) z" | |
| 268 | apply (rule_tac z = z in eq_Abs_hcomplex) | |
| 269 | apply (simp add: starfunC hcomplex_add) | |
| 270 | done | |
| 271 | declare starfunC_add [symmetric, simp] | |
| 272 | ||
| 273 | lemma starfunRC_add: "( *fRc* f) z + ( *fRc* g) z = ( *fRc* (%x. f x + g x)) z" | |
| 14469 | 274 | apply (cases z) | 
| 14407 | 275 | apply (simp add: starfunRC hcomplex_add) | 
| 276 | done | |
| 277 | declare starfunRC_add [symmetric, simp] | |
| 278 | ||
| 279 | lemma starfunCR_add: "( *fcR* f) z + ( *fcR* g) z = ( *fcR* (%x. f x + g x)) z" | |
| 280 | apply (rule_tac z = z in eq_Abs_hcomplex) | |
| 281 | apply (simp add: starfunCR hypreal_add) | |
| 282 | done | |
| 283 | declare starfunCR_add [symmetric, simp] | |
| 284 | ||
| 285 | (** uminus **) | |
| 286 | lemma starfunC_minus [simp]: "( *fc* (%x. - f x)) x = - ( *fc* f) x" | |
| 287 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 288 | apply (simp add: starfunC hcomplex_minus) | |
| 289 | done | |
| 290 | ||
| 291 | lemma starfunRC_minus [simp]: "( *fRc* (%x. - f x)) x = - ( *fRc* f) x" | |
| 14469 | 292 | apply (cases x) | 
| 14407 | 293 | apply (simp add: starfunRC hcomplex_minus) | 
| 294 | done | |
| 295 | ||
| 296 | lemma starfunCR_minus [simp]: "( *fcR* (%x. - f x)) x = - ( *fcR* f) x" | |
| 297 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 298 | apply (simp add: starfunCR hypreal_minus) | |
| 299 | done | |
| 300 | ||
| 301 | (** addition: ( *f) - ( *g) = *(f - g) **) | |
| 302 | ||
| 303 | lemma starfunC_diff: "( *fc* f) y - ( *fc* g) y = ( *fc* (%x. f x - g x)) y" | |
| 304 | by (simp add: diff_minus) | |
| 305 | declare starfunC_diff [symmetric, simp] | |
| 306 | ||
| 307 | lemma starfunRC_diff: | |
| 308 | "( *fRc* f) y - ( *fRc* g) y = ( *fRc* (%x. f x - g x)) y" | |
| 309 | by (simp add: diff_minus) | |
| 310 | declare starfunRC_diff [symmetric, simp] | |
| 311 | ||
| 312 | lemma starfunCR_diff: | |
| 313 | "( *fcR* f) y - ( *fcR* g) y = ( *fcR* (%x. f x - g x)) y" | |
| 314 | by (simp add: diff_minus) | |
| 315 | declare starfunCR_diff [symmetric, simp] | |
| 316 | ||
| 317 | (** composition: ( *f) o ( *g) = *(f o g) **) | |
| 318 | ||
| 319 | lemma starfunC_o2: "(%x. ( *fc* f) (( *fc* g) x)) = *fc* (%x. f (g x))" | |
| 320 | apply (rule ext) | |
| 321 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 322 | apply (simp add: starfunC) | |
| 323 | done | |
| 324 | ||
| 325 | lemma starfunC_o: "( *fc* f) o ( *fc* g) = ( *fc* (f o g))" | |
| 326 | by (simp add: o_def starfunC_o2) | |
| 327 | ||
| 328 | lemma starfunC_starfunRC_o2: | |
| 329 | "(%x. ( *fc* f) (( *fRc* g) x)) = *fRc* (%x. f (g x))" | |
| 330 | apply (rule ext) | |
| 331 | apply (rule_tac z = x in eq_Abs_hypreal) | |
| 332 | apply (simp add: starfunRC starfunC) | |
| 333 | done | |
| 334 | ||
| 335 | lemma starfun_starfunCR_o2: | |
| 336 | "(%x. ( *f* f) (( *fcR* g) x)) = *fcR* (%x. f (g x))" | |
| 337 | apply (rule ext) | |
| 338 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 339 | apply (simp add: starfunCR starfun) | |
| 340 | done | |
| 341 | ||
| 342 | lemma starfunC_starfunRC_o: "( *fc* f) o ( *fRc* g) = ( *fRc* (f o g))" | |
| 343 | by (simp add: o_def starfunC_starfunRC_o2) | |
| 344 | ||
| 345 | lemma starfun_starfunCR_o: "( *f* f) o ( *fcR* g) = ( *fcR* (f o g))" | |
| 346 | by (simp add: o_def starfun_starfunCR_o2) | |
| 347 | ||
| 348 | lemma starfunC_const_fun [simp]: "( *fc* (%x. k)) z = hcomplex_of_complex k" | |
| 14469 | 349 | apply (cases z) | 
| 14407 | 350 | apply (simp add: starfunC hcomplex_of_complex_def) | 
| 351 | done | |
| 352 | ||
| 353 | lemma starfunRC_const_fun [simp]: "( *fRc* (%x. k)) z = hcomplex_of_complex k" | |
| 14469 | 354 | apply (cases z) | 
| 14407 | 355 | apply (simp add: starfunRC hcomplex_of_complex_def) | 
| 356 | done | |
| 357 | ||
| 358 | lemma starfunCR_const_fun [simp]: "( *fcR* (%x. k)) z = hypreal_of_real k" | |
| 14469 | 359 | apply (cases z) | 
| 14407 | 360 | apply (simp add: starfunCR hypreal_of_real_def) | 
| 361 | done | |
| 362 | ||
| 363 | lemma starfunC_inverse: "inverse (( *fc* f) x) = ( *fc* (%x. inverse (f x))) x" | |
| 14469 | 364 | apply (cases x) | 
| 14407 | 365 | apply (simp add: starfunC hcomplex_inverse) | 
| 366 | done | |
| 367 | declare starfunC_inverse [symmetric, simp] | |
| 368 | ||
| 369 | lemma starfunRC_inverse: | |
| 370 | "inverse (( *fRc* f) x) = ( *fRc* (%x. inverse (f x))) x" | |
| 14469 | 371 | apply (cases x) | 
| 14407 | 372 | apply (simp add: starfunRC hcomplex_inverse) | 
| 373 | done | |
| 374 | declare starfunRC_inverse [symmetric, simp] | |
| 375 | ||
| 376 | lemma starfunCR_inverse: | |
| 377 | "inverse (( *fcR* f) x) = ( *fcR* (%x. inverse (f x))) x" | |
| 14469 | 378 | apply (cases x) | 
| 14407 | 379 | apply (simp add: starfunCR hypreal_inverse) | 
| 380 | done | |
| 381 | declare starfunCR_inverse [symmetric, simp] | |
| 382 | ||
| 383 | lemma starfunC_eq [simp]: | |
| 384 | "( *fc* f) (hcomplex_of_complex a) = hcomplex_of_complex (f a)" | |
| 385 | by (simp add: starfunC hcomplex_of_complex_def) | |
| 386 | ||
| 387 | lemma starfunRC_eq [simp]: | |
| 388 | "( *fRc* f) (hypreal_of_real a) = hcomplex_of_complex (f a)" | |
| 389 | by (simp add: starfunRC hcomplex_of_complex_def hypreal_of_real_def) | |
| 13957 | 390 | |
| 14407 | 391 | lemma starfunCR_eq [simp]: | 
| 392 | "( *fcR* f) (hcomplex_of_complex a) = hypreal_of_real (f a)" | |
| 393 | by (simp add: starfunCR hcomplex_of_complex_def hypreal_of_real_def) | |
| 394 | ||
| 395 | lemma starfunC_capprox: | |
| 396 | "( *fc* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)" | |
| 397 | by auto | |
| 398 | ||
| 399 | lemma starfunRC_capprox: | |
| 400 | "( *fRc* f) (hypreal_of_real a) @c= hcomplex_of_complex (f a)" | |
| 401 | by auto | |
| 402 | ||
| 403 | lemma starfunCR_approx: | |
| 404 | "( *fcR* f) (hcomplex_of_complex a) @= hypreal_of_real (f a)" | |
| 405 | by auto | |
| 406 | ||
| 407 | (* | |
| 408 | Goal "( *fcNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N" | |
| 409 | *) | |
| 410 | ||
| 411 | lemma starfunC_hcpow: "( *fc* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n" | |
| 14469 | 412 | apply (cases Z) | 
| 14407 | 413 | apply (simp add: hcpow starfunC hypnat_of_nat_eq) | 
| 414 | done | |
| 415 | ||
| 416 | lemma starfunC_lambda_cancel: | |
| 417 | "( *fc* (%h. f (x + h))) y = ( *fc* f) (hcomplex_of_complex x + y)" | |
| 14469 | 418 | apply (cases y) | 
| 14407 | 419 | apply (simp add: starfunC hcomplex_of_complex_def hcomplex_add) | 
| 420 | done | |
| 421 | ||
| 422 | lemma starfunCR_lambda_cancel: | |
| 423 | "( *fcR* (%h. f (x + h))) y = ( *fcR* f) (hcomplex_of_complex x + y)" | |
| 14469 | 424 | apply (cases y) | 
| 14407 | 425 | apply (simp add: starfunCR hcomplex_of_complex_def hcomplex_add) | 
| 426 | done | |
| 427 | ||
| 428 | lemma starfunRC_lambda_cancel: | |
| 429 | "( *fRc* (%h. f (x + h))) y = ( *fRc* f) (hypreal_of_real x + y)" | |
| 14469 | 430 | apply (cases y) | 
| 14407 | 431 | apply (simp add: starfunRC hypreal_of_real_def hypreal_add) | 
| 432 | done | |
| 433 | ||
| 434 | lemma starfunC_lambda_cancel2: | |
| 435 | "( *fc* (%h. f(g(x + h)))) y = ( *fc* (f o g)) (hcomplex_of_complex x + y)" | |
| 14469 | 436 | apply (cases y) | 
| 14407 | 437 | apply (simp add: starfunC hcomplex_of_complex_def hcomplex_add) | 
| 438 | done | |
| 439 | ||
| 440 | lemma starfunCR_lambda_cancel2: | |
| 441 | "( *fcR* (%h. f(g(x + h)))) y = ( *fcR* (f o g)) (hcomplex_of_complex x + y)" | |
| 14469 | 442 | apply (cases y) | 
| 14407 | 443 | apply (simp add: starfunCR hcomplex_of_complex_def hcomplex_add) | 
| 444 | done | |
| 445 | ||
| 446 | lemma starfunRC_lambda_cancel2: | |
| 447 | "( *fRc* (%h. f(g(x + h)))) y = ( *fRc* (f o g)) (hypreal_of_real x + y)" | |
| 14469 | 448 | apply (cases y) | 
| 14407 | 449 | apply (simp add: starfunRC hypreal_of_real_def hypreal_add) | 
| 450 | done | |
| 451 | ||
| 452 | lemma starfunC_mult_CFinite_capprox: | |
| 453 | "[| ( *fc* f) y @c= l; ( *fc* g) y @c= m; l: CFinite; m: CFinite |] | |
| 454 | ==> ( *fc* (%x. f x * g x)) y @c= l * m" | |
| 455 | apply (drule capprox_mult_CFinite, assumption+) | |
| 456 | apply (auto intro: capprox_sym [THEN [2] capprox_CFinite]) | |
| 457 | done | |
| 458 | ||
| 459 | lemma starfunCR_mult_HFinite_capprox: | |
| 460 | "[| ( *fcR* f) y @= l; ( *fcR* g) y @= m; l: HFinite; m: HFinite |] | |
| 461 | ==> ( *fcR* (%x. f x * g x)) y @= l * m" | |
| 462 | apply (drule approx_mult_HFinite, assumption+) | |
| 463 | apply (auto intro: approx_sym [THEN [2] approx_HFinite]) | |
| 464 | done | |
| 465 | ||
| 466 | lemma starfunRC_mult_CFinite_capprox: | |
| 467 | "[| ( *fRc* f) y @c= l; ( *fRc* g) y @c= m; l: CFinite; m: CFinite |] | |
| 468 | ==> ( *fRc* (%x. f x * g x)) y @c= l * m" | |
| 469 | apply (drule capprox_mult_CFinite, assumption+) | |
| 470 | apply (auto intro: capprox_sym [THEN [2] capprox_CFinite]) | |
| 471 | done | |
| 472 | ||
| 473 | lemma starfunC_add_capprox: | |
| 474 | "[| ( *fc* f) y @c= l; ( *fc* g) y @c= m |] | |
| 475 | ==> ( *fc* (%x. f x + g x)) y @c= l + m" | |
| 476 | by (auto intro: capprox_add) | |
| 477 | ||
| 478 | lemma starfunRC_add_capprox: | |
| 479 | "[| ( *fRc* f) y @c= l; ( *fRc* g) y @c= m |] | |
| 480 | ==> ( *fRc* (%x. f x + g x)) y @c= l + m" | |
| 481 | by (auto intro: capprox_add) | |
| 482 | ||
| 483 | lemma starfunCR_add_approx: | |
| 484 | "[| ( *fcR* f) y @= l; ( *fcR* g) y @= m | |
| 485 | |] ==> ( *fcR* (%x. f x + g x)) y @= l + m" | |
| 486 | by (auto intro: approx_add) | |
| 487 | ||
| 488 | lemma starfunCR_cmod: "*fcR* cmod = hcmod" | |
| 489 | apply (rule ext) | |
| 490 | apply (rule_tac z = x in eq_Abs_hcomplex) | |
| 491 | apply (simp add: starfunCR hcmod) | |
| 492 | done | |
| 493 | ||
| 494 | lemma starfunC_inverse_inverse: "( *fc* inverse) x = inverse(x)" | |
| 14469 | 495 | apply (cases x) | 
| 14407 | 496 | apply (simp add: starfunC hcomplex_inverse) | 
| 497 | done | |
| 498 | ||
| 499 | lemma starfunC_divide: "( *fc* f) y / ( *fc* g) y = ( *fc* (%x. f x / g x)) y" | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14407diff
changeset | 500 | by (simp add: divide_inverse) | 
| 14407 | 501 | declare starfunC_divide [symmetric, simp] | 
| 502 | ||
| 503 | lemma starfunCR_divide: | |
| 504 | "( *fcR* f) y / ( *fcR* g) y = ( *fcR* (%x. f x / g x)) y" | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14407diff
changeset | 505 | by (simp add: divide_inverse) | 
| 14407 | 506 | declare starfunCR_divide [symmetric, simp] | 
| 507 | ||
| 508 | lemma starfunRC_divide: | |
| 509 | "( *fRc* f) y / ( *fRc* g) y = ( *fRc* (%x. f x / g x)) y" | |
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14407diff
changeset | 510 | by (simp add: divide_inverse) | 
| 14407 | 511 | declare starfunRC_divide [symmetric, simp] | 
| 512 | ||
| 513 | ||
| 514 | subsection{*Internal Functions - Some Redundancy With *Fc* Now*}
 | |
| 515 | ||
| 516 | lemma starfunC_n_congruent: | |
| 15169 | 517 |       "(%X. hcomplexrel``{%n. f n (X n)}) respects hcomplexrel"
 | 
| 14407 | 518 | by (auto simp add: congruent_def hcomplexrel_iff, ultra) | 
| 519 | ||
| 520 | lemma starfunC_n: | |
| 521 |      "( *fcn* f) (Abs_hcomplex(hcomplexrel``{%n. X n})) =
 | |
| 522 |       Abs_hcomplex(hcomplexrel `` {%n. f n (X n)})"
 | |
| 523 | apply (simp add: starfunC_n_def) | |
| 524 | apply (rule arg_cong [where f = Abs_hcomplex]) | |
| 525 | apply (auto iff add: hcomplexrel_iff, ultra) | |
| 526 | done | |
| 527 | ||
| 528 | (** multiplication: ( *fn) x ( *gn) = *(fn x gn) **) | |
| 529 | ||
| 530 | lemma starfunC_n_mult: | |
| 531 | "( *fcn* f) z * ( *fcn* g) z = ( *fcn* (% i x. f i x * g i x)) z" | |
| 14469 | 532 | apply (cases z) | 
| 14407 | 533 | apply (simp add: starfunC_n hcomplex_mult) | 
| 534 | done | |
| 535 | ||
| 536 | (** addition: ( *fn) + ( *gn) = *(fn + gn) **) | |
| 537 | ||
| 538 | lemma starfunC_n_add: | |
| 539 | "( *fcn* f) z + ( *fcn* g) z = ( *fcn* (%i x. f i x + g i x)) z" | |
| 14469 | 540 | apply (cases z) | 
| 14407 | 541 | apply (simp add: starfunC_n hcomplex_add) | 
| 542 | done | |
| 543 | ||
| 544 | (** uminus **) | |
| 545 | ||
| 546 | lemma starfunC_n_minus: "- ( *fcn* g) z = ( *fcn* (%i x. - g i x)) z" | |
| 14469 | 547 | apply (cases z) | 
| 14407 | 548 | apply (simp add: starfunC_n hcomplex_minus) | 
| 549 | done | |
| 550 | ||
| 551 | (** subtraction: ( *fn) - ( *gn) = *(fn - gn) **) | |
| 13957 | 552 | |
| 14407 | 553 | lemma starfunNat_n_diff: | 
| 554 | "( *fcn* f) z - ( *fcn* g) z = ( *fcn* (%i x. f i x - g i x)) z" | |
| 555 | by (simp add: diff_minus starfunC_n_add starfunC_n_minus) | |
| 556 | ||
| 557 | (** composition: ( *fn) o ( *gn) = *(fn o gn) **) | |
| 558 | ||
| 559 | lemma starfunC_n_const_fun [simp]: | |
| 560 | "( *fcn* (%i x. k)) z = hcomplex_of_complex k" | |
| 14469 | 561 | apply (cases z) | 
| 14407 | 562 | apply (simp add: starfunC_n hcomplex_of_complex_def) | 
| 563 | done | |
| 564 | ||
| 565 | lemma starfunC_n_eq [simp]: | |
| 566 |     "( *fcn* f) (hcomplex_of_complex n) = Abs_hcomplex(hcomplexrel `` {%i. f i n})"
 | |
| 567 | by (simp add: starfunC_n hcomplex_of_complex_def) | |
| 568 | ||
| 569 | lemma starfunC_eq_iff: "(( *fc* f) = ( *fc* g)) = (f = g)" | |
| 570 | apply auto | |
| 571 | apply (rule ext, rule ccontr) | |
| 572 | apply (drule_tac x = "hcomplex_of_complex (x) " in fun_cong) | |
| 573 | apply (simp add: starfunC hcomplex_of_complex_def) | |
| 574 | done | |
| 575 | ||
| 576 | lemma starfunRC_eq_iff: "(( *fRc* f) = ( *fRc* g)) = (f = g)" | |
| 577 | apply auto | |
| 578 | apply (rule ext, rule ccontr) | |
| 579 | apply (drule_tac x = "hypreal_of_real (x) " in fun_cong) | |
| 580 | apply auto | |
| 581 | done | |
| 582 | ||
| 583 | lemma starfunCR_eq_iff: "(( *fcR* f) = ( *fcR* g)) = (f = g)" | |
| 584 | apply auto | |
| 585 | apply (rule ext, rule ccontr) | |
| 586 | apply (drule_tac x = "hcomplex_of_complex (x) " in fun_cong) | |
| 587 | apply auto | |
| 588 | done | |
| 589 | ||
| 590 | lemma starfunC_eq_Re_Im_iff: | |
| 591 | "(( *fc* f) x = z) = ((( *fcR* (%x. Re(f x))) x = hRe (z)) & | |
| 592 | (( *fcR* (%x. Im(f x))) x = hIm (z)))" | |
| 14469 | 593 | apply (cases x, cases z) | 
| 14407 | 594 | apply (auto simp add: starfunCR starfunC hIm hRe complex_Re_Im_cancel_iff, ultra+) | 
| 595 | done | |
| 596 | ||
| 597 | lemma starfunC_approx_Re_Im_iff: | |
| 598 | "(( *fc* f) x @c= z) = ((( *fcR* (%x. Re(f x))) x @= hRe (z)) & | |
| 599 | (( *fcR* (%x. Im(f x))) x @= hIm (z)))" | |
| 14469 | 600 | apply (cases x, cases z) | 
| 14407 | 601 | apply (simp add: starfunCR starfunC hIm hRe capprox_approx_iff) | 
| 602 | done | |
| 603 | ||
| 604 | lemma starfunC_Idfun_capprox: | |
| 605 | "x @c= hcomplex_of_complex a ==> ( *fc* (%x. x)) x @c= hcomplex_of_complex a" | |
| 14469 | 606 | apply (cases x) | 
| 14407 | 607 | apply (simp add: starfunC) | 
| 608 | done | |
| 609 | ||
| 610 | lemma starfunC_Id [simp]: "( *fc* (%x. x)) x = x" | |
| 14469 | 611 | apply (cases x) | 
| 14407 | 612 | apply (simp add: starfunC) | 
| 613 | done | |
| 13957 | 614 | |
| 14407 | 615 | ML | 
| 616 | {*
 | |
| 617 | val STARC_complex_set = thm "STARC_complex_set"; | |
| 618 | val STARC_empty_set = thm "STARC_empty_set"; | |
| 619 | val STARC_Un = thm "STARC_Un"; | |
| 620 | val starsetC_n_Un = thm "starsetC_n_Un"; | |
| 621 | val InternalCSets_Un = thm "InternalCSets_Un"; | |
| 622 | val STARC_Int = thm "STARC_Int"; | |
| 623 | val starsetC_n_Int = thm "starsetC_n_Int"; | |
| 624 | val InternalCSets_Int = thm "InternalCSets_Int"; | |
| 625 | val STARC_Compl = thm "STARC_Compl"; | |
| 626 | val starsetC_n_Compl = thm "starsetC_n_Compl"; | |
| 627 | val InternalCSets_Compl = thm "InternalCSets_Compl"; | |
| 628 | val STARC_mem_Compl = thm "STARC_mem_Compl"; | |
| 629 | val STARC_diff = thm "STARC_diff"; | |
| 630 | val starsetC_n_diff = thm "starsetC_n_diff"; | |
| 631 | val InternalCSets_diff = thm "InternalCSets_diff"; | |
| 632 | val STARC_subset = thm "STARC_subset"; | |
| 633 | val STARC_mem = thm "STARC_mem"; | |
| 634 | val STARC_hcomplex_of_complex_image_subset = thm "STARC_hcomplex_of_complex_image_subset"; | |
| 635 | val STARC_SComplex_subset = thm "STARC_SComplex_subset"; | |
| 636 | val STARC_hcomplex_of_complex_Int = thm "STARC_hcomplex_of_complex_Int"; | |
| 637 | val lemma_not_hcomplexA = thm "lemma_not_hcomplexA"; | |
| 638 | val starsetC_starsetC_n_eq = thm "starsetC_starsetC_n_eq"; | |
| 639 | val InternalCSets_starsetC_n = thm "InternalCSets_starsetC_n"; | |
| 640 | val InternalCSets_UNIV_diff = thm "InternalCSets_UNIV_diff"; | |
| 641 | val starsetC_n_starsetC = thm "starsetC_n_starsetC"; | |
| 642 | val starfunC_n_starfunC = thm "starfunC_n_starfunC"; | |
| 643 | val starfunRC_n_starfunRC = thm "starfunRC_n_starfunRC"; | |
| 644 | val starfunCR_n_starfunCR = thm "starfunCR_n_starfunCR"; | |
| 645 | val starfunC_congruent = thm "starfunC_congruent"; | |
| 646 | val starfunC = thm "starfunC"; | |
| 647 | val starfunRC = thm "starfunRC"; | |
| 648 | val starfunCR = thm "starfunCR"; | |
| 649 | val starfunC_mult = thm "starfunC_mult"; | |
| 650 | val starfunRC_mult = thm "starfunRC_mult"; | |
| 651 | val starfunCR_mult = thm "starfunCR_mult"; | |
| 652 | val starfunC_add = thm "starfunC_add"; | |
| 653 | val starfunRC_add = thm "starfunRC_add"; | |
| 654 | val starfunCR_add = thm "starfunCR_add"; | |
| 655 | val starfunC_minus = thm "starfunC_minus"; | |
| 656 | val starfunRC_minus = thm "starfunRC_minus"; | |
| 657 | val starfunCR_minus = thm "starfunCR_minus"; | |
| 658 | val starfunC_diff = thm "starfunC_diff"; | |
| 659 | val starfunRC_diff = thm "starfunRC_diff"; | |
| 660 | val starfunCR_diff = thm "starfunCR_diff"; | |
| 661 | val starfunC_o2 = thm "starfunC_o2"; | |
| 662 | val starfunC_o = thm "starfunC_o"; | |
| 663 | val starfunC_starfunRC_o2 = thm "starfunC_starfunRC_o2"; | |
| 664 | val starfun_starfunCR_o2 = thm "starfun_starfunCR_o2"; | |
| 665 | val starfunC_starfunRC_o = thm "starfunC_starfunRC_o"; | |
| 666 | val starfun_starfunCR_o = thm "starfun_starfunCR_o"; | |
| 667 | val starfunC_const_fun = thm "starfunC_const_fun"; | |
| 668 | val starfunRC_const_fun = thm "starfunRC_const_fun"; | |
| 669 | val starfunCR_const_fun = thm "starfunCR_const_fun"; | |
| 670 | val starfunC_inverse = thm "starfunC_inverse"; | |
| 671 | val starfunRC_inverse = thm "starfunRC_inverse"; | |
| 672 | val starfunCR_inverse = thm "starfunCR_inverse"; | |
| 673 | val starfunC_eq = thm "starfunC_eq"; | |
| 674 | val starfunRC_eq = thm "starfunRC_eq"; | |
| 675 | val starfunCR_eq = thm "starfunCR_eq"; | |
| 676 | val starfunC_capprox = thm "starfunC_capprox"; | |
| 677 | val starfunRC_capprox = thm "starfunRC_capprox"; | |
| 678 | val starfunCR_approx = thm "starfunCR_approx"; | |
| 679 | val starfunC_hcpow = thm "starfunC_hcpow"; | |
| 680 | val starfunC_lambda_cancel = thm "starfunC_lambda_cancel"; | |
| 681 | val starfunCR_lambda_cancel = thm "starfunCR_lambda_cancel"; | |
| 682 | val starfunRC_lambda_cancel = thm "starfunRC_lambda_cancel"; | |
| 683 | val starfunC_lambda_cancel2 = thm "starfunC_lambda_cancel2"; | |
| 684 | val starfunCR_lambda_cancel2 = thm "starfunCR_lambda_cancel2"; | |
| 685 | val starfunRC_lambda_cancel2 = thm "starfunRC_lambda_cancel2"; | |
| 686 | val starfunC_mult_CFinite_capprox = thm "starfunC_mult_CFinite_capprox"; | |
| 687 | val starfunCR_mult_HFinite_capprox = thm "starfunCR_mult_HFinite_capprox"; | |
| 688 | val starfunRC_mult_CFinite_capprox = thm "starfunRC_mult_CFinite_capprox"; | |
| 689 | val starfunC_add_capprox = thm "starfunC_add_capprox"; | |
| 690 | val starfunRC_add_capprox = thm "starfunRC_add_capprox"; | |
| 691 | val starfunCR_add_approx = thm "starfunCR_add_approx"; | |
| 692 | val starfunCR_cmod = thm "starfunCR_cmod"; | |
| 693 | val starfunC_inverse_inverse = thm "starfunC_inverse_inverse"; | |
| 694 | val starfunC_divide = thm "starfunC_divide"; | |
| 695 | val starfunCR_divide = thm "starfunCR_divide"; | |
| 696 | val starfunRC_divide = thm "starfunRC_divide"; | |
| 697 | val starfunC_n_congruent = thm "starfunC_n_congruent"; | |
| 698 | val starfunC_n = thm "starfunC_n"; | |
| 699 | val starfunC_n_mult = thm "starfunC_n_mult"; | |
| 700 | val starfunC_n_add = thm "starfunC_n_add"; | |
| 701 | val starfunC_n_minus = thm "starfunC_n_minus"; | |
| 702 | val starfunNat_n_diff = thm "starfunNat_n_diff"; | |
| 703 | val starfunC_n_const_fun = thm "starfunC_n_const_fun"; | |
| 704 | val starfunC_n_eq = thm "starfunC_n_eq"; | |
| 705 | val starfunC_eq_iff = thm "starfunC_eq_iff"; | |
| 706 | val starfunRC_eq_iff = thm "starfunRC_eq_iff"; | |
| 707 | val starfunCR_eq_iff = thm "starfunCR_eq_iff"; | |
| 708 | val starfunC_eq_Re_Im_iff = thm "starfunC_eq_Re_Im_iff"; | |
| 709 | val starfunC_approx_Re_Im_iff = thm "starfunC_approx_Re_Im_iff"; | |
| 710 | val starfunC_Idfun_capprox = thm "starfunC_Idfun_capprox"; | |
| 711 | val starfunC_Id = thm "starfunC_Id"; | |
| 712 | *} | |
| 713 | ||
| 714 | end |