Sigmoid function #
In this file we define the sigmoid function x : ℝ ↦ (1 + exp (-x))⁻¹
and prove some of
its analytic properties.
We then show that the sigmoid function can be seen as an order embedding from ℝ
to I = [0, 1]
and that this embedding is both a topological embedding and a measurable embedding. We also prove
that the composition of this embedding with the measurable embedding from a standard Borel space
α
to ℝ
is a measurable embedding from α
to I
.
Main definitions and results #
Sigmoid as a function from ℝ
to ℝ
#
Real.sigmoid
: the sigmoid function fromℝ
toℝ
.Real.sigmoid_strictMono
: the sigmoid function is strictly monotone.Real.continuous_sigmoid
: the sigmoid function is continuous.Real.sigmoid_tendsto_nhds_1_atTop
: the sigmoid function tends to1
at+∞
.Real.sigmoid_tendsto_nhds_0_atBot
: the sigmoid function tends to0
at-∞
.Real.hasDerivAt_sigmoid
: the derivative of the sigmoid function.Real.analyticAt_sigmoid
: the sigmoid function is analytic at every point.
Sigmoid as a function from ℝ
to I
#
unitInterval.sigmoid
: the sigmoid function fromℝ
toI
.unitInterval.sigmoid_strictMono
: the sigmoid function is strictly monotone.unitInterval.continuous_sigmoid
: the sigmoid function is continuous.unitInterval.sigmoid_tendsto_nhds_1_atTop
: the sigmoid function tends to1
at+∞
.unitInterval.sigmoid_tendsto_nhds_0_atBot
: the sigmoid function tends to0
at-∞
.
Sigmoid as an OrderEmbedding
from ℝ
to I
#
OrderEmbedding.sigmoid
: the sigmoid function as anOrderEmbedding
fromℝ
toI
.OrderEmbedding.isEmbedding_sigmoid
: the sigmoid function fromℝ
toI
is a topological embedding.OrderEmbedding.measurableEmbedding_sigmoid
: the sigmoid function fromℝ
toI
is a measurable embedding.OrderEmbedding.measurableEmbedding_sigmoid_comp_embeddingReal
: the composition of the sigmoid function fromℝ
toI
with the measurable embedding from a standard Borel spaceα
toℝ
is a measurable embedding fromα
toI
.
Tags #
sigmoid, embedding, measurable embedding, topological embedding
theorem
AnalyticAt.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
(fa : AnalyticAt ℝ f x)
:
AnalyticAt ℝ (Real.sigmoid ∘ f) x
theorem
AnalyticAt.sigmoid'
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
(fa : AnalyticAt ℝ f x)
:
AnalyticAt ℝ (fun (z : E) => (f z).sigmoid) x
theorem
AnalyticOnNhd.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{s : Set E}
(fs : AnalyticOnNhd ℝ f s)
:
AnalyticOnNhd ℝ (Real.sigmoid ∘ f) s
theorem
AnalyticOn.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{s : Set E}
(fs : AnalyticOn ℝ f s)
:
AnalyticOn ℝ (Real.sigmoid ∘ f) s
theorem
AnalyticWithinAt.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{s : Set E}
{x : E}
(fa : AnalyticWithinAt ℝ f s x)
:
AnalyticWithinAt ℝ (Real.sigmoid ∘ f) s x
theorem
ContDiff.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
(hf : ContDiff ℝ ⊤ f)
:
ContDiff ℝ ⊤ (Real.sigmoid ∘ f)
theorem
Differentiable.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
(hf : Differentiable ℝ f)
:
Differentiable ℝ (Real.sigmoid ∘ f)
theorem
DifferentiableAt.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{f : E → ℝ}
{x : E}
(hf : DifferentiableAt ℝ f x)
:
DifferentiableAt ℝ (Real.sigmoid ∘ f) x
theorem
Continuous.sigmoid
{E : Type u_1}
[NormedAddCommGroup E]
{f : E → ℝ}
(hf : Continuous f)
:
Continuous (Real.sigmoid ∘ f)
The sigmoid function from ℝ
to I
.