Documentation

Mathlib.Algebra.Category.ModuleCat.Basic

The category of R-modules #

ModuleCat.{v} R is the category of bundled R-modules with carrier in the universe v. We show that it is preadditive and show that being an isomorphism, monomorphism and epimorphism is equivalent to being a linear equivalence, an injective linear map and a surjective linear map, respectively.

Implementation details #

To construct an object in the category of R-modules from a type M with an instance of the Module typeclass, write of R M. There is a coercion in the other direction. The roundtrip ↑(of R M) is definitionally equal to M itself (when M is a type with Module instance), and so is of R ↑M (when M : ModuleCat R M).

The morphisms are given their own type, not identified with LinearMap. There is a cast from morphisms in Module R to linear maps, written f.hom (ModuleCat.Hom.hom). To go from linear maps to morphisms in Module R, use ModuleCat.ofHom.

Similarly, given an isomorphism f : M ≅ N use f.toLinearEquiv and given a linear equiv f : M ≃ₗ[R] N, use f.toModuleIso.

structure ModuleCat (R : Type u) [Ring R] :
Type (max u (v + 1))

The category of R-modules and their morphisms.

Note that in the case of R = ℤ, we can not impose here that the -multiplication field from the module structure is defeq to the one coming from the isAddCommGroup structure (contrary to what we do for all module structures in mathlib), which creates some difficulties down the road.

@[reducible, inline]
abbrev ModuleCat.of (R : Type u) [Ring R] (X : Type v) [AddCommGroup X] [Module R X] :

The object in the category of R-algebras associated to a type equipped with the appropriate typeclasses. This is the preferred way to construct a term of ModuleCat R.

Equations
  • ModuleCat.of R X = { carrier := X, isAddCommGroup := inst✝¹, isModule := inst✝ }
theorem ModuleCat.coe_of (R : Type u) [Ring R] (X : Type v) [Ring X] [Module R X] :
(of R X) = X
structure ModuleCat.Hom {R : Type u} [Ring R] (M N : ModuleCat R) :

The type of morphisms in ModuleCat R.

  • hom' : M →ₗ[R] N

    The underlying linear map.

theorem ModuleCat.Hom.ext {R : Type u} {inst✝ : Ring R} {M N : ModuleCat R} {x y : M.Hom N} (hom' : x.hom' = y.hom') :
x = y
theorem ModuleCat.Hom.ext_iff {R : Type u} {inst✝ : Ring R} {M N : ModuleCat R} {x y : M.Hom N} :
x = y x.hom' = y.hom'
Equations
  • One or more equations did not get rendered due to their size.
Equations
  • One or more equations did not get rendered due to their size.
@[reducible, inline]
abbrev ModuleCat.Hom.hom {R : Type u} [Ring R] {A B : ModuleCat R} (f : A.Hom B) :
A →ₗ[R] B

Turn a morphism in ModuleCat back into a LinearMap.

Equations
@[reducible, inline]
abbrev ModuleCat.ofHom {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) :
of R X of R Y

Typecheck a LinearMap as a morphism in ModuleCat.

Equations
def ModuleCat.Hom.Simps.hom {R : Type u} [Ring R] (A B : ModuleCat R) (f : A.Hom B) :
A →ₗ[R] B

Use the ConcreteCategory.hom projection for @[simps] lemmas.

Equations

The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

@[simp]
theorem ModuleCat.hom_comp {R : Type u} [Ring R] {M N O : ModuleCat R} (f : M N) (g : N O) :
theorem ModuleCat.hom_ext {R : Type u} [Ring R] {M N : ModuleCat R} {f g : M N} (hf : Hom.hom f = Hom.hom g) :
f = g
theorem ModuleCat.hom_ext_iff {R : Type u} [Ring R] {M N : ModuleCat R} {f g : M N} :

Convenience shortcut for ModuleCat.hom_bijective.injective.

Convenience shortcut for ModuleCat.hom_bijective.surjective.

@[simp]
theorem ModuleCat.hom_ofHom {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) :
@[simp]
theorem ModuleCat.ofHom_hom {R : Type u} [Ring R] {M N : ModuleCat R} (f : M N) :
@[simp]
theorem ModuleCat.ofHom_comp {R : Type u} [Ring R] {M N O : Type v} [AddCommGroup M] [AddCommGroup N] [AddCommGroup O] [Module R M] [Module R N] [Module R O] (f : M →ₗ[R] N) (g : N →ₗ[R] O) :
theorem ModuleCat.ofHom_apply {R : Type u} [Ring R] {M N : Type v} [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] (f : M →ₗ[R] N) (x : M) :
def ModuleCat.homEquiv {R : Type u} [Ring R] {M N : ModuleCat R} :
(M N) (M →ₗ[R] N)

ModuleCat.Hom.hom bundled as an Equiv.

Equations
Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem ModuleCat.of_coe (R : Type u) [Ring R] (X : ModuleCat R) :
of R X = X
@[deprecated CategoryTheory.Iso.refl (since := "2025-05-15")]
def ModuleCat.ofSelfIso {R : Type u} [Ring R] (M : ModuleCat R) :
of R M M

Forgetting to the underlying type and then building the bundled object returns the original module.

Equations

Reinterpreting a linear map in the category of R-modules

Equations
@[deprecated ModuleCat.ofHom (since := "2024-11-29")]
def ModuleCat.asHomRight {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) :
of R X of R Y

Alias of ModuleCat.ofHom.


Typecheck a LinearMap as a morphism in ModuleCat.

Equations
@[deprecated ModuleCat.ofHom (since := "2024-11-29")]
def ModuleCat.asHomLeft {R : Type u} [Ring R] {X Y : Type v} [AddCommGroup X] [Module R X] [AddCommGroup Y] [Module R Y] (f : X →ₗ[R] Y) :
of R X of R Y

Alias of ModuleCat.ofHom.


Typecheck a LinearMap as a morphism in ModuleCat.

Equations

Reinterpreting a linear map in the category of R-modules. This notation is deprecated: use instead.

Equations

Reinterpreting a linear map in the category of R-modules. This notation is deprecated: use instead.

Equations
def LinearEquiv.toModuleIso {R : Type u} [Ring R] {X₁ X₂ : Type v} {g₁ : AddCommGroup X₁} {g₂ : AddCommGroup X₂} {m₁ : Module R X₁} {m₂ : Module R X₂} (e : X₁ ≃ₗ[R] X₂) :

Build an isomorphism in the category Module R from a LinearEquiv between Modules.

Equations
@[simp]
theorem LinearEquiv.toModuleIso_inv {R : Type u} [Ring R] {X₁ X₂ : Type v} {g₁ : AddCommGroup X₁} {g₂ : AddCommGroup X₂} {m₁ : Module R X₁} {m₂ : Module R X₂} (e : X₁ ≃ₗ[R] X₂) :
@[simp]
theorem LinearEquiv.toModuleIso_hom {R : Type u} [Ring R] {X₁ X₂ : Type v} {g₁ : AddCommGroup X₁} {g₂ : AddCommGroup X₂} {m₁ : Module R X₁} {m₂ : Module R X₂} (e : X₁ ≃ₗ[R] X₂) :
@[reducible, inline, deprecated LinearEquiv.toModuleIso (since := "2024-11-29")]
abbrev LinearEquiv.toModuleIso' {R : Type u} [Ring R] {M N : ModuleCat R} (i : M ≃ₗ[R] N) :
M N

Build an isomorphism in the category Module R from a LinearEquiv between Modules.

Equations
@[reducible, inline, deprecated LinearEquiv.toModuleIso (since := "2024-11-29")]
abbrev LinearEquiv.toModuleIso'Left {R : Type u} [Ring R] {X₂ : Type v} {X₁ : ModuleCat R} [AddCommGroup X₂] [Module R X₂] (e : X₁ ≃ₗ[R] X₂) :
X₁ ModuleCat.of R X₂

Build an isomorphism in the category ModuleCat R from a LinearEquiv between Modules.

Equations
@[reducible, inline, deprecated LinearEquiv.toModuleIso (since := "2024-11-29")]
abbrev LinearEquiv.toModuleIso'Right {R : Type u} [Ring R] {X₁ : Type v} [AddCommGroup X₁] [Module R X₁] {X₂ : ModuleCat R} (e : X₁ ≃ₗ[R] X₂) :
ModuleCat.of R X₁ X₂

Build an isomorphism in the category ModuleCat R from a LinearEquiv between Modules.

Equations
def CategoryTheory.Iso.toLinearEquiv {R : Type u} [Ring R] {X Y : ModuleCat R} (i : X Y) :
X ≃ₗ[R] Y

Build a LinearEquiv from an isomorphism in the category ModuleCat R.

Equations
def linearEquivIsoModuleIso {R : Type u} [Ring R] {X Y : Type u} [AddCommGroup X] [AddCommGroup Y] [Module R X] [Module R Y] :

linear equivalences between Modules are the same as (isomorphic to) isomorphisms in ModuleCat

Equations
@[simp]
instance ModuleCat.instAddHom {R : Type u} [Ring R] {M N : ModuleCat R} :
Add (M N)
Equations
@[simp]
theorem ModuleCat.hom_add {R : Type u} [Ring R] {M N : ModuleCat R} (f g : M N) :
instance ModuleCat.instZeroHom {R : Type u} [Ring R] {M N : ModuleCat R} :
Zero (M N)
Equations
@[simp]
theorem ModuleCat.hom_zero {R : Type u} [Ring R] {M N : ModuleCat R} :
instance ModuleCat.instSMulNatHom {R : Type u} [Ring R] {M N : ModuleCat R} :
SMul (M N)
Equations
@[simp]
theorem ModuleCat.hom_nsmul {R : Type u} [Ring R] {M N : ModuleCat R} (n : ) (f : M N) :
Hom.hom (n f) = n Hom.hom f
instance ModuleCat.instNegHom {R : Type u} [Ring R] {M N : ModuleCat R} :
Neg (M N)
Equations
@[simp]
theorem ModuleCat.hom_neg {R : Type u} [Ring R] {M N : ModuleCat R} (f : M N) :
instance ModuleCat.instSubHom {R : Type u} [Ring R] {M N : ModuleCat R} :
Sub (M N)
Equations
@[simp]
theorem ModuleCat.hom_sub {R : Type u} [Ring R] {M N : ModuleCat R} (f g : M N) :
instance ModuleCat.instSMulIntHom {R : Type u} [Ring R] {M N : ModuleCat R} :
SMul (M N)
Equations
@[simp]
theorem ModuleCat.hom_zsmul {R : Type u} [Ring R] {M N : ModuleCat R} (n : ) (f : M N) :
Hom.hom (n f) = n Hom.hom f
@[simp]
theorem ModuleCat.hom_sum {R : Type u} [Ring R] {M N : ModuleCat R} {ι : Type u_1} (f : ι → (M N)) (s : Finset ι) :
Hom.hom (∑ is, f i) = is, Hom.hom (f i)
Equations
def ModuleCat.homAddEquiv {R : Type u} [Ring R] {M N : ModuleCat R} :
(M N) ≃+ (M →ₗ[R] N)

ModuleCat.Hom.hom bundled as an additive equivalence.

Equations
@[simp]
theorem ModuleCat.homAddEquiv_apply {R : Type u} [Ring R] {M N : ModuleCat R} (f : M.Hom N) :
@[simp]
theorem ModuleCat.homAddEquiv_symm_apply_hom {R : Type u} [Ring R] {M N : ModuleCat R} (f : M →ₗ[R] N) :
instance ModuleCat.instSMulHom {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Monoid S] [DistribMulAction S N] [SMulCommClass R S N] :
SMul S (M N)
Equations
@[simp]
theorem ModuleCat.hom_smul {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Monoid S] [DistribMulAction S N] [SMulCommClass R S N] (s : S) (f : M N) :
Hom.hom (s f) = s Hom.hom f
instance ModuleCat.Hom.instModule {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Semiring S] [Module S N] [SMulCommClass R S N] :
Module S (M N)
Equations
def ModuleCat.homLinearEquiv {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Semiring S] [Module S N] [SMulCommClass R S N] :
(M N) ≃ₗ[S] M →ₗ[R] N

ModuleCat.Hom.hom bundled as a linear equivalence.

Equations
@[simp]
theorem ModuleCat.homLinearEquiv_symm_apply {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Semiring S] [Module S N] [SMulCommClass R S N] (a✝ : M →ₗ[R] N) :
@[simp]
theorem ModuleCat.homLinearEquiv_apply {R : Type u} [Ring R] {M N : ModuleCat R} {S : Type u_1} [Semiring S] [Module S N] [SMulCommClass R S N] (a✝ : M N) :
def ModuleCat.Algebra.instModuleCarrier {S₀ : Type u₀} [CommSemiring S₀] {S : Type u} [Ring S] [Algebra S₀ S] {M : ModuleCat S} :
Module S₀ M

Let S be an S₀-algebra. Then S-modules are modules over S₀.

Equations
theorem ModuleCat.Algebra.instIsScalarTowerCarrier {S₀ : Type u₀} [CommSemiring S₀] {S : Type u} [Ring S] [Algebra S₀ S] {M : ModuleCat S} :
IsScalarTower S₀ S M
theorem ModuleCat.Algebra.instSMulCommClassCarrier {S₀ : Type u₀} [CommSemiring S₀] {S : Type u} [Ring S] [Algebra S₀ S] {M : ModuleCat S} :
SMulCommClass S S₀ M
def ModuleCat.Algebra.instLinear {S₀ : Type u₀} [CommSemiring S₀] {S : Type u} [Ring S] [Algebra S₀ S] :

Let S be an S₀-algebra. Then the category of S-modules is S₀-linear.

Equations
theorem ModuleCat.Iso.homCongr_eq_arrowCongr {S : Type u} [CommRing S] {X Y X' Y' : ModuleCat S} (i : X X') (j : Y Y') (f : X Y) :
theorem ModuleCat.Iso.conj_eq_conj {S : Type u} [CommRing S] {X X' : ModuleCat S} (i : X X') (f : CategoryTheory.End X) :
i.conj f = { hom' := i.toLinearEquiv.conj (Hom.hom f) }

ModuleCat.Hom.hom as an isomorphism of rings.

Equations
@[simp]
theorem ModuleCat.endRingEquiv_symm_apply_hom {R : Type u} [Ring R] (M : ModuleCat R) (f : M →ₗ[R] M) :
@[simp]
theorem ModuleCat.endRingEquiv_apply {R : Type u} [Ring R] (M : ModuleCat R) (f : M.Hom M) :
@[deprecated ModuleCat.endRingEquiv (since := "2025-01-23")]
def ModuleCat.endMulEquiv {R : Type u} [Ring R] (M : ModuleCat R) :

ModuleCat.Hom.hom as an isomorphism of monoids.

Equations

The scalar multiplication on an object of ModuleCat R considered as a morphism of rings from R to the endomorphisms of the underlying abelian group.

Equations
  • M.smul = { toFun := fun (r : R) => AddCommGrp.ofHom { toFun := fun (m : M) => r m, map_zero' := , map_add' := }, map_one' := , map_mul' := , map_zero' := , map_add' := }

The scalar multiplication on ModuleCat R considered as a morphism of rings to the endomorphisms of the forgetful functor to AddCommGrp).

Equations
@[simp]
theorem ModuleCat.smulNatTrans_apply_app (R : Type u) [Ring R] (r : R) (M : ModuleCat R) :
((smulNatTrans R) r).app M = M.smul r

Given A : AddCommGrp and a ring morphism R →+* End A, this is a type synonym for A, on which we shall define a structure of R-module.

Equations
Equations
@[simp]
theorem ModuleCat.mkOfSMul'_smul {R : Type u} [Ring R] {A : AddCommGrp} (φ : R →+* CategoryTheory.End A) (r : R) (x : (mkOfSMul' φ)) :
r x = (CategoryTheory.ConcreteCategory.hom (have this := φ r; this)) x
Equations
@[reducible, inline]

Given A : AddCommGrp and a ring morphism R →+* End A, this is an object in ModuleCat R, whose underlying abelian group is A and whose scalar multiplication is given by R.

Equations
theorem ModuleCat.mkOfSMul_smul {R : Type u} [Ring R] {A : AddCommGrp} (φ : R →+* CategoryTheory.End A) (r : R) :
(mkOfSMul φ).smul r = φ r

Constructor for morphisms in ModuleCat R which takes as inputs a morphism between the underlying objects in AddCommGrp and the compatibility with the scalar multiplication.

Equations
def ModuleCat.ofHom₂ {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M →ₗ[R] N →ₗ[R] P) :
M of R (N P)

Turn a bilinear map into a homomorphism.

Equations
@[simp]
theorem ModuleCat.ofHom₂_hom_apply_hom {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M →ₗ[R] N →ₗ[R] P) (a✝ : M) :
Hom.hom ((Hom.hom (ofHom₂ f)) a✝) = f a✝
def ModuleCat.Hom.hom₂ {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M.Hom (of R (N P))) :
M →ₗ[R] N →ₗ[R] P

Turn a homomorphism into a bilinear map.

Equations
@[simp]
theorem ModuleCat.Hom.hom₂_apply {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M.Hom (of R (N P))) (a✝ : M) :
f.hom₂ a✝ = (ofHom homLinearEquiv).hom' (f.hom' a✝)
@[simp]
theorem ModuleCat.Hom.hom₂_ofHom₂ {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M →ₗ[R] N →ₗ[R] P) :
@[simp]
theorem ModuleCat.ofHom₂_hom₂ {R : Type u_1} [CommRing R] {M N P : ModuleCat R} (f : M of R (N P)) :

@[simp] lemmas for LinearMap.comp and categorical identities.