Documentation

Mathlib.Combinatorics.Additive.Convolution

Convolution #

This file defines convolution of finite subsets A and B of group G as the map A ⋆ B : G → ℕ that maps x ∈ G to the number of distinct representations of x in the form x = ab for a ∈ A, b ∈ B. It is shown how convolution behaves under the change of order of A and B, as well as under the left and right actions on A, B, and the function argument.

def Finset.convolution {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) :
G

Given finite subsets A and B of a group G, convolution of A and B is a map G → ℕ that maps x ∈ G to the number of distinct representations of x in the form x = ab, where a ∈ A, b ∈ B.

Equations
Instances For
    def Finset.addConvolution {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) :
    G

    Given finite subsets A and B of an additive group G, convolution of A and B is a map G → ℕ that maps x ∈ G to the number of distinct representations of x in the form x = a + b, where a ∈ A, b ∈ B.

    Equations
    Instances For
      theorem Finset.card_smul_inter_smul {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (x y : G) :
      (x A y B).card = A.convolution B⁻¹ (x⁻¹ * y)
      theorem Finset.card_vadd_inter_vadd {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) (x y : G) :
      ((x +ᵥ A) (y +ᵥ B)).card = A.addConvolution (-B) (-x + y)
      theorem Finset.card_inter_smul {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (x : G) :
      (A x B).card = A.convolution B⁻¹ x
      theorem Finset.card_inter_vadd {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) (x : G) :
      (A (x +ᵥ B)).card = A.addConvolution (-B) x
      theorem Finset.card_smul_inter {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (x : G) :
      theorem Finset.card_vadd_inter {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) (x : G) :
      ((x +ᵥ A) B).card = A.addConvolution (-B) (-x)
      theorem Finset.card_mul_inv_eq_convolution_inv {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (x : G) :
      {abA ×ˢ B | ab.1 * ab.2⁻¹ = x}.card = A.convolution B⁻¹ x
      theorem Finset.card_add_neg_eq_addConvolution_neg {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) (x : G) :
      {abA ×ˢ B | ab.1 + -ab.2 = x}.card = A.addConvolution (-B) x
      @[simp]
      theorem Finset.convolution_pos {G : Type u_1} [Group G] [DecidableEq G] {A B : Finset G} {x : G} :
      0 < A.convolution B x x A * B
      @[simp]
      theorem Finset.addConvolution_pos {G : Type u_1} [AddGroup G] [DecidableEq G] {A B : Finset G} {x : G} :
      0 < A.addConvolution B x x A + B
      theorem Finset.convolution_ne_zero {G : Type u_1} [Group G] [DecidableEq G] {A B : Finset G} {x : G} :
      A.convolution B x 0 x A * B
      theorem Finset.addConvolution_ne_zero {G : Type u_1} [AddGroup G] [DecidableEq G] {A B : Finset G} {x : G} :
      A.addConvolution B x 0 x A + B
      @[simp]
      theorem Finset.convolution_eq_zero {G : Type u_1} [Group G] [DecidableEq G] {A B : Finset G} {x : G} :
      A.convolution B x = 0 xA * B
      @[simp]
      theorem Finset.addConvolution_eq_zero {G : Type u_1} [AddGroup G] [DecidableEq G] {A B : Finset G} {x : G} :
      A.addConvolution B x = 0 xA + B
      theorem Finset.convolution_le_card_left {G : Type u_1} [Group G] [DecidableEq G] {A B : Finset G} {x : G} :
      theorem Finset.addConvolution_le_card_left {G : Type u_1} [AddGroup G] [DecidableEq G] {A B : Finset G} {x : G} :
      theorem Finset.convolution_le_card_right {G : Type u_1} [Group G] [DecidableEq G] {A B : Finset G} {x : G} :
      @[simp]
      theorem Finset.convolution_inv {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (x : G) :
      @[simp]
      theorem Finset.addConvolution_neg {G : Type u_1} [AddGroup G] [DecidableEq G] (A B : Finset G) (x : G) :
      @[simp]
      theorem Finset.smul_convolution_eq_convolution_inv_mul {G : Type u_1} [Group G] [DecidableEq G] (A B : Finset G) (s x : G) :
      (s A).convolution B x = A.convolution B (s⁻¹ * x)
      @[simp]
      @[simp]