Lemmas about different kinds of "lifts" to SkewMonoidAlgebra
. #
liftNCRingHom
as an AlgHom
, for when f
is an AlgHom
Equations
- SkewMonoidAlgebra.liftNCAlgHom f g h_comm = { toRingHom := SkewMonoidAlgebra.liftNCRingHom (↑f) g ⋯, commutes' := ⋯ }
Instances For
Any monoid homomorphism G →* A
can be lifted to an algebra homomorphism
SkewMonoidAlgebra k G →ₐ[k] A
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Decomposition of a k
-algebra homomorphism from SkewMonoidAlgebra k G
by
its values on F (single a 1)
.
If f : G → H
is a multiplicative homomorphism between two monoids, then
mapDomain f
is an algebra homomorphism between their monoid algebras.
Equations
- SkewMonoidAlgebra.mapDomainAlgHom k A hf = { toRingHom := SkewMonoidAlgebra.mapDomainRingHom hf, commutes' := ⋯ }
Instances For
Given f : G ≃ H
, we can map l : SkewMonoidAlgebra k G
to
equivMapDomain f l : SkewMonoidAlgebra k H
(computably) by mapping the support forwards
and the function backwards.
Equations
- SkewMonoidAlgebra.equivMapDomain f l = { toFinsupp := { support := Finset.map f.toEmbedding l.support, toFun := fun (a : H) => l.coeff (f.symm a), mem_support_toFun := ⋯ } }
Instances For
Given AddCommMonoid A
and e : G ≃ H
, domCongr e
is the corresponding Equiv
between
SkewMonoidAlgebra A G
and SkewMonoidAlgebra A H
.
Equations
- SkewMonoidAlgebra.domCongr e = { toFun := SkewMonoidAlgebra.equivMapDomain e, invFun := SkewMonoidAlgebra.equivMapDomain e.symm, left_inv := ⋯, right_inv := ⋯, map_add' := ⋯ }
Instances For
An equivalence of domains induces a linear equivalence of finitely supported functions.
This is domCongr
as a LinearEquiv
.
Equations
Instances For
If e : G ≃* H
is a multiplicative equivalence between two monoids and
∀ (a : G) (x : A), a • x = (e a) • x
, then SkewMonoidAlgebra.domCongr e
is an
algebra equivalence between their skew monoid algebras.
Equations
Instances For
A submodule over k
which is stable under scalar multiplication by elements of G
is a
submodule over SkewMonoidAlgebra k G
Equations
- SkewMonoidAlgebra.submoduleOfSmulMem W h = { carrier := ↑W, add_mem' := ⋯, zero_mem' := ⋯, smul_mem' := ⋯ }