| author | wenzelm | 
| Sat, 29 Mar 2008 19:13:58 +0100 | |
| changeset 26479 | 3a2efce3e992 | 
| parent 22966 | 9dc4f5048353 | 
| child 27435 | b3f8e9bdf9a7 | 
| permissions | -rw-r--r-- | 
| 10751 | 1 | (* Title : Star.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 14468 | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4 | 
| 14370 | 5 | *) | 
| 10751 | 6 | |
| 14370 | 7 | header{*Star-Transforms in Non-Standard Analysis*}
 | 
| 8 | ||
| 15131 | 9 | theory Star | 
| 15140 | 10 | imports NSA | 
| 15131 | 11 | begin | 
| 10751 | 12 | |
| 19765 | 13 | definition | 
| 10751 | 14 | (* internal sets *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 15 |   starset_n :: "(nat => 'a set) => 'a star set" ("*sn* _" [80] 80) where
 | 
| 19765 | 16 | "*sn* As = Iset (star_n As)" | 
| 14370 | 17 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 18 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 19 | InternalSets :: "'a star set set" where | 
| 19765 | 20 |   "InternalSets = {X. \<exists>As. X = *sn* As}"
 | 
| 10751 | 21 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 22 | definition | 
| 19765 | 23 | (* nonstandard extension of function *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 24 | is_starext :: "['a star => 'a star, 'a => 'a] => bool" where | 
| 19765 | 25 | "is_starext F f = (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y). | 
| 10751 | 26 |                         ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
 | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 27 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 28 | definition | 
| 19765 | 29 | (* internal functions *) | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 30 |   starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"   ("*fn* _" [80] 80) where
 | 
| 19765 | 31 | "*fn* F = Ifun (star_n F)" | 
| 10751 | 32 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 33 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20732diff
changeset | 34 |   InternalFuns :: "('a star => 'b star) set" where
 | 
| 19765 | 35 |   "InternalFuns = {X. \<exists>F. X = *fn* F}"
 | 
| 10751 | 36 | |
| 37 | ||
| 14370 | 38 | (*-------------------------------------------------------- | 
| 39 | Preamble - Pulling "EX" over "ALL" | |
| 40 | ---------------------------------------------------------*) | |
| 41 | ||
| 42 | (* This proof does not need AC and was suggested by the | |
| 43 | referee for the JCM Paper: let f(x) be least y such | |
| 44 | that Q(x,y) | |
| 45 | *) | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 46 | lemma no_choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>(f :: 'a => nat). \<forall>x. Q x (f x)" | 
| 14370 | 47 | apply (rule_tac x = "%x. LEAST y. Q x y" in exI) | 
| 48 | apply (blast intro: LeastI) | |
| 49 | done | |
| 50 | ||
| 15169 | 51 | subsection{*Properties of the Star-transform Applied to Sets of Reals*}
 | 
| 14370 | 52 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 53 | lemma STAR_star_of_image_subset: "star_of ` A <= *s* A" | 
| 20732 | 54 | by auto | 
| 14370 | 55 | |
| 56 | lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X" | |
| 20732 | 57 | by (auto simp add: SReal_def) | 
| 14370 | 58 | |
| 22966 | 59 | lemma STAR_star_of_Int: "*s* X Int Standard = star_of ` X" | 
| 60 | by (auto simp add: Standard_def) | |
| 61 | ||
| 14468 | 62 | lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A ==> \<forall>y \<in> A. x \<noteq> hypreal_of_real y" | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 63 | by auto | 
| 14370 | 64 | |
| 22966 | 65 | lemma lemma_not_starA: "x \<notin> star_of ` A ==> \<forall>y \<in> A. x \<noteq> star_of y" | 
| 66 | by auto | |
| 67 | ||
| 14370 | 68 | lemma lemma_Compl_eq: "- {n. X n = xa} = {n. X n \<noteq> xa}"
 | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 69 | by auto | 
| 14370 | 70 | |
| 71 | lemma STAR_real_seq_to_hypreal: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 72 | "\<forall>n. (X n) \<notin> M ==> star_n X \<notin> *s* M" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 73 | apply (unfold starset_def star_of_def) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 74 | apply (simp add: Iset_star_n) | 
| 14370 | 75 | done | 
| 76 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 77 | lemma STAR_singleton: "*s* {x} = {star_of x}"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 78 | by simp | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 79 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 80 | lemma STAR_not_mem: "x \<notin> F ==> star_of x \<notin> *s* F" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 81 | by transfer | 
| 14370 | 82 | |
| 83 | lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B" | |
| 20732 | 84 | by (erule rev_subsetD, simp) | 
| 14370 | 85 | |
| 15169 | 86 | text{*Nonstandard extension of a set (defined using a constant
 | 
| 87 | sequence) as a special case of an internal set*} | |
| 14370 | 88 | |
| 14468 | 89 | lemma starset_n_starset: "\<forall>n. (As n = A) ==> *sn* As = *s* A" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 90 | apply (drule expand_fun_eq [THEN iffD2]) | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 91 | apply (simp add: starset_n_def starset_def star_of_def) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 92 | done | 
| 14370 | 93 | |
| 94 | ||
| 95 | (*----------------------------------------------------------------*) | |
| 96 | (* Theorems about nonstandard extensions of functions *) | |
| 97 | (*----------------------------------------------------------------*) | |
| 98 | ||
| 99 | (*----------------------------------------------------------------*) | |
| 100 | (* Nonstandard extension of a function (defined using a *) | |
| 101 | (* constant sequence) as a special case of an internal function *) | |
| 102 | (*----------------------------------------------------------------*) | |
| 103 | ||
| 14468 | 104 | lemma starfun_n_starfun: "\<forall>n. (F n = f) ==> *fn* F = *f* f" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 105 | apply (drule expand_fun_eq [THEN iffD2]) | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 106 | apply (simp add: starfun_n_def starfun_def star_of_def) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 107 | done | 
| 14370 | 108 | |
| 109 | ||
| 110 | (* | |
| 15003 | 111 | Prove that abs for hypreal is a nonstandard extension of abs for real w/o | 
| 14370 | 112 | use of congruence property (proved after this for general | 
| 15003 | 113 | nonstandard extensions of real valued functions). | 
| 14370 | 114 | |
| 15003 | 115 | Proof now Uses the ultrafilter tactic! | 
| 14370 | 116 | *) | 
| 117 | ||
| 118 | lemma hrabs_is_starext_rabs: "is_starext abs abs" | |
| 14468 | 119 | apply (simp add: is_starext_def, safe) | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 120 | apply (rule_tac x=x in star_cases) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 121 | apply (rule_tac x=y in star_cases) | 
| 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 122 | apply (unfold star_n_def, auto) | 
| 17298 | 123 | apply (rule bexI, rule_tac [2] lemma_starrel_refl) | 
| 124 | apply (rule bexI, rule_tac [2] lemma_starrel_refl) | |
| 17302 | 125 | apply (fold star_n_def) | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 126 | apply (unfold star_abs_def starfun_def star_of_def) | 
| 17302 | 127 | apply (simp add: Ifun_star_n star_n_eq_iff) | 
| 14370 | 128 | done | 
| 129 | ||
| 15169 | 130 | text{*Nonstandard extension of functions*}
 | 
| 14370 | 131 | |
| 132 | lemma starfun: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 133 | "( *f* f) (star_n X) = star_n (%n. f (X n))" | 
| 20732 | 134 | by (rule starfun_star_n) | 
| 14370 | 135 | |
| 14477 | 136 | lemma starfun_if_eq: | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 137 | "!!w. w \<noteq> star_of x | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 138 | ==> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 139 | by (transfer, simp) | 
| 14477 | 140 | |
| 14370 | 141 | (*------------------------------------------- | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 142 | multiplication: ( *f) x ( *g) = *(f x g) | 
| 14370 | 143 | ------------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 144 | lemma starfun_mult: "!!x. ( *f* f) x * ( *f* g) x = ( *f* (%x. f x * g x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 145 | by (transfer, rule refl) | 
| 14370 | 146 | declare starfun_mult [symmetric, simp] | 
| 147 | ||
| 148 | (*--------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 149 | addition: ( *f) + ( *g) = *(f + g) | 
| 14370 | 150 | ---------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 151 | lemma starfun_add: "!!x. ( *f* f) x + ( *f* g) x = ( *f* (%x. f x + g x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 152 | by (transfer, rule refl) | 
| 14370 | 153 | declare starfun_add [symmetric, simp] | 
| 154 | ||
| 155 | (*-------------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 156 | subtraction: ( *f) + -( *g) = *(f + -g) | 
| 14370 | 157 | -------------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 158 | lemma starfun_minus: "!!x. - ( *f* f) x = ( *f* (%x. - f x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 159 | by (transfer, rule refl) | 
| 14370 | 160 | declare starfun_minus [symmetric, simp] | 
| 161 | ||
| 162 | (*FIXME: delete*) | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 163 | lemma starfun_add_minus: "!!x. ( *f* f) x + -( *f* g) x = ( *f* (%x. f x + -g x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 164 | by (transfer, rule refl) | 
| 14370 | 165 | declare starfun_add_minus [symmetric, simp] | 
| 166 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 167 | lemma starfun_diff: "!!x. ( *f* f) x - ( *f* g) x = ( *f* (%x. f x - g x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 168 | by (transfer, rule refl) | 
| 14370 | 169 | declare starfun_diff [symmetric, simp] | 
| 170 | ||
| 171 | (*-------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 172 | composition: ( *f) o ( *g) = *(f o g) | 
| 14370 | 173 | ---------------------------------------*) | 
| 174 | ||
| 175 | lemma starfun_o2: "(%x. ( *f* f) (( *f* g) x)) = *f* (%x. f (g x))" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 176 | by (transfer, rule refl) | 
| 14370 | 177 | |
| 178 | lemma starfun_o: "( *f* f) o ( *f* g) = ( *f* (f o g))" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 179 | by (transfer o_def, rule refl) | 
| 14370 | 180 | |
| 15169 | 181 | text{*NS extension of constant function*}
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 182 | lemma starfun_const_fun [simp]: "!!x. ( *f* (%x. k)) x = star_of k" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 183 | by (transfer, rule refl) | 
| 14370 | 184 | |
| 15169 | 185 | text{*the NS extension of the identity function*}
 | 
| 14370 | 186 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 187 | lemma starfun_Id [simp]: "!!x. ( *f* (%x. x)) x = x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 188 | by (transfer, rule refl) | 
| 14370 | 189 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 190 | (* this is trivial, given starfun_Id *) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 191 | lemma starfun_Idfun_approx: | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 192 | "x @= star_of a ==> ( *f* (%x. x)) x @= star_of a" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 193 | by (simp only: starfun_Id) | 
| 14370 | 194 | |
| 15169 | 195 | text{*The Star-function is a (nonstandard) extension of the function*}
 | 
| 14370 | 196 | |
| 197 | lemma is_starext_starfun: "is_starext ( *f* f) f" | |
| 14468 | 198 | apply (simp add: is_starext_def, auto) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 199 | apply (rule_tac x = x in star_cases) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 200 | apply (rule_tac x = y in star_cases) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 201 | apply (auto intro!: bexI [OF _ Rep_star_star_n] | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 202 | simp add: starfun star_n_eq_iff) | 
| 14370 | 203 | done | 
| 204 | ||
| 15169 | 205 | text{*Any nonstandard extension is in fact the Star-function*}
 | 
| 14370 | 206 | |
| 207 | lemma is_starfun_starext: "is_starext F f ==> F = *f* f" | |
| 14468 | 208 | apply (simp add: is_starext_def) | 
| 14370 | 209 | apply (rule ext) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 210 | apply (rule_tac x = x in star_cases) | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 211 | apply (drule_tac x = x in spec) | 
| 14370 | 212 | apply (drule_tac x = "( *f* f) x" in spec) | 
| 21850 | 213 | apply (auto simp add: starfun_star_n) | 
| 214 | apply (simp add: star_n_eq_iff [symmetric]) | |
| 215 | apply (simp add: starfun_star_n [of f, symmetric]) | |
| 14370 | 216 | done | 
| 217 | ||
| 218 | lemma is_starext_starfun_iff: "(is_starext F f) = (F = *f* f)" | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 219 | by (blast intro: is_starfun_starext is_starext_starfun) | 
| 14370 | 220 | |
| 15169 | 221 | text{*extented function has same solution as its standard
 | 
| 14370 | 222 | version for real arguments. i.e they are the same | 
| 15169 | 223 | for all real arguments*} | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 224 | lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)" | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 225 | by (rule starfun_star_of) | 
| 14370 | 226 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 227 | lemma starfun_approx: "( *f* f) (star_of a) @= star_of (f a)" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 228 | by simp | 
| 14370 | 229 | |
| 230 | (* useful for NS definition of derivatives *) | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 231 | lemma starfun_lambda_cancel: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 232 | "!!x'. ( *f* (%h. f (x + h))) x' = ( *f* f) (star_of x + x')" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 233 | by (transfer, rule refl) | 
| 14370 | 234 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 235 | lemma starfun_lambda_cancel2: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 236 | "( *f* (%h. f(g(x + h)))) x' = ( *f* (f o g)) (star_of x + x')" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 237 | by (unfold o_def, rule starfun_lambda_cancel) | 
| 14370 | 238 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 239 | lemma starfun_mult_HFinite_approx: | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 240 | fixes l m :: "'a::real_normed_algebra star" | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 241 | shows "[| ( *f* f) x @= l; ( *f* g) x @= m; | 
| 14370 | 242 | l: HFinite; m: HFinite | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 243 | |] ==> ( *f* (%x. f x * g x)) x @= l * m" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 244 | apply (drule (3) approx_mult_HFinite) | 
| 14370 | 245 | apply (auto intro: approx_HFinite [OF _ approx_sym]) | 
| 246 | done | |
| 247 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 248 | lemma starfun_add_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 249 | |] ==> ( *f* (%x. f x + g x)) x @= l + m" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 250 | by (auto intro: approx_add) | 
| 14370 | 251 | |
| 15169 | 252 | text{*Examples: hrabs is nonstandard extension of rabs
 | 
| 253 | inverse is nonstandard extension of inverse*} | |
| 14370 | 254 | |
| 255 | (* can be proved easily using theorem "starfun" and *) | |
| 256 | (* properties of ultrafilter as for inverse below we *) | |
| 257 | (* use the theorem we proved above instead *) | |
| 258 | ||
| 259 | lemma starfun_rabs_hrabs: "*f* abs = abs" | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 260 | by (simp only: star_abs_def) | 
| 14370 | 261 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 262 | lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)" | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 263 | by (simp only: star_inverse_def) | 
| 14370 | 264 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 265 | lemma starfun_inverse: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 266 | by (transfer, rule refl) | 
| 14370 | 267 | declare starfun_inverse [symmetric, simp] | 
| 268 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 269 | lemma starfun_divide: "!!x. ( *f* f) x / ( *f* g) x = ( *f* (%x. f x / g x)) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 270 | by (transfer, rule refl) | 
| 14370 | 271 | declare starfun_divide [symmetric, simp] | 
| 272 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 273 | lemma starfun_inverse2: "!!x. inverse (( *f* f) x) = ( *f* (%x. inverse (f x))) x" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 274 | by (transfer, rule refl) | 
| 14370 | 275 | |
| 15169 | 276 | text{*General lemma/theorem needed for proofs in elementary
 | 
| 277 | topology of the reals*} | |
| 14370 | 278 | lemma starfun_mem_starset: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 279 |       "!!x. ( *f* f) x : *s* A ==> x : *s* {x. f x  \<in> A}"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 280 | by (transfer, simp) | 
| 14370 | 281 | |
| 15169 | 282 | text{*Alternative definition for hrabs with rabs function
 | 
| 14370 | 283 | applied entrywise to equivalence class representative. | 
| 15169 | 284 | This is easily proved using starfun and ns extension thm*} | 
| 15170 | 285 | lemma hypreal_hrabs: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 286 | "abs (star_n X) = star_n (%n. abs (X n))" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 287 | by (simp only: starfun_rabs_hrabs [symmetric] starfun) | 
| 14370 | 288 | |
| 15169 | 289 | text{*nonstandard extension of set through nonstandard extension
 | 
| 14370 | 290 | of rabs function i.e hrabs. A more general result should be | 
| 291 | where we replace rabs by some arbitrary function f and hrabs | |
| 15169 | 292 | by its NS extenson. See second NS set extension below.*} | 
| 14370 | 293 | lemma STAR_rabs_add_minus: | 
| 294 |    "*s* {x. abs (x + - y) < r} =
 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 295 |      {x. abs(x + -star_of y) < star_of r}"
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 296 | by (transfer, rule refl) | 
| 14370 | 297 | |
| 298 | lemma STAR_starfun_rabs_add_minus: | |
| 299 |   "*s* {x. abs (f x + - y) < r} =
 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 300 |        {x. abs(( *f* f) x + -star_of y) < star_of r}"
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 301 | by (transfer, rule refl) | 
| 14370 | 302 | |
| 15169 | 303 | text{*Another characterization of Infinitesimal and one of @= relation.
 | 
| 15170 | 304 |    In this theory since @{text hypreal_hrabs} proved here. Maybe
 | 
| 305 | move both theorems??*} | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14371diff
changeset | 306 | lemma Infinitesimal_FreeUltrafilterNat_iff2: | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 307 | "(star_n X \<in> Infinitesimal) = | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 308 |       (\<forall>m. {n. norm(X n) < inverse(real(Suc m))}
 | 
| 14468 | 309 | \<in> FreeUltrafilterNat)" | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 310 | by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 311 | hnorm_def star_of_nat_def starfun_star_n | 
| 20730 | 312 | star_n_inverse star_n_less real_of_nat_def) | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 313 | |
| 21865 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 314 | lemma HNatInfinite_inverse_Infinitesimal [simp]: | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 315 | "n \<in> HNatInfinite ==> inverse (hypreal_of_hypnat n) \<in> Infinitesimal" | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 316 | apply (cases n) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 317 | apply (auto simp add: of_hypnat_def starfun_star_n real_of_nat_def [symmetric] star_n_inverse real_norm_def | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 318 | HNatInfinite_FreeUltrafilterNat_iff | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 319 | Infinitesimal_FreeUltrafilterNat_iff2) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 320 | apply (drule_tac x="Suc m" in spec) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 321 | apply (erule ultra, simp) | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 322 | done | 
| 
55cc354fd2d9
moved several theorems; rearranged theory dependencies
 huffman parents: 
21850diff
changeset | 323 | |
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 324 | lemma approx_FreeUltrafilterNat_iff: "star_n X @= star_n Y = | 
| 20563 | 325 |       (\<forall>r>0. {n. norm (X n - Y n) < r} : FreeUltrafilterNat)"
 | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 326 | apply (subst approx_minus_iff) | 
| 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 327 | apply (rule mem_infmal_iff [THEN subst]) | 
| 20563 | 328 | apply (simp add: star_n_diff) | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 329 | apply (simp add: Infinitesimal_FreeUltrafilterNat_iff) | 
| 14370 | 330 | done | 
| 331 | ||
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 332 | lemma approx_FreeUltrafilterNat_iff2: "star_n X @= star_n Y = | 
| 20563 | 333 |       (\<forall>m. {n. norm (X n - Y n) <
 | 
| 14370 | 334 | inverse(real(Suc m))} : FreeUltrafilterNat)" | 
| 335 | apply (subst approx_minus_iff) | |
| 336 | apply (rule mem_infmal_iff [THEN subst]) | |
| 20563 | 337 | apply (simp add: star_n_diff) | 
| 20552 
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
 huffman parents: 
19765diff
changeset | 338 | apply (simp add: Infinitesimal_FreeUltrafilterNat_iff2) | 
| 14370 | 339 | done | 
| 340 | ||
| 341 | lemma inj_starfun: "inj starfun" | |
| 342 | apply (rule inj_onI) | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 343 | apply (rule ext, rule ccontr) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 344 | apply (drule_tac x = "star_n (%n. xa)" in fun_cong) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 345 | apply (auto simp add: starfun star_n_eq_iff) | 
| 14370 | 346 | done | 
| 347 | ||
| 348 | end |