| author | wenzelm | 
| Fri, 10 Mar 2006 19:49:58 +0100 | |
| changeset 19240 | 3a73cb17a707 | 
| parent 17318 | bc1c75855f3d | 
| child 19765 | dfe940911617 | 
| 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 | |
| 14407 | 13 | subsection{*Properties of the *-Transform Applied to Sets of Reals*}
 | 
| 14 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 15 | lemma STARC_SComplex_subset: "SComplex \<subseteq> *s* (UNIV:: complex set)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 16 | by simp | 
| 14407 | 17 | |
| 18 | lemma STARC_hcomplex_of_complex_Int: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 19 | "*s* X Int SComplex = hcomplex_of_complex ` X" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 20 | by (auto simp add: SComplex_def STAR_mem_iff) | 
| 14407 | 21 | |
| 22 | lemma lemma_not_hcomplexA: | |
| 23 | "x \<notin> hcomplex_of_complex ` A ==> \<forall>y \<in> A. x \<noteq> hcomplex_of_complex y" | |
| 24 | by auto | |
| 25 | ||
| 26 | subsection{*Theorems about Nonstandard Extensions of Functions*}
 | |
| 27 | ||
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 28 | lemma cstarfun_if_eq: | 
| 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15169diff
changeset | 29 | "w \<noteq> hcomplex_of_complex x | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 30 | ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 31 | apply (cases w) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 32 | apply (simp add: star_of_def starfun star_n_eq_iff, ultra) | 
| 14407 | 33 | done | 
| 34 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 35 | lemma starfun_capprox: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 36 | "( *f* f) (hcomplex_of_complex a) @c= hcomplex_of_complex (f a)" | 
| 14407 | 37 | by auto | 
| 38 | ||
| 39 | (* | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 40 | Goal "( *fNat* (%n. z ^ n)) N = (hcomplex_of_complex z) hcpow N" | 
| 14407 | 41 | *) | 
| 42 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 43 | lemma starfunC_hcpow: "( *f* (%z. z ^ n)) Z = Z hcpow hypnat_of_nat n" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 44 | apply (cases Z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 45 | apply (simp add: hcpow starfun hypnat_of_nat_eq) | 
| 14407 | 46 | done | 
| 47 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 48 | lemma starfun_mult_CFinite_capprox: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 49 | "[| ( *f* f) y @c= l; ( *f* g) y @c= m; l: CFinite; m: CFinite |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 50 | ==> ( *f* (%x. f x * g x)) y @c= l * m" | 
| 14407 | 51 | apply (drule capprox_mult_CFinite, assumption+) | 
| 52 | apply (auto intro: capprox_sym [THEN [2] capprox_CFinite]) | |
| 53 | done | |
| 54 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 55 | lemma starfun_add_capprox: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 56 | "[| ( *f* f) y @c= l; ( *f* g) y @c= m |] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 57 | ==> ( *f* (%x. f x + g x)) y @c= l + m" | 
| 14407 | 58 | by (auto intro: capprox_add) | 
| 59 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 60 | lemma starfunCR_cmod: "*f* cmod = hcmod" | 
| 14407 | 61 | apply (rule ext) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 62 | apply (rule_tac x = x in star_cases) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 63 | apply (simp add: starfun hcmod) | 
| 14407 | 64 | done | 
| 65 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 66 | subsection{*Internal Functions - Some Redundancy With *f* Now*}
 | 
| 14407 | 67 | |
| 68 | (** subtraction: ( *fn) - ( *gn) = *(fn - gn) **) | |
| 13957 | 69 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 70 | lemma starfun_n_diff: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 71 | "( *fn* f) z - ( *fn* g) z = ( *fn* (%i x. f i x - g i x)) z" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 72 | apply (cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 73 | apply (simp add: starfun_n star_n_diff) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 74 | done | 
| 14407 | 75 | |
| 76 | (** composition: ( *fn) o ( *gn) = *(fn o gn) **) | |
| 77 | ||
| 78 | lemma starfunC_eq_Re_Im_iff: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 79 | "(( *f* f) x = z) = ((( *f* (%x. Re(f x))) x = hRe (z)) & | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 80 | (( *f* (%x. Im(f x))) x = hIm (z)))" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 81 | apply (cases x, cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 82 | apply (auto simp add: starfun hIm hRe complex_Re_Im_cancel_iff star_n_eq_iff, ultra+) | 
| 14407 | 83 | done | 
| 84 | ||
| 85 | lemma starfunC_approx_Re_Im_iff: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 86 | "(( *f* f) x @c= z) = ((( *f* (%x. Re(f x))) x @= hRe (z)) & | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 87 | (( *f* (%x. Im(f x))) x @= hIm (z)))" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 88 | apply (cases x, cases z) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 89 | apply (simp add: starfun hIm hRe capprox_approx_iff) | 
| 14407 | 90 | done | 
| 91 | ||
| 92 | lemma starfunC_Idfun_capprox: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 93 | "x @c= hcomplex_of_complex a ==> ( *f* (%x. x)) x @c= hcomplex_of_complex a" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 94 | by simp | 
| 13957 | 95 | |
| 14407 | 96 | ML | 
| 97 | {*
 | |
| 98 | val STARC_SComplex_subset = thm "STARC_SComplex_subset"; | |
| 99 | val STARC_hcomplex_of_complex_Int = thm "STARC_hcomplex_of_complex_Int"; | |
| 100 | val lemma_not_hcomplexA = thm "lemma_not_hcomplexA"; | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 101 | val starfun_capprox = thm "starfun_capprox"; | 
| 14407 | 102 | val starfunC_hcpow = thm "starfunC_hcpow"; | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 103 | val starfun_mult_CFinite_capprox = thm "starfun_mult_CFinite_capprox"; | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 104 | val starfun_add_capprox = thm "starfun_add_capprox"; | 
| 14407 | 105 | val starfunCR_cmod = thm "starfunCR_cmod"; | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 106 | val starfun_inverse_inverse = thm "starfun_inverse_inverse"; | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17300diff
changeset | 107 | val starfun_n_diff = thm "starfun_n_diff"; | 
| 14407 | 108 | val starfunC_eq_Re_Im_iff = thm "starfunC_eq_Re_Im_iff"; | 
| 109 | val starfunC_approx_Re_Im_iff = thm "starfunC_approx_Re_Im_iff"; | |
| 110 | val starfunC_Idfun_capprox = thm "starfunC_Idfun_capprox"; | |
| 111 | *} | |
| 112 | ||
| 113 | end |