Documentation

Mathlib.RingTheory.GradedAlgebra.Homogeneous.Submodule

Homogeneous submodules of a graded module #

This file defines homogeneous submodule of a graded module ⨁ᵢ ℳᵢ over graded ring ⨁ᵢ 𝒜ᵢ and operations on them.

Main definitions #

For any p : Submodule A M:

Implementation notes #

The notion of homogeneous submodule does not rely on a graded ring, only a decomposition of the the module. However, most interesting properties of homogeneous submodules do rely on the base ring being a graded ring. For technical reasons, we make HomogeneousSubmodule depend on a graded ring. For example, if the definition of a homogeneous submodule does not depend on a graded ring, the instance that HomogeneousSubmodule is a complete lattice cannot be synthesized due to synthesation order.

Tags #

graded algebra, homogeneous

def Submodule.IsHomogeneous {ιM : Type u_2} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (p : Submodule A M) (ℳ : ιMσM) [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] :

An A-submodule p ⊆ M is homogeneous if for every m ∈ p, all homogeneous components of m are in p.

Equations
Instances For
    theorem Submodule.IsHomogeneous.mem_iff {ιM : Type u_2} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] {p : Submodule A M} (ℳ : ιMσM) [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] (hp : p.IsHomogeneous ) {x : M} :
    x p ∀ (i : ιM), (((DirectSum.decompose ) x) i) p
    structure HomogeneousSubmodule {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] extends Submodule A M :
    Type u_6

    For any Semiring A, we collect the homogeneous submodule of A-modules into a type.

    Instances For
      instance instSetLikeHomogeneousSubmodule {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] :
      Equations
      instance instAddSubmonoidClassHomogeneousSubmodule {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] :
      instance instSMulMemClassHomogeneousSubmodule {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] :
      theorem HomogeneousSubmodule.isHomogeneous {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] {𝒜 : ιAσA} {ℳ : ιMσM} [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] (p : HomogeneousSubmodule 𝒜 ) :
      theorem HomogeneousSubmodule.toSubmodule_injective {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] :
      instance HomogeneousSubmodule.setLike {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] :
      Equations
      theorem HomogeneousSubmodule.ext {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] {I J : HomogeneousSubmodule 𝒜 } (h : I.toSubmodule = J.toSubmodule) :
      I = J
      theorem HomogeneousSubmodule.ext' {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] {I J : HomogeneousSubmodule 𝒜 } (h : ∀ (i : ιM), x i, x I x J) :
      I = J

      The set-theoretic extensionality of HomogeneousSubmodule.

      @[simp]
      theorem HomogeneousSubmodule.mem_toSubmodule_iff {ιA : Type u_1} {ιM : Type u_2} {σA : Type u_3} {σM : Type u_4} {A : Type u_5} {M : Type u_6} [Semiring A] [AddCommMonoid M] [Module A M] (𝒜 : ιAσA) (ℳ : ιMσM) [DecidableEq ιA] [AddMonoid ιA] [SetLike σA A] [AddSubmonoidClass σA A] [GradedRing 𝒜] [DecidableEq ιM] [SetLike σM M] [AddSubmonoidClass σM M] [DirectSum.Decomposition ] [VAdd ιA ιM] [SetLike.GradedSMul 𝒜 ] {I : HomogeneousSubmodule 𝒜 } {x : M} :