Documentation

Mathlib.Algebra.SMulWithZero

Introduce SMulWithZero #

In analogy with the usual monoid action on a Type M, we introduce an action of a MonoidWithZero on a Type with 0.

In particular, for Types R and M, both containing 0, we define SMulWithZero R M to be the typeclass where the products r • 0 and 0 • m vanish for all r : R and all m : M.

Moreover, in the case in which R is a MonoidWithZero, we introduce the typeclass MulActionWithZero R M, mimicking group actions and having an absorbing 0 in R. Thus, the action is required to be compatible with

We also add an instance:

Main declarations #

class SMulWithZero (R : Type u_1) (M : Type u_3) [Zero R] [Zero M] extends SMulZeroClass , SMul :
Type (max u_1 u_3)

SMulWithZero is a class consisting of a Type R with 0 ∈ R and a scalar multiplication of R on a Type M with 0, such that the equality r • m = 0 holds if at least one among r or m equals 0.

    Instances
      theorem SMulWithZero.zero_smul {R : Type u_1} {M : Type u_3} :
      ∀ {inst : Zero R} {inst_1 : Zero M} [self : SMulWithZero R M] (m : M), 0 m = 0

      Scalar multiplication by the scalar 0 is 0.

      @[simp]
      theorem zero_smul (R : Type u_1) {M : Type u_3} [Zero R] [Zero M] [SMulWithZero R M] (m : M) :
      0 m = 0
      theorem smul_eq_zero_of_left {R : Type u_1} {M : Type u_3} [Zero R] [Zero M] [SMulWithZero R M] {a : R} (h : a = 0) (b : M) :
      a b = 0
      theorem left_ne_zero_of_smul {R : Type u_1} {M : Type u_3} [Zero R] [Zero M] [SMulWithZero R M] {a : R} {b : M} :
      a b 0a 0
      @[reducible, inline]
      abbrev Function.Injective.smulWithZero {R : Type u_1} {M : Type u_3} {M' : Type u_4} [Zero R] [Zero M] [SMulWithZero R M] [Zero M'] [SMul R M'] (f : ZeroHom M' M) (hf : Function.Injective f) (smul : ∀ (a : R) (b : M'), f (a b) = a f b) :

      Pullback a SMulWithZero structure along an injective zero-preserving homomorphism. See note [reducible non-instances].

      Equations
      Instances For
        @[reducible, inline]
        abbrev Function.Surjective.smulWithZero {R : Type u_1} {M : Type u_3} {M' : Type u_4} [Zero R] [Zero M] [SMulWithZero R M] [Zero M'] [SMul R M'] (f : ZeroHom M M') (hf : Function.Surjective f) (smul : ∀ (a : R) (b : M), f (a b) = a f b) :

        Pushforward a SMulWithZero structure along a surjective zero-preserving homomorphism. See note [reducible non-instances].

        Equations
        Instances For
          def SMulWithZero.compHom {R : Type u_1} {R' : Type u_2} (M : Type u_3) [Zero R] [Zero M] [SMulWithZero R M] [Zero R'] (f : ZeroHom R' R) :

          Compose a SMulWithZero with a ZeroHom, with action f r' • m

          Equations
          Instances For
            Equations
            Equations
            class MulActionWithZero (R : Type u_1) (M : Type u_3) [MonoidWithZero R] [Zero M] extends MulAction , SMul :
            Type (max u_1 u_3)

            An action of a monoid with zero R on a Type M, also with 0, extends MulAction and is compatible with 0 (both in R and in M), with 1 ∈ R, and with associativity of multiplication on the monoid M.

            • smul : RMM
            • one_smul : ∀ (b : M), 1 b = b
            • mul_smul : ∀ (x y : R) (b : M), (x * y) b = x y b
            • smul_zero : ∀ (r : R), r 0 = 0

              Scalar multiplication by any element send 0 to 0.

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

              Scalar multiplication by the scalar 0 is 0.

            Instances
              theorem MulActionWithZero.smul_zero {R : Type u_1} {M : Type u_3} :
              ∀ {inst : MonoidWithZero R} {inst_1 : Zero M} [self : MulActionWithZero R M] (r : R), r 0 = 0

              Scalar multiplication by any element send 0 to 0.

              theorem MulActionWithZero.zero_smul {R : Type u_1} {M : Type u_3} :
              ∀ {inst : MonoidWithZero R} {inst_1 : Zero M} [self : MulActionWithZero R M] (m : M), 0 m = 0

              Scalar multiplication by the scalar 0 is 0.

              @[instance 100]
              instance MulActionWithZero.toSMulWithZero (R : Type u_5) (M : Type u_6) :
              {x : MonoidWithZero R} → {x_1 : Zero M} → [m : MulActionWithZero R M] → SMulWithZero R M
              Equations
              theorem ite_zero_smul {R : Type u_1} {M : Type u_3} [MonoidWithZero R] [Zero M] [MulActionWithZero R M] (p : Prop) [Decidable p] (a : R) (b : M) :
              (if p then a else 0) b = if p then a b else 0
              theorem boole_smul {R : Type u_1} {M : Type u_3} [MonoidWithZero R] [Zero M] [MulActionWithZero R M] (p : Prop) [Decidable p] (a : M) :
              (if p then 1 else 0) a = if p then a else 0
              theorem Pi.single_apply_smul {R : Type u_1} {M : Type u_3} [MonoidWithZero R] [Zero M] [MulActionWithZero R M] {ι : Type u_5} [DecidableEq ι] (x : M) (i : ι) (j : ι) :
              Pi.single i 1 j x = Pi.single i x j
              @[reducible, inline]
              abbrev Function.Injective.mulActionWithZero {R : Type u_1} {M : Type u_3} {M' : Type u_4} [MonoidWithZero R] [Zero M] [MulActionWithZero R M] [Zero M'] [SMul R M'] (f : ZeroHom M' M) (hf : Function.Injective f) (smul : ∀ (a : R) (b : M'), f (a b) = a f b) :

              Pullback a MulActionWithZero structure along an injective zero-preserving homomorphism. See note [reducible non-instances].

              Equations
              Instances For
                @[reducible, inline]
                abbrev Function.Surjective.mulActionWithZero {R : Type u_1} {M : Type u_3} {M' : Type u_4} [MonoidWithZero R] [Zero M] [MulActionWithZero R M] [Zero M'] [SMul R M'] (f : ZeroHom M M') (hf : Function.Surjective f) (smul : ∀ (a : R) (b : M), f (a b) = a f b) :

                Pushforward a MulActionWithZero structure along a surjective zero-preserving homomorphism. See note [reducible non-instances].

                Equations
                Instances For
                  def MulActionWithZero.compHom {R : Type u_1} {R' : Type u_2} (M : Type u_3) [MonoidWithZero R] [MonoidWithZero R'] [Zero M] [MulActionWithZero R M] (f : R' →*₀ R) :

                  Compose a MulActionWithZero with a MonoidWithZeroHom, with action f r' • m

                  Equations
                  Instances For
                    theorem smul_inv₀ {α : Type u_5} {β : Type u_6} [GroupWithZero α] [GroupWithZero β] [MulActionWithZero α β] [SMulCommClass α β β] [IsScalarTower α β β] (c : α) (x : β) :