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.
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
.
Instances For
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
.
Instances For
theorem
Finset.card_smul_inter_smul
{G : Type u_1}
[Group G]
[DecidableEq G]
(A B : Finset G)
(x y : G)
:
theorem
Finset.card_vadd_inter_vadd
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(x y : G)
:
theorem
Finset.card_inter_vadd
{G : Type u_1}
[AddGroup 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)
:
theorem
Finset.card_mul_inv_eq_convolution_inv
{G : Type u_1}
[Group G]
[DecidableEq G]
(A B : Finset G)
(x : G)
:
theorem
Finset.card_add_neg_eq_addConvolution_neg
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(x : G)
:
@[simp]
@[simp]
theorem
Finset.addConvolution_pos
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
theorem
Finset.convolution_ne_zero
{G : Type u_1}
[Group G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
theorem
Finset.addConvolution_ne_zero
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
@[simp]
theorem
Finset.convolution_eq_zero
{G : Type u_1}
[Group G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
@[simp]
theorem
Finset.addConvolution_eq_zero
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
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}
:
theorem
Finset.addConvolution_le_card_right
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
{A B : Finset G}
{x : G}
:
@[simp]
@[simp]
theorem
Finset.addConvolution_neg
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(x : G)
:
@[simp]
theorem
Finset.op_smul_convolution_eq_convolution_smul
{G : Type u_1}
[Group G]
[DecidableEq G]
(A B : Finset G)
(s : G)
:
@[simp]
theorem
Finset.op_vadd_addConvolution_eq_addConvolution_vadd
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(s : 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)
:
@[simp]
theorem
Finset.vadd_addConvolution_eq_addConvolution_neg_add
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(s x : G)
:
@[simp]
theorem
Finset.convolution_op_smul_eq_convolution_mul_inv
{G : Type u_1}
[Group G]
[DecidableEq G]
(A B : Finset G)
(s x : G)
:
@[simp]
theorem
Finset.addConvolution_op_vadd_eq_addConvolution_add_neg
{G : Type u_1}
[AddGroup G]
[DecidableEq G]
(A B : Finset G)
(s x : G)
: