Documentation

Mathlib.RingTheory.Valuation.ValuativeRel

Valuative Relations #

In this file we introduce a class called ValuativeRel R for a ring R. This bundles a relation rel : R → R → Prop on R which mimics a preorder on R arising from a valuation. We introduce the notation x ≤ᵥ y for this relation.

Recall that the equivalence class of a valuation is completely characterized by such a preorder. Thus, we can think of ValuativeRel R as a way of saying that R is endowed with an equivalence class of valuations.

Main Definitions #

Remark #

The last two axioms in ValuativeRel, namely rel_mul_cancel and not_rel_one_zero, are used to ensure that we have a well-behaved valuation taking values in a value group (with zero). In principle, it should be possible to drop these two axioms and obtain a value monoid, however, such a value monoid would not necessarily embed into an ordered abelian group with zero. Similarly, without these axioms, the support of the valuation need not be a prime ideal. We have thus opted to include these two axioms and obtain a ValueGroupWithZero associated to a ValuativeRel in order to best align with the literature about valuations on commutative rings.

Future work could refactor ValuativeRel by dropping the rel_mul_cancel and not_rel_one_zero axioms, opting to make these mixins instead.

Projects #

The ValuativeRel class should eventually replace the existing Valued typeclass. Once such a refactor happens, ValuativeRel could be renamed to Valued.

class ValuativeRel (R : Type u_1) [CommRing R] :
Type u_1

The class [ValuativeRel R] class introduces an operator x ≤ᵥ y : Prop for x y : R which is the natural relation arising from (the equivalence class of) a valuation on R. More precisely, if v is a valuation on R then the associated relation is x ≤ᵥ y ↔ v x ≤ v y. Use this class to talk about the case where R is equipped with an equivalence class of valuations.

Instances

    The relation operator arising from ValuativeRel.

    Equations
    Instances For

      We say that a valuation v is Compatible if the relation x ≤ᵥ y is equivalent to v x ≤ x y.

      Instances

        A preorder on a ring is said to be "valuative" if it agrees with the valuative relation.

        Instances
          @[simp]
          theorem ValuativeRel.rel_refl {R : Type u_1} [CommRing R] [ValuativeRel R] (x : R) :
          x ≤ᵥ x
          theorem ValuativeRel.rel_rfl {R : Type u_1} [CommRing R] [ValuativeRel R] {x : R} :
          x ≤ᵥ x
          theorem ValuativeRel.rel.refl {R : Type u_1} [CommRing R] [ValuativeRel R] (x : R) :
          x ≤ᵥ x

          Alias of ValuativeRel.rel_refl.

          theorem ValuativeRel.rel.rfl {R : Type u_1} [CommRing R] [ValuativeRel R] {x : R} :
          x ≤ᵥ x

          Alias of ValuativeRel.rel_rfl.

          @[simp]
          theorem ValuativeRel.zero_rel {R : Type u_1} [CommRing R] [ValuativeRel R] (x : R) :
          0 ≤ᵥ x
          theorem ValuativeRel.rel_mul_left {R : Type u_1} [CommRing R] [ValuativeRel R] {x y : R} (z : R) :
          x ≤ᵥ yz * x ≤ᵥ z * y
          theorem ValuativeRel.rel.trans {R : Type u_1} {inst✝ : CommRing R} [self : ValuativeRel R] {z y x : R} :
          x ≤ᵥ yy ≤ᵥ zx ≤ᵥ z

          Alias of ValuativeRel.rel_trans.

          theorem ValuativeRel.rel_trans' {R : Type u_1} [CommRing R] [ValuativeRel R] {x y z : R} (h1 : y ≤ᵥ z) (h2 : x ≤ᵥ y) :
          x ≤ᵥ z
          theorem ValuativeRel.rel.trans' {R : Type u_1} [CommRing R] [ValuativeRel R] {x y z : R} (h1 : y ≤ᵥ z) (h2 : x ≤ᵥ y) :
          x ≤ᵥ z

          Alias of ValuativeRel.rel_trans'.

          theorem ValuativeRel.rel_mul {R : Type u_1} [CommRing R] [ValuativeRel R] {x x' y y' : R} (h1 : x ≤ᵥ y) (h2 : x' ≤ᵥ y') :
          x * x' ≤ᵥ y * y'
          theorem ValuativeRel.rel_add_cases {R : Type u_1} [CommRing R] [ValuativeRel R] (x y : R) :
          x + y ≤ᵥ x x + y ≤ᵥ y

          The submonoid of elements x : R whose valuation is positive.

          Equations
          Instances For
            @[simp]
            @[simp]
            theorem ValuativeRel.right_cancel_posSubmonoid {R : Type u_1} [CommRing R] [ValuativeRel R] (x y : R) (u : (posSubmonoid R)) :
            x * u ≤ᵥ y * u x ≤ᵥ y
            @[simp]
            theorem ValuativeRel.left_cancel_posSubmonoid {R : Type u_1} [CommRing R] [ValuativeRel R] (x y : R) (u : (posSubmonoid R)) :
            u * x ≤ᵥ u * y x ≤ᵥ y

            The setoid used to construct ValueGroupWithZero R.

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

              The "canonical" value group-with-zero of a ring with a valuative relation.

              Equations
              Instances For

                Construct an element of the value group-with-zero from an element r : R and y : posSubmonoid R. This should be thought of as v r / v y.

                Equations
                Instances For
                  theorem ValuativeRel.ValueGroupWithZero.sound {R : Type u_1} [CommRing R] [ValuativeRel R] {x y : R} {t s : (posSubmonoid R)} (h₁ : x * s ≤ᵥ y * t) (h₂ : y * t ≤ᵥ x * s) :
                  theorem ValuativeRel.ValueGroupWithZero.exact {R : Type u_1} [CommRing R] [ValuativeRel R] {x y : R} {t s : (posSubmonoid R)} (h : ValueGroupWithZero.mk x t = ValueGroupWithZero.mk y s) :
                  x * s ≤ᵥ y * t y * t ≤ᵥ x * s
                  theorem ValuativeRel.ValueGroupWithZero.ind {R : Type u_1} [CommRing R] [ValuativeRel R] {motive : ValueGroupWithZero RProp} (mk : ∀ (x : R) (y : (posSubmonoid R)), motive (ValueGroupWithZero.mk x y)) (t : ValueGroupWithZero R) :
                  motive t
                  def ValuativeRel.ValueGroupWithZero.lift {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) (t : ValueGroupWithZero R) :
                  α

                  Lifts a function R → posSubmonoid R → α to the value group-with-zero of R.

                  Equations
                  Instances For
                    @[simp]
                    theorem ValuativeRel.ValueGroupWithZero.lift_mk {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) (x : R) (y : (posSubmonoid R)) :
                    def ValuativeRel.ValueGroupWithZero.lift₂ {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)R(posSubmonoid R)α) (hf : ∀ (x y z w : R) (t s u v : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tz * u ≤ᵥ w * vw * v ≤ᵥ z * uf x s z v = f y t w u) (t₁ t₂ : ValueGroupWithZero R) :
                    α

                    Lifts a function R → posSubmonoid R → R → posSubmonoid R → α to the value group-with-zero of R.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem ValuativeRel.ValueGroupWithZero.lift₂_mk {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)R(posSubmonoid R)α) (hf : ∀ (x y z w : R) (t s u v : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tz * u ≤ᵥ w * vw * v ≤ᵥ z * uf x s z v = f y t w u) (x y : R) (z w : (posSubmonoid R)) :
                      theorem ValuativeRel.ValueGroupWithZero.mk_eq_mk {R : Type u_1} [CommRing R] [ValuativeRel R] {x y : R} {t s : (posSubmonoid R)} :
                      ValueGroupWithZero.mk x t = ValueGroupWithZero.mk y s x * s ≤ᵥ y * t y * t ≤ᵥ x * s
                      @[simp]
                      theorem ValuativeRel.ValueGroupWithZero.mk_eq_one {R : Type u_1} [CommRing R] [ValuativeRel R] (x : R) (y : (posSubmonoid R)) :
                      theorem ValuativeRel.ValueGroupWithZero.lift_zero {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) :
                      @[simp]
                      theorem ValuativeRel.ValueGroupWithZero.lift_one {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) :
                      Equations
                      • One or more equations did not get rendered due to their size.
                      theorem ValuativeRel.ValueGroupWithZero.lift_mul {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Type u_2} [Mul α] (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) (hdist : ∀ (a b : R) (r s : (posSubmonoid R)), f (a * b) (r * s) = f a r * f b s) (a b : ValueGroupWithZero R) :
                      Equations
                      • One or more equations did not get rendered due to their size.
                      Equations
                      • One or more equations did not get rendered due to their size.
                      @[simp]
                      Equations
                      • One or more equations did not get rendered due to their size.
                      @[simp]
                      theorem ValuativeRel.ValueGroupWithZero.mk_lt_mk {R : Type u_1} [CommRing R] [ValuativeRel R] (x y : R) (t s : (posSubmonoid R)) :
                      ValueGroupWithZero.mk x t < ValueGroupWithZero.mk y s x * s ≤ᵥ y * t ¬y * t ≤ᵥ x * s
                      Equations
                      • One or more equations did not get rendered due to their size.

                      The value group-with-zero is a linearly ordered commutative group with zero.

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

                      The "canonical" valuation associated to a valuative relation.

                      Equations
                      Instances For
                        @[simp]
                        theorem ValuativeRel.ValueGroupWithZero.lift_valuation {R : Type u_1} [CommRing R] [ValuativeRel R] {α : Sort u_2} (f : R(posSubmonoid R)α) (hf : ∀ (x y : R) (t s : (posSubmonoid R)), x * t ≤ᵥ y * sy * s ≤ᵥ x * tf x s = f y t) (x : R) :

                        Construct a valuative relation on a ring using a valuation.

                        Equations
                        • ValuativeRel.ofValuation v = { rel := fun (x y : S) => v x v y, rel_total := , rel_trans := , rel_add := , rel_mul_right := , rel_mul_cancel := , not_rel_one_zero := }
                        Instances For
                          theorem ValuativeRel.isEquiv {R : Type u_1} [CommRing R] [ValuativeRel R] {Γ₁ : Type u_2} {Γ₂ : Type u_3} [LinearOrderedCommMonoidWithZero Γ₁] [LinearOrderedCommMonoidWithZero Γ₂] (v₁ : Valuation R Γ₁) (v₂ : Valuation R Γ₂) [v₁.Compatible] [v₂.Compatible] :
                          v₁.IsEquiv v₂

                          An alias for endowing a ring with a preorder defined as the valuative relation.

                          Equations
                          Instances For

                            The ring instance on WithPreorder R arising from the ring structure on R.

                            Equations

                            The preorder on WithPreorder R arising from the valuative relation on R.

                            Equations

                            The valutaive relation on WithPreorder R arising from the valuative relation on R. This is defined as the preorder itself.

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

                            The support of the valuation on R.

                            Equations
                            Instances For
                              @[simp]
                              theorem ValuativeRel.supp_def {R : Type u_1} [CommRing R] [ValuativeRel R] (x : R) :
                              x supp R x ≤ᵥ 0

                              An auxiliary structure used to define IsRankOne.

                              Instances For

                                We say that a ring with a valuative relation is of rank one if there exists a strictly monotone embedding of the "canonical" value group-with-zero into the nonnegative reals, and the image of this embedding contains some element different from 0 and 1.

                                Instances

                                  We say that a valuative relation on a ring is nontrivial if the value group-with-zero is nontrivial, meaning that it has an element which is different from 0 and 1.

                                  Instances

                                    A ring with a valuative relation is discrete if its value group-with-zero has a maximal element < 1.

                                    • has_maximal_element : γ < 1, δ < 1, δ γ
                                    Instances
                                      theorem ValuativeRel.valuation_surjective {R : Type u_1} [CommRing R] [ValuativeRel R] (γ : ValueGroupWithZero R) :
                                      ∃ (a : R) (b : (posSubmonoid R)), (valuation R) a / (valuation R) b = γ

                                      We say that a topology on R is valuative if the neighborhoods of 0 in R are determined by the relation · ≤ᵥ ·.

                                      Instances
                                        class ValuativeExtension (A : Type u_1) (B : Type u_2) [CommRing A] [CommRing B] [ValuativeRel A] [ValuativeRel B] [Algebra A B] :

                                        If B is an A algebra and both A and B have valuative relations, we say that B|A is a valuative extension if the valuative relation on A is induced by the one on B.

                                        Instances

                                          The morphism of posSubmonoids associated to an algebra map. This is used in constructing ValuativeExtension.mapValueGroupWithZero.

                                          Equations
                                          Instances For
                                            @[simp]
                                            theorem ValuativeExtension.mapPosSubmonoid_apply_coe (A : Type u_1) (B : Type u_2) [CommRing A] [CommRing B] [ValuativeRel A] [ValuativeRel B] [Algebra A B] [ValuativeExtension A B] (x✝ : (ValuativeRel.posSubmonoid A)) :
                                            ((mapPosSubmonoid A B) x✝) = (algebraMap A B) x✝

                                            The map on value groups-with-zero associated to the structure morphism of an algebra.

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