Further basic results about modules. #
@[simp]
theorem
invOf_two_smul_add_invOf_two_smul
{M : Type u_3}
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
[Invertible 2]
(x : M)
:
theorem
map_inv_natCast_smul
{M : Type u_3}
{M₂ : Type u_4}
[AddCommMonoid M]
[AddCommMonoid M₂]
{F : Type u_5}
[FunLike F M M₂]
[AddMonoidHomClass F M M₂]
(f : F)
(R : Type u_6)
(S : Type u_7)
[DivisionSemiring R]
[DivisionSemiring S]
[Module R M]
[Module S M₂]
(n : ℕ)
(x : M)
:
theorem
map_inv_intCast_smul
{M : Type u_3}
{M₂ : Type u_4}
[AddCommGroup M]
[AddCommGroup M₂]
{F : Type u_5}
[FunLike F M M₂]
[AddMonoidHomClass F M M₂]
(f : F)
(R : Type u_6)
(S : Type u_7)
[DivisionRing R]
[DivisionRing S]
[Module R M]
[Module S M₂]
(z : ℤ)
(x : M)
:
theorem
inv_natCast_smul_eq
{E : Type u_5}
(R : Type u_6)
(S : Type u_7)
[AddCommMonoid E]
[DivisionSemiring R]
[DivisionSemiring S]
[Module R E]
[Module S E]
(n : ℕ)
(x : E)
:
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_intCast_smul_eq
{E : Type u_5}
(R : Type u_6)
(S : Type u_7)
[AddCommGroup E]
[DivisionRing R]
[DivisionRing S]
[Module R E]
[Module S E]
(n : ℤ)
(x : E)
:
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_natCast_smul_comm
{α : Type u_5}
{E : Type u_6}
(R : Type u_7)
[AddCommMonoid E]
[DivisionSemiring R]
[Monoid α]
[Module R E]
[DistribMulAction α E]
(n : ℕ)
(s : α)
(x : E)
:
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_intCast_smul_comm
{α : Type u_5}
{E : Type u_6}
(R : Type u_7)
[AddCommGroup E]
[DivisionRing R]
[Monoid α]
[Module R E]
[DistribMulAction α E]
(n : ℤ)
(s : α)
(x : E)
:
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
Function.support_const_smul_of_ne_zero
{α : Type u_1}
{R : Type u_2}
{M : Type u_3}
[Zero R]
[Zero M]
[SMulWithZero R M]
[NoZeroSMulDivisors R M]
(c : R)
(g : α → M)
(hc : c ≠ 0)
:
theorem
Function.support_smul
{α : Type u_1}
{R : Type u_2}
{M : Type u_3}
[Zero R]
[Zero M]
[SMulWithZero R M]
[NoZeroSMulDivisors R M]
(f : α → R)
(g : α → M)
:
theorem
Set.smul_indicator_one_apply
{α : Type u_1}
{R : Type u_2}
[MulZeroOneClass R]
(s : Set α)
(r : R)
(a : α)
: