Pointwise operations on filters #
This file defines pointwise operations on filters. This is useful because usual algebraic operations distribute over pointwise operations. For example,
(f₁ * f₂).map m = f₁.map m * f₂.map m
𝓝 (x * y) = 𝓝 x * 𝓝 y
Main declarations #
0
(Filter.instZero
): Pure filter at0 : α
, or alternatively principal filter at0 : Set α
.1
(Filter.instOne
): Pure filter at1 : α
, or alternatively principal filter at1 : Set α
.f + g
(Filter.instAdd
): Addition, filter generated by alls + t
wheres ∈ f
andt ∈ g
.f * g
(Filter.instMul
): Multiplication, filter generated by alls * t
wheres ∈ f
andt ∈ g
.-f
(Filter.instNeg
): Negation, filter of all-s
wheres ∈ f
.f⁻¹
(Filter.instInv
): Inversion, filter of alls⁻¹
wheres ∈ f
.f - g
(Filter.instSub
): Subtraction, filter generated by alls - t
wheres ∈ f
andt ∈ g
.f / g
(Filter.instDiv
): Division, filter generated by alls / t
wheres ∈ f
andt ∈ g
.f +ᵥ g
(Filter.instVAdd
): Scalar addition, filter generated by alls +ᵥ t
wheres ∈ f
andt ∈ g
.f -ᵥ g
(Filter.instVSub
): Scalar subtraction, filter generated by alls -ᵥ t
wheres ∈ f
andt ∈ g
.f • g
(Filter.instSMul
): Scalar multiplication, filter generated by alls • t
wheres ∈ f
andt ∈ g
.a +ᵥ f
(Filter.instVAddFilter
): Translation, filter of alla +ᵥ s
wheres ∈ f
.a • f
(Filter.instSMulFilter
): Scaling, filter of alla • s
wheres ∈ f
.
For α
a semigroup/monoid, Filter α
is a semigroup/monoid.
As an unfortunate side effect, this means that n • f
, where n : ℕ
, is ambiguous between
pointwise scaling and repeated pointwise addition. See note [pointwise nat action].
Implementation notes #
We put all instances in the locale Pointwise
, so that these instances are not available by
default. Note that we do not mark them as reducible (as argued by note [reducible non-instances])
since we expect the locale to be open whenever the instances are actually used (and making the
instances reducible changes the behavior of simp
).
Tags #
filter multiplication, filter addition, pointwise addition, pointwise multiplication,
0
/1
as filters #
1 : Filter α
is defined as the filter of sets containing 1 : α
in locale Pointwise
.
Equations
- Filter.instOne = { one := pure 1 }
Instances For
0 : Filter α
is defined as the filter of sets containing 0 : α
in locale Pointwise
.
Equations
- Filter.instZero = { zero := pure 0 }
Instances For
Filter negation/inversion #
The inverse of a filter is the pointwise preimage under ⁻¹
of its sets.
Equations
- Filter.instInv = { inv := Filter.map Inv.inv }
The negation of a filter is the pointwise preimage under -
of its sets.
Equations
- Filter.instNeg = { neg := Filter.map Neg.neg }
Inversion is involutive on Filter α
if it is on α
.
Equations
- Filter.instInvolutiveInv = { toInv := Filter.instInv, inv_inv := ⋯ }
Instances For
Negation is involutive on Filter α
if it is on α
.
Equations
- Filter.instInvolutiveNeg = { toNeg := Filter.instNeg, neg_neg := ⋯ }
Instances For
Filter addition/multiplication #
pure
operation as a MulHom
.
Equations
- Filter.pureMulHom = { toFun := pure, map_mul' := ⋯ }
Instances For
The singleton operation as an AddHom
.
Equations
- Filter.pureAddHom = { toFun := pure, map_add' := ⋯ }
Instances For
Filter subtraction/division #
Repeated pointwise multiplication/division (not the same as pointwise repeated
multiplication/division!) of a Filter
. See Note [pointwise nat action].
Instances For
Filter α
is an AddSemigroup
under pointwise operations if α
is.
Equations
- Filter.addSemigroup = { add := fun (x1 x2 : Filter α) => x1 + x2, add_assoc := ⋯ }
Instances For
Filter α
is a CommSemigroup
under pointwise operations if α
is.
Equations
- Filter.commSemigroup = { toSemigroup := Filter.semigroup, mul_comm := ⋯ }
Instances For
Filter α
is an AddCommSemigroup
under pointwise operations if α
is.
Equations
- Filter.addCommSemigroup = { toAddSemigroup := Filter.addSemigroup, add_comm := ⋯ }
Instances For
Filter α
is a MulOneClass
under pointwise operations if α
is.
Equations
- Filter.mulOneClass = { one := 1, mul := fun (x1 x2 : Filter α) => x1 * x2, one_mul := ⋯, mul_one := ⋯ }
Instances For
Filter α
is an AddZeroClass
under pointwise operations if α
is.
Equations
- Filter.addZeroClass = { zero := 0, add := fun (x1 x2 : Filter α) => x1 + x2, zero_add := ⋯, add_zero := ⋯ }
Instances For
If φ : α →* β
then mapMonoidHom φ
is the monoid homomorphism
Filter α →* Filter β
induced by map φ
.
Equations
- Filter.mapMonoidHom φ = { toFun := Filter.map ⇑φ, map_one' := ⋯, map_mul' := ⋯ }
Instances For
If φ : α →+ β
then mapAddMonoidHom φ
is the monoid homomorphism
Filter α →+ Filter β
induced by map φ
.
Equations
- Filter.mapAddMonoidHom φ = { toFun := Filter.map ⇑φ, map_zero' := ⋯, map_add' := ⋯ }
Instances For
pure
as a MonoidHom
.
Equations
- Filter.pureMonoidHom = { toFun := Filter.pureMulHom.toFun, map_one' := ⋯, map_mul' := ⋯ }
Instances For
pure
as an AddMonoidHom
.
Equations
- Filter.pureAddMonoidHom = { toFun := Filter.pureAddHom.toFun, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Filter α
is a Monoid
under pointwise operations if α
is.
Equations
- Filter.monoid = { toMul := MulOneClass.toMul, mul_assoc := ⋯, toOne := MulOneClass.toOne, one_mul := ⋯, mul_one := ⋯, npow := npowRecAuto, npow_zero := ⋯, npow_succ := ⋯ }
Instances For
Filter α
is an AddMonoid
under pointwise operations if α
is.
Equations
- Filter.addMonoid = { toAdd := AddZeroClass.toAdd, add_assoc := ⋯, toZero := AddZeroClass.toZero, zero_add := ⋯, add_zero := ⋯, nsmul := nsmulRecAuto, nsmul_zero := ⋯, nsmul_succ := ⋯ }
Instances For
Filter α
is a CommMonoid
under pointwise operations if α
is.
Equations
- Filter.commMonoid = { toMul := MulOneClass.toMul, mul_assoc := ⋯, toOne := MulOneClass.toOne, one_mul := ⋯, mul_one := ⋯, npow := npowRecAuto, npow_zero := ⋯, npow_succ := ⋯, mul_comm := ⋯ }
Instances For
Filter α
is an AddCommMonoid
under pointwise operations if α
is.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Filter α
is a division monoid under pointwise operations if α
is.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Filter α
is a subtraction monoid under pointwise operations if
α
is.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Filter α
is a commutative division monoid under pointwise operations if α
is.
Equations
- Filter.divisionCommMonoid = { toDivisionMonoid := Filter.divisionMonoid, mul_comm := ⋯ }
Instances For
Filter α
is a commutative subtraction monoid under pointwise operations if α
is.
Equations
- Filter.subtractionCommMonoid = { toSubtractionMonoid := Filter.subtractionMonoid, add_comm := ⋯ }
Instances For
Filter α
has distributive negation if α
has.
Equations
- Filter.instDistribNeg = { toInvolutiveNeg := Filter.instInvolutiveNeg, neg_mul := ⋯, mul_neg := ⋯ }
Instances For
Note that Filter
is not a MulZeroClass
because 0 * ⊥ ≠ 0
.
Note that Filter α
is not a group because f / f ≠ 1
in general
Scalar addition/multiplication of filters #
Scalar subtraction of filters #
Translation/scaling of filters #
a • f
is the map of f
under a •
in locale Pointwise
.
Equations
- Filter.instSMulFilter = { smul := fun (a : α) => Filter.map fun (x : β) => a • x }
Instances For
a +ᵥ f
is the map of f
under a +ᵥ
in locale Pointwise
.
Equations
- Filter.instVAddFilter = { vadd := fun (a : α) => Filter.map fun (x : β) => a +ᵥ x }
Instances For
A multiplicative action of a monoid α
on a type β
gives a multiplicative action of
Filter α
on Filter β
.
Equations
- Filter.mulAction = { toSMul := Filter.instSMul, one_smul := ⋯, mul_smul := ⋯ }
Instances For
An additive action of an additive monoid α
on a type β
gives an additive action
of Filter α
on Filter β
Equations
- Filter.addAction = { toVAdd := Filter.instVAdd, zero_vadd := ⋯, add_vadd := ⋯ }
Instances For
A multiplicative action of a monoid on a type β
gives a multiplicative action on Filter β
.
Equations
- Filter.mulActionFilter = { toSMul := Filter.instSMulFilter, one_smul := ⋯, mul_smul := ⋯ }
Instances For
An additive action of an additive monoid on a type β
gives an additive action on
Filter β
.
Equations
- Filter.addActionFilter = { toVAdd := Filter.instVAddFilter, zero_vadd := ⋯, add_vadd := ⋯ }
Instances For
A distributive multiplicative action of a monoid on an additive monoid β
gives a distributive
multiplicative action on Filter β
.
Equations
- Filter.distribMulActionFilter = { toMulAction := Filter.mulActionFilter, smul_zero := ⋯, smul_add := ⋯ }
Instances For
A multiplicative action of a monoid on a monoid β
gives a multiplicative action on Set β
.
Equations
- Filter.mulDistribMulActionFilter = { toMulAction := Set.mulActionSet, smul_mul := ⋯, smul_one := ⋯ }
Instances For
Note that we have neither SMulWithZero α (Filter β)
nor SMulWithZero (Filter α) (Filter β)
because 0 * ⊥ ≠ 0
.