src/HOL/Nonstandard_Analysis/Star.thy
author immler
Sun, 03 Nov 2019 21:46:46 -0500
changeset 71034 e0755162093f
parent 70224 3706106c2e0f
permissions -rw-r--r--
replace approximation oracle by less ad-hoc @{computation}s
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/Star.thy
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     2
    Author:     Jacques D. Fleuriot
716336f19aa9 clarified session;
wenzelm
parents: 61982
diff changeset
     3
    Copyright:  1998  University of Cambridge
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     6
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
     7
section \<open>Star-Transforms in Non-Standard Analysis\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     9
theory Star
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    10
  imports NSA
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    13
definition  \<comment> \<open>internal sets\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    14
  starset_n :: "(nat \<Rightarrow> 'a set) \<Rightarrow> 'a star set"  ("*sn* _" [80] 80)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    15
  where "*sn* As = Iset (star_n As)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    17
definition InternalSets :: "'a star set set"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    18
  where "InternalSets = {X. \<exists>As. X = *sn* As}"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    20
definition  \<comment> \<open>nonstandard extension of function\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    21
  is_starext :: "('a star \<Rightarrow> 'a star) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    22
  where "is_starext F f \<longleftrightarrow>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    23
    (\<forall>x y. \<exists>X \<in> Rep_star x. \<exists>Y \<in> Rep_star y. y = F x \<longleftrightarrow> eventually (\<lambda>n. Y n = f(X n)) \<U>)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    24
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    25
definition  \<comment> \<open>internal functions\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    26
  starfun_n :: "(nat \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a star \<Rightarrow> 'b star"  ("*fn* _" [80] 80)
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    27
  where "*fn* F = Ifun (star_n F)"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    28
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    29
definition InternalFuns :: "('a star => 'b star) set"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    30
  where "InternalFuns = {X. \<exists>F. X = *fn* F}"
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    31
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    33
subsection \<open>Preamble - Pulling \<open>\<exists>\<close> over \<open>\<forall>\<close>\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    34
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    35
text \<open>This proof does not need AC and was suggested by the
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    36
   referee for the JCM Paper: let \<open>f x\<close> be least \<open>y\<close> such
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    37
   that \<open>Q x y\<close>.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    38
lemma no_choice: "\<forall>x. \<exists>y. Q x y \<Longrightarrow> \<exists>f :: 'a \<Rightarrow> nat. \<forall>x. Q x (f x)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    39
  by (rule exI [where x = "\<lambda>x. LEAST y. Q x y"]) (blast intro: LeastI)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    40
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    41
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    42
subsection \<open>Properties of the Star-transform Applied to Sets of Reals\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    43
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    44
lemma STAR_star_of_image_subset: "star_of ` A \<subseteq> *s* A"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    45
  by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    46
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    47
lemma STAR_hypreal_of_real_Int: "*s* X \<inter> \<real> = hypreal_of_real ` X"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    48
  by (auto simp add: SReal_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    49
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    50
lemma STAR_star_of_Int: "*s* X \<inter> Standard = star_of ` X"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    51
  by (auto simp add: Standard_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    52
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    53
lemma lemma_not_hyprealA: "x \<notin> hypreal_of_real ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> hypreal_of_real y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    54
  by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    55
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    56
lemma lemma_not_starA: "x \<notin> star_of ` A \<Longrightarrow> \<forall>y \<in> A. x \<noteq> star_of y"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    57
  by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    58
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    59
lemma STAR_real_seq_to_hypreal: "\<forall>n. (X n) \<notin> M \<Longrightarrow> star_n X \<notin> *s* M"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    60
  by (simp add: starset_def star_of_def Iset_star_n FreeUltrafilterNat.proper)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    61
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    62
lemma STAR_singleton: "*s* {x} = {star_of x}"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    63
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    64
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    65
lemma STAR_not_mem: "x \<notin> F \<Longrightarrow> star_of x \<notin> *s* F"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    66
  by transfer
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    67
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    68
lemma STAR_subset_closed: "x \<in> *s* A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<in> *s* B"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    69
  by (erule rev_subsetD) simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    70
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    71
text \<open>Nonstandard extension of a set (defined using a constant
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    72
   sequence) as a special case of an internal set.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    73
lemma starset_n_starset: "\<forall>n. As n = A \<Longrightarrow> *sn* As = *s* A"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    74
  by (drule fun_eq_iff [THEN iffD2]) (simp add: starset_n_def starset_def star_of_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    75
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    76
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    77
subsection \<open>Theorems about nonstandard extensions of functions\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    78
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    79
text \<open>Nonstandard extension of a function (defined using a
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    80
  constant sequence) as a special case of an internal function.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    81
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    82
lemma starfun_n_starfun: "F = (\<lambda>n. f) \<Longrightarrow> *fn* F = *f* f"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    83
  by (simp add: starfun_n_def starfun_def star_of_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    85
text \<open>Prove that \<open>abs\<close> for hypreal is a nonstandard extension of abs for real w/o
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    86
  use of congruence property (proved after this for general
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    87
  nonstandard extensions of real valued functions).
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    88
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    89
  Proof now Uses the ultrafilter tactic!\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
lemma hrabs_is_starext_rabs: "is_starext abs abs"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    92
  proof -
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    93
  have "\<exists>f\<in>Rep_star (star_n h). \<exists>g\<in>Rep_star (star_n k). (star_n k = \<bar>star_n h\<bar>) = (\<forall>\<^sub>F n in \<U>. (g n::'a) = \<bar>f n\<bar>)"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    94
    for x y :: "'a star" and h k
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    95
    by (metis (full_types) Rep_star_star_n star_n_abs star_n_eq_iff)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    96
  then show ?thesis
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    97
    unfolding is_starext_def by (metis star_cases)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
    98
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   100
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   101
text \<open>Nonstandard extension of functions.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   102
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   103
lemma starfun: "( *f* f) (star_n X) = star_n (\<lambda>n. f (X n))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   104
  by (rule starfun_star_n)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   106
lemma starfun_if_eq: "\<And>w. w \<noteq> star_of x \<Longrightarrow> ( *f* (\<lambda>z. if z = x then a else g z)) w = ( *f* g) w"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   107
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   108
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   109
text \<open>Multiplication: \<open>( *f) x ( *g) = *(f x g)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   110
lemma starfun_mult: "\<And>x. ( *f* f) x * ( *f* g) x = ( *f* (\<lambda>x. f x * g x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   111
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   112
declare starfun_mult [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   114
text \<open>Addition: \<open>( *f) + ( *g) = *(f + g)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   115
lemma starfun_add: "\<And>x. ( *f* f) x + ( *f* g) x = ( *f* (\<lambda>x. f x + g x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   116
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   117
declare starfun_add [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   118
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   119
text \<open>Subtraction: \<open>( *f) + -( *g) = *(f + -g)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   120
lemma starfun_minus: "\<And>x. - ( *f* f) x = ( *f* (\<lambda>x. - f x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   121
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   122
declare starfun_minus [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   123
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   124
(*FIXME: delete*)
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   125
lemma starfun_add_minus: "\<And>x. ( *f* f) x + -( *f* g) x = ( *f* (\<lambda>x. f x + -g x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   126
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   127
declare starfun_add_minus [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   129
lemma starfun_diff: "\<And>x. ( *f* f) x  - ( *f* g) x = ( *f* (\<lambda>x. f x - g x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   130
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   131
declare starfun_diff [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   132
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   133
text \<open>Composition: \<open>( *f) \<circ> ( *g) = *(f \<circ> g)\<close>\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   134
lemma starfun_o2: "(\<lambda>x. ( *f* f) (( *f* g) x)) = *f* (\<lambda>x. f (g x))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   135
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   136
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   137
lemma starfun_o: "( *f* f) \<circ> ( *f* g) = ( *f* (f \<circ> g))"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   138
  by (transfer o_def) (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   139
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   140
text \<open>NS extension of constant function.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   141
lemma starfun_const_fun [simp]: "\<And>x. ( *f* (\<lambda>x. k)) x = star_of k"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   142
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   143
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   144
text \<open>The NS extension of the identity function.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   145
lemma starfun_Id [simp]: "\<And>x. ( *f* (\<lambda>x. x)) x = x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   146
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   147
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   148
text \<open>The Star-function is a (nonstandard) extension of the function.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   149
lemma is_starext_starfun: "is_starext ( *f* f) f"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   150
proof -
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   151
  have "\<exists>X\<in>Rep_star x. \<exists>Y\<in>Rep_star y. (y = (*f* f) x) = (\<forall>\<^sub>F n in \<U>. Y n = f (X n))"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   152
    for x y
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   153
    by (metis (mono_tags) Rep_star_star_n star_cases star_n_eq_iff starfun_star_n)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   154
  then show ?thesis
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   155
    by (auto simp: is_starext_def)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   156
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   157
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   158
text \<open>Any nonstandard extension is in fact the Star-function.\<close>
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   159
lemma is_starfun_starext:
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   160
  assumes "is_starext F f"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   161
  shows "F = *f* f"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   162
  proof -
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   163
  have "F x = (*f* f) x"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   164
    if "\<forall>x y. \<exists>X\<in>Rep_star x. \<exists>Y\<in>Rep_star y. (y = F x) = (\<forall>\<^sub>F n in \<U>. Y n = f (X n))" for x
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   165
    by (metis that mem_Rep_star_iff star_n_eq_iff starfun_star_n)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   166
  with assms show ?thesis
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   167
    by (force simp add: is_starext_def)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   168
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   169
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   170
lemma is_starext_starfun_iff: "is_starext F f \<longleftrightarrow> F = *f* f"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   171
  by (blast intro: is_starfun_starext is_starext_starfun)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   172
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   173
text \<open>Extended function has same solution as its standard version
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   174
  for real arguments. i.e they are the same for all real arguments.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   175
lemma starfun_eq: "( *f* f) (star_of a) = star_of (f a)"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   176
  by (rule starfun_star_of)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   177
61982
3af5a06577c7 more symbols;
wenzelm
parents: 61975
diff changeset
   178
lemma starfun_approx: "( *f* f) (star_of a) \<approx> star_of (f a)"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   179
  by simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   180
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   181
text \<open>Useful for NS definition of derivatives.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   182
lemma starfun_lambda_cancel: "\<And>x'. ( *f* (\<lambda>h. f (x + h))) x'  = ( *f* f) (star_of x + x')"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   183
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   184
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   185
lemma starfun_lambda_cancel2: "( *f* (\<lambda>h. f (g (x + h)))) x' = ( *f* (f \<circ> g)) (star_of x + x')"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   186
  unfolding o_def by (rule starfun_lambda_cancel)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   187
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   188
lemma starfun_mult_HFinite_approx:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   189
  "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> l \<in> HFinite \<Longrightarrow> m \<in> HFinite \<Longrightarrow>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   190
    ( *f* (\<lambda>x. f x * g x)) x \<approx> l * m"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   191
  for l m :: "'a::real_normed_algebra star"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   192
  using approx_mult_HFinite by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   193
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   194
lemma starfun_add_approx: "( *f* f) x \<approx> l \<Longrightarrow> ( *f* g) x \<approx> m \<Longrightarrow> ( *f* (%x. f x + g x)) x \<approx> l + m"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   195
  by (auto intro: approx_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   196
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   197
text \<open>Examples: \<open>hrabs\<close> is nonstandard extension of \<open>rabs\<close>,
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   198
  \<open>inverse\<close> is nonstandard extension of \<open>inverse\<close>.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   199
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   200
text \<open>Can be proved easily using theorem \<open>starfun\<close> and
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   201
  properties of ultrafilter as for inverse below we
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   202
  use the theorem we proved above instead.\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   203
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   204
lemma starfun_rabs_hrabs: "*f* abs = abs"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   205
  by (simp only: star_abs_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   206
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   207
lemma starfun_inverse_inverse [simp]: "( *f* inverse) x = inverse x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   208
  by (simp only: star_inverse_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   209
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   210
lemma starfun_inverse: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   211
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   212
declare starfun_inverse [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   213
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   214
lemma starfun_divide: "\<And>x. ( *f* f) x / ( *f* g) x = ( *f* (\<lambda>x. f x / g x)) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   215
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   216
declare starfun_divide [symmetric, simp]
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   217
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   218
lemma starfun_inverse2: "\<And>x. inverse (( *f* f) x) = ( *f* (\<lambda>x. inverse (f x))) x"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   219
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   220
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   221
text \<open>General lemma/theorem needed for proofs in elementary topology of the reals.\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 64435
diff changeset
   222
lemma starfun_mem_starset: "\<And>x. ( *f* f) x \<in> *s* A \<Longrightarrow> x \<in> *s* {x. f x \<in> A}"
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   223
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   224
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   225
text \<open>Alternative definition for \<open>hrabs\<close> with \<open>rabs\<close> function applied
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   226
  entrywise to equivalence class representative.
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   227
  This is easily proved using @{thm [source] starfun} and ns extension thm.\<close>
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   228
lemma hypreal_hrabs: "\<bar>star_n X\<bar> = star_n (\<lambda>n. \<bar>X n\<bar>)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   229
  by (simp only: starfun_rabs_hrabs [symmetric] starfun)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   230
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   231
text \<open>Nonstandard extension of set through nonstandard extension
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   232
   of \<open>rabs\<close> function i.e. \<open>hrabs\<close>. A more general result should be
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   233
   where we replace \<open>rabs\<close> by some arbitrary function \<open>f\<close> and \<open>hrabs\<close>
61975
b4b11391c676 isabelle update_cartouches -c -t;
wenzelm
parents: 61945
diff changeset
   234
   by its NS extenson. See second NS set extension below.\<close>
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   235
lemma STAR_rabs_add_minus: "*s* {x. \<bar>x + - y\<bar> < r} = {x. \<bar>x + -star_of y\<bar> < star_of r}"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   236
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   237
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   238
lemma STAR_starfun_rabs_add_minus:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   239
  "*s* {x. \<bar>f x + - y\<bar> < r} = {x. \<bar>( *f* f) x + -star_of y\<bar> < star_of r}"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   240
  by transfer (rule refl)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   241
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   242
text \<open>Another characterization of Infinitesimal and one of \<open>\<approx>\<close> relation.
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   243
  In this theory since \<open>hypreal_hrabs\<close> proved here. Maybe move both theorems??\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   244
lemma Infinitesimal_FreeUltrafilterNat_iff2:
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   245
  "star_n X \<in> Infinitesimal \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n) < inverse (real (Suc m))) \<U>)"
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   246
  by (simp add: Infinitesimal_hypreal_of_nat_iff star_of_def hnorm_def
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   247
      star_of_nat_def starfun_star_n star_n_inverse star_n_less)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   248
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   249
lemma HNatInfinite_inverse_Infinitesimal [simp]:
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   250
  assumes "n \<in> HNatInfinite"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   251
  shows "inverse (hypreal_of_hypnat n) \<in> Infinitesimal"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   252
proof (cases n)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   253
  case (star_n X)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   254
  then have *: "\<And>k. \<forall>\<^sub>F n in \<U>. k < X n"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   255
    using HNatInfinite_FreeUltrafilterNat assms by blast
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   256
  have "\<forall>\<^sub>F n in \<U>. inverse (real (X n)) < inverse (1 + real m)" for m
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   257
    using * [of "Suc m"] by (auto elim!: eventually_mono)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   258
  then show ?thesis
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   259
    using star_n by (auto simp: of_hypnat_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff2)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   260
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   261
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   262
lemma approx_FreeUltrafilterNat_iff:
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   263
  "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>r>0. eventually (\<lambda>n. norm (X n - Y n) < r) \<U>)"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   264
  (is "?lhs = ?rhs")
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   265
proof -
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   266
  have "?lhs = (star_n X - star_n Y \<approx> 0)"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   267
    using approx_minus_iff by blast
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   268
  also have "... = ?rhs"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   269
    by (metis (full_types) Infinitesimal_FreeUltrafilterNat_iff mem_infmal_iff star_n_diff)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   270
  finally show ?thesis .
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   271
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   272
64435
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   273
lemma approx_FreeUltrafilterNat_iff2:
c93b0e6131c3 misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   274
  "star_n X \<approx> star_n Y \<longleftrightarrow> (\<forall>m. eventually (\<lambda>n. norm (X n - Y n) < inverse (real (Suc m))) \<U>)"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   275
  (is "?lhs = ?rhs")
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   276
proof -
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   277
  have "?lhs = (star_n X - star_n Y \<approx> 0)"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   278
    using approx_minus_iff by blast
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   279
  also have "... = ?rhs"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   280
    by (metis (full_types) Infinitesimal_FreeUltrafilterNat_iff2 mem_infmal_iff star_n_diff)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   281
  finally show ?thesis .
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   282
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   283
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   284
lemma inj_starfun: "inj starfun"
70218
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   285
proof (rule inj_onI)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   286
  show "\<phi> = \<psi>" if eq: "*f* \<phi> = *f* \<psi>" for \<phi> \<psi> :: "'a \<Rightarrow> 'b"
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   287
  proof (rule ext, rule ccontr)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   288
    show False
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   289
      if "\<phi> x \<noteq> \<psi> x" for x
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   290
      by (metis eq that star_of_inject starfun_eq)
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   291
  qed
e48c0b5897a6 more de-applying
paulson <lp15@cam.ac.uk>
parents: 67613
diff changeset
   292
qed
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   293
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   294
end