Documentation

Mathlib.Algebra.Module.Basic

Modules over a ring #

In this file we define

Implementation notes #

In typical mathematical usage, our definition of Module corresponds to "semimodule", and the word "module" is reserved for Module R M where R is a Ring and M an AddCommGroup. If R is a Field and M an AddCommGroup, M would be called an R-vector space. Since those assumptions can be made by changing the typeclasses applied to R and M, without changing the axioms in Module, mathlib calls everything a Module.

In older versions of mathlib3, we had separate semimodule and vector_space abbreviations. This caused inference issues in some cases, while not providing any real advantages, so we decided to use a canonical Module typeclass throughout.

Tags #

semimodule, module, vector space

theorem Module.ext {R : Type u} {M : Type v} :
∀ {inst : Semiring R} {inst_1 : AddCommMonoid M} (x y : Module R M), SMul.smul = SMul.smulx = y
theorem Module.ext_iff {R : Type u} {M : Type v} :
∀ {inst : Semiring R} {inst_1 : AddCommMonoid M} (x y : Module R M), x = y SMul.smul = SMul.smul
class Module (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] extends DistribMulAction :
Type (max u v)
  • smul : RMM
  • one_smul : ∀ (b : M), 1 b = b
  • mul_smul : ∀ (x y : R) (b : M), (x * y) b = x y b
  • smul_zero : ∀ (a : R), a 0 = 0
  • smul_add : ∀ (a : R) (x y : M), a (x + y) = a x + a y
  • add_smul : ∀ (r s_1 : R) (x : M), (r + s_1) x = r x + s_1 x

    Scalar multiplication distributes over addition from the right.

  • zero_smul : ∀ (x : M), 0 x = 0

    Scalar multiplication by zero gives zero.

A module is a generalization of vector spaces to a scalar semiring. It consists of a scalar semiring R and an additive monoid of "vectors" M, connected by a "scalar multiplication" operation r • x : M (where r : R and x : M) with some natural associativity and distributivity axioms similar to those on a ring.

Instances
    instance Module.toMulActionWithZero {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] :

    A module over a semiring automatically inherits a MulActionWithZero structure.

    Equations
    Equations
    theorem add_smul {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (r : R) (s : R) (x : M) :
    (r + s) x = r x + s x
    theorem Convex.combo_self {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] {a : R} {b : R} (h : a + b = 1) (x : M) :
    a x + b x = x
    theorem two_smul (R : Type u_2) {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (x : M) :
    2 x = x + x
    @[deprecated]
    theorem two_smul' (R : Type u_2) {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (x : M) :
    2 x = bit0 x
    @[simp]
    theorem invOf_two_smul_add_invOf_two_smul (R : Type u_2) {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] [Invertible 2] (x : M) :
    2 x + 2 x = x
    @[reducible]
    def Function.Injective.module (R : Type u_2) {M : Type u_5} {M₂ : Type u_6} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [SMul R M₂] (f : M₂ →+ M) (hf : Function.Injective f) (smul : ∀ (c : R) (x : M₂), f (c x) = c f x) :
    Module R M₂

    Pullback a Module structure along an injective additive monoid homomorphism. See note [reducible non-instances].

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def Function.Surjective.module (R : Type u_2) {M : Type u_5} {M₂ : Type u_6} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid M₂] [SMul R M₂] (f : M →+ M₂) (hf : Function.Surjective f) (smul : ∀ (c : R) (x : M), f (c x) = c f x) :
      Module R M₂

      Pushforward a Module structure along a surjective additive monoid homomorphism.

      Equations
      Instances For
        @[reducible]
        def Function.Surjective.moduleLeft {R : Type u_9} {S : Type u_10} {M : Type u_11} [Semiring R] [AddCommMonoid M] [Module R M] [Semiring S] [SMul S M] (f : R →+* S) (hf : Function.Surjective f) (hsmul : ∀ (c : R) (x : M), f c x = c x) :
        Module S M

        Push forward the action of R on M along a compatible surjective map f : R →+* S.

        See also Function.Surjective.mulActionLeft and Function.Surjective.distribMulActionLeft.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[reducible]
          def Module.compHom {R : Type u_2} {S : Type u_4} (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [Semiring S] (f : S →+* R) :
          Module S M

          Compose a Module with a RingHom, with action f s • m.

          See note [reducible non-instances].

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem Module.toAddMonoidEnd_apply_apply (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] (x : R) :
            ∀ (x : M), ↑(↑(Module.toAddMonoidEnd R M) x) x = x x
            def Module.toAddMonoidEnd (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] :

            (•) as an AddMonoidHom.

            This is a stronger version of DistribMulAction.toAddMonoidEnd

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              def smulAddHom (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] :
              R →+ M →+ M

              A convenience alias for Module.toAddMonoidEnd as an AddMonoidHom, usually to allow the use of AddMonoidHom.flip.

              Equations
              Instances For
                @[simp]
                theorem smulAddHom_apply {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (r : R) (x : M) :
                ↑(↑(smulAddHom R M) r) x = r x
                theorem Module.eq_zero_of_zero_eq_one {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (x : M) (zero_eq_one : 0 = 1) :
                x = 0
                @[simp]
                theorem smul_add_one_sub_smul {M : Type u_5} [AddCommMonoid M] {R : Type u_9} [Ring R] [Module R M] {r : R} {m : M} :
                r m + (1 - r) m = m
                Equations
                structure Module.Core (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommGroup M] extends SMul :
                Type (max u_2 u_5)
                • smul : RMM
                • smul_add : ∀ (r : R) (x y : M), r (x + y) = r x + r y

                  Scalar multiplication distributes over addition from the left.

                • add_smul : ∀ (r s_1 : R) (x : M), (r + s_1) x = r x + s_1 x

                  Scalar multiplication distributes over addition from the right.

                • mul_smul : ∀ (r s_1 : R) (x : M), (r * s_1) x = r s_1 x

                  Scalar multiplication distributes over multiplication from the right.

                • one_smul : ∀ (x : M), 1 x = x

                  Scalar multiplication by one is the identity.

                A structure containing most informations as in a module, except the fields zero_smul and smul_zero. As these fields can be deduced from the other ones when M is an AddCommGroup, this provides a way to construct a module structure by checking less properties, in Module.ofCore.

                Instances For
                  def Module.ofCore {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommGroup M] (H : Module.Core R M) :
                  Module R M

                  Define Module without proving zero_smul and smul_zero by using an auxiliary structure Module.Core, when the underlying space is an AddCommGroup.

                  Equations
                  Instances For
                    theorem Convex.combo_eq_smul_sub_add {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommGroup M] [Module R M] {x : M} {y : M} {a : R} {b : R} (h : a + b = 1) :
                    a x + b y = b (y - x) + x
                    theorem Module.ext' {R : Type u_9} [Semiring R] {M : Type u_10} [AddCommMonoid M] (P : Module R M) (Q : Module R M) (w : ∀ (r : R) (m : M), r m = r m) :
                    P = Q

                    A variant of Module.ext that's convenient for term-mode.

                    @[simp]
                    theorem neg_smul {R : Type u_2} {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) :
                    -r x = -(r x)
                    theorem neg_smul_neg {R : Type u_2} {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (r : R) (x : M) :
                    -r -x = r x
                    @[simp]
                    theorem Units.neg_smul {R : Type u_2} {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (u : Rˣ) (x : M) :
                    -u x = -(u x)
                    theorem neg_one_smul (R : Type u_2) {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (x : M) :
                    -1 x = -x
                    theorem sub_smul {R : Type u_2} {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (r : R) (s : R) (y : M) :
                    (r - s) y = r y - s y
                    @[reducible]

                    An AddCommMonoid that is a Module over a Ring carries a natural AddCommGroup structure. See note [reducible non-instances].

                    Equations
                    Instances For
                      theorem Module.subsingleton (R : Type u_9) (M : Type u_10) [Semiring R] [Subsingleton R] [AddCommMonoid M] [Module R M] :

                      A module over a Subsingleton semiring is a Subsingleton. We cannot register this as an instance because Lean has no way to guess R.

                      theorem Module.nontrivial (R : Type u_9) (M : Type u_10) [Semiring R] [Nontrivial M] [AddCommMonoid M] [Module R M] :

                      A semiring is Nontrivial provided that there exists a nontrivial module over this semiring.

                      instance Semiring.toModule {R : Type u_2} [Semiring R] :
                      Module R R
                      Equations
                      • Semiring.toModule = Module.mk (_ : ∀ (a b c : R), (a + b) * c = a * c + b * c) (_ : ∀ (a : R), 0 * a = 0)

                      Like Semiring.toModule, but multiplies on the right.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      def RingHom.toModule {R : Type u_2} {S : Type u_4} [Semiring R] [Semiring S] (f : R →+* S) :
                      Module R S

                      A ring homomorphism f : R →+* M defines a module structure by r • x = f r * x.

                      Equations
                      Instances For

                        The tautological action by R →+* R on R.

                        This generalizes Function.End.applyMulAction.

                        Equations
                        @[simp]
                        theorem RingHom.smul_def {R : Type u_2} [Semiring R] (f : R →+* R) (a : R) :
                        f a = f a
                        theorem nsmul_eq_smul_cast (R : Type u_2) {M : Type u_5} [Semiring R] [AddCommMonoid M] [Module R M] (n : ) (b : M) :
                        n b = n b

                        nsmul is equal to any other module structure via a cast.

                        theorem nat_smul_eq_nsmul {M : Type u_5} [AddCommMonoid M] (h : Module M) (n : ) (x : M) :
                        SMul.smul n x = n x

                        Convert back any exotic -smul to the canonical instance. This should not be needed since in mathlib all AddCommMonoids should normally have exactly one -module structure by design.

                        All -module structures are equal. Not an instance since in mathlib all AddCommMonoid should normally have exactly one -module structure by design.

                        Equations
                        • AddCommMonoid.natModule.unique = { toInhabited := { default := inferInstance }, uniq := (_ : ∀ (P : Module M), P = default) }
                        Instances For
                          theorem zsmul_eq_smul_cast (R : Type u_2) {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] (n : ) (b : M) :
                          n b = n b

                          zsmul is equal to any other module structure via a cast.

                          theorem int_smul_eq_zsmul {M : Type u_5} [AddCommGroup M] (h : Module M) (n : ) (x : M) :
                          SMul.smul n x = n x

                          Convert back any exotic -smul to the canonical instance. This should not be needed since in mathlib all AddCommGroups should normally have exactly one -module structure by design.

                          All -module structures are equal. Not an instance since in mathlib all AddCommGroup should normally have exactly one -module structure by design.

                          Equations
                          • AddCommGroup.intModule.unique = { toInhabited := { default := inferInstance }, uniq := (_ : ∀ (P : Module M), P = default) }
                          Instances For
                            theorem map_int_cast_smul {M : Type u_5} {M₂ : Type u_6} [AddCommGroup M] [AddCommGroup M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (R : Type u_10) (S : Type u_11) [Ring R] [Ring S] [Module R M] [Module S M₂] (x : ) (a : M) :
                            f (x a) = x f a
                            theorem map_nat_cast_smul {M : Type u_5} {M₂ : Type u_6} [AddCommMonoid M] [AddCommMonoid M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (R : Type u_10) (S : Type u_11) [Semiring R] [Semiring S] [Module R M] [Module S M₂] (x : ) (a : M) :
                            f (x a) = x f a
                            theorem map_inv_nat_cast_smul {M : Type u_5} {M₂ : Type u_6} [AddCommMonoid M] [AddCommMonoid M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (R : Type u_10) (S : Type u_11) [DivisionSemiring R] [DivisionSemiring S] [Module R M] [Module S M₂] (n : ) (x : M) :
                            f ((n)⁻¹ x) = (n)⁻¹ f x
                            theorem map_inv_int_cast_smul {M : Type u_5} {M₂ : Type u_6} [AddCommGroup M] [AddCommGroup M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (R : Type u_10) (S : Type u_11) [DivisionRing R] [DivisionRing S] [Module R M] [Module S M₂] (z : ) (x : M) :
                            f ((z)⁻¹ x) = (z)⁻¹ f x
                            theorem map_rat_cast_smul {M : Type u_5} {M₂ : Type u_6} [AddCommGroup M] [AddCommGroup M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (R : Type u_10) (S : Type u_11) [DivisionRing R] [DivisionRing S] [Module R M] [Module S M₂] (c : ) (x : M) :
                            f (c x) = c f x
                            theorem map_rat_smul {M : Type u_5} {M₂ : Type u_6} [AddCommGroup M] [AddCommGroup M₂] [Module M] [Module M₂] {F : Type u_9} [AddMonoidHomClass F M M₂] (f : F) (c : ) (x : M) :
                            f (c x) = c f x

                            There can be at most one Module ℚ E structure on an additive commutative group.

                            Equations
                            theorem inv_nat_cast_smul_eq {E : Type u_9} (R : Type u_10) (S : Type u_11) [AddCommMonoid E] [DivisionSemiring R] [DivisionSemiring S] [Module R E] [Module S E] (n : ) (x : E) :
                            (n)⁻¹ x = (n)⁻¹ x

                            If E is a vector space over two division semirings R and S, then scalar multiplications agree on inverses of natural numbers in R and S.

                            theorem inv_int_cast_smul_eq {E : Type u_9} (R : Type u_10) (S : Type u_11) [AddCommGroup E] [DivisionRing R] [DivisionRing S] [Module R E] [Module S E] (n : ) (x : E) :
                            (n)⁻¹ x = (n)⁻¹ x

                            If E is a vector space over two division rings R and S, then scalar multiplications agree on inverses of integer numbers in R and S.

                            theorem inv_nat_cast_smul_comm {α : Type u_9} {E : Type u_10} (R : Type u_11) [AddCommMonoid E] [DivisionSemiring R] [Monoid α] [Module R E] [DistribMulAction α E] (n : ) (s : α) (x : E) :
                            (n)⁻¹ s x = s (n)⁻¹ x

                            If E is a vector space over a division semiring R and has a monoid action by α, then that action commutes by scalar multiplication of inverses of natural numbers in R.

                            theorem inv_int_cast_smul_comm {α : Type u_9} {E : Type u_10} (R : Type u_11) [AddCommGroup E] [DivisionRing R] [Monoid α] [Module R E] [DistribMulAction α E] (n : ) (s : α) (x : E) :
                            (n)⁻¹ s x = s (n)⁻¹ x

                            If E is a vector space over a division ring R and has a monoid action by α, then that action commutes by scalar multiplication of inverses of integers in R

                            theorem rat_cast_smul_eq {E : Type u_9} (R : Type u_10) (S : Type u_11) [AddCommGroup E] [DivisionRing R] [DivisionRing S] [Module R E] [Module S E] (r : ) (x : E) :
                            r x = r x

                            If E is a vector space over two division rings R and S, then scalar multiplications agree on rational numbers in R and S.

                            NoZeroSMulDivisors #

                            This section defines the NoZeroSMulDivisors class, and includes some tests for the vanishing of elements (especially in modules over division rings).

                            class NoZeroSMulDivisors (R : Type u_9) (M : Type u_10) [Zero R] [Zero M] [SMul R M] :
                            • eq_zero_or_eq_zero_of_smul_eq_zero : ∀ {c : R} {x : M}, c x = 0c = 0 x = 0

                              If scalar multiplication yields zero, either the scalar or the vector was zero.

                            NoZeroSMulDivisors R M states that a scalar multiple is 0 only if either argument is 0. This is a version of saying that M is torsion free, without assuming R is zero-divisor free.

                            The main application of NoZeroSMulDivisors R M, when M is a module, is the result smul_eq_zero: a scalar multiple is 0 iff either argument is 0.

                            It is a generalization of the NoZeroDivisors class to heterogeneous multiplication.

                            Instances
                              theorem Function.Injective.noZeroSMulDivisors {R : Type u_9} {M : Type u_10} {N : Type u_11} [Zero R] [Zero M] [Zero N] [SMul R M] [SMul R N] [NoZeroSMulDivisors R N] (f : MN) (hf : Function.Injective f) (h0 : f 0 = 0) (hs : ∀ (c : R) (x : M), f (c x) = c f x) :

                              Pullback a NoZeroSMulDivisors instance along an injective function.

                              theorem smul_ne_zero {R : Type u_2} {M : Type u_5} [Zero R] [Zero M] [SMul R M] [NoZeroSMulDivisors R M] {c : R} {x : M} (hc : c 0) (hx : x 0) :
                              c x 0
                              @[simp]
                              theorem smul_eq_zero {R : Type u_2} {M : Type u_5} [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} :
                              c x = 0 c = 0 x = 0
                              theorem smul_ne_zero_iff {R : Type u_2} {M : Type u_5} [Zero R] [Zero M] [SMulWithZero R M] [NoZeroSMulDivisors R M] {c : R} {x : M} :
                              c x 0 c 0 x 0
                              theorem two_nsmul_eq_zero (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommMonoid M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] {v : M} :
                              2 v = 0 v = 0
                              theorem CharZero.of_module (R : Type u_2) [Semiring R] (M : Type u_9) [AddCommMonoidWithOne M] [CharZero M] [Module R M] :

                              If M is an R-module with one and M has characteristic zero, then R has characteristic zero as well. Usually M is an R-algebra.

                              theorem smul_right_injective {R : Type u_2} (M : Type u_5) [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] {c : R} (hc : c 0) :
                              Function.Injective ((fun x x_1 => x x_1) c)
                              theorem smul_right_inj {R : Type u_2} {M : Type u_5} [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] {c : R} (hc : c 0) {x : M} {y : M} :
                              c x = c y x = y
                              theorem self_eq_neg (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] {v : M} :
                              v = -v v = 0
                              theorem neg_eq_self (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] {v : M} :
                              -v = v v = 0
                              theorem self_ne_neg (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] {v : M} :
                              v -v v 0
                              theorem neg_ne_self (R : Type u_2) (M : Type u_5) [Semiring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] [CharZero R] {v : M} :
                              -v v v 0
                              theorem smul_left_injective (R : Type u_2) {M : Type u_5} [Ring R] [AddCommGroup M] [Module R M] [NoZeroSMulDivisors R M] {x : M} (hx : x 0) :
                              Function.Injective fun c => c x
                              theorem Nat.smul_one_eq_coe {R : Type u_9} [Semiring R] (m : ) :
                              m 1 = m
                              theorem Int.smul_one_eq_coe {R : Type u_9} [Ring R] (m : ) :
                              m 1 = m