Documentation

Mathlib.RingTheory.IntegralClosure.IsIntegralClosure.Basic

# Integral closure as a characteristic predicate #

We prove basic properties of IsIntegralClosure.

theorem Algebra.IsIntegral.finite {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] [Algebra.IsIntegral R A] [h' : Algebra.FiniteType R A] :

The Kurosh problem asks to show that this is still true when A is not necessarily commutative and R is a field, but it has been solved in the negative. See https://arxiv.org/pdf/1706.02383.pdf for criteria for a finitely generated algebraic (= integral) algebra over a field to be finite dimensional.

This could be an instance, but we tend to go from Module.Finite to IsIntegral/IsAlgebraic, and making it an instance will cause the search to be complicated a lot.

finite = integral + finite type

theorem RingHom.IsIntegral.to_finite {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] {f : R →+* S} (h : f.IsIntegral) (h' : f.FiniteType) :
f.Finite
theorem RingHom.Finite.of_isIntegral_of_finiteType {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] {f : R →+* S} (h : f.IsIntegral) (h' : f.FiniteType) :
f.Finite

Alias of RingHom.IsIntegral.to_finite.

theorem RingHom.finite_iff_isIntegral_and_finiteType {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] {f : R →+* S} :
f.Finite f.IsIntegral f.FiniteType

finite = integral + finite type

theorem mem_integralClosure_iff_mem_fg (R : Type u_1) (A : Type u_2) [CommRing R] [CommRing A] [Algebra R A] {r : A} :
r integralClosure R A ∃ (M : Subalgebra R A), (Subalgebra.toSubmodule M).FG r M
theorem adjoin_le_integralClosure {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {x : A} (hx : IsIntegral R x) :
theorem Algebra.isIntegral_sup {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {S : Subalgebra R A} {T : Subalgebra R A} :
theorem integralClosure_map_algEquiv {R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A ≃ₐ[R] S) :

Mapping an integral closure along an AlgEquiv gives the integral closure.

def AlgHom.mapIntegralClosure {R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A →ₐ[R] S) :

An AlgHom between two rings restrict to an AlgHom between the integral closures inside them.

Equations
Instances For
    @[simp]
    theorem AlgHom.coe_mapIntegralClosure {R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A →ₐ[R] S) (x : (integralClosure R A)) :
    (f.mapIntegralClosure x) = f x
    def AlgEquiv.mapIntegralClosure {R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A ≃ₐ[R] S) :

    An AlgEquiv between two rings restrict to an AlgEquiv between the integral closures inside them.

    Equations
    • f.mapIntegralClosure = AlgEquiv.ofAlgHom (↑f).mapIntegralClosure (↑f.symm).mapIntegralClosure
    Instances For
      @[simp]
      theorem AlgEquiv.coe_mapIntegralClosure {R : Type u_1} {A : Type u_2} {S : Type u_4} [CommRing R] [CommRing A] [CommRing S] [Algebra R A] [Algebra R S] (f : A ≃ₐ[R] S) (x : (integralClosure R A)) :
      (f.mapIntegralClosure x) = f x
      theorem integralClosure.isIntegral {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] (x : (integralClosure R A)) :
      Equations
      • =
      theorem IsIntegral.of_mul_unit {R : Type u_1} {B : Type u_3} [CommRing R] [Ring B] [Algebra R B] {x : B} {y : B} {r : R} (hr : (algebraMap R B) r * y = 1) (hx : IsIntegral R (x * y)) :
      theorem RingHom.IsIntegralElem.of_mul_unit {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) (x : S) (y : S) (r : R) (hr : f r * y = 1) (hx : f.IsIntegralElem (x * y)) :
      f.IsIntegralElem x
      theorem IsIntegral.of_mem_closure' {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] (G : Set A) (hG : xG, IsIntegral R x) (x : A) :

      Generalization of IsIntegral.of_mem_closure bootstrapped up from that lemma

      theorem IsIntegral.of_mem_closure'' {R : Type u_1} [CommRing R] {S : Type u_5} [CommRing S] {f : R →+* S} (G : Set S) (hG : xG, f.IsIntegralElem x) (x : S) :
      x Subring.closure Gf.IsIntegralElem x
      theorem IsIntegral.pow {R : Type u_1} {B : Type u_3} [CommRing R] [Ring B] [Algebra R B] {x : B} (h : IsIntegral R x) (n : ) :
      IsIntegral R (x ^ n)
      theorem IsIntegral.nsmul {R : Type u_1} {B : Type u_3} [CommRing R] [Ring B] [Algebra R B] {x : B} (h : IsIntegral R x) (n : ) :
      IsIntegral R (n x)
      theorem IsIntegral.zsmul {R : Type u_1} {B : Type u_3} [CommRing R] [Ring B] [Algebra R B] {x : B} (h : IsIntegral R x) (n : ) :
      IsIntegral R (n x)
      theorem IsIntegral.multiset_prod {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {s : Multiset A} (h : xs, IsIntegral R x) :
      IsIntegral R s.prod
      theorem IsIntegral.multiset_sum {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {s : Multiset A} (h : xs, IsIntegral R x) :
      IsIntegral R s.sum
      theorem IsIntegral.prod {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {α : Type u_5} {s : Finset α} (f : αA) (h : xs, IsIntegral R (f x)) :
      IsIntegral R (∏ xs, f x)
      theorem IsIntegral.sum {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {α : Type u_5} {s : Finset α} (f : αA) (h : xs, IsIntegral R (f x)) :
      IsIntegral R (∑ xs, f x)
      theorem IsIntegral.det {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {n : Type u_5} [Fintype n] [DecidableEq n] {M : Matrix n n A} (h : ∀ (i j : n), IsIntegral R (M i j)) :
      IsIntegral R M.det
      @[simp]
      theorem IsIntegral.pow_iff {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {x : A} {n : } (hn : 0 < n) :
      theorem IsIntegral.tmul {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommRing R] [CommRing A] [Ring B] [Algebra R A] [Algebra R B] (x : A) {y : B} (h : IsIntegral R y) :
      noncomputable def normalizeScaleRoots {R : Type u_1} [CommRing R] (p : Polynomial R) :

      The monic polynomial whose roots are p.leadingCoeff * x for roots x of p.

      Equations
      Instances For
        theorem normalizeScaleRoots_coeff_mul_leadingCoeff_pow {R : Type u_1} [CommRing R] (p : Polynomial R) (i : ) (hp : 1 p.natDegree) :
        (normalizeScaleRoots p).coeff i * p.leadingCoeff ^ i = p.coeff i * p.leadingCoeff ^ (p.natDegree - 1)
        theorem leadingCoeff_smul_normalizeScaleRoots {R : Type u_1} [CommRing R] (p : Polynomial R) :
        p.leadingCoeff normalizeScaleRoots p = p.scaleRoots p.leadingCoeff
        theorem normalizeScaleRoots_support {R : Type u_1} [CommRing R] (p : Polynomial R) :
        (normalizeScaleRoots p).support p.support
        theorem normalizeScaleRoots_degree {R : Type u_1} [CommRing R] (p : Polynomial R) :
        (normalizeScaleRoots p).degree = p.degree
        theorem normalizeScaleRoots_eval₂_leadingCoeff_mul {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (p : Polynomial R) (h : 1 p.natDegree) (f : R →+* S) (x : S) :
        Polynomial.eval₂ f (f p.leadingCoeff * x) (normalizeScaleRoots p) = f p.leadingCoeff ^ (p.natDegree - 1) * Polynomial.eval₂ f x p
        theorem normalizeScaleRoots_monic {R : Type u_1} [CommRing R] (p : Polynomial R) (h : p 0) :
        theorem RingHom.isIntegralElem_leadingCoeff_mul {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) (p : Polynomial R) (x : S) (h : Polynomial.eval₂ f x p = 0) :
        f.IsIntegralElem (f p.leadingCoeff * x)

        Given a p : R[X] and a x : S such that p.eval₂ f x = 0, f p.leadingCoeff * x is integral.

        theorem isIntegral_leadingCoeff_smul {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (p : Polynomial R) (x : S) [Algebra R S] (h : (Polynomial.aeval x) p = 0) :
        IsIntegral R (p.leadingCoeff x)

        Given a p : R[X] and a root x : S, then p.leadingCoeff • x : S is integral over R.

        instance integralClosure.isIntegralClosure (R : Type u_1) (A : Type u_2) [CommRing R] [CommRing A] [Algebra R A] :
        Equations
        • =
        theorem IsIntegralClosure.isIntegral (R : Type u_1) {A : Type u_2} (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] [Algebra R A] [IsScalarTower R A B] (x : A) :
        theorem IsIntegralClosure.isIntegral_algebra (R : Type u_1) {A : Type u_2} (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] [Algebra R A] [IsScalarTower R A B] :
        noncomputable def IsIntegralClosure.mk' {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (x : B) (hx : IsIntegral R x) :
        A

        If x : B is integral over R, then it is an element of the integral closure of R in B.

        Equations
        Instances For
          @[simp]
          theorem IsIntegralClosure.algebraMap_mk' {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (x : B) (hx : IsIntegral R x) :
          @[simp]
          theorem IsIntegralClosure.mk'_one {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (h : optParam (IsIntegral R 1) ) :
          @[simp]
          theorem IsIntegralClosure.mk'_zero {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (h : optParam (IsIntegral R 0) ) :
          theorem IsIntegralClosure.mk'_add {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (x : B) (y : B) (hx : IsIntegral R x) (hy : IsIntegral R y) :
          theorem IsIntegralClosure.mk'_mul {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (x : B) (y : B) (hx : IsIntegral R x) (hy : IsIntegral R y) :
          @[simp]
          theorem IsIntegralClosure.mk'_algebraMap {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] [Algebra R A] [IsScalarTower R A B] (x : R) (h : optParam (IsIntegral R ((algebraMap R B) x)) ) :
          noncomputable def IsIntegralClosure.lift (R : Type u_1) (A : Type u_2) (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] {S : Type u_4} [CommRing S] [Algebra R S] [Algebra S B] [IsScalarTower R S B] [Algebra R A] [IsScalarTower R A B] [isIntegral : Algebra.IsIntegral R S] :

          If B / S / R is a tower of ring extensions where S is integral over R, then S maps (uniquely) into an integral closure B / A / R.

          Equations
          Instances For
            @[simp]
            theorem IsIntegralClosure.algebraMap_lift (R : Type u_1) (A : Type u_2) (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] {S : Type u_4} [CommRing S] [Algebra R S] [Algebra S B] [IsScalarTower R S B] [Algebra R A] [IsScalarTower R A B] [isIntegral : Algebra.IsIntegral R S] (x : S) :
            (algebraMap A B) ((IsIntegralClosure.lift R A B) x) = (algebraMap S B) x
            noncomputable def IsIntegralClosure.equiv (R : Type u_1) (A : Type u_2) (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (A' : Type u_4) [CommRing A'] [Algebra A' B] [IsIntegralClosure A' R B] [Algebra R A] [Algebra R A'] [IsScalarTower R A B] [IsScalarTower R A' B] :
            A ≃ₐ[R] A'

            Integral closures are all isomorphic to each other.

            Equations
            Instances For
              @[simp]
              theorem IsIntegralClosure.algebraMap_equiv (R : Type u_1) (A : Type u_2) (B : Type u_3) [CommRing R] [CommRing A] [CommRing B] [Algebra R B] [Algebra A B] [IsIntegralClosure A R B] (A' : Type u_4) [CommRing A'] [Algebra A' B] [IsIntegralClosure A' R B] [Algebra R A] [Algebra R A'] [IsScalarTower R A B] [IsScalarTower R A' B] (x : A) :
              (algebraMap A' B) ((IsIntegralClosure.equiv R A B A') x) = (algebraMap A B) x
              theorem isIntegral_trans {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommRing R] [CommRing A] [Ring B] [Algebra A B] [Algebra R B] [Algebra R A] [IsScalarTower R A B] [Algebra.IsIntegral R A] (x : B) (hx : IsIntegral A x) :

              If A is an R-algebra all of whose elements are integral over R, and x is an element of an A-algebra that is integral over A, then x is integral over R.

              theorem Algebra.IsIntegral.trans {R : Type u_1} (A : Type u_2) {B : Type u_3} [CommRing R] [CommRing A] [Ring B] [Algebra A B] [Algebra R B] [Algebra R A] [IsScalarTower R A B] [Algebra.IsIntegral R A] [Algebra.IsIntegral A B] :

              If A is an R-algebra all of whose elements are integral over R, and B is an A-algebra all of whose elements are integral over A, then all elements of B are integral over R.

              theorem RingHom.IsIntegral.trans {R : Type u_1} {S : Type u_4} {T : Type u_5} [CommRing R] [CommRing S] [CommRing T] (f : R →+* S) (g : S →+* T) (hf : f.IsIntegral) (hg : g.IsIntegral) :
              (g.comp f).IsIntegral
              theorem IsIntegralClosure.tower_top {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {B : Type u_6} {C : Type u_7} [CommRing C] [CommRing B] [Algebra R B] [Algebra A B] [Algebra C B] [IsScalarTower R A B] [IsIntegralClosure C R B] [Algebra.IsIntegral R A] :

              If R → A → B is an algebra tower, C is the integral closure of R in B and A is integral over R, then C is the integral closure of A in B.

              theorem RingHom.isIntegral_of_surjective {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) (hf : Function.Surjective f) :
              f.IsIntegral
              theorem IsIntegral.tower_bot {R : Type u_1} {A : Type u_2} {B : Type u_3} [CommRing R] [CommRing A] [Ring B] [Algebra A B] [Algebra R B] [Algebra R A] [IsScalarTower R A B] (H : Function.Injective (algebraMap A B)) {x : A} (h : IsIntegral R ((algebraMap A B) x)) :

              If R → A → B is an algebra tower with A → B injective, then if the entire tower is an integral extension so is R → A

              theorem RingHom.IsIntegral.tower_bot {R : Type u_1} {S : Type u_4} {T : Type u_5} [CommRing R] [CommRing S] [CommRing T] (f : R →+* S) (g : S →+* T) (hg : Function.Injective g) (hfg : (g.comp f).IsIntegral) :
              f.IsIntegral
              theorem IsIntegral.tower_bot_of_field {R : Type u_6} {A : Type u_7} {B : Type u_8} [CommRing R] [Field A] [CommRing B] [Nontrivial B] [Algebra R A] [Algebra A B] [Algebra R B] [IsScalarTower R A B] {x : A} (h : IsIntegral R ((algebraMap A B) x)) :
              theorem RingHom.isIntegralElem.of_comp {R : Type u_1} {S : Type u_4} {T : Type u_5} [CommRing R] [CommRing S] [CommRing T] (f : R →+* S) (g : S →+* T) {x : T} (h : (g.comp f).IsIntegralElem x) :
              g.IsIntegralElem x
              theorem RingHom.IsIntegral.tower_top {R : Type u_1} {S : Type u_4} {T : Type u_5} [CommRing R] [CommRing S] [CommRing T] (f : R →+* S) (g : S →+* T) (h : (g.comp f).IsIntegral) :
              g.IsIntegral
              theorem RingHom.IsIntegral.quotient {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal S} (hf : f.IsIntegral) :
              (Ideal.quotientMap I f ).IsIntegral
              instance Algebra.IsIntegral.quotient {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] {I : Ideal A} [Algebra.IsIntegral R A] :
              Equations
              • =
              theorem isIntegral_quotientMap_iff {R : Type u_1} {S : Type u_4} [CommRing R] [CommRing S] (f : R →+* S) {I : Ideal S} :
              (Ideal.quotientMap I f ).IsIntegral ((Ideal.Quotient.mk I).comp f).IsIntegral
              theorem isField_of_isIntegral_of_isField {R : Type u_6} {S : Type u_7} [CommRing R] [CommRing S] [Algebra R S] [Algebra.IsIntegral R S] (hRS : Function.Injective (algebraMap R S)) (hS : IsField S) :

              If the integral extension R → S is injective, and S is a field, then R is also a field.

              theorem integralClosure_idem {R : Type u_1} {A : Type u_2} [CommRing R] [CommRing A] [Algebra R A] :
              Equations
              • =
              theorem roots_mem_integralClosure {R : Type u_1} {S : Type u_2} [CommRing R] [CommRing S] [IsDomain S] [Algebra R S] {f : Polynomial R} (hf : f.Monic) {a : S} (ha : a f.aroots S) :