Topology on extended natural numbers #
Topology on ℕ∞
.
Note: this is different from the EMetricSpace
topology. The EMetricSpace
topology has
IsOpen {∞}
, but all neighborhoods of ∞
in ℕ∞
contain infinite intervals.
Equations
@[deprecated ENat.isEmbedding_natCast]
Alias of ENat.isEmbedding_natCast
.
@[deprecated ENat.isOpenEmbedding_natCast]
Alias of ENat.isOpenEmbedding_natCast
.
theorem
ENat.continuousAt_sub
{a b : ℕ∞}
(h : a ≠ ⊤ ∨ b ≠ ⊤)
:
ContinuousAt (Function.uncurry fun (x1 x2 : ℕ∞) => x1 - x2) (a, b)
theorem
Filter.Tendsto.enatSub
{α : Type u_1}
{l : Filter α}
{f g : α → ℕ∞}
{a b : ℕ∞}
(hf : Filter.Tendsto f l (nhds a))
(hg : Filter.Tendsto g l (nhds b))
(h : a ≠ ⊤ ∨ b ≠ ⊤)
:
Filter.Tendsto (fun (x : α) => f x - g x) l (nhds (a - b))
theorem
ContinuousWithinAt.enatSub
{X : Type u_1}
[TopologicalSpace X]
{f g : X → ℕ∞}
{s : Set X}
{x : X}
(hf : ContinuousWithinAt f s x)
(hg : ContinuousWithinAt g s x)
(h : f x ≠ ⊤ ∨ g x ≠ ⊤)
:
ContinuousWithinAt (fun (x : X) => f x - g x) s x
theorem
ContinuousAt.enatSub
{X : Type u_1}
[TopologicalSpace X]
{f g : X → ℕ∞}
{x : X}
(hf : ContinuousAt f x)
(hg : ContinuousAt g x)
(h : f x ≠ ⊤ ∨ g x ≠ ⊤)
:
ContinuousAt (fun (x : X) => f x - g x) x
theorem
ContinuousOn.enatSub
{X : Type u_1}
[TopologicalSpace X]
{f g : X → ℕ∞}
{s : Set X}
(hf : ContinuousOn f s)
(hg : ContinuousOn g s)
(h : ∀ x ∈ s, f x ≠ ⊤ ∨ g x ≠ ⊤)
:
ContinuousOn (fun (x : X) => f x - g x) s
theorem
Continuous.enatSub
{X : Type u_1}
[TopologicalSpace X]
{f g : X → ℕ∞}
(hf : Continuous f)
(hg : Continuous g)
(h : ∀ (x : X), f x ≠ ⊤ ∨ g x ≠ ⊤)
:
Continuous fun (x : X) => f x - g x