Documentation

Mathlib.RingTheory.AdicCompletion.RingHom

Lift of ring homomorphisms to adic completions #

Let R, S be rings, I be an ideal of S. In this file we prove that a compatible family of ring homomorphisms from a ring R to S ⧸ I ^ n can be lifted to a ring homomorphism R →+* AdicCompletion I S. If S is I-adically complete, then this compatible family of ring homomorphisms can be lifted to a ring homomorphism R →+* S.

Main definitions #

noncomputable def IsAdicComplete.liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) :
R →+* S

Universal property of IsAdicComplete for rings. The lift ring map lift I f hf : R →+* S of a sequence of compatible ring maps f n : R →+* S ⧸ (I ^ n).

Equations
Instances For
    @[simp]
    theorem IsAdicComplete.of_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) (x : R) :
    @[simp]
    theorem IsAdicComplete.ofAlgEquiv_comp_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) :
    @[simp]
    theorem IsAdicComplete.mk_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) (n : ) (x : R) :
    (Ideal.Quotient.mk (I ^ n)) ((liftRingHom I f ) x) = (f n) x

    The composition of lift linear map lift I f hf : R →+* S with the canonical projection S →+* S ⧸ (I ^ n) is f n .

    @[simp]
    theorem IsAdicComplete.mk_comp_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) (n : ) :
    (Ideal.Quotient.mk (I ^ n)).comp (liftRingHom I f ) = f n
    theorem IsAdicComplete.eq_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) [IsAdicComplete I S] (f : (n : ) → R →+* S I ^ n) (hf : ∀ {m n : } (hle : m n), (Ideal.Quotient.factorPow I hle).comp (f n) = f m) (F : R →+* S) (hF : ∀ (n : ), (Ideal.Quotient.mk (I ^ n)).comp F = f n) :
    F = liftRingHom I f

    Uniqueness of the lift. Given a compatible family of linear maps f n : R →ₗ[R] S ⧸ (I ^ n). If F : R →+* S makes the following diagram commutes

      R
      | \
     F|  \ f n
      |   \
      v    v
      S --> S ⧸ (I ^ n)
    

    Then it is the map IsAdicComplete.lift.

    theorem IsAdicComplete.StrictMono.factorPow_comp_eq_of_factorPow_comp_succ_eq' {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] {I : Ideal S} {a : } (ha : StrictMono a) (f : (n : ) → R →+* S I ^ a n) (hf : ∀ {m : }, (Ideal.Quotient.factorPow I ).comp (f (m + 1)) = f m) {m n : } (hle : m n) :
    (Ideal.Quotient.factorPow I ).comp (f n) = f m

    RingHom variant of IsAdicComplete.StrictMono.factorPow_comp_eq_of_factorPow_comp_succ_eq.

    noncomputable def IsAdicComplete.StrictMono.liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) {a : } (ha : StrictMono a) (f : (n : ) → R →+* S I ^ a n) (hf : ∀ {m : }, (Ideal.Quotient.factorPow I ).comp (f (m + 1)) = f m) [IsAdicComplete I S] :
    R →+* S

    A variant of IsAdicComplete.liftRingHom. Only takes f n : R →+* S ⧸ I ^ (a n) from a strictly increasing sequence a n.

    Equations
    Instances For
      @[simp]
      theorem IsAdicComplete.StrictMono.mk_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) {a : } (ha : StrictMono a) (f : (n : ) → R →+* S I ^ a n) (hf : ∀ {m : }, (Ideal.Quotient.factorPow I ).comp (f (m + 1)) = f m) [IsAdicComplete I S] {n : } (x : R) :
      (Ideal.Quotient.mk (I ^ a n)) ((liftRingHom I ha f ) x) = (f n) x
      @[simp]
      theorem IsAdicComplete.StrictMono.mk_comp_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) {a : } (ha : StrictMono a) (f : (n : ) → R →+* S I ^ a n) (hf : ∀ {m : }, (Ideal.Quotient.factorPow I ).comp (f (m + 1)) = f m) [IsAdicComplete I S] {n : } :
      (Ideal.Quotient.mk (I ^ a n)).comp (liftRingHom I ha f ) = f n
      theorem IsAdicComplete.StrictMono.eq_liftRingHom {R : Type u_1} {S : Type u_2} [NonAssocSemiring R] [CommRing S] (I : Ideal S) {a : } (ha : StrictMono a) (f : (n : ) → R →+* S I ^ a n) (hf : ∀ {m : }, (Ideal.Quotient.factorPow I ).comp (f (m + 1)) = f m) [IsAdicComplete I S] {F : R →+* S} (hF : ∀ (n : ), (Ideal.Quotient.mk (I ^ a n)).comp F = f n) :
      F = liftRingHom I ha f