Type tags for right action on the domain of a function #
By default, M acts on α → β if it acts on β, and the action is given by
(c • f) a = c • (f a).
In some cases, it is useful to consider another action: if M acts on α on the left, then it acts
on α → β on the right so that (c • f) a = f (c • a). E.g., this action is used to reformulate
the Mean Ergodic Theorem in terms of an operator on (L^2).
Main definitions #
- DomMulAct M(notation:- Mᵈᵐᵃ): type synonym for- Mᵐᵒᵖ; if- Mmultiplicatively acts on- α, then- Mᵈᵐᵃacts on- α → βfor any type- β;
- DomAddAct M(notation:- Mᵈᵃᵃ): the additive version.
We also define actions of Mᵈᵐᵃ on:
- α → βprovided that- Macts on- α;
- A →* Bprovided that- Macts on- Aby a- MulDistribMulAction;
- A →+ Bprovided that- Macts on- Aby a- DistribMulAction.
Implementation details #
Motivation #
Right action can be represented in mathlib as an action of the opposite group Mᵐᵒᵖ. However,
this "domain shift" action cannot be an instance because this would create a "diamond"
(a.k.a. ambiguous notation): if M is a monoid, then how does Mᵐᵒᵖ act on M → M? On the one
hand, Mᵐᵒᵖ acts on M by c • a = a * c.unop, thus we have an action
(c • f) a = f a * c.unop. On the other hand, M acts on itself by multiplication on the left, so
with this new instance we would have (c • f) a = f (c.unop * a). Clearly, these are two different
actions, so both of them cannot be instances in the library.
To overcome this difficulty, we introduce a type synonym DomMulAct M := Mᵐᵒᵖ (notation:
Mᵈᵐᵃ). This new type carries the same algebraic structures as Mᵐᵒᵖ but acts on α → β by this
new action. So, e.g., Mᵈᵐᵃ acts on (M → M) → M by DomMulAct.mk c • F f = F (fun a ↦ c • f a)
while (Mᵈᵐᵃ)ᵈᵐᵃ (which is isomorphic to M) acts on (M → M) → M by
DomMulAct.mk (DomMulAct.mk c) • F f = F (fun a ↦ f (c • a)).
Action on bundled homomorphisms #
If the action of M on A preserves some structure, then Mᵈᵐᵃ acts on bundled homomorphisms from
A to any type B that preserve the same structure. Examples (some of them are not yet in the
library) include:
- a MulDistribMulActiongenerates an action onA →* B;
- a DistribMulActiongenerates an action onA →+ B;
- an action on αthat commutes with action of some other monoidNgenerates an action onα →[N] β;
- a DistribMulActionon anR-module that commutes with scalar multiplications byc : Rgenerates an action onR-linear maps from this module;
- a continuous action on Xgenerates an action onC(X, Y);
- a measurable action on Xgenerates an action on{ f : X → Y // Measurable f };
- a quasi measure preserving action on Xgenerates an action onX →ₘ[μ] Y;
- a measure preserving action generates an isometric action on MeasureTheory.Lp _ _ _.
Left action vs right action #
It is common in the literature to consider the left action given by (c • f) a = f (c⁻¹ • a)
instead of the action defined in this file. However, this left action is defined only if c belongs
to a group, not to a monoid, so we decided to go with the right action.
The left action can be written in terms of DomMulAct as (DomMulAct.mk c)⁻¹ • f. As for higher
level dynamics objects (orbits, invariant functions etc), they coincide for the left and for the
right action, so lemmas can be formulated in terms of DomMulAct.
Keywords #
group action, function, domain
If M multiplicatively acts on α, then DomMulAct M acts on α → β as well as some
bundled maps from α. This is a type synonym for MulOpposite M, so this corresponds to a right
action of M.
Equations
- «term_ᵈᵐᵃ» = Lean.ParserDescr.trailingNode `term_ᵈᵐᵃ 1024 1024 (Lean.ParserDescr.symbol "ᵈᵐᵃ")
Instances For
If M additively acts on α, then DomAddAct M acts on α → β as
well as some bundled maps from α. This is a type synonym for AddOpposite M, so this corresponds
to a right action of M.
Equations
- «term_ᵈᵃᵃ» = Lean.ParserDescr.trailingNode `term_ᵈᵃᵃ 1024 1024 (Lean.ParserDescr.symbol "ᵈᵃᵃ")
Instances For
Copy instances from Mᵐᵒᵖ #
Equations
- DomAddAct.instAddSemigroupOfAddOpposite = inst
Equations
- DomAddAct.instAddRightCancelMonoidOfAddOpposite = inst
Equations
- DomAddAct.instNonAssocSemiringOfAddOpposite = inst
Equations
- DomAddAct.instAddRightCancelSemigroupOfAddOpposite = inst
Equations
- DomMulAct.instCommMonoidOfMulOpposite = inst
Equations
- DomAddAct.instNonUnitalSemiringOfAddOpposite = inst
Equations
- DomMulAct.instInvolutiveInvOfMulOpposite = inst
Equations
- DomAddAct.instAddZeroClassOfAddOpposite = inst
Equations
- DomMulAct.instDivInvOneMonoidOfMulOpposite = inst
Equations
- DomAddAct.instDivisionAddMonoidOfAddOpposite = inst
Equations
- DomAddAct.instInvolutiveNegOfAddOpposite = inst
Equations
- DomMulAct.instInvOneClassOfMulOpposite = inst
Equations
- DomAddAct.instNegZeroClassOfAddOpposite = inst
Equations
- DomAddAct.instAddCommGroupOfAddOpposite = inst
Equations
- DomMulAct.instRightCancelMonoidOfMulOpposite = inst
Equations
- DomAddAct.instAddCommSemigroupOfAddOpposite = inst
Equations
- DomMulAct.instMulOneClassOfMulOpposite = inst
Equations
- DomMulAct.instNonAssocSemiringOfMulOpposite = inst
Equations
- DomAddAct.instSubNegAddMonoidOfAddOpposite = inst
Equations
- DomAddAct.instAddCancelMonoidOfAddOpposite = inst
Equations
- DomAddAct.instDivisionAddCommMonoidOfAddOpposite = inst
Equations
- DomAddAct.instAddLeftCancelSemigroupOfAddOpposite = inst
Equations
- DomMulAct.instCancelMonoidOfMulOpposite = inst
Equations
- DomMulAct.instDivisionMonoidOfMulOpposite = inst
Equations
- DomMulAct.instLeftCancelMonoidOfMulOpposite = inst
Equations
- DomMulAct.instLeftCancelSemigroupOfMulOpposite = inst
Equations
- DomMulAct.instDivisionCommMonoidOfMulOpposite = inst
Equations
- DomAddAct.instSubNegZeroAddMonoidOfAddOpposite = inst
Equations
- DomMulAct.instCommSemigroupOfMulOpposite = inst
Equations
- DomAddAct.instAddCancelCommMonoidOfAddOpposite = inst
Equations
- DomMulAct.instRightCancelSemigroupOfMulOpposite = inst
Equations
- DomAddAct.instNonAssocSemiringOfAddOpposite_1 = inst
Equations
- DomAddAct.instAddLeftCancelMonoidOfAddOpposite = inst
Equations
- DomMulAct.instDivInvMonoidOfMulOpposite = inst
Equations
- DomMulAct.instCancelCommMonoidOfMulOpposite = inst
Equations
- DomAddAct.instAddCommMonoidOfAddOpposite = inst
Equations
- DomMulAct.instNonUnitalSemiringOfMulOpposite = inst
Equations
- DomMulAct.instNonAssocSemiringOfMulOpposite_1 = inst
Equations
- ⋯ = inst
Equations
- ⋯ = inst
Equations
- ⋯ = inst
Equations
- ⋯ = inst
Equations
- ⋯ = inst
Equations
- ⋯ = inst
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- DomMulAct.instSMulZeroClassForallOfSMul = SMulZeroClass.mk ⋯
Equations
- DomMulAct.instDistribSMulForallOfSMul = DistribSMul.mk ⋯
Equations
- DomMulAct.instDistribMulActionForallOfMulAction = DistribMulAction.mk ⋯ ⋯
Equations
- DomMulAct.instMulDistribMulActionForallOfMulAction = MulDistribMulAction.mk ⋯ ⋯
Equations
- DomMulAct.instSMulMonoidHom = { smul := fun (c : Mᵈᵐᵃ) (f : A →* B) => f.comp (MulDistribMulAction.toMonoidHom A (DomMulAct.mk.symm c)) }
Equations
- ⋯ = ⋯
Equations
- DomMulAct.instMulActionMonoidHom = Function.Injective.mulAction DFunLike.coe ⋯ ⋯
Equations
- DomMulAct.instSMulAddMonoidHom = { smul := fun (c : Mᵈᵐᵃ) (f : A →+ B) => f.comp (DistribSMul.toAddMonoidHom A (DomMulAct.mk.symm c)) }
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- DomMulAct.instMulActionAddMonoidHomOfDistribMulAction = Function.Injective.mulAction DFunLike.coe ⋯ ⋯
Equations
- DomMulAct.instDistribMulActionAddMonoidHom = Function.Injective.distribMulAction (AddMonoidHom.coeFn A B) ⋯ ⋯
Equations
- DomMulAct.instMulDistribMulActionMonoidHom = Function.Injective.mulDistribMulAction (MonoidHom.coeFn A B) ⋯ ⋯