Documentation

LeanAPAP.Prereqs.Discrete.Convolution.Basic

Convolution #

This file defines several versions of the discrete convolution of functions.

Main declarations #

Notation #

Notes #

Some lemmas could technically be generalised to a non-commutative semiring domain. Doesn't seem very useful given that the codomain in applications is either , ℝ≥0 or .

Similarly we could drop the commutativity assumption on the domain, but this is unneeded at this point in time.

TODO #

Multiplicativise? Probably ugly and not very useful.

Convolution of functions #

In this section, we define the convolution f ∗ g and difference convolution f ○ g of functions f g : α → β, and show how they interact.

def conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) :
αβ

Convolution

Equations
Instances For
    def dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
    αβ

    Difference convolution

    Equations
    Instances For
      def trivChar {α : Type u_1} {β : Type u_2} [DecidableEq α] [AddCommGroup α] [CommSemiring β] :
      αβ

      The trivial character.

      Equations
      Instances For
        theorem conv_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = (Finset.filter (fun (x : α × α) => x.1 + x.2 = a) Finset.univ).sum fun (x : α × α) => f x.1 * g x.2
        theorem dconv_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = (Finset.filter (fun (x : α × α) => x.1 - x.2 = a) Finset.univ).sum fun (x : α × α) => f x.1 * (starRingEnd (αβ)) g x.2
        @[simp]
        theorem trivChar_apply {α : Type u_1} {β : Type u_2} [DecidableEq α] [AddCommGroup α] [CommSemiring β] (a : α) :
        trivChar a = if a = 0 then 1 else 0
        @[simp]
        theorem conv_conjneg {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        f conjneg g = f g
        @[simp]
        theorem dconv_conjneg {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        f conjneg g = f g
        @[simp]
        theorem translate_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (a : α) (f : αβ) (g : αβ) :
        τ a f g = τ a (f g)
        @[simp]
        theorem translate_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (a : α) (f : αβ) (g : αβ) :
        τ a f g = τ a (f g)
        @[simp]
        theorem conv_translate {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (a : α) (f : αβ) (g : αβ) :
        f τ a g = τ a (f g)
        @[simp]
        theorem dconv_translate {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (a : α) (f : αβ) (g : αβ) :
        f τ a g = τ (-a) (f g)
        theorem conv_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) :
        f g = g f
        @[simp]
        theorem conj_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        (starRingEnd (αβ)) (f g) = (starRingEnd (αβ)) f (starRingEnd (αβ)) g
        @[simp]
        theorem conj_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        (starRingEnd (αβ)) (f g) = (starRingEnd (αβ)) f (starRingEnd (αβ)) g
        @[simp]
        theorem conj_trivChar {α : Type u_1} {β : Type u_2} [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] :
        (starRingEnd (αβ)) trivChar = trivChar
        theorem IsSelfAdjoint.conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] {f : αβ} {g : αβ} (hf : IsSelfAdjoint f) (hg : IsSelfAdjoint g) :
        theorem IsSelfAdjoint.dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] {f : αβ} {g : αβ} (hf : IsSelfAdjoint f) (hg : IsSelfAdjoint g) :
        @[simp]
        theorem isSelfAdjoint_trivChar {α : Type u_1} {β : Type u_2} [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] :
        IsSelfAdjoint trivChar
        @[simp]
        theorem conjneg_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        @[simp]
        theorem conjneg_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        conjneg (f g) = g f
        @[simp]
        theorem conjneg_trivChar {α : Type u_1} {β : Type u_2} [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] :
        conjneg trivChar = trivChar
        @[simp]
        theorem conv_zero {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
        f 0 = 0
        @[simp]
        theorem zero_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
        0 f = 0
        @[simp]
        theorem dconv_zero {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) :
        f 0 = 0
        @[simp]
        theorem zero_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) :
        0 f = 0
        theorem conv_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        f (g + h) = f g + f h
        theorem add_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        (f + g) h = f h + g h
        theorem dconv_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) :
        f (g + h) = f g + f h
        theorem add_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) :
        (f + g) h = f h + g h
        theorem smul_conv {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [DistribSMul γ β] [IsScalarTower γ β β] (c : γ) (f : αβ) (g : αβ) :
        c f g = c (f g)
        theorem smul_dconv {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] [DistribSMul γ β] [IsScalarTower γ β β] (c : γ) (f : αβ) (g : αβ) :
        c f g = c (f g)
        theorem conv_smul {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [DistribSMul γ β] [SMulCommClass γ β β] (c : γ) (f : αβ) (g : αβ) :
        f c g = c (f g)
        theorem dconv_smul {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] [Star γ] [DistribSMul γ β] [SMulCommClass γ β β] [StarModule γ β] (c : γ) (f : αβ) (g : αβ) :
        f c g = star c (f g)
        theorem smul_conv_assoc {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [DistribSMul γ β] [IsScalarTower γ β β] (c : γ) (f : αβ) (g : αβ) :
        c f g = c (f g)

        Alias of smul_conv.

        theorem smul_dconv_assoc {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] [DistribSMul γ β] [IsScalarTower γ β β] (c : γ) (f : αβ) (g : αβ) :
        c f g = c (f g)

        Alias of smul_dconv.

        theorem smul_conv_left_comm {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [DistribSMul γ β] [SMulCommClass γ β β] (c : γ) (f : αβ) (g : αβ) :
        f c g = c (f g)

        Alias of conv_smul.

        theorem smul_dconv_left_comm {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] [Star γ] [DistribSMul γ β] [SMulCommClass γ β β] [StarModule γ β] (c : γ) (f : αβ) (g : αβ) :
        f c g = star c (f g)

        Alias of dconv_smul.

        theorem mul_smul_conv_comm {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [Monoid γ] [DistribMulAction γ β] [IsScalarTower γ β β] [SMulCommClass γ β β] (c : γ) (d : γ) (f : αβ) (g : αβ) :
        (c * d) (f g) = c f d g
        theorem conv_assoc {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        f g h = f (g h)
        theorem conv_right_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        f g h = f h g
        theorem conv_left_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        f (g h) = g (f h)
        theorem conv_conv_conv_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) (i : αβ) :
        f g (h i) = f h (g i)
        theorem conv_dconv_conv_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) (i : αβ) :
        f g (h i) = f h (g i)
        theorem dconv_conv_dconv_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) (i : αβ) :
        f g (h i) = f h (g i)
        theorem dconv_dconv_dconv_comm {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) (i : αβ) :
        f g (h i) = f h (g i)
        theorem map_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] {γ : Type u_4} [CommSemiring γ] [StarRing γ] (m : β →+* γ) (f : αβ) (g : αβ) (a : α) :
        m ((f g) a) = (m f m g) a
        theorem comp_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] {γ : Type u_4} [CommSemiring γ] [StarRing γ] (m : β →+* γ) (f : αβ) (g : αβ) :
        m (f g) = m f m g
        theorem map_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (g : αNNReal) (a : α) :
        ((f g) a) = (NNReal.toReal f NNReal.toReal g) a
        theorem conv_eq_sum_sub {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (a - t) * g t
        theorem dconv_eq_sum_sub {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (a - t) * (starRingEnd β) (g (-t))
        theorem conv_eq_sum_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (a + t) * g (-t)
        theorem dconv_eq_sum_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (a + t) * (starRingEnd β) (g t)
        theorem conv_eq_sum_sub' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f t * g (a - t)
        theorem dconv_eq_sum_sub' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f t * (starRingEnd β) (g (t - a))
        theorem conv_eq_sum_add' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (-t) * g (a + t)
        theorem dconv_eq_sum_add' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) a = Finset.univ.sum fun (t : α) => f (-t) * (starRingEnd (αβ)) g (-(a + t))
        theorem conv_apply_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (a : α) (b : α) :
        (f g) (a + b) = Finset.univ.sum fun (t : α) => f (a + t) * g (b - t)
        theorem dconv_apply_neg {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) :
        (f g) (-a) = (starRingEnd β) ((g f) a)
        theorem dconv_apply_sub {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (a : α) (b : α) :
        (f g) (a - b) = Finset.univ.sum fun (t : α) => f (a + t) * (starRingEnd β) (g (b + t))
        theorem sum_conv_mul {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (h : αβ) :
        (Finset.univ.sum fun (a : α) => (f g) a * h a) = Finset.univ.sum fun (a : α) => Finset.univ.sum fun (b : α) => f a * g b * h (a + b)
        theorem sum_dconv_mul {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) :
        (Finset.univ.sum fun (a : α) => (f g) a * h a) = Finset.univ.sum fun (a : α) => Finset.univ.sum fun (b : α) => f a * (starRingEnd β) (g b) * h (a - b)
        theorem sum_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) :
        (Finset.univ.sum fun (a : α) => (f g) a) = (Finset.univ.sum fun (a : α) => f a) * Finset.univ.sum fun (a : α) => g a
        theorem sum_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        (Finset.univ.sum fun (a : α) => (f g) a) = (Finset.univ.sum fun (a : α) => f a) * Finset.univ.sum fun (a : α) => (starRingEnd β) (g a)
        @[simp]
        theorem conv_const {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (b : β) :
        f Function.const α b = Function.const α ((Finset.univ.sum fun (x : α) => f x) * b)
        @[simp]
        theorem const_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (b : β) (f : αβ) :
        Function.const α b f = Function.const α (b * Finset.univ.sum fun (x : α) => f x)
        @[simp]
        theorem dconv_const {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (b : β) :
        f Function.const α b = Function.const α ((Finset.univ.sum fun (x : α) => f x) * (starRingEnd β) b)
        @[simp]
        theorem const_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (b : β) (f : αβ) :
        Function.const α b f = Function.const α (b * Finset.univ.sum fun (x : α) => (starRingEnd β) (f x))
        @[simp]
        theorem conv_trivChar {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
        f trivChar = f
        @[simp]
        theorem trivChar_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
        trivChar f = f
        @[simp]
        theorem dconv_trivChar {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) :
        f trivChar = f
        @[simp]
        theorem trivChar_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) :
        trivChar f = conjneg f
        theorem support_conv_subset {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) :
        theorem support_dconv_subset {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) :
        theorem indicate_conv_indicate_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (s : Finset α) (t : Finset α) (a : α) :
        (𝟭 s 𝟭 t) a = (Finset.filter (fun (x : α × α) => x.1 + x.2 = a) (s ×ˢ t)).card
        theorem indicate_dconv_indicate_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (s : Finset α) (t : Finset α) (a : α) :
        (𝟭 s 𝟭 t) a = (Finset.filter (fun (x : α × α) => x.1 - x.2 = a) (s ×ˢ t)).card
        theorem indicate_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (s : Finset α) (f : αβ) :
        𝟭 s f = s.sum fun (a : α) => τ a f
        theorem conv_indicate {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (s : Finset α) :
        f 𝟭 s = s.sum fun (a : α) => τ a f
        theorem indicate_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (s : Finset α) (f : αβ) :
        𝟭 s f = s.sum fun (a : α) => τ a (conjneg f)
        theorem dconv_indicate {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (s : Finset α) :
        f 𝟭 s = s.sum fun (a : α) => τ (-a) f
        @[simp]
        theorem conv_neg {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] (f : αβ) (g : αβ) :
        f -g = -(f g)
        @[simp]
        theorem neg_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] (f : αβ) (g : αβ) :
        -f g = -(f g)
        @[simp]
        theorem dconv_neg {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] [StarRing β] (f : αβ) (g : αβ) :
        f -g = -(f g)
        @[simp]
        theorem neg_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] [StarRing β] (f : αβ) (g : αβ) :
        -f g = -(f g)
        theorem conv_sub {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] (f : αβ) (g : αβ) (h : αβ) :
        f (g - h) = f g - f h
        theorem sub_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] (f : αβ) (g : αβ) (h : αβ) :
        (f - g) h = f h - g h
        theorem dconv_sub {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) :
        f (g - h) = f g - f h
        theorem sub_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommRing β] [StarRing β] (f : αβ) (g : αβ) (h : αβ) :
        (f - g) h = f h - g h
        @[simp]
        theorem mu_univ_conv_mu_univ {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] :
        μ Finset.univ μ Finset.univ = μ Finset.univ
        @[simp]
        theorem mu_univ_dconv_mu_univ {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [StarRing β] :
        μ Finset.univ μ Finset.univ = μ Finset.univ
        theorem expect_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] (f : αβ) (g : αβ) :
        (Finset.univ.expect fun (a : α) => (f g) a) = (Finset.univ.sum fun (a : α) => f a) * Finset.univ.expect fun (a : α) => g a
        theorem expect_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [StarRing β] (f : αβ) (g : αβ) :
        (Finset.univ.expect fun (a : α) => (f g) a) = (Finset.univ.sum fun (a : α) => f a) * Finset.univ.expect fun (a : α) => (starRingEnd β) (g a)
        theorem expect_conv' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] (f : αβ) (g : αβ) :
        (Finset.univ.expect fun (a : α) => (f g) a) = (Finset.univ.expect fun (a : α) => f a) * Finset.univ.sum fun (a : α) => g a
        theorem expect_dconv' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [StarRing β] (f : αβ) (g : αβ) :
        (Finset.univ.expect fun (a : α) => (f g) a) = (Finset.univ.expect fun (a : α) => f a) * Finset.univ.sum fun (a : α) => (starRingEnd β) (g a)
        theorem mu_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] (s : Finset α) (f : αβ) :
        μ s f = (s.card)⁻¹ s.sum fun (a : α) => τ a f
        theorem conv_mu {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] (f : αβ) (s : Finset α) :
        f μ s = (s.card)⁻¹ s.sum fun (a : α) => τ a f
        theorem mu_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [StarRing β] (s : Finset α) (f : αβ) :
        μ s f = (s.card)⁻¹ s.sum fun (a : α) => τ a (conjneg f)
        theorem dconv_mu {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [StarRing β] (f : αβ) (s : Finset α) :
        f μ s = (s.card)⁻¹ s.sum fun (a : α) => τ (-a) f
        @[simp]
        theorem balance_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Field β] [CharZero β] (f : αβ) (g : αβ) :
        @[simp]
        theorem balance_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Field β] [StarRing β] [CharZero β] (f : αβ) (g : αβ) :
        @[simp]
        theorem RCLike.coe_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] {𝕜 : Type} [RCLike 𝕜] (f : α) (g : α) (a : α) :
        ((f g) a) = (RCLike.ofReal f RCLike.ofReal g) a
        @[simp]
        theorem RCLike.coe_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] {𝕜 : Type} [RCLike 𝕜] (f : α) (g : α) (a : α) :
        ((f g) a) = (RCLike.ofReal f RCLike.ofReal g) a
        @[simp]
        theorem RCLike.coe_comp_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] {𝕜 : Type} [RCLike 𝕜] (f : α) (g : α) :
        RCLike.ofReal (f g) = RCLike.ofReal f RCLike.ofReal g
        @[simp]
        theorem RCLike.coe_comp_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] {𝕜 : Type} [RCLike 𝕜] (f : α) (g : α) :
        RCLike.ofReal (f g) = RCLike.ofReal f RCLike.ofReal g
        @[simp]
        theorem Complex.coe_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : α) (g : α) (a : α) :
        @[simp]
        theorem Complex.coe_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : α) (g : α) (a : α) :
        @[simp]
        theorem Complex.coe_comp_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : α) (g : α) :
        @[simp]
        theorem Complex.coe_comp_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : α) (g : α) :
        @[simp]
        theorem NNReal.coe_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (g : αNNReal) (a : α) :
        ((f g) a) = (NNReal.toReal f NNReal.toReal g) a
        @[simp]
        theorem NNReal.coe_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (g : αNNReal) (a : α) :
        ((f g) a) = (NNReal.toReal f NNReal.toReal g) a
        @[simp]
        theorem NNReal.coe_comp_conv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (g : αNNReal) :
        @[simp]
        theorem NNReal.coe_comp_dconv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (g : αNNReal) :

        Iterated convolution #

        def iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
        αβ

        Iterated convolution.

        Equations
        Instances For
          @[simp]
          theorem iterConv_zero {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
          f ∗^ 0 = trivChar
          @[simp]
          theorem iterConv_one {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) :
          f ∗^ 1 = f
          theorem iterConv_succ {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (n : ) :
          f ∗^ (n + 1) = f ∗^ n f
          theorem iterConv_succ' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (n : ) :
          f ∗^ (n + 1) = f f ∗^ n
          theorem iterConv_add {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (m : ) (n : ) :
          f ∗^ (m + n) = f ∗^ m f ∗^ n
          theorem iterConv_mul {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (m : ) (n : ) :
          f ∗^ (m * n) = f ∗^ m ∗^ n
          theorem iterConv_mul' {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (m : ) (n : ) :
          f ∗^ (m * n) = f ∗^ n ∗^ m
          @[simp]
          theorem conj_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (n : ) :
          (starRingEnd (αβ)) (f ∗^ n) = (starRingEnd (αβ)) f ∗^ n
          @[simp]
          theorem conj_iterConv_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (n : ) (a : α) :
          (starRingEnd β) ((f ∗^ n) a) = ((starRingEnd (αβ)) f ∗^ n) a
          theorem IsSelfAdjoint.iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] {f : αβ} (hf : IsSelfAdjoint f) (n : ) :
          @[simp]
          theorem conjneg_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (n : ) :
          theorem iterConv_conv_distrib {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (g : αβ) (n : ) :
          (f g) ∗^ n = f ∗^ n g ∗^ n
          theorem iterConv_dconv_distrib {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (g : αβ) (n : ) :
          (f g) ∗^ n = f ∗^ n g ∗^ n
          @[simp]
          theorem zero_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] {n : } :
          n 00 ∗^ n = 0
          @[simp]
          theorem smul_iterConv {α : Type u_1} {β : Type u_2} {γ : Type u_3} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [Monoid γ] [DistribMulAction γ β] [IsScalarTower γ β β] [SMulCommClass γ β β] (c : γ) (f : αβ) (n : ) :
          (c f) ∗^ n = c ^ n f ∗^ n
          theorem comp_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] {γ : Type u_4} [CommSemiring γ] [StarRing γ] (m : β →+* γ) (f : αβ) (n : ) :
          m (f ∗^ n) = m f ∗^ n
          theorem map_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] {γ : Type u_4} [CommSemiring γ] [StarRing γ] (m : β →+* γ) (f : αβ) (a : α) (n : ) :
          m ((f ∗^ n) a) = (m f ∗^ n) a
          theorem sum_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (n : ) :
          (Finset.univ.sum fun (a : α) => (f ∗^ n) a) = (Finset.univ.sum fun (a : α) => f a) ^ n
          @[simp]
          theorem iterConv_trivChar {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (n : ) :
          trivChar ∗^ n = trivChar
          theorem support_iterConv_subset {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (n : ) :
          theorem indicate_iterConv_apply {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (s : Finset α) (n : ) (a : α) :
          (𝟭 s ∗^ n) a = (Finset.filter (fun (x : Fin nα) => (Finset.univ.sum fun (i : Fin n) => x i) = a) (s^^n)).card
          theorem indicate_iterConv_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (s : Finset α) (n : ) (f : αβ) :
          𝟭 s ∗^ n f = (s^^n).sum fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) f
          theorem conv_indicate_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] (f : αβ) (s : Finset α) (n : ) :
          f 𝟭 s ∗^ n = (s^^n).sum fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) f
          theorem indicate_iterConv_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (s : Finset α) (n : ) (f : αβ) :
          𝟭 s ∗^ n f = (s^^n).sum fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) (conjneg f)
          theorem dconv_indicate_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [CommSemiring β] [StarRing β] (f : αβ) (s : Finset α) (n : ) :
          f 𝟭 s ∗^ n = (s^^n).sum fun (a : Fin nα) => τ (-Finset.univ.sum fun (i : Fin n) => a i) f
          theorem mu_iterConv_conv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [Module ℚ≥0 β] (s : Finset α) (n : ) (f : αβ) :
          μ s ∗^ n f = (Fintype.piFinset fun (x : Fin n) => s).expect fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) f
          theorem conv_mu_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [Module ℚ≥0 β] (f : αβ) (s : Finset α) (n : ) :
          f μ s ∗^ n = (Fintype.piFinset fun (x : Fin n) => s).expect fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) f
          theorem mu_iterConv_dconv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [Module ℚ≥0 β] [StarRing β] (s : Finset α) (n : ) (f : αβ) :
          μ s ∗^ n f = (Fintype.piFinset fun (x : Fin n) => s).expect fun (a : Fin nα) => τ (Finset.univ.sum fun (i : Fin n) => a i) (conjneg f)
          theorem dconv_mu_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Semifield β] [CharZero β] [Module ℚ≥0 β] [StarRing β] (f : αβ) (s : Finset α) (n : ) :
          f μ s ∗^ n = (Fintype.piFinset fun (x : Fin n) => s).expect fun (a : Fin nα) => τ (-Finset.univ.sum fun (i : Fin n) => a i) f
          @[simp]
          theorem balance_iterConv {α : Type u_1} {β : Type u_2} [Fintype α] [DecidableEq α] [AddCommGroup α] [Field β] [CharZero β] (f : αβ) {n : } :
          @[simp]
          theorem NNReal.coe_iterConv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : αNNReal) (n : ) (a : α) :
          ((f ∗^ n) a) = (NNReal.toReal f ∗^ n) a
          @[simp]
          theorem Complex.coe_iterConv {α : Type u_1} [Fintype α] [DecidableEq α] [AddCommGroup α] (f : α) (n : ) (a : α) :
          ((f ∗^ n) a) = (Complex.ofReal' f ∗^ n) a