Pi instances for ring #
This file defines instances for ring, semiring and related structures on Pi Types
Equations
- Pi.distrib = { mul := fun (x1 x2 : (i : I) → f i) => x1 * x2, add := fun (x1 x2 : (i : I) → f i) => x1 + x2, left_distrib := ⋯, right_distrib := ⋯ }
Equations
- Pi.hasDistribNeg = { toInvolutiveNeg := Pi.involutiveNeg, neg_mul := ⋯, mul_neg := ⋯ }
Equations
- Pi.addMonoidWithOne = { natCast := fun (n : ℕ) (x : I) => ↑n, toAddMonoid := Pi.addMonoid, toOne := Pi.instOne, natCast_zero := ⋯, natCast_succ := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- Pi.nonUnitalSemiring = { toNonUnitalNonAssocSemiring := Pi.nonUnitalNonAssocSemiring, mul_assoc := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Pi.nonUnitalCommSemiring = { toNonUnitalSemiring := Pi.nonUnitalSemiring, mul_comm := ⋯ }
Equations
- Pi.commSemiring = { toSemiring := Pi.semiring, mul_comm := ⋯ }
Equations
- Pi.nonUnitalNonAssocRing = { toAddCommGroup := Pi.addCommGroup, toMul := NonUnitalNonAssocSemiring.toMul, left_distrib := ⋯, right_distrib := ⋯, zero_mul := ⋯, mul_zero := ⋯ }
Equations
- Pi.nonUnitalRing = { toNonUnitalNonAssocRing := Pi.nonUnitalNonAssocRing, mul_assoc := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Pi.nonUnitalCommRing = { toNonUnitalRing := Pi.nonUnitalRing, mul_comm := ⋯ }
Equations
- Pi.commRing = { toRing := Pi.ring, mul_comm := ⋯ }
A family of non-unital ring homomorphisms f a : γ →ₙ+* β a
defines a non-unital ring
homomorphism Pi.nonUnitalRingHom f : γ →+* Π a, β a
given by
Pi.nonUnitalRingHom f x b = f b x
.
Equations
- Pi.nonUnitalRingHom g = { toFun := fun (x : γ) (b : I) => (g b) x, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
A family of ring homomorphisms f a : γ →+* β a
defines a ring homomorphism
Pi.ringHom f : γ →+* Π a, β a
given by Pi.ringHom f x b = f b x
.
Equations
- Pi.ringHom g = { toFun := fun (x : γ) (b : I) => (g b) x, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Evaluation of functions into an indexed collection of non-unital rings at a point is a
non-unital ring homomorphism. This is Function.eval
as a NonUnitalRingHom
.
Equations
- Pi.evalNonUnitalRingHom f i = { toMulHom := Pi.evalMulHom f i, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Function.const
as a NonUnitalRingHom
.
Equations
- Pi.constNonUnitalRingHom α β = { toFun := Function.const α, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Non-unital ring homomorphism between the function spaces I → α
and I → β
, induced by a
non-unital ring homomorphism f
between α
and β
.
Equations
Instances For
Evaluation of functions into an indexed collection of rings at a point is a ring
homomorphism. This is Function.eval
as a RingHom
.
Equations
- Pi.evalRingHom f i = { toMonoidHom := Pi.evalMonoidHom f i, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Function.const
as a RingHom
.
Equations
- Pi.constRingHom α β = { toFun := Function.const α, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Ring homomorphism between the function spaces I → α
and I → β
, induced by a ring
homomorphism f
between α
and β
.