Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 4921be9

Browse files
committed
feat(analysis/special_functions/arsinh): inverse hyperbolic sine function (#3801)
Added the following lemmas and definitions: `cosh_def` `sinh_def` `cosh_pos` `sinh_strict_mono` `sinh_injective` `sinh_surjective` `sinh_bijective` `real.cosh_sq_sub_sinh_sq` `sqrt_one_add_sinh_sq` `sinh_arsinh` `arsinh_sin` This is from the list of UG not in lean. `cosh` coming soon.
1 parent 7a48761 commit 4921be9

File tree

4 files changed

+110
-1
lines changed

4 files changed

+110
-1
lines changed
Lines changed: 81 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,81 @@
1+
/-
2+
Copyright (c) 2020 James Arthur. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: James Arthur, Chris Hughes, Shing Tak Lam
5+
-/
6+
import analysis.special_functions.trigonometric
7+
noncomputable theory
8+
9+
/-!
10+
# Inverse of the sinh function
11+
12+
In this file we prove that sinh is bijective and hence has an
13+
inverse, arsinh.
14+
15+
## Main Results
16+
17+
- `sinh_injective`: The proof that `sinh` is injective
18+
- `sinh_surjective`: The proof that `sinh` is surjective
19+
- `sinh_bijective`: The proof `sinh` is bijective
20+
- `arsinh`: The inverse function of `sinh`
21+
22+
## Tags
23+
24+
arsinh, arcsinh, argsinh, asinh, sinh injective, sinh bijective, sinh surjective
25+
-/
26+
27+
namespace real
28+
29+
/-- `arsinh` is defined using a logarithm, `arsinh x = log (x + sqrt(1 + x^2))`. -/
30+
@[pp_nodot] def arsinh (x : ℝ) := log (x + sqrt (1 + x^2))
31+
32+
/-- `sinh` is injective, `∀ a b, sinh a = sinh b → a = b`. -/
33+
lemma sinh_injective : function.injective sinh := sinh_strict_mono.injective
34+
35+
private lemma aux_lemma (x : ℝ) : 1 / (x + sqrt (1 + x ^ 2)) = -x + sqrt (1 + x ^ 2) :=
36+
begin
37+
refine (eq_one_div_of_mul_eq_one _).symm,
38+
have : 01 + x ^ 2 := add_nonneg zero_le_one (pow_two_nonneg x),
39+
rw [add_comm, ← sub_eq_neg_add, ← mul_self_sub_mul_self,
40+
mul_self_sqrt this, pow_two, add_sub_cancel]
41+
end
42+
43+
private lemma b_lt_sqrt_b_one_add_sq (b : ℝ) : b < sqrt (1 + b ^ 2) :=
44+
calc b
45+
≤ sqrt (b ^ 2) : le_sqrt_of_sqr_le $ le_refl _
46+
... < sqrt (1 + b ^ 2) : (sqrt_lt (pow_two_nonneg _) (add_nonneg zero_le_one (pow_two_nonneg _))).2
47+
(lt_one_add _)
48+
49+
private lemma add_sqrt_one_add_pow_two_pos (b : ℝ) : 0 < b + sqrt (1 + b ^ 2) :=
50+
by { rw [← neg_neg b, ← sub_eq_neg_add, sub_pos, pow_two, neg_mul_neg, ← pow_two],
51+
exact b_lt_sqrt_b_one_add_sq (-b) }
52+
53+
/-- `arsinh` is the right inverse of `sinh`. -/
54+
lemma sinh_arsinh (x : ℝ) : sinh (arsinh x) = x :=
55+
by rw [sinh_eq, arsinh, ← log_inv, exp_log (add_sqrt_one_add_pow_two_pos x),
56+
exp_log (inv_pos.2 (add_sqrt_one_add_pow_two_pos x)),
57+
inv_eq_one_div, aux_lemma x, sub_eq_add_neg, neg_add, neg_neg, ← sub_eq_add_neg,
58+
add_add_sub_cancel, add_self_div_two]
59+
60+
/-- `sinh` is surjective, `∀ b, ∃ a, sinh a = b`. In this case, we use `a = arsinh b`. -/
61+
lemma sinh_surjective : function.surjective sinh := function.left_inverse.surjective sinh_arsinh
62+
63+
/-- `sinh` is bijective, both injective and surjective. -/
64+
lemma sinh_bijective : function.bijective sinh :=
65+
⟨sinh_injective, sinh_surjective⟩
66+
67+
/-- A rearrangement and `sqrt` of `real.cosh_sq_sub_sinh_sq`. -/
68+
lemma sqrt_one_add_sinh_sq (x : ℝ): sqrt (1 + sinh x ^ 2) = cosh x :=
69+
begin
70+
have H := real.cosh_sq_sub_sinh_sq x,
71+
have G : cosh x ^ 2 - sinh x ^ 2 + sinh x ^ 2 = 1 + sinh x ^ 2 := by rw H,
72+
rw sub_add_cancel at G,
73+
rw [←G, sqrt_sqr],
74+
exact le_of_lt (cosh_pos x),
75+
end
76+
77+
/-- `arsinh` is the left inverse of `sinh`. -/
78+
lemma arsinh_sinh (x : ℝ) : arsinh (sinh x) = x :=
79+
function.right_inverse_of_injective_of_left_inverse sinh_injective sinh_arsinh x
80+
81+
end real

src/analysis/special_functions/exp_log.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -198,7 +198,7 @@ end
198198
to `log |x|` for `x < 0`, and to `0` for `0`. We use this unconventional extension to
199199
`(-∞, 0]` as it gives the formula `log (x * y) = log x + log y` for all nonzero `x` and `y`, and
200200
the derivative of `log` is `1/x` away from `0`. -/
201-
noncomputable def log (x : ℝ) : ℝ :=
201+
@[pp_nodot] noncomputable def log (x : ℝ) : ℝ :=
202202
if hx : x ≠ 0 then classical.some (exists_exp_eq_of_pos (abs_pos_iff.mpr hx)) else 0
203203

204204
lemma exp_log_eq_abs (hx : x ≠ 0) : exp (log x) = abs x :=

src/analysis/special_functions/trigonometric.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -348,6 +348,10 @@ funext $ λ x, (has_deriv_at_cosh x).deriv
348348
lemma continuous_cosh : continuous cosh :=
349349
differentiable_cosh.continuous
350350

351+
/-- `sinh` is strictly monotone. -/
352+
lemma sinh_strict_mono : strict_mono sinh :=
353+
strict_mono_of_deriv_pos differentiable_sinh (by { rw [real.deriv_sinh], exact cosh_pos })
354+
351355
end real
352356

353357
section

src/data/complex/exponential.lean

Lines changed: 24 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -861,6 +861,11 @@ of_real_inj.1 $ by simpa using cos_square x
861861
lemma sin_square : sin x ^ 2 = 1 - cos x ^ 2 :=
862862
eq_sub_iff_add_eq.2 $ sin_sq_add_cos_sq _
863863

864+
/-- The definition of `sinh` in terms of `exp`. -/
865+
lemma sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 :=
866+
eq_div_of_mul_eq two_ne_zero $ by rw [sinh, exp, exp, complex.of_real_neg, complex.sinh, mul_two,
867+
← complex.add_re, ← mul_two, div_mul_cancel _ (two_ne_zero' : (2 : ℂ) ≠ 0), complex.sub_re]
868+
864869
@[simp] lemma sinh_zero : sinh 0 = 0 := by simp [sinh]
865870

866871
@[simp] lemma sinh_neg : sinh (-x) = -sinh x :=
@@ -869,6 +874,11 @@ by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul]
869874
lemma sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y :=
870875
by rw ← of_real_inj; simp [sinh_add]
871876

877+
/-- The definition of `cosh` in terms of `exp`. -/
878+
lemma cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 :=
879+
eq_div_of_mul_eq two_ne_zero $ by rw [cosh, exp, exp, complex.of_real_neg, complex.cosh, mul_two,
880+
← complex.add_re, ← mul_two, div_mul_cancel _ (two_ne_zero' : (2 : ℂ) ≠ 0), complex.add_re]
881+
872882
@[simp] lemma cosh_zero : cosh 0 = 1 := by simp [cosh]
873883

874884
@[simp] lemma cosh_neg : cosh (-x) = cosh x :=
@@ -883,6 +893,16 @@ by simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg]
883893
lemma cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y :=
884894
by simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg]
885895

896+
lemma cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 :=
897+
begin
898+
rw [sinh, cosh],
899+
have := congr_arg complex.re (complex.cosh_sq_sub_sinh_sq x),
900+
rw [pow_two, pow_two] at this,
901+
change (⟨_, _⟩ : ℂ).re - (⟨_, _⟩ : ℂ).re = 1 at this,
902+
rw [complex.cosh_of_real_im x, complex.sinh_of_real_im x, mul_zero, sub_zero, sub_zero] at this,
903+
rwa [pow_two, pow_two],
904+
end
905+
886906
lemma tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x :=
887907
of_real_inj.1 $ by simp [tanh_eq_sinh_div_cosh]
888908

@@ -943,6 +963,10 @@ by rw [← exp_zero, exp_lt_exp]
943963
lemma exp_lt_one_iff {x : ℝ} : exp x < 1 ↔ x < 0 :=
944964
by rw [← exp_zero, exp_lt_exp]
945965

966+
/-- `real.cosh` is always positive -/
967+
lemma cosh_pos (x : ℝ) : 0 < real.cosh x :=
968+
(cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x)))
969+
946970
end real
947971

948972
namespace complex

0 commit comments

Comments
 (0)