Documentation

Mathlib.RingTheory.Ideal.Maps

Maps on modules and ideals #

def Ideal.map {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] (f : F) (I : Ideal R) :

I.map f is the span of the image of the ideal I under f, which may be bigger than the image itself.

Equations
Instances For
    def Ideal.comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal S) :

    I.comap f is the preimage of I under f.

    Equations
    • Ideal.comap f I = { carrier := f ⁻¹' I, add_mem' := , zero_mem' := , smul_mem' := }
    Instances For
      @[simp]
      theorem Ideal.coe_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal S) :
      (Ideal.comap f I) = f ⁻¹' I
      theorem Ideal.map_mono {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] {f : F} {I : Ideal R} {J : Ideal R} (h : I J) :
      theorem Ideal.mem_map_of_mem {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] (f : F) {I : Ideal R} {x : R} (h : x I) :
      f x Ideal.map f I
      theorem Ideal.apply_coe_mem_map {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] (f : F) (I : Ideal R) (x : I) :
      f x Ideal.map f I
      theorem Ideal.map_le_iff_le_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {I : Ideal R} {K : Ideal S} :
      @[simp]
      theorem Ideal.mem_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {K : Ideal S} {x : R} :
      x Ideal.comap f K f x K
      theorem Ideal.comap_mono {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {K : Ideal S} {L : Ideal S} (h : K L) :
      theorem Ideal.comap_ne_top {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} (hK : K ) :
      theorem Ideal.map_le_comap_of_inv_on {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] (f : F) {G : Type u_2} [FunLike G S R] [rcg : RingHomClass G S R] (g : G) (I : Ideal R) (hf : Set.LeftInvOn g f I) :
      theorem Ideal.comap_le_map_of_inv_on {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {G : Type u_2} [FunLike G S R] (g : G) (I : Ideal S) (hf : Set.LeftInvOn (g) (f) (f ⁻¹' I)) :
      theorem Ideal.map_le_comap_of_inverse {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] (f : F) {G : Type u_2} [FunLike G S R] [rcg : RingHomClass G S R] (g : G) (I : Ideal R) (h : Function.LeftInverse g f) :

      The Ideal version of Set.image_subset_preimage_of_inverse.

      theorem Ideal.comap_le_map_of_inverse {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {G : Type u_2} [FunLike G S R] (g : G) (I : Ideal S) (h : Function.LeftInverse g f) :

      The Ideal version of Set.preimage_subset_image_of_inverse.

      instance Ideal.IsPrime.comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} [hK : K.IsPrime] :
      (Ideal.comap f K).IsPrime
      Equations
      • =
      theorem Ideal.map_top {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) :
      theorem Ideal.gc_map_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) :
      @[simp]
      theorem Ideal.comap_id {R : Type u} [Semiring R] (I : Ideal R) :
      @[simp]
      theorem Ideal.map_id {R : Type u} [Semiring R] (I : Ideal R) :
      theorem Ideal.comap_comap {R : Type u} {S : Type v} [Semiring R] [Semiring S] {T : Type u_3} [Semiring T] {I : Ideal T} (f : R →+* S) (g : S →+* T) :
      Ideal.comap f (Ideal.comap g I) = Ideal.comap (g.comp f) I
      theorem Ideal.map_map {R : Type u} {S : Type v} [Semiring R] [Semiring S] {T : Type u_3} [Semiring T] {I : Ideal R} (f : R →+* S) (g : S →+* T) :
      Ideal.map g (Ideal.map f I) = Ideal.map (g.comp f) I
      theorem Ideal.map_span {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (s : Set R) :
      theorem Ideal.map_le_of_le_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {I : Ideal R} {K : Ideal S} :
      I Ideal.comap f KIdeal.map f I K
      theorem Ideal.le_comap_of_map_le {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {I : Ideal R} {K : Ideal S} :
      Ideal.map f I KI Ideal.comap f K
      theorem Ideal.le_comap_map {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {I : Ideal R} :
      theorem Ideal.map_comap_le {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {K : Ideal S} :
      @[simp]
      theorem Ideal.comap_top {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} :
      @[simp]
      theorem Ideal.comap_eq_top_iff {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} {I : Ideal S} :
      @[simp]
      theorem Ideal.map_bot {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} :
      @[simp]
      theorem Ideal.map_comap_map {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal R) :
      @[simp]
      theorem Ideal.comap_map_comap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (K : Ideal S) :
      theorem Ideal.map_sup {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal R) (J : Ideal R) :
      theorem Ideal.comap_inf {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (K : Ideal S) (L : Ideal S) :
      theorem Ideal.map_iSup {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {ι : Sort u_3} (K : ιIdeal R) :
      Ideal.map f (iSup K) = ⨆ (i : ι), Ideal.map f (K i)
      theorem Ideal.comap_iInf {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {ι : Sort u_3} (K : ιIdeal S) :
      Ideal.comap f (iInf K) = ⨅ (i : ι), Ideal.comap f (K i)
      theorem Ideal.map_sSup {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (s : Set (Ideal R)) :
      Ideal.map f (sSup s) = Is, Ideal.map f I
      theorem Ideal.comap_sInf {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (s : Set (Ideal S)) :
      Ideal.comap f (sInf s) = Is, Ideal.comap f I
      theorem Ideal.comap_sInf' {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (s : Set (Ideal S)) :
      Ideal.comap f (sInf s) = IIdeal.comap f '' s, I
      theorem Ideal.comap_isPrime {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (K : Ideal S) [H : K.IsPrime] :
      (Ideal.comap f K).IsPrime
      theorem Ideal.map_inf_le {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {I : Ideal R} {J : Ideal R} :
      theorem Ideal.le_comap_sup {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} {L : Ideal S} :
      @[simp]
      @[simp]
      theorem Ideal.coe_restrictScalars {R : Type u_4} {S : Type u_5} [CommSemiring R] [Semiring S] [Algebra R S] (I : Ideal S) :
      @[simp]

      The smallest S-submodule that contains all x ∈ I * y ∈ J is also the smallest R-submodule that does so.

      theorem Ideal.map_comap_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) (I : Ideal S) :
      def Ideal.giMapComap {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) :

      map and comap are adjoint, and the composition map f ∘ comap f is the identity

      Equations
      Instances For
        theorem Ideal.map_surjective_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) :
        theorem Ideal.comap_injective_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) :
        theorem Ideal.map_sup_comap_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) (I : Ideal S) (J : Ideal S) :
        theorem Ideal.map_iSup_comap_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {ι : Sort u_3} (hf : Function.Surjective f) (K : ιIdeal S) :
        Ideal.map f (⨆ (i : ι), Ideal.comap f (K i)) = iSup K
        theorem Ideal.map_inf_comap_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) (I : Ideal S) (J : Ideal S) :
        theorem Ideal.map_iInf_comap_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {ι : Sort u_3} (hf : Function.Surjective f) (K : ιIdeal S) :
        Ideal.map f (⨅ (i : ι), Ideal.comap f (K i)) = iInf K
        theorem Ideal.mem_image_of_mem_map_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) {I : Ideal R} {y : S} (H : y Ideal.map f I) :
        y f '' I
        theorem Ideal.mem_map_iff_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Surjective f) {I : Ideal R} {y : S} :
        y Ideal.map f I xI, f x = y
        theorem Ideal.le_map_of_comap_le_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {I : Ideal R} {K : Ideal S} (hf : Function.Surjective f) :
        Ideal.comap f K IK Ideal.map f I
        theorem Ideal.map_eq_submodule_map {R : Type u} {S : Type v} [Semiring R] [Semiring S] (f : R →+* S) [h : RingHomSurjective f] (I : Ideal R) :
        Ideal.map f I = Submodule.map f.toSemilinearMap I
        theorem Ideal.comap_bot_le_of_injective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {I : Ideal R} (hf : Function.Injective f) :
        theorem Ideal.comap_bot_of_injective {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (hf : Function.Injective f) :
        @[simp]
        theorem Ideal.map_of_equiv {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal R) (f : R ≃+* S) :
        Ideal.map (f.symm) (Ideal.map (f) I) = I

        If f : R ≃+* S is a ring isomorphism and I : Ideal R, then map f.symm (map f I) = I.

        @[simp]
        theorem Ideal.comap_of_equiv {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal R) (f : R ≃+* S) :
        Ideal.comap (f) (Ideal.comap (f.symm) I) = I

        If f : R ≃+* S is a ring isomorphism and I : Ideal R, then comap f (comap f.symm I) = I.

        theorem Ideal.map_comap_of_equiv {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal R) (f : R ≃+* S) :
        Ideal.map (f) I = Ideal.comap f.symm I

        If f : R ≃+* S is a ring isomorphism and I : Ideal R, then map f I = comap f.symm I.

        @[simp]
        theorem Ideal.comap_symm {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal R) (f : R ≃+* S) :
        Ideal.comap f.symm I = Ideal.map f I

        If f : R ≃+* S is a ring isomorphism and I : Ideal R, then comap f.symm I = map f I.

        @[simp]
        theorem Ideal.map_symm {R : Type u} {S : Type v} [Semiring R] [Semiring S] (I : Ideal S) (f : R ≃+* S) :
        Ideal.map f.symm I = Ideal.comap f I

        If f : R ≃+* S is a ring isomorphism and I : Ideal R, then map f.symm I = comap f I.

        theorem Ideal.comap_map_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) (I : Ideal R) :
        def Ideal.relIsoOfSurjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) :
        Ideal S ≃o { p : Ideal R // Ideal.comap f p }

        Correspondence theorem

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          def Ideal.orderEmbeddingOfSurjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) :

          The map on ideals induced by a surjective map preserves inclusion.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            theorem Ideal.map_eq_top_or_isMaximal_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) {I : Ideal R} (H : I.IsMaximal) :
            Ideal.map f I = (Ideal.map f I).IsMaximal
            theorem Ideal.comap_isMaximal_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) {K : Ideal S} [H : K.IsMaximal] :
            (Ideal.comap f K).IsMaximal
            theorem Ideal.comap_le_comap_iff_of_surjective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Surjective f) (I : Ideal S) (J : Ideal S) :
            def Ideal.relIsoOfBijective {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Bijective f) :

            Special case of the correspondence theorem for isomorphic rings

            Equations
            Instances For
              theorem Ideal.comap_le_iff_le_map {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Bijective f) {I : Ideal R} {K : Ideal S} :
              theorem Ideal.map.isMaximal {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [RingHomClass F R S] (f : F) (hf : Function.Bijective f) {I : Ideal R} (H : I.IsMaximal) :
              (Ideal.map f I).IsMaximal
              theorem Ideal.RingEquiv.bot_maximal_iff {R : Type u} {S : Type v} [Ring R] [Ring S] (e : R ≃+* S) :
              .IsMaximal .IsMaximal
              theorem Ideal.map_mul {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal R) (J : Ideal R) :
              Ideal.map f (I * J) = Ideal.map f I * Ideal.map f J
              @[simp]
              theorem Ideal.mapHom_apply {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal R) :
              def Ideal.mapHom {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) :

              The pushforward Ideal.map as a monoid-with-zero homomorphism.

              Equations
              Instances For
                theorem Ideal.map_pow {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (I : Ideal R) (n : ) :
                Ideal.map f (I ^ n) = Ideal.map f I ^ n
                theorem Ideal.comap_radical {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) (K : Ideal S) :
                Ideal.comap f K.radical = (Ideal.comap f K).radical
                theorem Ideal.IsRadical.comap {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} (hK : K.IsRadical) :
                (Ideal.comap f K).IsRadical
                theorem Ideal.map_radical_le {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {I : Ideal R} :
                Ideal.map f I.radical (Ideal.map f I).radical
                theorem Ideal.le_comap_mul {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} {L : Ideal S} :
                theorem Ideal.le_comap_pow {R : Type u} {S : Type v} {F : Type u_1} [CommRing R] [CommRing S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {K : Ideal S} (n : ) :
                Ideal.comap f K ^ n Ideal.comap f (K ^ n)
                def RingHom.ker {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] (f : F) :

                Kernel of a ring homomorphism as an ideal of the domain.

                Equations
                Instances For
                  theorem RingHom.mem_ker {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] (f : F) {r : R} :
                  r RingHom.ker f f r = 0

                  An element is in the kernel if and only if it maps to zero.

                  theorem RingHom.ker_eq {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] (f : F) :
                  (RingHom.ker f) = f ⁻¹' {0}
                  theorem RingHom.ker_eq_comap_bot {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] (f : F) :
                  theorem RingHom.comap_ker {R : Type u} {S : Type v} {T : Type w} [Semiring R] [Semiring S] [Semiring T] (f : S →+* R) (g : T →+* S) :
                  theorem RingHom.not_one_mem_ker {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] [Nontrivial S] (f : F) :
                  1RingHom.ker f

                  If the target is not the zero ring, then one is not in the kernel.

                  theorem RingHom.ker_ne_top {R : Type u} {S : Type v} {F : Type u_1} [Semiring R] [Semiring S] [FunLike F R S] [rcf : RingHomClass F R S] [Nontrivial S] (f : F) :
                  theorem Pi.ker_ringHom {S : Type v} [Semiring S] {ι : Type u_3} {R : ιType u_4} [(i : ι) → Semiring (R i)] (φ : (i : ι) → S →+* R i) :
                  RingHom.ker (Pi.ringHom φ) = ⨅ (i : ι), RingHom.ker (φ i)
                  @[simp]
                  theorem RingHom.ker_rangeSRestrict {R : Type u} {S : Type v} [Semiring R] [Semiring S] (f : R →+* S) :
                  RingHom.ker f.rangeSRestrict = RingHom.ker f
                  theorem RingHom.injective_iff_ker_eq_bot {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) :
                  theorem RingHom.ker_eq_bot_iff_eq_zero {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) :
                  RingHom.ker f = ∀ (x : R), f x = 0x = 0
                  @[simp]
                  theorem RingHom.ker_coe_equiv {R : Type u} {S : Type v} [Ring R] [Semiring S] (f : R ≃+* S) :
                  @[simp]
                  theorem RingHom.ker_equiv {R : Type u} {S : Type v} [Ring R] [Semiring S] {F' : Type u_2} [EquivLike F' R S] [RingEquivClass F' R S] (f : F') :
                  theorem RingHom.sub_mem_ker_iff {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [FunLike F R S] [rc : RingHomClass F R S] (f : F) {x : R} {y : R} :
                  x - y RingHom.ker f f x = f y
                  @[simp]
                  theorem RingHom.ker_rangeRestrict {R : Type u} {S : Type v} [Ring R] [Ring S] (f : R →+* S) :
                  RingHom.ker f.rangeRestrict = RingHom.ker f
                  theorem RingHom.ker_isPrime {R : Type u} {S : Type v} {F : Type u_1} [Ring R] [Ring S] [IsDomain S] [FunLike F R S] [RingHomClass F R S] (f : F) :
                  (RingHom.ker f).IsPrime

                  The kernel of a homomorphism to a domain is a prime ideal.

                  theorem RingHom.ker_isMaximal_of_surjective {R : Type u_1} {K : Type u_2} {F : Type u_3} [Ring R] [Field K] [FunLike F R K] [RingHomClass F R K] (f : F) (hf : Function.Surjective f) :
                  (RingHom.ker f).IsMaximal

                  The kernel of a homomorphism to a field is a maximal ideal.

                  theorem Ideal.map_eq_bot_iff_le_ker {R : Type u_1} {S : Type u_2} {F : Type u_3} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {I : Ideal R} (f : F) :
                  theorem Ideal.ker_le_comap {R : Type u_1} {S : Type u_2} {F : Type u_3} [Semiring R] [Semiring S] [FunLike F R S] [rc : RingHomClass F R S] {K : Ideal S} (f : F) :
                  theorem Ideal.map_isPrime_of_equiv {R : Type u_1} {S : Type u_2} [Semiring R] [Semiring S] {F' : Type u_4} [EquivLike F' R S] [RingEquivClass F' R S] (f : F') {I : Ideal R} [I.IsPrime] :
                  (Ideal.map f I).IsPrime
                  theorem Ideal.map_sInf {R : Type u_1} {S : Type u_2} {F : Type u_3} [Ring R] [Ring S] [FunLike F R S] [rc : RingHomClass F R S] {A : Set (Ideal R)} {f : F} (hf : Function.Surjective f) :
                  (JA, RingHom.ker f J)Ideal.map f (sInf A) = sInf (Ideal.map f '' A)
                  theorem Ideal.map_isPrime_of_surjective {R : Type u_1} {S : Type u_2} {F : Type u_3} [Ring R] [Ring S] [FunLike F R S] [rc : RingHomClass F R S] {f : F} (hf : Function.Surjective f) {I : Ideal R} [H : I.IsPrime] (hk : RingHom.ker f I) :
                  (Ideal.map f I).IsPrime
                  theorem Ideal.map_eq_bot_iff_of_injective {R : Type u_1} {S : Type u_2} {F : Type u_3} [Ring R] [Ring S] [FunLike F R S] [rc : RingHomClass F R S] {I : Ideal R} {f : F} (hf : Function.Injective f) :
                  theorem Ideal.map_eq_iff_sup_ker_eq_of_surjective {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] {I : Ideal R} {J : Ideal R} (f : R →+* S) (hf : Function.Surjective f) :
                  theorem Ideal.map_radical_of_surjective {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] {f : R →+* S} (hf : Function.Surjective f) {I : Ideal R} (h : RingHom.ker f I) :
                  Ideal.map f I.radical = (Ideal.map f I).radical
                  def RingHom.liftOfRightInverseAux {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) (g : A →+* C) (hg : RingHom.ker f RingHom.ker g) :
                  B →+* C

                  Auxiliary definition used to define liftOfRightInverse

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem RingHom.liftOfRightInverseAux_comp_apply {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) (g : A →+* C) (hg : RingHom.ker f RingHom.ker g) (a : A) :
                    (f.liftOfRightInverseAux f_inv hf g hg) (f a) = g a
                    def RingHom.liftOfRightInverse {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) :
                    { g : A →+* C // RingHom.ker f RingHom.ker g } (B →+* C)

                    liftOfRightInverse f hf g hg is the unique ring homomorphism φ

                    See RingHom.eq_liftOfRightInverse for the uniqueness lemma.

                       A .
                       |  \
                     f |   \ g
                       |    \
                       v     \⌟
                       B ----> C
                          ∃!φ
                    
                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[reducible, inline]
                      noncomputable abbrev RingHom.liftOfSurjective {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (hf : Function.Surjective f) :
                      { g : A →+* C // RingHom.ker f RingHom.ker g } (B →+* C)

                      A non-computable version of RingHom.liftOfRightInverse for when no computable right inverse is available, that uses Function.surjInv.

                      Equations
                      Instances For
                        theorem RingHom.liftOfRightInverse_comp_apply {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) (g : { g : A →+* C // RingHom.ker f RingHom.ker g }) (x : A) :
                        ((f.liftOfRightInverse f_inv hf) g) (f x) = g x
                        theorem RingHom.liftOfRightInverse_comp {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) (g : { g : A →+* C // RingHom.ker f RingHom.ker g }) :
                        ((f.liftOfRightInverse f_inv hf) g).comp f = g
                        theorem RingHom.eq_liftOfRightInverse {A : Type u_1} {B : Type u_2} {C : Type u_3} [Ring A] [Ring B] [Ring C] (f : A →+* B) (f_inv : BA) (hf : Function.RightInverse f_inv f) (g : A →+* C) (hg : RingHom.ker f RingHom.ker g) (h : B →+* C) (hh : h.comp f = g) :
                        h = (f.liftOfRightInverse f_inv hf) g, hg
                        theorem AlgHom.coe_ker {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) :
                        theorem AlgHom.coe_ideal_map {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommSemiring R] [Semiring A] [Semiring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) (I : Ideal A) :
                        Ideal.map f I = Ideal.map (f) I