| author | wenzelm | 
| Tue, 08 Nov 2005 10:43:08 +0100 | |
| changeset 18117 | 61a430a67d7c | 
| parent 17429 | e8d6ed3aacfe | 
| child 19765 | dfe940911617 | 
| 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 | |
| 13 | constdefs | |
| 14 | (* internal sets *) | |
| 17302 | 15 |     starset_n :: "(nat => 'a set) => 'a star set"        ("*sn* _" [80] 80)
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 16 | "*sn* As == Iset (star_n As)" | 
| 14370 | 17 | |
| 17302 | 18 | InternalSets :: "'a star set set" | 
| 14468 | 19 |     "InternalSets == {X. \<exists>As. X = *sn* As}"
 | 
| 10751 | 20 | |
| 21 | (* nonstandard extension of function *) | |
| 17302 | 22 | is_starext :: "['a star => 'a star, 'a => 'a] => bool" | 
| 17298 | 23 | "is_starext F f == (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y). | 
| 10751 | 24 |                         ((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
 | 
| 25 | (* internal functions *) | |
| 17303 | 26 |     starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star"
 | 
| 14370 | 27 |                  ("*fn* _" [80] 80)
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 28 | "*fn* F == Ifun (star_n F)" | 
| 10751 | 29 | |
| 17303 | 30 |     InternalFuns :: "('a star => 'b star) set"
 | 
| 14468 | 31 |     "InternalFuns == {X. \<exists>F. X = *fn* F}"
 | 
| 10751 | 32 | |
| 33 | ||
| 34 | ||
| 14370 | 35 | (*-------------------------------------------------------- | 
| 36 | Preamble - Pulling "EX" over "ALL" | |
| 37 | ---------------------------------------------------------*) | |
| 38 | ||
| 39 | (* This proof does not need AC and was suggested by the | |
| 40 | referee for the JCM Paper: let f(x) be least y such | |
| 41 | that Q(x,y) | |
| 42 | *) | |
| 14468 | 43 | lemma no_choice: "\<forall>x. \<exists>y. Q x y ==> \<exists>(f :: nat => nat). \<forall>x. Q x (f x)" | 
| 14370 | 44 | apply (rule_tac x = "%x. LEAST y. Q x y" in exI) | 
| 45 | apply (blast intro: LeastI) | |
| 46 | done | |
| 47 | ||
| 15169 | 48 | subsection{*Properties of the Star-transform Applied to Sets of Reals*}
 | 
| 14370 | 49 | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 50 | lemma STAR_UNIV_set: "*s*(UNIV::'a set) = (UNIV::'a star set)" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 51 | by (transfer UNIV_def, rule refl) | 
| 14370 | 52 | |
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 53 | lemma STAR_empty_set: "*s* {} = {}"
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 54 | by (transfer empty_def, rule refl) | 
| 14370 | 55 | |
| 56 | lemma STAR_Un: "*s* (A Un B) = *s* A Un *s* B" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 57 | by (transfer Un_def, rule refl) | 
| 14370 | 58 | |
| 59 | lemma STAR_Int: "*s* (A Int B) = *s* A Int *s* B" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 60 | by (transfer Int_def, rule refl) | 
| 14370 | 61 | |
| 62 | lemma STAR_Compl: "*s* -A = -( *s* A)" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 63 | by (transfer Compl_def, rule refl) | 
| 14370 | 64 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 65 | lemma STAR_mem_Compl: "!!x. x \<notin> *s* F ==> x : *s* (- F)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 66 | by (transfer Compl_def, simp) | 
| 14370 | 67 | |
| 68 | lemma STAR_diff: "*s* (A - B) = *s* A - *s* B" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 69 | by (transfer set_diff_def, rule refl) | 
| 14370 | 70 | |
| 71 | lemma STAR_subset: "A <= B ==> *s* A <= *s* B" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 72 | by (transfer subset_def, simp) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 73 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 74 | lemma STAR_mem: "a \<in> A ==> star_of a : *s* A" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 75 | by transfer | 
| 14370 | 76 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 77 | lemma STAR_mem_iff: "(star_of x \<in> *s* A) = (x \<in> A)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 78 | by (transfer, rule refl) | 
| 14370 | 79 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 80 | lemma STAR_star_of_image_subset: "star_of ` A <= *s* A" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 81 | by (auto simp add: STAR_mem) | 
| 14370 | 82 | |
| 83 | lemma STAR_hypreal_of_real_Int: "*s* X Int Reals = hypreal_of_real ` X" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 84 | by (auto simp add: SReal_def STAR_mem_iff) | 
| 14370 | 85 | |
| 14468 | 86 | 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 | 87 | by auto | 
| 14370 | 88 | |
| 89 | 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 | 90 | by auto | 
| 14370 | 91 | |
| 92 | lemma STAR_real_seq_to_hypreal: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 93 | "\<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 | 94 | apply (unfold starset_def star_of_def) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 95 | apply (simp add: Iset_star_n) | 
| 14370 | 96 | done | 
| 97 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 98 | lemma STAR_singleton: "*s* {x} = {star_of x}"
 | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 99 | by simp | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 100 | |
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 101 | 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 | 102 | by transfer | 
| 14370 | 103 | |
| 104 | lemma STAR_subset_closed: "[| x : *s* A; A <= B |] ==> x : *s* B" | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 105 | by (blast dest: STAR_subset) | 
| 14370 | 106 | |
| 15169 | 107 | text{*Nonstandard extension of a set (defined using a constant
 | 
| 108 | sequence) as a special case of an internal set*} | |
| 14370 | 109 | |
| 14468 | 110 | 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 | 111 | 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 | 112 | 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 | 113 | done | 
| 14370 | 114 | |
| 115 | ||
| 116 | (*----------------------------------------------------------------*) | |
| 117 | (* Theorems about nonstandard extensions of functions *) | |
| 118 | (*----------------------------------------------------------------*) | |
| 119 | ||
| 120 | (*----------------------------------------------------------------*) | |
| 121 | (* Nonstandard extension of a function (defined using a *) | |
| 122 | (* constant sequence) as a special case of an internal function *) | |
| 123 | (*----------------------------------------------------------------*) | |
| 124 | ||
| 14468 | 125 | 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 | 126 | 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 | 127 | 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 | 128 | done | 
| 14370 | 129 | |
| 130 | ||
| 131 | (* | |
| 15003 | 132 | Prove that abs for hypreal is a nonstandard extension of abs for real w/o | 
| 14370 | 133 | use of congruence property (proved after this for general | 
| 15003 | 134 | nonstandard extensions of real valued functions). | 
| 14370 | 135 | |
| 15003 | 136 | Proof now Uses the ultrafilter tactic! | 
| 14370 | 137 | *) | 
| 138 | ||
| 139 | lemma hrabs_is_starext_rabs: "is_starext abs abs" | |
| 14468 | 140 | 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 | 141 | 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 | 142 | 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 | 143 | apply (unfold star_n_def, auto) | 
| 17298 | 144 | apply (rule bexI, rule_tac [2] lemma_starrel_refl) | 
| 145 | apply (rule bexI, rule_tac [2] lemma_starrel_refl) | |
| 17302 | 146 | 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 | 147 | apply (unfold star_abs_def starfun_def star_of_def) | 
| 17302 | 148 | apply (simp add: Ifun_star_n star_n_eq_iff) | 
| 14370 | 149 | done | 
| 150 | ||
| 17298 | 151 | lemma Rep_star_FreeUltrafilterNat: | 
| 152 | "[| X \<in> Rep_star z; Y \<in> Rep_star z |] | |
| 14370 | 153 |       ==> {n. X n = Y n} : FreeUltrafilterNat"
 | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 154 | by (rule FreeUltrafilterNat_Rep_hypreal) | 
| 14370 | 155 | |
| 15169 | 156 | text{*Nonstandard extension of functions*}
 | 
| 14370 | 157 | |
| 158 | lemma starfun: | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 159 | "( *f* f) (star_n X) = star_n (%n. f (X n))" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17318diff
changeset | 160 | by (simp add: starfun_def star_of_def Ifun_star_n) | 
| 14370 | 161 | |
| 14477 | 162 | 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 | 163 | "!!w. w \<noteq> star_of x | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 164 | ==> ( *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 | 165 | by (transfer, simp) | 
| 14477 | 166 | |
| 14370 | 167 | (*------------------------------------------- | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 168 | multiplication: ( *f) x ( *g) = *(f x g) | 
| 14370 | 169 | ------------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 170 | 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 | 171 | by (transfer, rule refl) | 
| 14370 | 172 | declare starfun_mult [symmetric, simp] | 
| 173 | ||
| 174 | (*--------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 175 | addition: ( *f) + ( *g) = *(f + g) | 
| 14370 | 176 | ---------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 177 | 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 | 178 | by (transfer, rule refl) | 
| 14370 | 179 | declare starfun_add [symmetric, simp] | 
| 180 | ||
| 181 | (*-------------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 182 | subtraction: ( *f) + -( *g) = *(f + -g) | 
| 14370 | 183 | -------------------------------------------*) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 184 | 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 | 185 | by (transfer, rule refl) | 
| 14370 | 186 | declare starfun_minus [symmetric, simp] | 
| 187 | ||
| 188 | (*FIXME: delete*) | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 189 | 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 | 190 | by (transfer, rule refl) | 
| 14370 | 191 | declare starfun_add_minus [symmetric, simp] | 
| 192 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 193 | 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 | 194 | by (transfer, rule refl) | 
| 14370 | 195 | declare starfun_diff [symmetric, simp] | 
| 196 | ||
| 197 | (*-------------------------------------- | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 198 | composition: ( *f) o ( *g) = *(f o g) | 
| 14370 | 199 | ---------------------------------------*) | 
| 200 | ||
| 201 | 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 | 202 | by (transfer, rule refl) | 
| 14370 | 203 | |
| 204 | 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 | 205 | by (transfer o_def, rule refl) | 
| 14370 | 206 | |
| 15169 | 207 | text{*NS extension of constant function*}
 | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 208 | 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 | 209 | by (transfer, rule refl) | 
| 14370 | 210 | |
| 15169 | 211 | text{*the NS extension of the identity function*}
 | 
| 14370 | 212 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 213 | 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 | 214 | by (transfer, rule refl) | 
| 14370 | 215 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 216 | (* this is trivial, given starfun_Id *) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 217 | lemma starfun_Idfun_approx: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 218 | "x @= hypreal_of_real a ==> ( *f* (%x. x)) x @= hypreal_of_real a" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 219 | by (simp only: starfun_Id) | 
| 14370 | 220 | |
| 15169 | 221 | text{*The Star-function is a (nonstandard) extension of the function*}
 | 
| 14370 | 222 | |
| 223 | lemma is_starext_starfun: "is_starext ( *f* f) f" | |
| 14468 | 224 | apply (simp add: is_starext_def, auto) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 225 | apply (rule_tac x = x in star_cases) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 226 | apply (rule_tac x = y in star_cases) | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 227 | 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 | 228 | simp add: starfun star_n_eq_iff) | 
| 14370 | 229 | done | 
| 230 | ||
| 15169 | 231 | text{*Any nonstandard extension is in fact the Star-function*}
 | 
| 14370 | 232 | |
| 233 | lemma is_starfun_starext: "is_starext F f ==> F = *f* f" | |
| 14468 | 234 | apply (simp add: is_starext_def) | 
| 14370 | 235 | apply (rule ext) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 236 | 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 | 237 | apply (drule_tac x = x in spec) | 
| 14370 | 238 | apply (drule_tac x = "( *f* f) x" in spec) | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 239 | apply (auto dest!: FreeUltrafilterNat_Compl_mem simp add: starfun, ultra) | 
| 14370 | 240 | done | 
| 241 | ||
| 242 | 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 | 243 | by (blast intro: is_starfun_starext is_starext_starfun) | 
| 14370 | 244 | |
| 15169 | 245 | text{*extented function has same solution as its standard
 | 
| 14370 | 246 | version for real arguments. i.e they are the same | 
| 15169 | 247 | for all real arguments*} | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 248 | lemma starfun_eq [simp]: "( *f* f) (star_of a) = star_of (f a)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 249 | by (transfer, rule refl) | 
| 14370 | 250 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 251 | lemma starfun_approx: "( *f* f) (star_of a) @= hypreal_of_real (f a)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 252 | by simp | 
| 14370 | 253 | |
| 254 | (* useful for NS definition of derivatives *) | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 255 | lemma starfun_lambda_cancel: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 256 | "!!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 | 257 | by (transfer, rule refl) | 
| 14370 | 258 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 259 | lemma starfun_lambda_cancel2: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 260 | "( *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 | 261 | by (unfold o_def, rule starfun_lambda_cancel) | 
| 14370 | 262 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 263 | lemma starfun_mult_HFinite_approx: "[| ( *f* f) x @= l; ( *f* g) x @= m; | 
| 14370 | 264 | l: HFinite; m: HFinite | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 265 | |] ==> ( *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 | 266 | apply (drule (3) approx_mult_HFinite) | 
| 14370 | 267 | apply (auto intro: approx_HFinite [OF _ approx_sym]) | 
| 268 | done | |
| 269 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 270 | 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 | 271 | |] ==> ( *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 | 272 | by (auto intro: approx_add) | 
| 14370 | 273 | |
| 15169 | 274 | text{*Examples: hrabs is nonstandard extension of rabs
 | 
| 275 | inverse is nonstandard extension of inverse*} | |
| 14370 | 276 | |
| 277 | (* can be proved easily using theorem "starfun" and *) | |
| 278 | (* properties of ultrafilter as for inverse below we *) | |
| 279 | (* use the theorem we proved above instead *) | |
| 280 | ||
| 281 | lemma starfun_rabs_hrabs: "*f* abs = abs" | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 282 | by (rule hrabs_is_starext_rabs [THEN is_starext_starfun_iff [THEN iffD1], symmetric]) | 
| 14370 | 283 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 284 | lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse(x)" | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 285 | by (unfold star_inverse_def, rule refl) | 
| 14370 | 286 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 287 | 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 | 288 | by (transfer, rule refl) | 
| 14370 | 289 | declare starfun_inverse [symmetric, simp] | 
| 290 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 291 | 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 | 292 | by (transfer, rule refl) | 
| 14370 | 293 | declare starfun_divide [symmetric, simp] | 
| 294 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 295 | 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 | 296 | by (transfer, rule refl) | 
| 14370 | 297 | |
| 15169 | 298 | text{*General lemma/theorem needed for proofs in elementary
 | 
| 299 | topology of the reals*} | |
| 14370 | 300 | lemma starfun_mem_starset: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 301 |       "!!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 | 302 | by (transfer, simp) | 
| 14370 | 303 | |
| 15169 | 304 | text{*Alternative definition for hrabs with rabs function
 | 
| 14370 | 305 | applied entrywise to equivalence class representative. | 
| 15169 | 306 | This is easily proved using starfun and ns extension thm*} | 
| 15170 | 307 | lemma hypreal_hrabs: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 308 | "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 | 309 | by (simp only: starfun_rabs_hrabs [symmetric] starfun) | 
| 14370 | 310 | |
| 15169 | 311 | text{*nonstandard extension of set through nonstandard extension
 | 
| 14370 | 312 | of rabs function i.e hrabs. A more general result should be | 
| 313 | where we replace rabs by some arbitrary function f and hrabs | |
| 15169 | 314 | by its NS extenson. See second NS set extension below.*} | 
| 14370 | 315 | lemma STAR_rabs_add_minus: | 
| 316 |    "*s* {x. abs (x + - y) < r} =
 | |
| 317 |      {x. abs(x + -hypreal_of_real y) < hypreal_of_real r}"
 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 318 | by (transfer, rule refl) | 
| 14370 | 319 | |
| 320 | lemma STAR_starfun_rabs_add_minus: | |
| 321 |   "*s* {x. abs (f x + - y) < r} =
 | |
| 322 |        {x. abs(( *f* f) x + -hypreal_of_real y) < hypreal_of_real r}"
 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 323 | by (transfer, rule refl) | 
| 14370 | 324 | |
| 15169 | 325 | text{*Another characterization of Infinitesimal and one of @= relation.
 | 
| 15170 | 326 |    In this theory since @{text hypreal_hrabs} proved here. Maybe
 | 
| 327 | move both theorems??*} | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14371diff
changeset | 328 | lemma Infinitesimal_FreeUltrafilterNat_iff2: | 
| 14468 | 329 | "(x \<in> Infinitesimal) = | 
| 17298 | 330 | (\<exists>X \<in> Rep_star(x). | 
| 14468 | 331 |         \<forall>m. {n. abs(X n) < inverse(real(Suc m))}
 | 
| 332 | \<in> FreeUltrafilterNat)" | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 333 | apply (cases x) | 
| 17298 | 334 | apply (auto intro!: bexI lemma_starrel_refl | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 335 | simp add: Infinitesimal_hypreal_of_nat_iff star_of_def | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 336 | star_n_inverse star_n_abs star_n_less hypreal_of_nat_eq) | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 337 | apply (drule_tac x = n in spec, ultra) | 
| 14370 | 338 | done | 
| 339 | ||
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 340 | lemma approx_FreeUltrafilterNat_iff: "star_n X @= star_n Y = | 
| 14468 | 341 |       (\<forall>m. {n. abs (X n + - Y n) <
 | 
| 14370 | 342 | inverse(real(Suc m))} : FreeUltrafilterNat)" | 
| 343 | apply (subst approx_minus_iff) | |
| 344 | apply (rule mem_infmal_iff [THEN subst]) | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 345 | apply (auto simp add: star_n_minus star_n_add Infinitesimal_FreeUltrafilterNat_iff2) | 
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 346 | apply (drule_tac x = m in spec, ultra) | 
| 14370 | 347 | done | 
| 348 | ||
| 349 | lemma inj_starfun: "inj starfun" | |
| 350 | apply (rule inj_onI) | |
| 14371 
c78c7da09519
Conversion of HyperNat to Isar format and its declaration as a semiring
 paulson parents: 
14370diff
changeset | 351 | apply (rule ext, rule ccontr) | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 352 | 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 | 353 | apply (auto simp add: starfun star_n_eq_iff) | 
| 14370 | 354 | done | 
| 355 | ||
| 356 | ML | |
| 357 | {*
 | |
| 358 | val starset_n_def = thm"starset_n_def"; | |
| 359 | val InternalSets_def = thm"InternalSets_def"; | |
| 360 | val is_starext_def = thm"is_starext_def"; | |
| 361 | val starfun_n_def = thm"starfun_n_def"; | |
| 362 | val InternalFuns_def = thm"InternalFuns_def"; | |
| 363 | ||
| 364 | val no_choice = thm "no_choice"; | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 365 | val STAR_UNIV_set = thm "STAR_UNIV_set"; | 
| 14370 | 366 | val STAR_empty_set = thm "STAR_empty_set"; | 
| 367 | val STAR_Un = thm "STAR_Un"; | |
| 368 | val STAR_Int = thm "STAR_Int"; | |
| 369 | val STAR_Compl = thm "STAR_Compl"; | |
| 370 | val STAR_mem_Compl = thm "STAR_mem_Compl"; | |
| 371 | val STAR_diff = thm "STAR_diff"; | |
| 372 | val STAR_subset = thm "STAR_subset"; | |
| 373 | val STAR_mem = thm "STAR_mem"; | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17303diff
changeset | 374 | val STAR_star_of_image_subset = thm "STAR_star_of_image_subset"; | 
| 14370 | 375 | val STAR_hypreal_of_real_Int = thm "STAR_hypreal_of_real_Int"; | 
| 376 | val STAR_real_seq_to_hypreal = thm "STAR_real_seq_to_hypreal"; | |
| 377 | val STAR_singleton = thm "STAR_singleton"; | |
| 378 | val STAR_not_mem = thm "STAR_not_mem"; | |
| 379 | val STAR_subset_closed = thm "STAR_subset_closed"; | |
| 380 | val starset_n_starset = thm "starset_n_starset"; | |
| 381 | val starfun_n_starfun = thm "starfun_n_starfun"; | |
| 382 | val hrabs_is_starext_rabs = thm "hrabs_is_starext_rabs"; | |
| 17298 | 383 | val Rep_star_FreeUltrafilterNat = thm "Rep_star_FreeUltrafilterNat"; | 
| 14370 | 384 | val starfun = thm "starfun"; | 
| 385 | val starfun_mult = thm "starfun_mult"; | |
| 386 | val starfun_add = thm "starfun_add"; | |
| 387 | val starfun_minus = thm "starfun_minus"; | |
| 388 | val starfun_add_minus = thm "starfun_add_minus"; | |
| 389 | val starfun_diff = thm "starfun_diff"; | |
| 390 | val starfun_o2 = thm "starfun_o2"; | |
| 391 | val starfun_o = thm "starfun_o"; | |
| 392 | val starfun_const_fun = thm "starfun_const_fun"; | |
| 393 | val starfun_Idfun_approx = thm "starfun_Idfun_approx"; | |
| 394 | val starfun_Id = thm "starfun_Id"; | |
| 395 | val is_starext_starfun = thm "is_starext_starfun"; | |
| 396 | val is_starfun_starext = thm "is_starfun_starext"; | |
| 397 | val is_starext_starfun_iff = thm "is_starext_starfun_iff"; | |
| 398 | val starfun_eq = thm "starfun_eq"; | |
| 399 | val starfun_approx = thm "starfun_approx"; | |
| 400 | val starfun_lambda_cancel = thm "starfun_lambda_cancel"; | |
| 401 | val starfun_lambda_cancel2 = thm "starfun_lambda_cancel2"; | |
| 402 | val starfun_mult_HFinite_approx = thm "starfun_mult_HFinite_approx"; | |
| 403 | val starfun_add_approx = thm "starfun_add_approx"; | |
| 404 | val starfun_rabs_hrabs = thm "starfun_rabs_hrabs"; | |
| 405 | val starfun_inverse_inverse = thm "starfun_inverse_inverse"; | |
| 406 | val starfun_inverse = thm "starfun_inverse"; | |
| 407 | val starfun_divide = thm "starfun_divide"; | |
| 408 | val starfun_inverse2 = thm "starfun_inverse2"; | |
| 409 | val starfun_mem_starset = thm "starfun_mem_starset"; | |
| 410 | val hypreal_hrabs = thm "hypreal_hrabs"; | |
| 411 | val STAR_rabs_add_minus = thm "STAR_rabs_add_minus"; | |
| 412 | val STAR_starfun_rabs_add_minus = thm "STAR_starfun_rabs_add_minus"; | |
| 413 | val Infinitesimal_FreeUltrafilterNat_iff2 = thm "Infinitesimal_FreeUltrafilterNat_iff2"; | |
| 414 | val approx_FreeUltrafilterNat_iff = thm "approx_FreeUltrafilterNat_iff"; | |
| 415 | val inj_starfun = thm "inj_starfun"; | |
| 416 | *} | |
| 417 | ||
| 418 | end |