The basics of valuation theory. #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
The basic theory of valuations (non-archimedean norms) on a commutative ring, following T. Wedhorn's unpublished notes “Adic Spaces” ([Wed19]).
The definition of a valuation we use here is Definition 1.22 of [Wed19].
A valuation on a ring R is a monoid homomorphism v to a linearly ordered
commutative monoid with zero, that in addition satisfies the following two axioms:
v 0 = 0∀ x y, v (x + y) ≤ max (v x) (v y)
valuation R Γ₀is the type of valuations R → Γ₀, with a coercion to the underlying
function. If v is a valuation from R to Γ₀ then the induced group
homomorphism units(R) → Γ₀ is called unit_map v.
The equivalence "relation" is_equiv v₁ v₂ : Prop defined in 1.27 of [Wed19] is not strictly
speaking a relation, because v₁ : valuation R Γ₁ and v₂ : valuation R Γ₂ might
not have the same type. This corresponds in ZFC to the set-theoretic difficulty
that the class of all valuations (as Γ₀ varies) on a ring R is not a set.
The "relation" is however reflexive, symmetric and transitive in the obvious
sense. Note that we use 1.27(iii) of [Wed19] as the definition of equivalence.
Main definitions #
-
valuation R Γ₀, the type of valuations onRwith values inΓ₀ -
valuation.is_equiv, the heterogeneous equivalence relation on valuations -
valuation.supp, the support of a valuation -
add_valuation R Γ₀, the type of additive valuations onRwith values in a linearly ordered additive commutative group with a top element,Γ₀.
Implementation Details #
add_valuation R Γ₀ is implemented as valuation R (multiplicative Γ₀)ᵒᵈ.
Notation #
In the discrete_valuation locale:
ℕₘ₀is a shorthand forwith_zero (multiplicative ℕ)ℤₘ₀is a shorthand forwith_zero (multiplicative ℤ)
TODO #
If ever someone extends valuation, we should fully comply to the fun_like by migrating the
boilerplate lemmas to valuation_class.
- to_monoid_with_zero_hom : R →*₀ Γ₀
- map_add_le_max' : ∀ (x y : R), self.to_monoid_with_zero_hom.to_fun (x + y) ≤ linear_order.max (self.to_monoid_with_zero_hom.to_fun x) (self.to_monoid_with_zero_hom.to_fun y)
The type of Γ₀-valued valuations on R.
When you extend this structure, make sure to extend valuation_class.
Instances for valuation
- valuation.has_sizeof_inst
- valuation.has_coe_t
- valuation.valuation_class
- valuation.has_coe_to_fun
- to_monoid_with_zero_hom_class : monoid_with_zero_hom_class F R Γ₀
- map_add_le_max : ∀ (f : F) (x y : R), ⇑f (x + y) ≤ linear_order.max (⇑f x) (⇑f y)
valuation_class F α β states that F is a type of valuations.
You should also extend this typeclass when you extend valuation.
Instances of this typeclass
Instances of other typeclasses for valuation_class
- valuation_class.has_sizeof_inst
Equations
- valuation.has_coe_t F R Γ₀ = {coe := λ (f : F), {to_monoid_with_zero_hom := {to_fun := ⇑f, map_zero' := _, map_one' := _, map_mul' := _}, map_add_le_max' := _}}
Equations
- valuation.valuation_class = {to_monoid_with_zero_hom_class := {coe := λ (f : valuation R Γ₀), f.to_monoid_with_zero_hom.to_fun, coe_injective' := _, map_mul := _, map_one := _, map_zero := _}, map_add_le_max := _}
Helper instance for when there's too many metavariables to apply fun_like.has_coe_to_fun
directly.
Equations
Deprecated. Use fun_like.ext_iff.
A valuation gives a preorder on the underlying ring.
Equations
If v is a valuation on a division ring then v(x) = 0 iff x = 0.
A ring homomorphism S → R induces a map valuation R Γ₀ → valuation S Γ₀.
Equations
- valuation.comap f v = {to_monoid_with_zero_hom := {to_fun := ⇑v ∘ ⇑f, map_zero' := _, map_one' := _, map_mul' := _}, map_add_le_max' := _}
A ≤-preserving group homomorphism Γ₀ → Γ'₀ induces a map valuation R Γ₀ → valuation R Γ'₀.
Equations
- valuation.map f hf v = {to_monoid_with_zero_hom := {to_fun := ⇑f ∘ ⇑v, map_zero' := _, map_one' := _, map_mul' := _}, map_add_le_max' := _}
Two valuations on R are defined to be equivalent if they induce the same preorder on R.
The subgroup of elements whose valuation is less than a certain unit.
comap preserves equivalence.
The support of a valuation v : R → Γ₀ is the ideal of R where v vanishes.
Instances for valuation.supp
The support of a valuation is a prime ideal.
The type of Γ₀-valued additive valuations on R.
Equations
- add_valuation R Γ₀ = valuation R (multiplicative Γ₀ᵒᵈ)
Instances for add_valuation
A valuation is coerced to the underlying function R → Γ₀.
Equations
- add_valuation.has_coe_to_fun R Γ₀ = {coe := λ (v : add_valuation R Γ₀), v.to_monoid_with_zero_hom.to_fun}
An alternate constructor of add_valuation, that doesn't reference multiplicative Γ₀ᵒᵈ
Equations
- add_valuation.of f h0 h1 hadd hmul = {to_monoid_with_zero_hom := {to_fun := f, map_zero' := h0, map_one' := h1, map_mul' := hmul}, map_add_le_max' := hadd}
The valuation associated to an add_valuation (useful if the latter is constructed using
add_valuation.of).
A valuation gives a preorder on the underlying ring.
Equations
If v is an additive valuation on a division ring then v(x) = ⊤ iff x = 0.
A ring homomorphism S → R induces a map add_valuation R Γ₀ → add_valuation S Γ₀.
Equations
- add_valuation.comap f v = valuation.comap f v
A ≤-preserving, ⊤-preserving group homomorphism Γ₀ → Γ'₀ induces a map
add_valuation R Γ₀ → add_valuation R Γ'₀.
Equations
- add_valuation.map f ht hf v = valuation.map {to_fun := ⇑f, map_zero' := ht, map_one' := _, map_mul' := _} _ v
Two additive valuations on R are defined to be equivalent if they induce the same
preorder on R.
Equations
- v₁.is_equiv v₂ = valuation.is_equiv v₁ v₂
comap preserves equivalence.
The support of an additive valuation v : R → Γ₀ is the ideal of R where v x = ⊤
Equations
- v.supp = valuation.supp v