Construction of the hyperreal numbers as an ultraproduct of real sequences. #
Hyperreal numbers on the ultrafilter extending the cofinite filter
Equations
- Hyperreal.«termℝ*» = Lean.ParserDescr.node `Hyperreal.«termℝ*» 1024 (Lean.ParserDescr.symbol "ℝ*")
Instances For
Equations
Natural embedding ℝ → ℝ*
.
Equations
Instances For
Equations
- Hyperreal.instCoeTCReal = { coe := Hyperreal.ofReal }
Construct a hyperreal number from a sequence of real numbers.
Equations
- Hyperreal.ofSeq f = ↑f
Instances For
A sample infinitesimal hyperreal
Equations
- Hyperreal.epsilon = Hyperreal.ofSeq fun (n : ℕ) => (↑n)⁻¹
Instances For
A sample infinite hyperreal
Equations
Instances For
A sample infinitesimal hyperreal
Equations
- Hyperreal.termε = Lean.ParserDescr.node `Hyperreal.termε 1024 (Lean.ParserDescr.symbol "ε")
Instances For
A sample infinite hyperreal
Equations
- Hyperreal.termω = Lean.ParserDescr.node `Hyperreal.termω 1024 (Lean.ParserDescr.symbol "ω")
Instances For
theorem
Hyperreal.lt_of_tendsto_zero_of_pos
{f : ℕ → ℝ}
(hf : Filter.Tendsto f Filter.atTop (nhds 0))
{r : ℝ}
:
theorem
Hyperreal.neg_lt_of_tendsto_zero_of_pos
{f : ℕ → ℝ}
(hf : Filter.Tendsto f Filter.atTop (nhds 0))
{r : ℝ}
:
theorem
Hyperreal.gt_of_tendsto_zero_of_neg
{f : ℕ → ℝ}
(hf : Filter.Tendsto f Filter.atTop (nhds 0))
{r : ℝ}
:
Standard part function: like a "round" to ℝ instead of ℤ
Equations
- x.st = if h : ∃ (r : ℝ), x.IsSt r then Classical.choose h else 0
Instances For
A hyperreal number is infinitesimal if its standard part is 0
Equations
- x.Infinitesimal = x.IsSt 0
Instances For
A hyperreal number is positive infinite if it is larger than all real numbers
Equations
- x.InfinitePos = ∀ (r : ℝ), ↑r < x
Instances For
A hyperreal number is negative infinite if it is smaller than all real numbers
Equations
- x.InfiniteNeg = ∀ (r : ℝ), x < ↑r
Instances For
A hyperreal number is infinite if it is infinite positive or infinite negative
Equations
- x.Infinite = (x.InfinitePos ∨ x.InfiniteNeg)
Instances For
theorem
Hyperreal.isSt_of_tendsto
{f : ℕ → ℝ}
{r : ℝ}
(hf : Filter.Tendsto f Filter.atTop (nhds r))
:
theorem
Hyperreal.IsSt.map
{x : ℝ*}
{r : ℝ}
(hxr : x.IsSt r)
{f : ℝ → ℝ}
(hf : ContinuousAt f r)
:
IsSt (Filter.Germ.map f x) (f r)
theorem
Hyperreal.IsSt.map₂
{x y : ℝ*}
{r s : ℝ}
(hxr : x.IsSt r)
(hys : y.IsSt s)
{f : ℝ → ℝ → ℝ}
(hf : ContinuousAt (Function.uncurry f) (r, s))
:
IsSt (Filter.Germ.map₂ f x y) (f r s)
Basic lemmas about infinite #
theorem
Hyperreal.infinitePos_add_not_infiniteNeg
{x y : ℝ*}
:
x.InfinitePos → ¬y.InfiniteNeg → (x + y).InfinitePos
theorem
Hyperreal.not_infiniteNeg_add_infinitePos
{x y : ℝ*}
:
¬x.InfiniteNeg → y.InfinitePos → (x + y).InfinitePos
theorem
Hyperreal.infiniteNeg_add_not_infinitePos
{x y : ℝ*}
:
x.InfiniteNeg → ¬y.InfinitePos → (x + y).InfiniteNeg
theorem
Hyperreal.not_infinitePos_add_infiniteNeg
{x y : ℝ*}
:
¬x.InfinitePos → y.InfiniteNeg → (x + y).InfiniteNeg
theorem
Hyperreal.infinitePos_add_infinitePos
{x y : ℝ*}
:
x.InfinitePos → y.InfinitePos → (x + y).InfinitePos
theorem
Hyperreal.infiniteNeg_add_infiniteNeg
{x y : ℝ*}
:
x.InfiniteNeg → y.InfiniteNeg → (x + y).InfiniteNeg
theorem
Hyperreal.infinitePos_add_not_infinite
{x y : ℝ*}
:
x.InfinitePos → ¬y.Infinite → (x + y).InfinitePos
theorem
Hyperreal.infiniteNeg_add_not_infinite
{x y : ℝ*}
:
x.InfiniteNeg → ¬y.Infinite → (x + y).InfiniteNeg
theorem
Hyperreal.infinitePos_of_tendsto_top
{f : ℕ → ℝ}
(hf : Filter.Tendsto f Filter.atTop Filter.atTop)
:
(ofSeq f).InfinitePos
theorem
Hyperreal.infiniteNeg_of_tendsto_bot
{f : ℕ → ℝ}
(hf : Filter.Tendsto f Filter.atTop Filter.atBot)
:
(ofSeq f).InfiniteNeg
Basic lemmas about infinitesimal #
theorem
Hyperreal.lt_neg_of_pos_of_infinitesimal
{x : ℝ*}
:
x.Infinitesimal → ∀ (r : ℝ), 0 < r → -↑r < x
theorem
Hyperreal.Infinitesimal.add
{x y : ℝ*}
(hx : x.Infinitesimal)
(hy : y.Infinitesimal)
:
(x + y).Infinitesimal
theorem
Hyperreal.Infinitesimal.mul
{x y : ℝ*}
(hx : x.Infinitesimal)
(hy : y.Infinitesimal)
:
(x * y).Infinitesimal
theorem
Hyperreal.infinitesimal_of_tendsto_zero
{f : ℕ → ℝ}
(h : Filter.Tendsto f Filter.atTop (nhds 0))
:
(ofSeq f).Infinitesimal
theorem
Hyperreal.not_real_of_infinitesimal_ne_zero
(x : ℝ*)
:
x.Infinitesimal → x ≠ 0 → ∀ (r : ℝ), x ≠ ↑r
theorem
Hyperreal.infinite_of_infinitesimal_inv
{x : ℝ*}
(h0 : x ≠ 0)
(hi : x⁻¹.Infinitesimal)
:
x.Infinite
Hyperreal.st
stuff that requires infinitesimal machinery #
Infinite stuff that requires infinitesimal machinery #
theorem
Hyperreal.infinitePos_mul_of_infinitePos_not_infinitesimal_pos
{x y : ℝ*}
:
x.InfinitePos → ¬y.Infinitesimal → 0 < y → (x * y).InfinitePos
theorem
Hyperreal.infinitePos_mul_of_not_infinitesimal_pos_infinitePos
{x y : ℝ*}
:
¬x.Infinitesimal → 0 < x → y.InfinitePos → (x * y).InfinitePos
theorem
Hyperreal.infinitePos_mul_of_infiniteNeg_not_infinitesimal_neg
{x y : ℝ*}
:
x.InfiniteNeg → ¬y.Infinitesimal → y < 0 → (x * y).InfinitePos
theorem
Hyperreal.infinitePos_mul_of_not_infinitesimal_neg_infiniteNeg
{x y : ℝ*}
:
¬x.Infinitesimal → x < 0 → y.InfiniteNeg → (x * y).InfinitePos
theorem
Hyperreal.infiniteNeg_mul_of_infinitePos_not_infinitesimal_neg
{x y : ℝ*}
:
x.InfinitePos → ¬y.Infinitesimal → y < 0 → (x * y).InfiniteNeg
theorem
Hyperreal.infiniteNeg_mul_of_not_infinitesimal_neg_infinitePos
{x y : ℝ*}
:
¬x.Infinitesimal → x < 0 → y.InfinitePos → (x * y).InfiniteNeg
theorem
Hyperreal.infiniteNeg_mul_of_infiniteNeg_not_infinitesimal_pos
{x y : ℝ*}
:
x.InfiniteNeg → ¬y.Infinitesimal → 0 < y → (x * y).InfiniteNeg
theorem
Hyperreal.infiniteNeg_mul_of_not_infinitesimal_pos_infiniteNeg
{x y : ℝ*}
:
¬x.Infinitesimal → 0 < x → y.InfiniteNeg → (x * y).InfiniteNeg
theorem
Hyperreal.infinitePos_mul_infinitePos
{x y : ℝ*}
:
x.InfinitePos → y.InfinitePos → (x * y).InfinitePos
theorem
Hyperreal.infiniteNeg_mul_infiniteNeg
{x y : ℝ*}
:
x.InfiniteNeg → y.InfiniteNeg → (x * y).InfinitePos
theorem
Hyperreal.infinitePos_mul_infiniteNeg
{x y : ℝ*}
:
x.InfinitePos → y.InfiniteNeg → (x * y).InfiniteNeg
theorem
Hyperreal.infiniteNeg_mul_infinitePos
{x y : ℝ*}
:
x.InfiniteNeg → y.InfinitePos → (x * y).InfiniteNeg
theorem
Hyperreal.infinite_mul_of_infinite_not_infinitesimal
{x y : ℝ*}
:
x.Infinite → ¬y.Infinitesimal → (x * y).Infinite
theorem
Hyperreal.infinite_mul_of_not_infinitesimal_infinite
{x y : ℝ*}
:
¬x.Infinitesimal → y.Infinite → (x * y).Infinite