|
| 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 : 0 ≤ 1 + 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 |
0 commit comments