Morphisms of non-unital algebras #
This file defines morphisms between two types, each of which carries:
- an addition,
- an additive zero,
- a multiplication,
- a scalar action.
The multiplications are not assumed to be associative or unital, or even to be compatible with the scalar actions. In a typical application, the operations will satisfy compatibility conditions making them into algebras (albeit possibly non-associative and/or non-unital) but such conditions are not required to make this definition.
This notion of morphism should be useful for any category of non-unital algebras. The motivating
application at the time it was introduced was to be able to state the adjunction property for
magma algebras. These are non-unital, non-associative algebras obtained by applying the
group-algebra construction except where we take a type carrying just Mul
instead of Group
.
For a plausible future application, one could take the non-unital algebra of compactly-supported
functions on a non-compact topological space. A proper map between a pair of such spaces
(contravariantly) induces a morphism between their algebras of compactly-supported functions which
will be a NonUnitalAlgHom
.
TODO: add NonUnitalAlgEquiv
when needed.
Main definitions #
Tags #
non-unital, algebra, morphism
- toFun : A → B
- map_smul' : ∀ (m : R) (x : A), MulActionHom.toFun s.toMulActionHom (m • x) = m • MulActionHom.toFun s.toMulActionHom x
- map_zero' : MulActionHom.toFun s.toMulActionHom 0 = 0
- map_add' : ∀ (x y : A), MulActionHom.toFun s.toMulActionHom (x + y) = MulActionHom.toFun s.toMulActionHom x + MulActionHom.toFun s.toMulActionHom y
- map_mul' : ∀ (x y : A), MulActionHom.toFun s.toMulActionHom (x * y) = MulActionHom.toFun s.toMulActionHom x * MulActionHom.toFun s.toMulActionHom y
The proposition that the function preserves multiplication
A morphism respecting addition, multiplication, and scalar multiplication. When these arise from algebra structures, this is the same as a not-necessarily-unital morphism of algebras.
Instances For
A morphism respecting addition, multiplication, and scalar multiplication. When these arise from algebra structures, this is the same as a not-necessarily-unital morphism of algebras.
Instances For
A morphism respecting addition, multiplication, and scalar multiplication. When these arise from algebra structures, this is the same as a not-necessarily-unital morphism of algebras.
Instances For
- coe : F → A → B
- coe_injective' : Function.Injective FunLike.coe
- map_zero : ∀ (f : F), ↑f 0 = 0
The proposition that the function preserves multiplication
NonUnitalAlgHomClass F R A B
asserts F
is a type of bundled algebra homomorphisms
from A
to B
.
Instances
Turn an element of a type F
satisfying NonUnitalAlgHomClass F R A B
into an actual
NonUnitalAlgHom
. This is declared as the default coercion from F
to A →ₙₐ[R] B
.
Instances For
See Note [custom simps projection]
Instances For
The identity map as a NonUnitalAlgHom
.
Instances For
The composition of morphisms is a morphism.
Instances For
The inverse of a bijective morphism is a morphism.
Instances For
Operations on the product type #
Note that much of this is copied from LinearAlgebra/Prod
.
The first projection of a product is a non-unital alg_hom.
Instances For
The second projection of a product is a non-unital alg_hom.
Instances For
The prod of two morphisms is a morphism.
Instances For
Taking the product of two maps with the same domain is equivalent to taking the product of their codomains.
Instances For
The left injection into a product is a non-unital algebra homomorphism.
Instances For
The right injection into a product is a non-unital algebra homomorphism.
Instances For
A unital morphism of algebras is a NonUnitalAlgHom
.