Documentation

PrimeNumberTheoremAnd.MellinCalculus

theorem MeasureTheory.setIntegral_integral_swap {α : Type u_1} {β : Type u_2} {E : Type u_3} [MeasurableSpace α] [MeasurableSpace β] {μ : Measure α} {ν : Measure β} [NormedAddCommGroup E] [SigmaFinite ν] [NormedSpace E] [SigmaFinite μ] (f : αβE) {s : Set α} {t : Set β} (hf : IntegrableOn (Function.uncurry f) (s ×ˢ t) (μ.prod ν)) :
(x : α) in s, (y : β) in t, f x y ν μ = (y : β) in t, (x : α) in s, f x y μ ν
theorem MeasureTheory.integral_comp_mul_right_I0i_haar {𝕂 : Type u_1} [RCLike 𝕂] (f : 𝕂) {a : } (ha : 0 < a) :
(y : ) in Set.Ioi 0, f (y * a) / y = (y : ) in Set.Ioi 0, f y / y
theorem MeasureTheory.integral_comp_mul_right_I0i_haar_real (f : ) {a : } (ha : 0 < a) :
(y : ) in Set.Ioi 0, f (y * a) / y = (y : ) in Set.Ioi 0, f y / y
theorem MeasureTheory.integral_comp_mul_left_I0i_haar {𝕂 : Type u_1} [RCLike 𝕂] (f : 𝕂) {a : } (ha : 0 < a) :
(y : ) in Set.Ioi 0, f (a * y) / y = (y : ) in Set.Ioi 0, f y / y
theorem MeasureTheory.integral_comp_rpow_I0i_haar_real (f : ) {p : } (hp : p 0) :
(y : ) in Set.Ioi 0, |p| * f (y ^ p) / y = (y : ) in Set.Ioi 0, f y / y
theorem MeasureTheory.integral_comp_inv_I0i_haar {𝕂 : Type u_1} [RCLike 𝕂] (f : 𝕂) :
(y : ) in Set.Ioi 0, f (1 / y) / y = (y : ) in Set.Ioi 0, f y / y
theorem MeasureTheory.integral_comp_div_I0i_haar {𝕂 : Type u_1} [RCLike 𝕂] (f : 𝕂) {a : } (ha : 0 < a) :
(y : ) in Set.Ioi 0, f (a / y) / y = (y : ) in Set.Ioi 0, f y / y
theorem Complex.ofReal_rpow {x : } (h : x > 0) (y : ) :
↑(x ^ y) = x ^ y
@[simp]
theorem Function.support_abs {𝕂 : Type u_1} [RCLike 𝕂] {α : Type u_2} (f : α𝕂) :
(support fun (x : α) => f x) = support f
@[simp]
theorem Function.support_ofReal {f : } :
(support fun (x : ) => (f x)) = support f
theorem Function.support_id :
(support fun (x : ) => x) = Set.Iio 0 Set.Ioi 0
theorem Function.support_mul_subset_of_subset {𝕂 : Type u_1} [RCLike 𝕂] {s : Set } {f g : 𝕂} (fSupp : support f s) :
support (f * g) s
theorem Function.support_of_along_fiber_subset_subset {α : Type u_2} {β : Type u_3} {M : Type u_4} [Zero M] {f : α × βM} {s : Set α} {t : Set β} (hx : ∀ (y : β), (support fun (x : α) => f (x, y)) s) (hy : ∀ (x : α), (support fun (y : β) => f (x, y)) t) :
theorem Function.support_deriv_subset_Icc {𝕂 : Type u_1} [RCLike 𝕂] {a b : } {f : 𝕂} (fSupp : support f Set.Icc a b) :
theorem SetIntegral.integral_eq_integral_inter_of_support_subset_Icc {a b : } {μ : MeasureTheory.Measure } {E : Type u_2} [NormedAddCommGroup E] [NormedSpace E] {s : Set } {f : E} (h : Function.support f Set.Icc a b) (hs : Set.Icc a b s) :
(x : ) in s, f x μ = (x : ) in Set.Icc a b, f x μ
theorem intervalIntegral.norm_integral_le_of_norm_le_const' {a b C : } {E : Type u_2} [NormedAddCommGroup E] [NormedSpace E] {f : E} (hab : a b) (h : xSet.Icc a b, f x C) :
(x : ) in a..b, f x C * |b - a|
theorem Filter.TendstoAtZero_of_support_in_Icc {𝕂 : Type u_1} [RCLike 𝕂] {a b : } (f : 𝕂) (ha : 0 < a) (fSupp : Function.support f Set.Icc a b) :
theorem Filter.TendstoAtTop_of_support_in_Icc {𝕂 : Type u_1} [RCLike 𝕂] {a b : } (f : 𝕂) (fSupp : Function.support f Set.Icc a b) :
theorem Filter.BigO_zero_atZero_of_support_in_Icc {𝕂 : Type u_1} [RCLike 𝕂] {a b : } (f : 𝕂) (ha : 0 < a) (fSupp : Function.support f Set.Icc a b) :
f =O[nhdsWithin 0 (Set.Ioi 0)] fun (x : ) => 0
theorem Filter.BigO_zero_atTop_of_support_in_Icc {𝕂 : Type u_1} [RCLike 𝕂] {a b : } (f : 𝕂) (fSupp : Function.support f Set.Icc a b) :
f =O[atTop] fun (x : ) => 0
theorem deriv.ofReal_comp' {f : } :
(deriv fun (x : ) => (f x)) = fun (x : ) => (deriv f x)
theorem deriv.comp_ofReal' {e : } (hf : Differentiable e) :
(deriv fun (x : ) => e x) = fun (x : ) => deriv e x
theorem PartialIntegration (f g : ) (fDiff : DifferentiableOn f (Set.Ioi 0)) (gDiff : DifferentiableOn g (Set.Ioi 0)) (fDerivgInt : MeasureTheory.IntegrableOn (f * deriv g) (Set.Ioi 0) MeasureTheory.volume) (gDerivfInt : MeasureTheory.IntegrableOn (deriv f * g) (Set.Ioi 0) MeasureTheory.volume) (lim_at_zero : Filter.Tendsto (f * g) (nhdsWithin 0 (Set.Ioi 0)) (nhds 0)) (lim_at_inf : Filter.Tendsto (f * g) Filter.atTop (nhds 0)) :
(x : ) in Set.Ioi 0, f x * deriv g x = - (x : ) in Set.Ioi 0, deriv f x * g x

Need differentiability, and decay at 0 and

theorem PartialIntegration_of_support_in_Icc {a b : } (f g : ) (ha : 0 < a) (h : a b) (fSupp : Function.support f Set.Icc a b) (fDiff : DifferentiableOn f (Set.Ioi 0)) (gDiff : DifferentiableOn g (Set.Ioi 0)) (fderivCont : ContinuousOn (deriv f) (Set.Ioi 0)) (gderivCont : ContinuousOn (deriv g) (Set.Ioi 0)) :
(x : ) in Set.Ioi 0, f x * deriv g x = - (x : ) in Set.Ioi 0, deriv f x * g x
noncomputable def MellinTransform (f : ) (s : ) :
Equations
Instances For
    noncomputable def MellinInverseTransform (F : ) (σ x : ) :
    Equations
    Instances For
      theorem PerronInverseMellin_lt {t x : } (tpos : 0 < t) (t_lt_x : t < x) {σ : } (σpos : 0 < σ) :
      theorem PerronInverseMellin_gt {t x : } (xpos : 0 < x) (x_lt_t : x < t) {σ : } (σpos : 0 < σ) :
      MellinInverseTransform (Perron.f t) σ x = 1 - x / t
      theorem MellinInversion (σ : ) {f : } {x : } (hx : 0 < x) (hf : MellinConvergent f σ) (hFf : Complex.VerticalIntegrable (mellin f) σ MeasureTheory.volume) (hfx : ContinuousAt f x) :
      noncomputable def MellinConvolution {𝕂 : Type u_1} [RCLike 𝕂] (f g : 𝕂) (x : ) :
      𝕂
      Equations
      Instances For
        theorem MellinConvolutionSymmetric {𝕂 : Type u_1} [RCLike 𝕂] (f g : 𝕂) {x : } (xpos : 0 < x) :
        theorem support_MellinConvolution_subsets {𝕂 : Type u_1} [RCLike 𝕂] {f g : 𝕂} {A B : Set } (hf : Function.support f A) (hg : Function.support g B) :
        theorem SmoothExistence :
        ∃ (ν : ), ContDiff (↑) ν (∀ (x : ), 0 ν x) Function.support ν Set.Icc (1 / 2) 2 (x : ) in Set.Ici 0, ν x / x = 1
        theorem mem_within_strip (σ₁ σ₂ : ) :
        {s : | σ₁ s.re s.re σ₂} Filter.principal {s : | σ₁ s.re s.re σ₂}
        theorem MellinOfPsi_aux {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) {s : } (hs : s 0) :
        (x : ) in Set.Ioi 0, (ν x) * x ^ (s - 1) = -(1 / s) * (x : ) in Set.Ioi 0, (deriv ν x) * x ^ s
        theorem MellinOfPsi {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) :
        C > 0, ∀ (σ₁ : ), 0 < σ₁∀ (s : ), σ₁ s.res.re 2MellinTransform (fun (x : ) => (ν x)) s C * s⁻¹
        noncomputable def DeltaSpike (ν : ) (ε : ) :
        Equations
        Instances For
          theorem DeltaSpikeMass {ν : } (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) {ε : } (εpos : 0 < ε) :
          (x : ) in Set.Ioi 0, DeltaSpike ν ε x / x = 1
          theorem DeltaSpikeSupport_aux {ν : } {ε : } (εpos : 0 < ε) (suppν : Function.support ν Set.Icc (1 / 2) 2) :
          (Function.support fun (x : ) => if x < 0 then 0 else DeltaSpike ν ε x) Set.Icc (2 ^ (-ε)) (2 ^ ε)
          theorem DeltaSpikeSupport' {ν : } {ε x : } (εpos : 0 < ε) (xnonneg : 0 x) (suppν : Function.support ν Set.Icc (1 / 2) 2) :
          DeltaSpike ν ε x 0x Set.Icc (2 ^ (-ε)) (2 ^ ε)
          theorem DeltaSpikeSupport {ν : } {ε x : } (εpos : 0 < ε) (xnonneg : 0 x) (suppν : Function.support ν Set.Icc (1 / 2) 2) :
          xSet.Icc (2 ^ (-ε)) (2 ^ ε)DeltaSpike ν ε x = 0
          theorem DeltaSpikeContinuous {ν : } {ε : } (εpos : 0 < ε) (diffν : ContDiff 1 ν) :
          Continuous fun (x : ) => DeltaSpike ν ε x
          theorem DeltaSpikeOfRealContinuous {ν : } {ε : } (εpos : 0 < ε) (diffν : ContDiff 1 ν) :
          Continuous fun (x : ) => (DeltaSpike ν ε x)
          theorem MellinOfDeltaSpike (ν : ) {ε : } (εpos : ε > 0) (s : ) :
          MellinTransform (fun (x : ) => (DeltaSpike ν ε x)) s = MellinTransform (fun (x : ) => (ν x)) (ε * s)
          theorem MellinOfDeltaSpikeAt1 (ν : ) {ε : } (εpos : ε > 0) :
          MellinTransform (fun (x : ) => (DeltaSpike ν ε x)) 1 = MellinTransform (fun (x : ) => (ν x)) ε
          theorem MellinOfDeltaSpikeAt1_asymp {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) :
          (fun (ε : ) => MellinTransform (fun (x : ) => (ν x)) ε - 1) =O[nhdsWithin 0 (Set.Ioi 0)] id
          theorem MellinOf1 (s : ) (h : s.re > 0) :
          MellinTransform (fun (x : ) => if 0 < x x 1 then 1 else 0) s = 1 / s
          noncomputable def Smooth1 (ν : ) (ε : ) :
          Equations
          Instances For
            theorem Smooth1_def_ite {ν : } {ε x : } (xpos : 0 < x) :
            Smooth1 ν ε x = MellinConvolution (fun (x : ) => if 0 < x x 1 then 1 else 0) (fun (x : ) => if x < 0 then 0 else DeltaSpike ν ε x) x
            theorem Smooth1Properties_estimate {ε : } (εpos : 0 < ε) :
            (1 - 2 ^ (-ε)) / ε < Real.log 2
            theorem Smooth1Properties_below_aux {x ε : } (hx : x 1 - Real.log 2 * ε) (εpos : 0 < ε) :
            x < 2 ^ (-ε)
            theorem Smooth1Properties_below {ν : } (suppν : Function.support ν Set.Icc (1 / 2) 2) (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) :
            ∃ (c : ), 0 < c c = Real.log 2 ∀ (ε x : ), 0 < ε0 < xx 1 - c * εSmooth1 ν ε x = 1
            theorem Smooth1Properties_above_aux {x ε : } (hx : 1 + 2 * Real.log 2 * ε x) ( : ε Set.Ioo 0 1) :
            2 ^ ε < x
            theorem Smooth1Properties_above_aux2 {x y ε : } ( : ε Set.Ioo 0 1) (hy : y Set.Ioc 0 1) (hx2 : 2 ^ ε < x) :
            2 < (x / y) ^ (1 / ε)
            theorem Smooth1Properties_above {ν : } (suppν : Function.support ν Set.Icc (1 / 2) 2) :
            ∃ (c : ), 0 < c c = 2 * Real.log 2 ∀ (ε x : ), ε Set.Ioo 0 11 + c * ε xSmooth1 ν ε x = 0
            theorem DeltaSpikeNonNeg_of_NonNeg {ν : } (νnonneg : x > 0, 0 ν x) {x ε : } (xpos : 0 < x) (εpos : 0 < ε) :
            0 DeltaSpike ν ε x
            theorem MellinConvNonNeg_of_NonNeg {f g : } (f_nonneg : x > 0, 0 f x) (g_nonneg : x > 0, 0 g x) {x : } (xpos : 0 < x) :
            theorem Smooth1Nonneg {ν : } (νnonneg : x > 0, 0 ν x) {ε x : } (xpos : 0 < x) (εpos : 0 < ε) :
            0 Smooth1 ν ε x
            theorem Smooth1LeOne_aux {x ε : } {ν : } (xpos : 0 < x) (εpos : 0 < ε) (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) :
            (y : ) in Set.Ioi 0, ν ((x / y) ^ (1 / ε)) / ε / y = 1
            theorem Smooth1LeOne {ν : } (νnonneg : x > 0, 0 ν x) (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) {ε : } (εpos : 0 < ε) {x : } (xpos : 0 < x) :
            Smooth1 ν ε x 1
            theorem MellinOfSmooth1a {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) {ε : } (εpos : 0 < ε) {s : } (hs : 0 < s.re) :
            MellinTransform (fun (x : ) => (Smooth1 ν ε x)) s = s⁻¹ * MellinTransform (fun (x : ) => (ν x)) (ε * s)
            theorem MellinOfSmooth1b {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) :
            ∃ (C : ) (_ : 0 < C), ∀ (σ₁ : ), 0 < σ₁∀ (s : ), σ₁ s.res.re 2∀ (ε : ), 0 < εε < 1MellinTransform (fun (x : ) => (Smooth1 ν ε x)) s C * (ε * s ^ 2)⁻¹
            theorem MellinOfSmooth1c {ν : } (diffν : ContDiff 1 ν) (suppν : Function.support ν Set.Icc (1 / 2) 2) (mass_one : (x : ) in Set.Ioi 0, ν x / x = 1) :
            (fun (ε : ) => MellinTransform (fun (x : ) => (Smooth1 ν ε x)) 1 - 1) =O[nhdsWithin 0 (Set.Ioi 0)] id
            theorem Smooth1ContinuousAt {SmoothingF : } (diffSmoothingF : ContDiff 1 SmoothingF) (SmoothingFpos : x > 0, 0 SmoothingF x) (suppSmoothingF : Function.support SmoothingF Set.Icc (1 / 2) 2) {ε : } (εpos : 0 < ε) {y : } (ypos : 0 < y) :
            ContinuousAt (fun (x : ) => Smooth1 SmoothingF ε x) y
            theorem Smooth1MellinConvergent {Ψ : } {ε : } (diffΨ : ContDiff 1 Ψ) (suppΨ : Function.support Ψ Set.Icc (1 / 2) 2) ( : ε Set.Ioo 0 1) (Ψnonneg : x > 0, 0 Ψ x) (mass_one : (x : ) in Set.Ioi 0, Ψ x / x = 1) {s : } (hs : 0 < s.re) :
            MellinConvergent (fun (x : ) => (Smooth1 Ψ ε x)) s
            theorem Smooth1MellinDifferentiable {Ψ : } {ε : } (diffΨ : ContDiff 1 Ψ) (suppΨ : Function.support Ψ Set.Icc (1 / 2) 2) ( : ε Set.Ioo 0 1) (Ψnonneg : x > 0, 0 Ψ x) (mass_one : (x : ) in Set.Ioi 0, Ψ x / x = 1) {s : } (hs : 0 < s.re) :
            DifferentiableAt (MellinTransform fun (x : ) => (Smooth1 Ψ ε x)) s