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