Documentation

Mathlib.RingTheory.Spectrum.Maximal.Localization

Maximal spectrum of a commutative (semi)ring #

Localization results.

theorem MaximalSpectrum.iInf_localization_eq_bot (R : Type u_4) [CommRing R] [IsDomain R] (K : Type u_5) [Field K] [Algebra R K] [IsFractionRing R K] :
⨅ (v : MaximalSpectrum R), Localization.subalgebra.ofField K v.asIdeal.primeCompl =

An integral domain is equal to the intersection of its localizations at all its maximal ideals viewed as subalgebras of its field of fractions.

theorem PrimeSpectrum.iInf_localization_eq_bot (R : Type u_4) [CommRing R] [IsDomain R] (K : Type u_5) [Field K] [Algebra R K] [IsFractionRing R K] :
⨅ (v : PrimeSpectrum R), Localization.subalgebra.ofField K v.asIdeal.primeCompl =

An integral domain is equal to the intersection of its localizations at all its prime ideals viewed as subalgebras of its field of fractions.

@[reducible, inline]

The product of localizations at all maximal ideals of a commutative semiring.

Equations
Instances For

    The canonical ring homomorphism from a commutative semiring to the product of its localizations at all maximal ideals. It is always injective.

    Equations
    Instances For
      noncomputable def MaximalSpectrum.mapPiLocalization {R : Type u_1} {S : Type u_2} [CommSemiring R] [CommSemiring S] (f : R →+* S) (hf : Function.Bijective f) :

      Functoriality of PiLocalization but restricted to bijective ring homs. If R and S are commutative rings, surjectivity would be enough.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        theorem MaximalSpectrum.mapPiLocalization_comp {R : Type u_1} {S : Type u_2} (P : Type u_3) [CommSemiring R] [CommSemiring S] [CommSemiring P] (f : R →+* S) (g : S →+* P) (hf : Function.Bijective f) (hg : Function.Bijective g) :
        mapPiLocalization (g.comp f) = (mapPiLocalization g hg).comp (mapPiLocalization f hf)
        theorem MaximalSpectrum.toPiLocalization_not_surjective_of_infinite {ι : Type u_5} (R : ιType u_4) [(i : ι) → CommSemiring (R i)] [∀ (i : ι), Nontrivial (R i)] [Infinite ι] :
        ¬Function.Surjective (toPiLocalization ((i : ι) → R i))
        theorem MaximalSpectrum.finite_of_toPiLocalization_pi_surjective {ι : Type u_5} {R : ιType u_4} [(i : ι) → CommSemiring (R i)] [∀ (i : ι), Nontrivial (R i)] (h : Function.Surjective (toPiLocalization ((i : ι) → R i))) :
        @[reducible, inline]

        The product of localizations at all prime ideals of a commutative semiring.

        Equations
        Instances For

          The canonical ring homomorphism from a commutative semiring to the product of its localizations at all prime ideals. It is always injective.

          Equations
          Instances For

            The projection from the product of localizations at primes to the product of localizations at maximal ideals.

            Equations
            Instances For

              If R has Krull dimension ≤ 0, then piLocalizationToIsMaximal R is an isomorphism.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                theorem PrimeSpectrum.piLocalizationToMaximal_bijective {R : Type u_1} [CommSemiring R] (h : ∀ (I : Ideal R), I.IsPrimeI.IsMaximal) :
                noncomputable def PrimeSpectrum.mapPiLocalization {R : Type u_1} {S : Type u_2} [CommSemiring R] [CommSemiring S] (f : R →+* S) :

                A ring homomorphism induces a homomorphism between the products of localizations at primes.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  theorem PrimeSpectrum.mapPiLocalization_comp {R : Type u_1} {S : Type u_2} (P : Type u_3) [CommSemiring R] [CommSemiring S] [CommSemiring P] (f : R →+* S) (g : S →+* P) :
                  theorem PrimeSpectrum.toPiLocalization_not_surjective_of_infinite {ι : Type u_5} (R : ιType u_4) [(i : ι) → CommSemiring (R i)] [∀ (i : ι), Nontrivial (R i)] [Infinite ι] :
                  ¬Function.Surjective (toPiLocalization ((i : ι) → R i))
                  theorem PrimeSpectrum.finite_of_toPiLocalization_pi_surjective {ι : Type u_5} {R : ιType u_4} [(i : ι) → CommSemiring (R i)] [∀ (i : ι), Nontrivial (R i)] (h : Function.Surjective (toPiLocalization ((i : ι) → R i))) :