@[simp]
Alias of the reverse direction of Complex.countable_preimage_exp.
@[simp]
@[simp]
theorem
Filter.Tendsto.clog
{α : Type u_1}
{l : Filter α}
{f : α → ℂ}
{x : ℂ}
(h : Tendsto f l (nhds x))
(hx : x ∈ Complex.slitPlane)
:
Tendsto (fun (t : α) => Complex.log (f t)) l (nhds (Complex.log x))
theorem
ContinuousAt.clog
{α : Type u_1}
[TopologicalSpace α]
{f : α → ℂ}
{x : α}
(h₁ : ContinuousAt f x)
(h₂ : f x ∈ Complex.slitPlane)
:
ContinuousAt (fun (t : α) => Complex.log (f t)) x
theorem
ContinuousWithinAt.clog
{α : Type u_1}
[TopologicalSpace α]
{f : α → ℂ}
{s : Set α}
{x : α}
(h₁ : ContinuousWithinAt f s x)
(h₂ : f x ∈ Complex.slitPlane)
:
ContinuousWithinAt (fun (t : α) => Complex.log (f t)) s x
theorem
ContinuousOn.clog
{α : Type u_1}
[TopologicalSpace α]
{f : α → ℂ}
{s : Set α}
(h₁ : ContinuousOn f s)
(h₂ : ∀ x ∈ s, f x ∈ Complex.slitPlane)
:
ContinuousOn (fun (t : α) => Complex.log (f t)) s
theorem
Continuous.clog
{α : Type u_1}
[TopologicalSpace α]
{f : α → ℂ}
(h₁ : Continuous f)
(h₂ : ∀ (x : α), f x ∈ Complex.slitPlane)
:
Continuous fun (t : α) => Complex.log (f t)
Complex.exp as an OpenPartialHomeomorph with source = {z | -π < im z < π} and
target = {z | 0 < re z} ∪ {z | im z ≠ 0} (a.k.a. slitPlane).
This definition is used to prove that Complex.log
is complex differentiable at all points but the negative real semi-axis.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[deprecated Complex.expOpenPartialHomeomorph (since := "2026-01-13")]
Alias of Complex.expOpenPartialHomeomorph.
Complex.exp as an OpenPartialHomeomorph with source = {z | -π < im z < π} and
target = {z | 0 < re z} ∪ {z | im z ≠ 0} (a.k.a. slitPlane).
This definition is used to prove that Complex.log
is complex differentiable at all points but the negative real semi-axis.