Idempotents in rings #
The predicate IsIdempotentElem
is defined for general monoids in Algebra/Ring/Idempotents.lean
.
In this file we provide various results regarding idempotent elements in rings.
Main definitions #
OrthogonalIdempotents
: A family{ eᵢ }
of idempotent elements is orthogonal ifeᵢ * eⱼ = 0
for alli ≠ j
.CompleteOrthogonalIdempotents
: A family{ eᵢ }
of orthogonal idempotent elements is complete if∑ eᵢ = 1
.
Main results #
CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker
: If the kernel off : R →+* S
consists of nilpotent elements, and{ eᵢ }
is a family of complete orthogonal idempotents in the range off
, then{ eᵢ }
is the image of some complete orthogonal idempotents inR
.existsUnique_isIdempotentElem_eq_of_ker_isNilpotent
: IfR
is commutative and the kernel off : R →+* S
consists of nilpotent elements, then every idempotent in the range off
lifts to a unique idempotent inR
.CompleteOrthogonalIdempotents.bijective_pi
: IfR
is commutative, then a family{ eᵢ }
of complete orthogonal idempotent elements induces a ring isomorphismR ≃ ∏ R ⧸ ⟨1 - eᵢ⟩
.
A family { eᵢ }
of idempotent elements is orthogonal if eᵢ * eⱼ = 0
for all i ≠ j
.
- idem (i : I) : IsIdempotentElem (e i)
Instances For
A family { eᵢ }
of idempotent elements is complete orthogonal if
- (orthogonal)
eᵢ * eⱼ = 0
for alli ≠ j
. - (complete)
∑ eᵢ = 1
- complete : ∑ i : I, e i = 1
Instances For
Orthogonal idempotents lift along nil ideals.
Idempotents lift along nil ideals.
A family of orthogonal idempotents lift along nil ideals.
A system of complete orthogonal idempotents lift along nil ideals.
A family of orthogonal idempotents induces an surjection R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩
A family of complete orthogonal idempotents induces an isomorphism R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩
The corner associated to an element e
in a semigroup
is the subsemigroup of all elements of the form e * r * e
.
Equations
- Subsemigroup.corner e = { carrier := Set.range fun (x : R) => e * x * e, mul_mem' := ⋯ }
Instances For
The corner associated to an idempotent e
in a semiring without 1
is the semiring with e
as 1 consisting of all element of the form e * r * e
.
Equations
- x✝.Corner = ↥(Subsemigroup.corner e)
Instances For
The corner associated to an element e
in a semiring without 1
is the subsemiring without 1 of all elements of the form e * r * e
.
Equations
- NonUnitalSubsemiring.corner e = { carrier := (Subsemigroup.corner e).carrier, add_mem' := ⋯, zero_mem' := ⋯, mul_mem' := ⋯ }
Instances For
The corner associated to an element e
in a ring without is the subring without 1 of all elements of the form
e * r * e`.
Equations
- NonUnitalRing.corner e = { toNonUnitalSubsemiring := NonUnitalSubsemiring.corner e, neg_mem' := ⋯ }
Instances For
Equations
- instSemiringCorner e idem = Semiring.mk ⋯ ⋯ ⋯ ⋯ npowRecAuto ⋯ ⋯
Equations
- instCommSemiringCorner e idem = CommSemiring.mk ⋯
Equations
- instRingCorner e idem = Ring.mk ⋯ SubNegMonoid.zsmul ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- instCommRingCorner e idem = CommRing.mk ⋯
A complete orthogonal family of central idempotents in a semiring give rise to a direct product decomposition.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A complete orthogonal family of idempotents in a commutative semiring give rise to a direct product decomposition.
Equations
- he.mulEquivOfComm = he.mulEquivOfIsMulCentral ⋯