Algebras over commutative semirings #
In this file we define associative unital algebra
s over commutative (semi)rings, algebra
homomorphisms alg_hom
, and algebra equivalences alg_equiv
.
subalgebra
s are defined in algebra.algebra.subalgebra
.
For the category of R
-algebras, denoted Algebra R
, see the file
algebra/category/Algebra/basic.lean
.
See the implementation notes for remarks about non-associative and non-unital algebras.
Main definitions: #
algebra R A
: the algebra typeclass.alg_hom R A B
: the type ofR
-algebra morphisms fromA
toB
.alg_equiv R A B
: the type ofR
-algebra isomorphisms betweenA
toB
.algebra_map R A : R →+* A
: the canonical map fromR
toA
, as aring_hom
. This is the preferred spelling of this map.algebra.linear_map R A : R →ₗ[R] A
: the canonical map fromR
toA
, as alinear_map
.algebra.of_id R A : R →ₐ[R] A
: the canonical map fromR
toA
, as nalg_hom
.- Instances of
algebra
in this file:
Notations #
A →ₐ[R] B
:R
-algebra homomorphism fromA
toB
.A ≃ₐ[R] B
:R
-algebra equivalence fromA
toB
.
Implementation notes #
Given a commutative (semi)ring R
, there are two ways to define an R
-algebra structure on a
(possibly noncommutative) (semi)ring A
:
- By endowing
A
with a morphism of ringsR →+* A
denotedalgebra_map R A
which lands in the center ofA
. - By requiring
A
be anR
-module such that the action associates and commutes with multiplication asr • (a₁ * a₂) = (r • a₁) * a₂ = a₁ * (r • a₂)
.
We define algebra R A
in a way that subsumes both definitions, by extending has_smul R A
and
requiring that this scalar action r • x
must agree with left multiplication by the image of the
structure morphism algebra_map R A r * x
.
As a result, there are two ways to talk about an R
-algebra A
when A
is a semiring:
-
variables [comm_semiring R] [semiring A] variables [algebra R A]
-
variables [comm_semiring R] [semiring A] variables [module R A] [smul_comm_class R A A] [is_scalar_tower R A A]
The first approach implies the second via typeclass search; so any lemma stated with the second set of arguments will automatically apply to the first set. Typeclass search does not know that the second approach implies the first, but this can be shown with:
example {R A : Type*} [comm_semiring R] [semiring A]
[module R A] [smul_comm_class R A A] [is_scalar_tower R A A] : algebra R A :=
algebra.of_module smul_mul_assoc mul_smul_comm
The advantage of the first approach is that algebra_map R A
is available, and alg_hom R A B
and
subalgebra R A
can be used. For concrete R
and A
, algebra_map R A
is often definitionally
convenient.
The advantage of the second approach is that comm_semiring R
, semiring A
, and module R A
can
all be relaxed independently; for instance, this allows us to:
- Replace
semiring A
withnon_unital_non_assoc_semiring A
in order to describe non-unital and/or non-associative algebras. - Replace
comm_semiring R
andmodule R A
withcomm_group R'
anddistrib_mul_action R' A
, which whenR' = Rˣ
lets us talk about the "algebra-like" action ofRˣ
on anR
-algebraA
.
While alg_hom R A B
cannot be used in the second approach, non_unital_alg_hom R A B
still can.
You should always use the first approach when working with associative unital algebras, and mimic
the second approach only when you need to weaken a condition on either R
or A
.
- to_has_smul : has_smul R A
- to_ring_hom : R →+* A
- commutes' : ∀ (r : R) (x : A), algebra.to_ring_hom.to_fun r * x = x * algebra.to_ring_hom.to_fun r
- smul_def' : ∀ (r : R) (x : A), r • x = algebra.to_ring_hom.to_fun r * x
An associative unital R
-algebra is a semiring A
equipped with a map into its center R → A
.
See the implementation notes in this file for discussion of the details of this definition.
Instances of this typeclass
- algebra_nat
- algebra_int
- ideal.quotient_algebra
- normed_algebra.to_algebra
- power_series.algebra_polynomial'
- algebra.id
- punit.algebra
- ulift.algebra
- prod.algebra
- algebra.of_subsemiring
- algebra.of_subring
- mul_opposite.algebra
- module.End.algebra
- algebra_rat
- pi.algebra
- function.algebra
- monoid_algebra.algebra
- add_monoid_algebra.algebra
- ideal.quotient.algebra
- subalgebra.algebra'
- subalgebra.algebra
- subalgebra.to_algebra
- mv_polynomial.algebra
- polynomial.algebra_of_algebra
- localization.algebra
- fraction_ring.algebra
- restrict_scalars.algebra
- matrix.algebra
- local_ring.residue_field.algebra
- algebra.tensor_product.left_algebra
- algebra.tensor_product.tensor_product.algebra
- subfield.to_algebra
- intermediate_field.algebra'
- intermediate_field.algebra
- intermediate_field.to_algebra
- adjoin_root.algebra
- polynomial.splitting_field_aux.algebra
- polynomial.splitting_field_aux.algebra'''
- polynomial.splitting_field_aux.algebra'
- polynomial.splitting_field_aux.algebra''
- polynomial.splitting_field.algebra'
- polynomial.splitting_field.algebra
- unitization.algebra
- free_algebra.algebra
- category_theory.linear.category_theory.End.algebra
- Algebra.is_algebra
- Algebra.algebra_obj
- Algebra.limit_algebra
- direct_sum.algebra
- triv_sq_zero_ext.algebra'
- triv_sq_zero_ext.algebra
- intermediate_field.algebra_over_bot
- fixed_points.algebra
- algebraic_closure.adjoin_monic.algebra
- algebraic_closure.step.algebra_succ
- algebraic_closure.step.algebra
- algebraic_closure.algebra_of_step
- algebraic_closure.algebra
- intermediate_field.fixed_field.algebra
- nnreal.algebra
- ennreal.algebra
- complex.algebra
- continuous_linear_map.algebra
- zmod.algebra
- mv_power_series.algebra
- mv_power_series.algebra_mv_polynomial
- mv_power_series.algebra_mv_power_series
- power_series.algebra
- power_series.algebra_polynomial
- power_series.algebra_power_series
- hahn_series.algebra
- hahn_series.power_series_algebra
- laurent_series.algebra
- ratfunc.algebra
- ratfunc.laurent_series.algebra
- Top.presheaf.stalk.algebra
- ideal.quotient.algebra_quotient_map_quotient
- continuous_map.algebra
- algebraic_geometry.structure_sheaf.localizations.algebra
- algebraic_geometry.structure_sheaf.stalk_algebra
- algebraic_geometry.structure_sheaf.open_algebra
- is_localization.localization.at_prime.algebra
- is_localization.localization.algebra
- algebraic_geometry.obj.algebra
- algebraic_geometry.Γ_restrict_algebra
- algebraic_geometry.Scheme.function_field.algebra
- algebraic_geometry.stalk_function_field_algebra
- algebraic_geometry.function_field.algebra
- ring_quot.algebra
- tensor_algebra.algebra
- universal_enveloping_algebra.algebra
- quaternion_algebra.algebra
- quaternion.algebra
- bounded_continuous_function.algebra
- uniform_space.completion.algebra
- uniform_space.completion.algebra'
- pi.algebra_of_normed_algebra
- pre_lp.algebra
- function.algebra_ring
- pi.matrix_algebra
- Module.Mon_Module_equivalence_Algebra.Mon_.X.algebra
- laurent_polynomial.algebra_polynomial
- nnrat.rat.algebra
- polynomial.gal.algebra
- galois_field.algebra
- smooth_map.algebra
- smooth_functions_algebra
- pointed_smooth_map.algebra
- pointed_smooth_map.cont_mdiff_map.algebra
- pointed_smooth_map.eval_algebra
- clifford_algebra.algebra
- cyclotomic_field.algebra
- cyclotomic_ring.cyclotomic_field.algebra
- ideal.quotient.algebra_quotient_pow_ramification_idx
- valuation.algebra
- valuation_subring.algebra
- valuation_subring.of_prime_algebra
- locally_constant.algebra
Instances of other typeclasses for algebra
- algebra.has_sizeof_inst
- nat_algebra_subsingleton
- algebra_rat_subsingleton
- int_algebra_subsingleton
- zmod.algebra.subsingleton
Embedding R →+* A
given by algebra
structure.
Equations
Creating an algebra from a morphism to the center of a semiring.
Equations
- i.to_algebra' h = {to_has_smul := {smul := λ (c : R) (x : S), ⇑i c * x}, to_ring_hom := i, commutes' := h, smul_def' := _}
Creating an algebra from a morphism to a commutative semiring.
Equations
- i.to_algebra = i.to_algebra' _
Let R
be a commutative semiring, let A
be a semiring with a module R
structure.
If (r • 1) * x = x * (r • 1) = r • x
for all r : R
and x : A
, then A
is an algebra
over R
.
See note [reducible non-instances].
Equations
- algebra.of_module' h₁ h₂ = {to_has_smul := smul_with_zero.to_has_smul (mul_action_with_zero.to_smul_with_zero R A), to_ring_hom := {to_fun := λ (r : R), r • 1, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}, commutes' := _, smul_def' := _}
Let R
be a commutative semiring, let A
be a semiring with a module R
structure.
If (r • x) * y = x * (r • y) = r • (x * y)
for all r : R
and x y : A
, then A
is an algebra
over R
.
See note [reducible non-instances].
Equations
- algebra.of_module h₁ h₂ = algebra.of_module' _ _
To prove two algebra structures on a fixed [comm_semiring R] [semiring A]
agree,
it suffices to check the algebra_map
s agree.
Equations
- algebra.to_module = {to_distrib_mul_action := {to_mul_action := {to_has_smul := algebra.to_has_smul _inst_4, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
mul_comm
for algebra
s when one element is from the base ring.
mul_left_comm
for algebra
s when one element is from the base ring.
mul_right_comm
for algebra
s when one element is from the base ring.
This is just a special case of the global mul_smul_comm
lemma that requires less typeclass
search (and was here first).
This is just a special case of the global smul_mul_assoc
lemma that requires less typeclass
search (and was here first).
The canonical ring homomorphism algebra_map R A : R →* A
for any R
-algebra A
,
packaged as an R
-linear map.
Equations
- algebra.linear_map R A = {to_fun := (algebra_map R A).to_fun, map_add' := _, map_smul' := _}
Equations
- algebra.id R = (ring_hom.id R).to_algebra
Equations
- punit.algebra = {to_has_smul := punit.has_smul R, to_ring_hom := {to_fun := λ (x : R), punit.star, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}, commutes' := _, smul_def' := _}
Equations
- ulift.algebra = {to_has_smul := mul_action.to_has_smul distrib_mul_action.to_mul_action, to_ring_hom := {to_fun := λ (r : R), {down := ⇑(algebra_map R A) r}, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}, commutes' := _, smul_def' := _}
Equations
- prod.algebra R A B = {to_has_smul := mul_action.to_has_smul distrib_mul_action.to_mul_action, to_ring_hom := {to_fun := ((algebra_map R A).prod (algebra_map R B)).to_fun, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}, commutes' := _, smul_def' := _}
Algebra over a subsemiring. This builds upon subsemiring.module
.
Equations
- algebra.of_subsemiring S = {to_has_smul := {smul := has_smul.smul S.has_smul}, to_ring_hom := {to_fun := ((algebra_map R A).comp S.subtype).to_fun, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}, commutes' := _, smul_def' := _}
Algebra over a subring. This builds upon subring.module
.
Equations
- algebra.of_subring S = {to_has_smul := {smul := has_smul.smul S.has_smul}, to_ring_hom := algebra.to_ring_hom (algebra.of_subsemiring S.to_subsemiring), commutes' := _, smul_def' := _}
Explicit characterization of the submonoid map in the case of an algebra.
S
is made explicit to help with type inference
Equations
- algebra.algebra_map_submonoid S M = submonoid.map ↑(algebra_map R S) M
A semiring
that is an algebra
over a commutative ring carries a natural ring
structure.
See note [reducible non-instances].
Equations
- algebra.semiring_to_ring R = {add := add_comm_group.add (module.add_comm_monoid_to_add_comm_group R), add_assoc := _, zero := add_comm_group.zero (module.add_comm_monoid_to_add_comm_group R), zero_add := _, add_zero := _, nsmul := add_comm_group.nsmul (module.add_comm_monoid_to_add_comm_group R), nsmul_zero' := _, nsmul_succ' := _, neg := add_comm_group.neg (module.add_comm_monoid_to_add_comm_group R), sub := add_comm_group.sub (module.add_comm_monoid_to_add_comm_group R), sub_eq_add_neg := _, zsmul := add_comm_group.zsmul (module.add_comm_monoid_to_add_comm_group R), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _, int_cast := add_comm_group_with_one.int_cast._default semiring.nat_cast add_comm_group.add _ add_comm_group.zero _ _ add_comm_group.nsmul _ _ semiring.one algebra.semiring_to_ring._proof_17 algebra.semiring_to_ring._proof_18 add_comm_group.neg add_comm_group.sub _ add_comm_group.zsmul _ _ _ _, nat_cast := semiring.nat_cast infer_instance, one := semiring.one infer_instance, nat_cast_zero := _, nat_cast_succ := _, int_cast_of_nat := _, int_cast_neg_succ_of_nat := _, mul := semiring.mul infer_instance, mul_assoc := _, one_mul := _, mul_one := _, npow := semiring.npow infer_instance, npow_zero' := _, npow_succ' := _, left_distrib := _, right_distrib := _}
Equations
- mul_opposite.algebra = {to_has_smul := {smul := has_smul.smul (mul_opposite.has_smul A R)}, to_ring_hom := (algebra_map R A).to_opposite mul_opposite.algebra._proof_1, commutes' := _, smul_def' := _}
Equations
- module.End.algebra R M = algebra.of_module _ _
An alternate statement of linear_map.map_smul
for when algebra_map
is more convenient to
work with than •
.
- to_fun : A → B
- map_one' : self.to_fun 1 = 1
- map_mul' : ∀ (x y : A), self.to_fun (x * y) = self.to_fun x * self.to_fun y
- map_zero' : self.to_fun 0 = 0
- map_add' : ∀ (x y : A), self.to_fun (x + y) = self.to_fun x + self.to_fun y
- commutes' : ∀ (r : R), self.to_fun (⇑(algebra_map R A) r) = ⇑(algebra_map R B) r
Defining the homomorphism in the category R-Alg.
Instances for alg_hom
- alg_hom.has_sizeof_inst
- alg_hom.has_coe_to_fun
- alg_hom.alg_hom_class
- alg_hom.coe_ring_hom
- alg_hom.coe_monoid_hom
- alg_hom.coe_add_monoid_hom
- alg_hom.End
- alg_equiv.has_coe_to_alg_hom
- alg_hom.non_unital_alg_hom.has_coe
- alg_hom.subsingleton
- minpoly.alg_hom.fintype
- alg_hom.lie_hom.has_coe
- alg_hom.continuous_linear_map_class
- coe : F → Π (a : A), (λ (_x : A), B) a
- coe_injective' : function.injective alg_hom_class.coe
- map_mul : ∀ (f : F) (x y : A), ⇑f (x * y) = ⇑f x * ⇑f y
- map_one : ∀ (f : F), ⇑f 1 = 1
- map_add : ∀ (f : F) (x y : A), ⇑f (x + y) = ⇑f x + ⇑f y
- map_zero : ∀ (f : F), ⇑f 0 = 0
- commutes : ∀ (f : F) (r : R), ⇑f (⇑(algebra_map R A) r) = ⇑(algebra_map R B) r
alg_hom_class F R A B
asserts F
is a type of bundled algebra homomorphisms
from A
to B
.
Instances of this typeclass
Instances of other typeclasses for alg_hom_class
- alg_hom_class.has_sizeof_inst
Equations
- alg_hom_class.linear_map_class = {coe := alg_hom_class.coe _inst_6, coe_injective' := _, map_add := _, map_smulₛₗ := _}
Equations
- alg_hom.has_coe_to_fun = {coe := alg_hom.to_fun _inst_7}
Equations
- alg_hom.alg_hom_class = {coe := alg_hom.to_fun _inst_7, coe_injective' := _, map_mul := _, map_one := _, map_add := _, map_zero := _, commutes := _}
Equations
- alg_hom.coe_ring_hom = {coe := alg_hom.to_ring_hom _inst_7}
Identity map as an alg_hom
.
Promote a linear_map
to an alg_hom
by supplying proofs about the behavior on 1
and *
.
Equations
- alg_hom.End = {mul := alg_hom.comp _inst_6, mul_assoc := _, one := alg_hom.id R A _inst_6, one_mul := _, mul_one := _, npow := npow_rec (mul_one_class.to_has_mul (A →ₐ[R] A)), npow_zero' := _, npow_succ' := _}
First projection as alg_hom
.
Second projection as alg_hom
.
Taking the product of two maps with the same domain is equivalent to taking the product of their codomains.
- to_fun : A → B
- inv_fun : B → A
- left_inv : function.left_inverse self.inv_fun self.to_fun
- right_inv : function.right_inverse self.inv_fun self.to_fun
- map_mul' : ∀ (x y : A), self.to_fun (x * y) = self.to_fun x * self.to_fun y
- map_add' : ∀ (x y : A), self.to_fun (x + y) = self.to_fun x + self.to_fun y
- commutes' : ∀ (r : R), self.to_fun (⇑(algebra_map R A) r) = ⇑(algebra_map R B) r
An equivalence of algebras is an equivalence of rings commuting with the actions of scalars.
Instances for alg_equiv
- alg_equiv.has_sizeof_inst
- alg_equiv.alg_equiv_class
- alg_equiv.has_coe_to_fun
- alg_equiv.has_coe_to_ring_equiv
- alg_equiv.has_coe_to_alg_hom
- alg_equiv.inhabited
- alg_equiv.aut
- alg_equiv.apply_mul_semiring_action
- alg_equiv.apply_has_faithful_smul
- alg_equiv.apply_smul_comm_class
- alg_equiv.apply_smul_comm_class'
- alg_equiv.subsingleton_left
- alg_equiv.subsingleton_right
- alg_equiv.fintype
- krull_topology
- alg_equiv.topological_group
- coe : F → A → B
- inv : F → B → A
- left_inv : ∀ (e : F), function.left_inverse (alg_equiv_class.inv e) (alg_equiv_class.coe e)
- right_inv : ∀ (e : F), function.right_inverse (alg_equiv_class.inv e) (alg_equiv_class.coe e)
- coe_injective' : ∀ (e g : F), alg_equiv_class.coe e = alg_equiv_class.coe g → alg_equiv_class.inv e = alg_equiv_class.inv g → e = g
- map_mul : ∀ (f : F) (a b : A), ⇑f (a * b) = ⇑f a * ⇑f b
- map_add : ∀ (f : F) (a b : A), ⇑f (a + b) = ⇑f a + ⇑f b
- commutes : ∀ (f : F) (r : R), ⇑f (⇑(algebra_map R A) r) = ⇑(algebra_map R B) r
alg_equiv_class F R A B
states that F
is a type of algebra structure preserving
equivalences. You should extend this class when you extend alg_equiv
.
Instances of this typeclass
Instances of other typeclasses for alg_equiv_class
- alg_equiv_class.has_sizeof_inst
Equations
- alg_equiv_class.to_alg_hom_class F R A B = {coe := coe_fn fun_like.has_coe_to_fun, coe_injective' := _, map_mul := _, map_one := _, map_add := _, map_zero := _, commutes := _}
Equations
- alg_equiv_class.to_linear_equiv_class F R A B = {coe := alg_equiv_class.coe h, inv := alg_equiv_class.inv h, left_inv := _, right_inv := _, coe_injective' := _, map_add := _, map_smulₛₗ := _}
Equations
- alg_equiv.alg_equiv_class = {coe := alg_equiv.to_fun _inst_6, inv := alg_equiv.inv_fun _inst_6, left_inv := _, right_inv := _, coe_injective' := _, map_mul := _, map_add := _, commutes := _}
Helper instance for when there's too many metavariables to apply
fun_like.has_coe_to_fun
directly.
Equations
- alg_equiv.has_coe_to_fun = {coe := alg_equiv.to_fun _inst_6}
Equations
- alg_equiv.has_coe_to_ring_equiv = {coe := alg_equiv.to_ring_equiv _inst_6}
Interpret an algebra equivalence as an algebra homomorphism.
This definition is included for symmetry with the other to_*_hom
projections.
The simp
normal form is to use the coercion of the has_coe_to_alg_hom
instance.
Equations
- alg_equiv.has_coe_to_alg_hom = {coe := alg_equiv.to_alg_hom _inst_6}
Algebra equivalences are reflexive.
Equations
- alg_equiv.inhabited = {default := alg_equiv.refl _inst_5}
Algebra equivalences are symmetric.
See Note [custom simps projection]
Equations
- alg_equiv.simps.symm_apply e = ⇑(e.symm)
Algebra equivalences are transitive.
Equations
- e₁.trans e₂ = {to_fun := (e₁.to_ring_equiv.trans e₂.to_ring_equiv).to_fun, inv_fun := (e₁.to_ring_equiv.trans e₂.to_ring_equiv).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _, commutes' := _}
If A₁
is equivalent to A₁'
and A₂
is equivalent to A₂'
, then the type of maps
A₁ →ₐ[R] A₂
is equivalent to the type of maps A₁' →ₐ[R] A₂'
.
Equations
- e₁.arrow_congr e₂ = {to_fun := λ (f : A₁ →ₐ[R] A₂), (e₂.to_alg_hom.comp f).comp e₁.symm.to_alg_hom, inv_fun := λ (f : A₁' →ₐ[R] A₂'), (e₂.symm.to_alg_hom.comp f).comp e₁.to_alg_hom, left_inv := _, right_inv := _}
If an algebra morphism has an inverse, it is a algebra isomorphism.
Promotes a bijective algebra homomorphism to an algebra equivalence.
Forgetting the multiplicative structures, an equivalence of algebras is a linear equivalence.
Interpret an algebra equivalence as a linear map.
Equations
Upgrade a linear equivalence to an algebra equivalence, given that it distributes over multiplication and action of scalars.
Equations
- alg_equiv.aut = {mul := λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ, mul_assoc := _, one := alg_equiv.refl _inst_5, one_mul := _, mul_one := _, npow := div_inv_monoid.npow._default alg_equiv.refl (λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ) alg_equiv.aut._proof_4 alg_equiv.aut._proof_5, npow_zero' := _, npow_succ' := _, inv := alg_equiv.symm _inst_5, div := div_inv_monoid.div._default (λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ) alg_equiv.aut._proof_8 alg_equiv.refl alg_equiv.aut._proof_9 alg_equiv.aut._proof_10 (div_inv_monoid.npow._default alg_equiv.refl (λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ) alg_equiv.aut._proof_4 alg_equiv.aut._proof_5) alg_equiv.aut._proof_11 alg_equiv.aut._proof_12 alg_equiv.symm, div_eq_mul_inv := _, zpow := div_inv_monoid.zpow._default (λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ) alg_equiv.aut._proof_14 alg_equiv.refl alg_equiv.aut._proof_15 alg_equiv.aut._proof_16 (div_inv_monoid.npow._default alg_equiv.refl (λ (ϕ ψ : A₁ ≃ₐ[R] A₁), ψ.trans ϕ) alg_equiv.aut._proof_4 alg_equiv.aut._proof_5) alg_equiv.aut._proof_17 alg_equiv.aut._proof_18 alg_equiv.symm, zpow_zero' := _, zpow_succ' := _, zpow_neg' := _, mul_left_inv := _}
An algebra isomorphism induces a group isomorphism between automorphism groups
The tautological action by A₁ ≃ₐ[R] A₁
on A₁
.
This generalizes function.End.apply_mul_action
.
Equations
- alg_equiv.apply_mul_semiring_action = {to_distrib_mul_action := {to_mul_action := {to_has_smul := {smul := λ (_x : A₁ ≃ₐ[R] A₁) (_y : A₁), ⇑_x _y}, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, smul_one := _, smul_mul := _}
Each element of the monoid defines a algebra homomorphism.
This is a stronger version of mul_semiring_action.to_ring_hom
and
distrib_mul_action.to_linear_map
.
Each element of the group defines a algebra equivalence.
This is a stronger version of mul_semiring_action.to_ring_equiv
and
distrib_mul_action.to_linear_equiv
.
Equations
- mul_semiring_action.to_alg_equiv R A g = {to_fun := (mul_semiring_action.to_ring_equiv G A g).to_fun, inv_fun := (mul_semiring_action.to_ring_equiv G A g).inv_fun, left_inv := _, right_inv := _, map_mul' := _, map_add' := _, commutes' := _}
Semiring ⥤ ℕ-Alg
Equations
- algebra_nat = {to_has_smul := add_monoid.has_smul_nat (add_monoid_with_one.to_add_monoid R), to_ring_hom := nat.cast_ring_hom R (semiring.to_non_assoc_semiring R), commutes' := _, smul_def' := _}
Reinterpret a ring_hom
as a ℚ
-algebra homomorphism. This actually yields an equivalence,
see ring_hom.equiv_rat_alg_hom
.