Documentation

Mathlib.Analysis.SpecialFunctions.Artanh

Inverse of the tanh function #

In this file we define an inverse of tanh as a function from ℝ to (-1, 1).

Main definitions #

Main Results #

Tags #

artanh, arctanh, argtanh, atanh

def Real.artanh (x : ) :

artanh is defined using a logarithm, artanh x = log √((1 + x) / (1 - x)).

Equations
Instances For
    theorem Real.artanh_eq_half_log {x : } (hx : x Set.Icc (-1) 1) :
    artanh x = 1 / 2 * log ((1 + x) / (1 - x))
    theorem Real.exp_artanh {x : } (hx : x Set.Ioo (-1) 1) :
    exp (artanh x) = ((1 + x) / (1 - x))
    @[simp]
    theorem Real.sinh_artanh {x : } (hx : x Set.Ioo (-1) 1) :
    sinh (artanh x) = x / (1 - x ^ 2)
    theorem Real.cosh_artanh {x : } (hx : x Set.Ioo (-1) 1) :
    cosh (artanh x) = 1 / (1 - x ^ 2)
    theorem Real.tanh_artanh {x : } (hx : x Set.Ioo (-1) 1) :
    tanh (artanh x) = x

    artanh is the right inverse of tanh over (-1, 1).

    theorem Real.artanh_tanh (x : ) :
    artanh (tanh x) = x

    artanh is the left inverse of tanh.

    theorem Real.strictMonoOn_one_add_div_one_sub :
    StrictMonoOn (fun (x : ) => (1 + x) / (1 - x)) (Set.Ioo (-1) 1)
    theorem Real.artanh_le_artanh_iff {x y : } (hx : x Set.Ioo (-1) 1) (hy : y Set.Ioo (-1) 1) :
    theorem Real.artanh_lt_artanh_iff {x y : } (hx : x Set.Ioo (-1) 1) (hy : y Set.Ioo (-1) 1) :
    artanh x < artanh y x < y
    theorem Real.artanh_le_artanh {x y : } (hx : -1 < x) (hy : y < 1) (hxy : x y) :
    theorem Real.artanh_lt_artanh {x y : } (hx : -1 < x) (hy : y < 1) (hxy : x < y) :
    theorem Real.artanh_eq_zero_iff {x : } :
    artanh x = 0 x -1 x = 0 1 x
    theorem Real.artanh_pos {x : } (hx : x Set.Ioo 0 1) :
    0 < artanh x
    theorem Real.artanh_neg {x : } (hx : x Set.Ioo (-1) 0) :
    artanh x < 0
    theorem Real.artanh_nonneg {x : } (hx : 0 x) :
    theorem Real.artanh_nonpos {x : } (hx : x 0) :

    Real.tanh as a PartialEquiv.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For