Pointwise instances on Submodules #
This file provides:
and the actions
which matches the action of Set.mulActionSet.
This file also provides:
- Submodule.pointwiseSetSMulSubmodule: for- R-module- M, a- s : Set Rcan act on- N : Submodule R Mby defining- s • Nto be the smallest submodule containing all- a • nwhere- a ∈ sand- n ∈ N.
These actions are available in the Pointwise locale.
Implementation notes #
For an R-module M, The action of a subset of R acting on a submodule of M introduced in
section set_acting_on_submodules does not have a counterpart in the file
Mathlib.Algebra.Group.Submonoid.Pointwise.
Other than section set_acting_on_submodules, most of the lemmas in this file are direct copies of
lemmas from the file Mathlib.Algebra.Group.Submonoid.Pointwise.
The submodule with every element negated. Note if R is a ring and not just a semiring, this
is a no-op, as shown by Submodule.neg_eq_self.
Recall that When R is the semiring corresponding to the nonnegative elements of R',
Submodule R' M is the type of cones of M. This instance reflects such cones about 0.
This is available as an instance in the Pointwise locale.
Equations
Instances For
Submodule.pointwiseNeg is involutive.
This is available as an instance in the Pointwise locale.
Equations
- Submodule.involutivePointwiseNeg = InvolutiveNeg.mk ⋯
Instances For
Submodule.pointwiseNeg as an order isomorphism.
Instances For
Equations
- Submodule.pointwiseAddCommMonoid = AddCommMonoid.mk ⋯
Equations
- Submodule.instCanonicallyOrderedAddCommMonoid = let __src := Submodule.pointwiseAddCommMonoid; let __src_1 := Submodule.completeLattice; CanonicallyOrderedAddCommMonoid.mk ⋯ ⋯
The action on a submodule corresponding to applying the action to every element.
This is available as an instance in the Pointwise locale.
Equations
- Submodule.pointwiseDistribMulAction = DistribMulAction.mk ⋯ ⋯
Instances For
Equations
- ⋯ = ⋯
See also Submodule.smul_bot.
See also Submodule.smul_sup.
Equations
- ⋯ = ⋯
The action on a submodule corresponding to applying the action to every element.
This is available as an instance in the Pointwise locale.
This is a stronger version of Submodule.pointwiseDistribMulAction. Note that add_smul does
not hold so this cannot be stated as a Module.
Equations
- Submodule.pointwiseMulActionWithZero = let __src := Submodule.pointwiseDistribMulAction; MulActionWithZero.mk ⋯ ⋯
Instances For
Sets acting on Submodules #
Let R be a (semi)ring and M an R-module. Let S be a monoid which acts on M distributively,
then subsets of S can act on submodules of M.
For subset s ⊆ S and submodule N ≤ M, we define s • N to be the smallest submodule containing
all r • n where r ∈ s and n ∈ N.
Results #
For arbitrary monoids S acting distributively on M, there is an induction principle for s • N:
To prove P holds for all s • N, it is enough
to prove:
- for all r ∈ sandn ∈ N,P (r • n);
- for all randm ∈ s • N,P (r • n);
- for all m₁, m₂,P m₁andP m₂impliesP (m₁ + m₂);
- P 0.
To invoke this induction principle, use induction x, hx using Submodule.set_smul_inductionOn where
x : M and hx : x ∈ s • N
When we consider subset of R acting on M
- Submodule.pointwiseSetDistribMulAction: the action described above is distributive.
- Submodule.mem_set_smul:- x ∈ s • Niff- xcan be written as- r₀ n₀ + ... + rₖ nₖwhere- rᵢ ∈ sand- nᵢ ∈ N.
- Submodule.coe_span_smul:- s • Nis the same as- ⟨s⟩ • Nwhere- ⟨s⟩is the ideal spanned by- s.
Notes #
- If we assume the addition on subsets of Ris the⊔and subtraction⊓i.e. useSetSemiring, then this action actually gives a module structure on submodules ofMover subsets ofR.
- If we generalize so that r • Nmakes sense for allr : S, thenSubmodule.singleton_set_smulandSubmodule.singleton_set_smulcan be generalized as well.
Let s ⊆ R be a set and N ≤ M be a submodule, then s • N is the smallest submodule containing
all r • n where r ∈ s and n ∈ N.
Equations
Instances For
Equations
- ⋯ = ⋯
Induction principle for set acting on submodules. To prove P holds for all s • N, it is enough
to prove:
- for all r ∈ sandn ∈ N,P (r • n);
- for all randm ∈ s • N,P (r • n);
- for all m₁, m₂,P m₁andP m₂impliesP (m₁ + m₂);
- P 0.
To invoke this induction principle, use induction x, hx using Submodule.set_smul_inductionOn where
x : M and hx : x ∈ s • N
A subset of a ring R has a multiplicative action on submodules of a module over R.
Equations
- Submodule.pointwiseSetMulAction = MulAction.mk ⋯ ⋯
Instances For
In a ring, sets acts on submodules.
Equations
- Submodule.pointwiseSetDistribMulAction = DistribMulAction.mk ⋯ ⋯