Documentation

Mathlib.RingTheory.Idempotents

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 #

Main results #

structure OrthogonalIdempotents {R : Type u_1} [Semiring R] {I : Type u_3} (e : IR) :

A family { eᵢ } of idempotent elements is orthogonal if eᵢ * eⱼ = 0 for all i ≠ j.

Instances For
    theorem orthogonalIdempotents_iff {R : Type u_1} [Semiring R] {I : Type u_3} (e : IR) :
    OrthogonalIdempotents e (∀ (i : I), IsIdempotentElem (e i)) Pairwise fun (x1 x2 : I) => e x1 * e x2 = 0
    theorem OrthogonalIdempotents.mul_eq {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} [DecidableEq I] (he : OrthogonalIdempotents e) (i j : I) :
    e i * e j = if i = j then e i else 0
    theorem OrthogonalIdempotents.iff_mul_eq {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} [DecidableEq I] :
    OrthogonalIdempotents e ∀ (i j : I), e i * e j = if i = j then e i else 0
    theorem OrthogonalIdempotents.isIdempotentElem_sum {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) {s : Finset I} :
    IsIdempotentElem (∑ is, e i)
    theorem OrthogonalIdempotents.mul_sum_of_mem {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) {i : I} {s : Finset I} (h : i s) :
    e i * js, e j = e i
    theorem OrthogonalIdempotents.mul_sum_of_not_mem {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) {i : I} {s : Finset I} (h : is) :
    e i * js, e j = 0
    theorem OrthogonalIdempotents.map {R : Type u_1} {S : Type u_2} [Semiring R] [Semiring S] (f : R →+* S) {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) :
    theorem OrthogonalIdempotents.map_injective_iff {R : Type u_1} {S : Type u_2} [Semiring R] [Semiring S] (f : R →+* S) {I : Type u_3} {e : IR} (hf : Function.Injective f) :
    theorem OrthogonalIdempotents.embedding {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) {J : Type u_4} (i : J I) :
    theorem OrthogonalIdempotents.equiv {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} {J : Type u_4} (i : J I) :
    theorem OrthogonalIdempotents.option {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) [Fintype I] (x : R) (hx : IsIdempotentElem x) (hx₁ : x * i : I, e i = 0) (hx₂ : (∑ i : I, e i) * x = 0) :
    OrthogonalIdempotents fun (x_1 : Option I) => x_1.elim x e
    structure CompleteOrthogonalIdempotents {R : Type u_1} [Semiring R] {I : Type u_3} [Fintype I] (e : IR) extends OrthogonalIdempotents e :

    A family { eᵢ } of idempotent elements is complete orthogonal if

    1. (orthogonal) eᵢ * eⱼ = 0 for all i ≠ j.
    2. (complete) ∑ eᵢ = 1
    Instances For
      theorem completeOrthogonalIdempotents_iff {R : Type u_1} [Semiring R] {I : Type u_3} [Fintype I] (e : IR) :
      theorem CompleteOrthogonalIdempotents.iff_ortho_complete {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} [Fintype I] :
      CompleteOrthogonalIdempotents e (Pairwise fun (x1 x2 : I) => e x1 * e x2 = 0) i : I, e i = 1

      If a family is complete orthogonal, it consists of idempotents.

      theorem CompleteOrthogonalIdempotents.pair_iff'ₛ {R : Type u_1} [Semiring R] {x y : R} :
      CompleteOrthogonalIdempotents ![x, y] x * y = 0 y * x = 0 x + y = 1
      theorem CompleteOrthogonalIdempotents.single {I : Type u_4} [Fintype I] [DecidableEq I] (R : IType u_5) [(i : I) → Semiring (R i)] :
      theorem CompleteOrthogonalIdempotents.map {R : Type u_1} {S : Type u_2} [Semiring R] [Semiring S] (f : R →+* S) {I : Type u_3} {e : IR} [Fintype I] (he : CompleteOrthogonalIdempotents e) :
      theorem CompleteOrthogonalIdempotents.equiv {R : Type u_1} [Semiring R] {I : Type u_3} {e : IR} [Fintype I] {J : Type u_4} [Fintype J] (i : J I) :
      theorem isIdempotentElem_one_sub_one_sub_pow_pow {R : Type u_1} [Ring R] (x : R) (n : ) (hx : (x - x ^ 2) ^ n = 0) :
      IsIdempotentElem (1 - (1 - x ^ n) ^ n)
      theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent_aux {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) (e₁ : S) (he : e₁ f.range) (he₁ : IsIdempotentElem e₁) (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) :
      ∃ (e' : R), IsIdempotentElem e' f e' = e₁ e' * e₂ = 0
      theorem exists_isIdempotentElem_mul_eq_zero_of_ker_isNilpotent {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) (e₁ : S) (he : e₁ f.range) (he₁ : IsIdempotentElem e₁) (e₂ : R) (he₂ : IsIdempotentElem e₂) (he₁e₂ : e₁ * f e₂ = 0) (he₂e₁ : f e₂ * e₁ = 0) :
      ∃ (e' : R), IsIdempotentElem e' f e' = e₁ e' * e₂ = 0 e₂ * e' = 0

      Orthogonal idempotents lift along nil ideals.

      theorem exists_isIdempotentElem_eq_of_ker_isNilpotent {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) (e : S) (he : e f.range) (he' : IsIdempotentElem e) :
      ∃ (e' : R), IsIdempotentElem e' f e' = e

      Idempotents lift along nil ideals.

      theorem OrthogonalIdempotents.lift_of_isNilpotent_ker_aux {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) {n : } {e : Fin nS} (he : OrthogonalIdempotents e) (he' : ∀ (i : Fin n), e i f.range) :
      ∃ (e' : Fin nR), OrthogonalIdempotents e' f e' = e
      theorem OrthogonalIdempotents.lift_of_isNilpotent_ker {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) {I : Type u_3} [Finite I] (h : xRingHom.ker f, IsNilpotent x) {e : IS} (he : OrthogonalIdempotents e) (he' : ∀ (i : I), e i f.range) :
      ∃ (e' : IR), OrthogonalIdempotents e' f e' = e

      A family of orthogonal idempotents lift along nil ideals.

      theorem CompleteOrthogonalIdempotents.option {R : Type u_1} [Ring R] {I : Type u_3} {e : IR} [Fintype I] (he : OrthogonalIdempotents e) :
      CompleteOrthogonalIdempotents fun (x : Option I) => x.elim (1 - i : I, e i) e
      theorem CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker_aux {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) {n : } {e : Fin nS} (he : CompleteOrthogonalIdempotents e) (he' : ∀ (i : Fin n), e i f.range) :
      ∃ (e' : Fin nR), CompleteOrthogonalIdempotents e' f e' = e
      theorem CompleteOrthogonalIdempotents.lift_of_isNilpotent_ker {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) {I : Type u_3} [Fintype I] (h : xRingHom.ker f, IsNilpotent x) {e : IS} (he : CompleteOrthogonalIdempotents e) (he' : ∀ (i : I), e i f.range) :
      ∃ (e' : IR), CompleteOrthogonalIdempotents e' f e' = e

      A system of complete orthogonal idempotents lift along nil ideals.

      theorem eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute {R : Type u_1} [Ring R] {e₁ e₂ : R} (he₁ : IsIdempotentElem e₁) (he₂ : IsIdempotentElem e₂) (H : IsNilpotent (e₁ - e₂)) (H' : Commute e₁ e₂) :
      e₁ = e₂
      theorem CompleteOrthogonalIdempotents.of_ker_isNilpotent_of_isMulCentral {R : Type u_1} {S : Type u_2} [Ring R] [Ring S] (f : R →+* S) {I : Type u_3} {e : IR} [Fintype I] (h : xRingHom.ker f, IsNilpotent x) (he : ∀ (i : I), IsIdempotentElem (e i)) (he' : ∀ (i : I), IsMulCentral (e i)) (he'' : CompleteOrthogonalIdempotents (f e)) :
      theorem eq_of_isNilpotent_sub_of_isIdempotentElem {R : Type u_1} [CommRing R] {e₁ e₂ : R} (he₁ : IsIdempotentElem e₁) (he₂ : IsIdempotentElem e₂) (H : IsNilpotent (e₁ - e₂)) :
      e₁ = e₂
      theorem existsUnique_isIdempotentElem_eq_of_ker_isNilpotent {R : Type u_1} {S : Type u_2} [CommRing R] [Ring S] (f : R →+* S) (h : xRingHom.ker f, IsNilpotent x) (e : S) (he : e f.range) (he' : IsIdempotentElem e) :
      ∃! e' : R, IsIdempotentElem e' f e' = e

      Stacks Tag 00J9

      theorem OrthogonalIdempotents.surjective_pi {R : Type u_1} [CommRing R] {I : Type u_3} [Finite I] {e : IR} (he : OrthogonalIdempotents e) :

      A family of orthogonal idempotents induces an surjection R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩

      theorem OrthogonalIdempotents.prod_one_sub {R : Type u_1} [CommRing R] {I : Type u_3} {e : IR} (he : OrthogonalIdempotents e) (s : Finset I) :
      is, (1 - e i) = 1 - is, e i
      theorem CompleteOrthogonalIdempotents.of_ker_isNilpotent {R : Type u_1} {S : Type u_2} [CommRing R] [Ring S] (f : R →+* S) {I : Type u_3} [Fintype I] {e : IR} (h : xRingHom.ker f, IsNilpotent x) (he : ∀ (i : I), IsIdempotentElem (e i)) (he' : CompleteOrthogonalIdempotents (f e)) :
      theorem CompleteOrthogonalIdempotents.prod_one_sub {R : Type u_1} [CommRing R] {I : Type u_3} [Fintype I] {e : IR} (he : CompleteOrthogonalIdempotents e) :
      i : I, (1 - e i) = 0
      theorem CompleteOrthogonalIdempotents.of_prod_one_sub {R : Type u_1} [CommRing R] {I : Type u_3} [Fintype I] {e : IR} (he : OrthogonalIdempotents e) (he' : i : I, (1 - e i) = 0) :
      theorem CompleteOrthogonalIdempotents.bijective_pi {R : Type u_1} [CommRing R] {I : Type u_3} [Fintype I] {e : IR} (he : CompleteOrthogonalIdempotents e) :

      A family of complete orthogonal idempotents induces an isomorphism R ≃+* ∏ R ⧸ ⟨1 - eᵢ⟩

      theorem CompleteOrthogonalIdempotents.bijective_pi' {R : Type u_1} [CommRing R] {I : Type u_3} [Fintype I] {e : IR} (he : CompleteOrthogonalIdempotents fun (x : I) => 1 - e x) :
      theorem bijective_pi_of_isIdempotentElem {R : Type u_1} [CommRing R] {I : Type u_3} [Fintype I] (e : IR) (he : ∀ (i : I), IsIdempotentElem (e i)) (he₁ : ∀ (i j : I), i j(1 - e i) * (1 - e j) = 0) (he₂ : i : I, e i = 0) :
      def Subsemigroup.corner {R : Type u_1} (e : R) [Semigroup R] :

      The corner associated to an element e in a semigroup is the subsemigroup of all elements of the form e * r * e.

      Equations
      Instances For
        theorem Subsemigroup.mem_corner_iff {R : Type u_1} {e : R} [Semigroup R] (idem : IsIdempotentElem e) {r : R} :
        r corner e e * r = r r * e = r
        theorem Subsemigroup.mem_corner_iff_mul_left {R : Type u_1} {e : R} [Semigroup R] (idem : IsIdempotentElem e) (hc : IsMulCentral e) {r : R} :
        r corner e e * r = r
        theorem Subsemigroup.mem_corner_iff_mul_right {R : Type u_1} {e : R} [Semigroup R] (idem : IsIdempotentElem e) (hc : IsMulCentral e) {r : R} :
        r corner e r * e = r
        theorem Subsemigroup.mem_corner_iff_mem_range_mul_left {R : Type u_1} {e : R} [Semigroup R] (idem : IsIdempotentElem e) (hc : IsMulCentral e) {r : R} :
        r corner e r Set.range fun (x : R) => e * x
        theorem Subsemigroup.mem_corner_iff_mem_range_mul_right {R : Type u_1} {e : R} [Semigroup R] (idem : IsIdempotentElem e) (hc : IsMulCentral e) {r : R} :
        r corner e r Set.range fun (x : R) => x * e
        def IsIdempotentElem.Corner {R : Type u_1} {e : R} [Semigroup R] :

        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
        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
          Instances For

            The corner associated to an element e in a ring without is the subring without 1 of all elements of the forme * r * e`.

            Equations
            Instances For
              instance instSemiringCorner {R : Type u_1} (e : R) [NonUnitalSemiring R] (idem : IsIdempotentElem e) :
              Semiring idem.Corner
              Equations
              instance instCommSemiringCorner {R : Type u_1} (e : R) [NonUnitalCommSemiring R] (idem : IsIdempotentElem e) :
              CommSemiring idem.Corner
              Equations
              instance instRingCorner {R : Type u_1} (e : R) [NonUnitalRing R] (idem : IsIdempotentElem e) :
              Ring idem.Corner
              Equations
              instance instCommRingCorner {R : Type u_1} (e : R) [NonUnitalCommRing R] (idem : IsIdempotentElem e) :
              CommRing idem.Corner
              Equations
              def CompleteOrthogonalIdempotents.mulEquivOfIsMulCentral {R : Type u_1} {I : Type u_2} [Fintype I] {e : IR} [Semiring R] (he : CompleteOrthogonalIdempotents e) (hc : ∀ (i : I), IsMulCentral (e i)) :
              R ≃+* ((i : I) → .Corner)

              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
                def CompleteOrthogonalIdempotents.mulEquivOfComm {R : Type u_1} {I : Type u_2} [Fintype I] {e : IR} [CommSemiring R] (he : CompleteOrthogonalIdempotents e) :
                R ≃+* ((i : I) → .Corner)

                A complete orthogonal family of idempotents in a commutative semiring give rise to a direct product decomposition.

                Equations
                • he.mulEquivOfComm = he.mulEquivOfIsMulCentral
                Instances For