Documentation

Mathlib.RingTheory.Valuation.ValuationSubring

Valuation subrings of a field #

Projects #

The order structure on ValuationSubring K.

structure ValuationSubring (K : Type u) [Field K] extends Subring K :

A valuation subring of a field K is a subring A such that for every x : K, either x ∈ A or x⁻¹ ∈ A.

This is equivalent to being maximal in the domination order of local subrings (the stacks project definition). See LocalSubring.isMax_iff.

Instances For
    Equations
    theorem ValuationSubring.mem_carrier {K : Type u} [Field K] (A : ValuationSubring K) (x : K) :
    x A.carrier x A
    @[simp]
    theorem ValuationSubring.mem_toSubring {K : Type u} [Field K] (A : ValuationSubring K) (x : K) :
    theorem ValuationSubring.ext {K : Type u} [Field K] (A B : ValuationSubring K) (h : ∀ (x : K), x A x B) :
    A = B
    theorem ValuationSubring.add_mem {K : Type u} [Field K] (A : ValuationSubring K) (x y : K) :
    x Ay Ax + y A
    theorem ValuationSubring.mul_mem {K : Type u} [Field K] (A : ValuationSubring K) (x y : K) :
    x Ay Ax * y A
    theorem ValuationSubring.neg_mem {K : Type u} [Field K] (A : ValuationSubring K) (x : K) :
    x A-x A
    theorem ValuationSubring.mem_or_inv_mem {K : Type u} [Field K] (A : ValuationSubring K) (x : K) :
    x A x⁻¹ A
    Equations
    theorem ValuationSubring.mem_top {K : Type u} [Field K] (x : K) :
    @[simp]
    theorem ValuationSubring.algebraMap_apply {K : Type u} [Field K] (A : ValuationSubring K) (a : A) :
    (algebraMap (↥A) K) a = a

    The value group of the valuation associated to A. Note: it is actually a group with zero.

    Equations
    Instances For

      Any valuation subring of K induces a natural valuation on K.

      Equations
      Instances For
        theorem ValuationSubring.valuation_le_one {K : Type u} [Field K] (A : ValuationSubring K) (a : A) :
        A.valuation a 1
        theorem ValuationSubring.mem_of_valuation_le_one {K : Type u} [Field K] (A : ValuationSubring K) (x : K) (h : A.valuation x 1) :
        x A
        theorem ValuationSubring.valuation_eq_iff {K : Type u} [Field K] (A : ValuationSubring K) (x y : K) :
        A.valuation x = A.valuation y ∃ (a : (↥A)ˣ), a * y = x
        theorem ValuationSubring.valuation_le_iff {K : Type u} [Field K] (A : ValuationSubring K) (x y : K) :
        A.valuation x A.valuation y ∃ (a : A), a * y = x
        theorem ValuationSubring.valuation_unit {K : Type u} [Field K] (A : ValuationSubring K) (a : (↥A)ˣ) :
        A.valuation a = 1
        theorem ValuationSubring.valuation_eq_one_iff {K : Type u} [Field K] (A : ValuationSubring K) (a : A) :
        IsUnit a A.valuation a = 1
        theorem ValuationSubring.valuation_lt_one_or_eq_one {K : Type u} [Field K] (A : ValuationSubring K) (a : A) :
        A.valuation a < 1 A.valuation a = 1
        def ValuationSubring.ofSubring {K : Type u} [Field K] (R : Subring K) (hR : ∀ (x : K), x R x⁻¹ R) :

        A subring R of K such that for all x : K either x ∈ R or x⁻¹ ∈ R is a valuation subring of K.

        Equations
        Instances For
          @[simp]
          theorem ValuationSubring.mem_ofSubring {K : Type u} [Field K] (R : Subring K) (hR : ∀ (x : K), x R x⁻¹ R) (x : K) :
          x ofSubring R hR x R

          An overring of a valuation ring is a valuation ring.

          Equations
          • R.ofLE S h = { toSubring := S, mem_or_inv_mem' := }
          Instances For
            def ValuationSubring.inclusion {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :
            R →+* S

            The ring homomorphism induced by the partial order.

            Equations
            Instances For
              def ValuationSubring.subtype {K : Type u} [Field K] (R : ValuationSubring K) :
              R →+* K

              The canonical ring homomorphism from a valuation ring to its field of fractions.

              Equations
              Instances For

                The canonical map on value groups induced by a coarsening of valuation rings.

                Equations
                Instances For
                  theorem ValuationSubring.monotone_mapOfLE {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :
                  Monotone (R.mapOfLE S h)
                  @[simp]
                  theorem ValuationSubring.mapOfLE_comp_valuation {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :
                  (R.mapOfLE S h) R.valuation = S.valuation
                  @[simp]
                  theorem ValuationSubring.mapOfLE_valuation_apply {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) (x : K) :
                  (R.mapOfLE S h) (R.valuation x) = S.valuation x
                  def ValuationSubring.idealOfLE {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :
                  Ideal R

                  The ideal corresponding to a coarsening of a valuation ring.

                  Equations
                  Instances For
                    instance ValuationSubring.prime_idealOfLE {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :

                    The coarsening of a valuation ring associated to a prime ideal.

                    Equations
                    Instances For
                      instance ValuationSubring.ofPrime_scalar_tower {K : Type u} [Field K] (A : ValuationSubring K) (P : Ideal A) [P.IsPrime] :
                      IsScalarTower (↥A) (↥(A.ofPrime P)) K
                      theorem ValuationSubring.le_ofPrime {K : Type u} [Field K] (A : ValuationSubring K) (P : Ideal A) [P.IsPrime] :
                      A A.ofPrime P
                      @[simp]
                      theorem ValuationSubring.idealOfLE_ofPrime {K : Type u} [Field K] (A : ValuationSubring K) (P : Ideal A) [P.IsPrime] :
                      A.idealOfLE (A.ofPrime P) = P
                      @[simp]
                      theorem ValuationSubring.ofPrime_idealOfLE {K : Type u} [Field K] (R S : ValuationSubring K) (h : R S) :
                      R.ofPrime (R.idealOfLE S h) = S
                      theorem ValuationSubring.ofPrime_le_of_le {K : Type u} [Field K] (A : ValuationSubring K) (P Q : Ideal A) [P.IsPrime] [Q.IsPrime] (h : P Q) :
                      theorem ValuationSubring.idealOfLE_le_of_le {K : Type u} [Field K] (A R S : ValuationSubring K) (hR : A R) (hS : A S) (h : R S) :
                      A.idealOfLE S hS A.idealOfLE R hR

                      The equivalence between coarsenings of a valuation ring and its prime ideals.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For

                        An ordered variant of primeSpectrumEquiv.

                        Equations
                        Instances For
                          @[simp]
                          theorem ValuationSubring.primeSpectrumOrderEquiv_apply_coe_carrier {K : Type u} [Field K] (A : ValuationSubring K) (P : PrimeSpectrum A) :
                          (A.primeSpectrumOrderEquiv P) = {x : K | aA, ∃ (a_1 : K), (∃ (x : a_1 A), a_1, P.asIdeal.primeCompl) x = a * a_1⁻¹}

                          The valuation subring associated to a valuation.

                          Equations
                          Instances For
                            @[simp]
                            theorem Valuation.mem_valuationSubring_iff {K : Type u} [Field K] {Γ : Type u_1} [LinearOrderedCommGroupWithZero Γ] (v : Valuation K Γ) (x : K) :
                            theorem Valuation.isEquiv_iff_valuationSubring {K : Type u} [Field K] {Γ₁ : Type u_2} {Γ₂ : Type u_3} [LinearOrderedCommGroupWithZero Γ₁] [LinearOrderedCommGroupWithZero Γ₂] (v₁ : Valuation K Γ₁) (v₂ : Valuation K Γ₂) :

                            The unit group of a valuation subring, as a subgroup of .

                            Equations
                            Instances For
                              @[simp]
                              theorem ValuationSubring.mem_unitGroup_iff {K : Type u} [Field K] (A : ValuationSubring K) (x : Kˣ) :
                              x A.unitGroup A.valuation x = 1

                              For a valuation subring A, A.unitGroup agrees with the units of A.

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]
                                theorem ValuationSubring.coe_unitGroupMulEquiv_apply {K : Type u} [Field K] (A : ValuationSubring K) (a : A.unitGroup) :
                                (A.unitGroupMulEquiv a) = a
                                @[simp]
                                theorem ValuationSubring.coe_unitGroupMulEquiv_symm_apply {K : Type u} [Field K] (A : ValuationSubring K) (a : (↥A)ˣ) :
                                (A.unitGroupMulEquiv.symm a) = a

                                The map on valuation subrings to their unit groups is an order embedding.

                                Equations
                                Instances For

                                  The nonunits of a valuation subring of K, as a subsemigroup of K

                                  Equations
                                  Instances For

                                    The map on valuation subrings to their nonunits is a dual order embedding.

                                    Equations
                                    Instances For

                                      The elements of A.nonunits are those of the maximal ideal of A after coercion to K.

                                      See also mem_nonunits_iff_exists_mem_maximalIdeal, which gets rid of the coercion to K, at the expense of a more complicated right hand side.

                                      The elements of A.nonunits are those of the maximal ideal of A.

                                      See also coe_mem_nonunits_iff, which has a simpler right hand side but requires the element to be in A already.

                                      A.nonunits agrees with the maximal ideal of A, after taking its image in K.

                                      The principal unit group of a valuation subring, as a subgroup of .

                                      Equations
                                      Instances For

                                        The map on valuation subrings to their principal unit groups is an order embedding.

                                        Equations
                                        Instances For

                                          The principal unit group agrees with the kernel of the canonical map from the units of A to the units of the residue field of A.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For

                                            The canonical map from the unit group of A to the units of the residue field of A.

                                            Equations
                                            Instances For

                                              Pointwise actions #

                                              This transfers the action from Subring.pointwiseMulAction, noting that it only applies when the action is by a group. Notably this provides an instances when G is K ≃+* K.

                                              These instances are in the Pointwise locale.

                                              The lemmas in this section are copied from the file Mathlib.Algebra.Ring.Subring.Pointwise; try to keep these in sync.

                                              The action on a valuation subring corresponding to applying the action to every element.

                                              This is available as an instance in the Pointwise locale.

                                              Equations
                                              Instances For
                                                @[simp]
                                                theorem ValuationSubring.coe_pointwise_smul {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] (g : G) (S : ValuationSubring K) :
                                                ↑(g S) = g S
                                                @[simp]
                                                theorem ValuationSubring.pointwise_smul_toSubring {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] (g : G) (S : ValuationSubring K) :

                                                The action on a valuation subring corresponding to applying the action to every element.

                                                This is available as an instance in the Pointwise locale.

                                                This is a stronger version of ValuationSubring.pointwiseSMul.

                                                Equations
                                                Instances For
                                                  theorem ValuationSubring.smul_mem_pointwise_smul {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] (g : G) (x : K) (S : ValuationSubring K) :
                                                  x Sg x g S
                                                  theorem ValuationSubring.mem_smul_pointwise_iff_exists {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] (g : G) (x : K) (S : ValuationSubring K) :
                                                  x g S sS, g s = x
                                                  @[simp]
                                                  theorem ValuationSubring.smul_mem_pointwise_smul_iff {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S : ValuationSubring K} {x : K} :
                                                  g x g S x S
                                                  theorem ValuationSubring.mem_pointwise_smul_iff_inv_smul_mem {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S : ValuationSubring K} {x : K} :
                                                  x g S g⁻¹ x S
                                                  theorem ValuationSubring.mem_inv_pointwise_smul_iff {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S : ValuationSubring K} {x : K} :
                                                  x g⁻¹ S g x S
                                                  @[simp]
                                                  theorem ValuationSubring.pointwise_smul_le_pointwise_smul_iff {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S T : ValuationSubring K} :
                                                  g S g T S T
                                                  theorem ValuationSubring.pointwise_smul_subset_iff {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S T : ValuationSubring K} :
                                                  g S T S g⁻¹ T
                                                  theorem ValuationSubring.subset_pointwise_smul_iff {K : Type u} [Field K] {G : Type u_1} [Group G] [MulSemiringAction G K] {g : G} {S T : ValuationSubring K} :
                                                  S g T g⁻¹ S T
                                                  def ValuationSubring.comap {K : Type u} [Field K] {L : Type u_1} [Field L] (A : ValuationSubring L) (f : K →+* L) :

                                                  The pullback of a valuation subring A along a ring homomorphism K →+* L.

                                                  Equations
                                                  Instances For
                                                    @[simp]
                                                    theorem ValuationSubring.coe_comap {K : Type u} [Field K] {L : Type u_1} [Field L] (A : ValuationSubring L) (f : K →+* L) :
                                                    (A.comap f) = f ⁻¹' A
                                                    @[simp]
                                                    theorem ValuationSubring.mem_comap {K : Type u} [Field K] {L : Type u_1} [Field L] {A : ValuationSubring L} {f : K →+* L} {x : K} :
                                                    x A.comap f f x A
                                                    theorem ValuationSubring.comap_comap {K : Type u} [Field K] {L : Type u_1} {J : Type u_2} [Field L] [Field J] (A : ValuationSubring J) (g : L →+* J) (f : K →+* L) :
                                                    (A.comap g).comap f = A.comap (g.comp f)