|
| 1 | +/- |
| 2 | +Copyright (c) 2024 Lawrence Wu. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Lawrence Wu |
| 5 | +-/ |
| 6 | +import Mathlib.Analysis.Fourier.Inversion |
| 7 | + |
| 8 | +/-! |
| 9 | +# Mellin inversion formula |
| 10 | +
|
| 11 | +We derive the Mellin inversion formula as a consequence of the Fourier inversion formula. |
| 12 | +
|
| 13 | +## Main results |
| 14 | +- `mellin_inversion`: The inverse Mellin transform of the Mellin transform applied to `x > 0` is x. |
| 15 | +-/ |
| 16 | + |
| 17 | +open Real Complex Set MeasureTheory |
| 18 | + |
| 19 | +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℂ E] |
| 20 | + |
| 21 | +open scoped FourierTransform |
| 22 | + |
| 23 | +private theorem rexp_neg_deriv_aux : |
| 24 | + ∀ x ∈ univ, HasDerivWithinAt (rexp ∘ Neg.neg) (-rexp (-x)) univ x := |
| 25 | + fun x _ ↦ mul_neg_one (rexp (-x)) ▸ |
| 26 | + ((Real.hasDerivAt_exp (-x)).comp x (hasDerivAt_neg x)).hasDerivWithinAt |
| 27 | + |
| 28 | +private theorem rexp_neg_image_aux : rexp ∘ Neg.neg '' univ = Ioi 0 := by |
| 29 | + rw [Set.image_comp, Set.image_univ_of_surjective neg_surjective, Set.image_univ, Real.range_exp] |
| 30 | + |
| 31 | +private theorem rexp_neg_injOn_aux : univ.InjOn (rexp ∘ Neg.neg) := |
| 32 | + (Real.exp_injective.injOn _).comp (neg_injective.injOn _) (univ.mapsTo_univ _) |
| 33 | + |
| 34 | +private theorem rexp_cexp_aux (x : ℝ) (s : ℂ) (f : E) : |
| 35 | + rexp (-x) • cexp (-↑x) ^ (s - 1) • f = cexp (-s * ↑x) • f := by |
| 36 | + show (rexp (-x) : ℂ) • _ = _ • f |
| 37 | + rw [← smul_assoc, smul_eq_mul] |
| 38 | + push_cast |
| 39 | + conv in cexp _ * _ => lhs; rw [← cpow_one (cexp _)] |
| 40 | + rw [← cpow_add _ _ (Complex.exp_ne_zero _), cpow_def_of_ne_zero (Complex.exp_ne_zero _), |
| 41 | + Complex.log_exp (by norm_num; exact pi_pos) (by simpa using pi_nonneg)] |
| 42 | + ring_nf |
| 43 | + |
| 44 | +theorem mellin_eq_fourierIntegral (f : ℝ → E) {s : ℂ} : |
| 45 | + mellin f s = 𝓕 (fun (u : ℝ) ↦ (Real.exp (-s.re * u) • f (Real.exp (-u)))) (s.im / (2 * π)) := |
| 46 | + calc |
| 47 | + mellin f s |
| 48 | + = ∫ (u : ℝ), Complex.exp (-s * u) • f (Real.exp (-u)) := by |
| 49 | + rw [mellin, ← rexp_neg_image_aux, integral_image_eq_integral_abs_deriv_smul |
| 50 | + MeasurableSet.univ rexp_neg_deriv_aux rexp_neg_injOn_aux] |
| 51 | + simp [rexp_cexp_aux] |
| 52 | + _ = ∫ (u : ℝ), Complex.exp (↑(-2 * π * (u * (s.im / (2 * π)))) * I) • |
| 53 | + (Real.exp (-s.re * u) • f (Real.exp (-u))) := by |
| 54 | + congr |
| 55 | + ext u |
| 56 | + trans Complex.exp (-s.im * u * I) • (Real.exp (-s.re * u) • f (Real.exp (-u))) |
| 57 | + · conv => lhs; rw [← re_add_im s] |
| 58 | + rw [neg_add, add_mul, Complex.exp_add, mul_comm, ← smul_eq_mul, smul_assoc] |
| 59 | + norm_cast |
| 60 | + push_cast |
| 61 | + ring_nf |
| 62 | + congr |
| 63 | + rw [mul_comm (-s.im : ℂ) (u : ℂ), mul_comm (-2 * π)] |
| 64 | + have : 2 * (π : ℂ) ≠ 0 := by norm_num; exact pi_ne_zero |
| 65 | + field_simp |
| 66 | + _ = 𝓕 (fun (u : ℝ) ↦ (Real.exp (-s.re * u) • f (Real.exp (-u)))) (s.im / (2 * π)) := by |
| 67 | + simp [fourierIntegral_eq'] |
| 68 | + |
| 69 | +theorem mellinInv_eq_fourierIntegralInv (σ : ℝ) (f : ℂ → E) {x : ℝ} (hx : 0 < x) : |
| 70 | + mellinInv σ f x = |
| 71 | + (x : ℂ) ^ (-σ : ℂ) • 𝓕⁻ (fun (y : ℝ) ↦ f (σ + 2 * π * y * I)) (-Real.log x) := calc |
| 72 | + mellinInv σ f x |
| 73 | + = (x : ℂ) ^ (-σ : ℂ) • |
| 74 | + (∫ (y : ℝ), Complex.exp (2 * π * (y * (-Real.log x)) * I) • f (σ + 2 * π * y * I)) := by |
| 75 | + rw [mellinInv, one_div, ← abs_of_pos (show 0 < (2 * π)⁻¹ by norm_num; exact pi_pos)] |
| 76 | + have hx0 : (x : ℂ) ≠ 0 := ofReal_ne_zero.mpr (ne_of_gt hx) |
| 77 | + simp_rw [neg_add, cpow_add _ _ hx0, mul_smul, integral_smul] |
| 78 | + rw [smul_comm, ← Measure.integral_comp_mul_left] |
| 79 | + congr! 3 |
| 80 | + rw [cpow_def_of_ne_zero hx0, ← Complex.ofReal_log hx.le] |
| 81 | + push_cast |
| 82 | + ring_nf |
| 83 | + _ = (x : ℂ) ^ (-σ : ℂ) • 𝓕⁻ (fun (y : ℝ) ↦ f (σ + 2 * π * y * I)) (-Real.log x) := by |
| 84 | + simp [fourierIntegralInv_eq'] |
| 85 | + |
| 86 | +variable [CompleteSpace E] |
| 87 | + |
| 88 | +/-- The inverse Mellin transform of the Mellin transform applied to `x > 0` is x. -/ |
| 89 | +theorem mellin_inversion (σ : ℝ) (f : ℝ → E) {x : ℝ} (hx : 0 < x) (hf : MellinConvergent f σ) |
| 90 | + (hFf : VerticalIntegrable (mellin f) σ) (hfx : ContinuousAt f x) : |
| 91 | + mellinInv σ (mellin f) x = f x := by |
| 92 | + let g := fun (u : ℝ) => Real.exp (-σ * u) • f (Real.exp (-u)) |
| 93 | + replace hf : Integrable g := by |
| 94 | + rw [MellinConvergent, ← rexp_neg_image_aux, integrableOn_image_iff_integrableOn_abs_deriv_smul |
| 95 | + MeasurableSet.univ rexp_neg_deriv_aux rexp_neg_injOn_aux] at hf |
| 96 | + replace hf : Integrable fun (x : ℝ) ↦ cexp (-↑σ * ↑x) • f (rexp (-x)) := by |
| 97 | + simpa [rexp_cexp_aux] using hf |
| 98 | + norm_cast at hf |
| 99 | + replace hFf : Integrable (𝓕 g) := by |
| 100 | + have h2π : 2 * π ≠ 0 := by norm_num; exact pi_ne_zero |
| 101 | + have : Integrable (𝓕 (fun u ↦ rexp (-(σ * u)) • f (rexp (-u)))) := by |
| 102 | + simpa [mellin_eq_fourierIntegral, mul_div_cancel _ h2π] using hFf.comp_mul_right' h2π |
| 103 | + simp_rw [neg_mul_eq_neg_mul] at this |
| 104 | + exact this |
| 105 | + replace hfx : ContinuousAt g (-Real.log x) := by |
| 106 | + refine ContinuousAt.smul (by fun_prop) (ContinuousAt.comp ?_ (by fun_prop)) |
| 107 | + simpa [Real.exp_log hx] using hfx |
| 108 | + calc |
| 109 | + mellinInv σ (mellin f) x |
| 110 | + = mellinInv σ (fun s ↦ 𝓕 g (s.im / (2 * π))) x := by |
| 111 | + simp [g, mellinInv, mellin_eq_fourierIntegral] |
| 112 | + _ = (x : ℂ) ^ (-σ : ℂ) • g (-Real.log x) := by |
| 113 | + rw [mellinInv_eq_fourierIntegralInv _ _ hx, ← hf.fourier_inversion hFf hfx] |
| 114 | + simp [mul_div_cancel_left _ (show 2 * π ≠ 0 by norm_num; exact pi_ne_zero)] |
| 115 | + _ = (x : ℂ) ^ (-σ : ℂ) • rexp (σ * Real.log x) • f (rexp (Real.log x)) := by simp [g] |
| 116 | + _ = f x := by |
| 117 | + norm_cast |
| 118 | + rw [mul_comm σ, ← rpow_def_of_pos hx, Real.exp_log hx, ← Complex.ofReal_cpow hx.le] |
| 119 | + norm_cast |
| 120 | + rw [← smul_assoc, smul_eq_mul, Real.rpow_neg hx.le, |
| 121 | + inv_mul_cancel (ne_of_gt (rpow_pos_of_pos hx σ)), one_smul] |
0 commit comments