mathlib3 documentation

number_theory.class_number.admissible_absolute_value

Admissible absolute values #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4. This file defines a structure absolute_value.is_admissible which we use to show the class number of the ring of integers of a global field is finite.

Main definitions #

Main results #

An absolute value R → ℤ is admissible if it respects the Euclidean domain structure and a large enough set of elements in R^n will contain a pair of elements whose remainders are pointwise close together.

Instances for absolute_value.is_admissible
theorem absolute_value.is_admissible.exists_partition {R : Type u_1} [euclidean_domain R] {abv : absolute_value R } {ι : Type u_2} [fintype ι] {ε : } (hε : 0 < ε) {b : R} (hb : b 0) (A : ι R) (h : abv.is_admissible) :
(t : ι fin (h.card ε)), (i₀ i₁ : ι), t i₀ = t i₁ (abv (A i₁ % b - A i₀ % b)) < abv b ε

For all ε > 0 and finite families A, we can partition the remainders of A mod b into abv.card ε sets, such that all elements in each part of remainders are close together.

theorem absolute_value.is_admissible.exists_approx_aux {R : Type u_1} [euclidean_domain R] {abv : absolute_value R } (n : ) (h : abv.is_admissible) {ε : } (hε : 0 < ε) {b : R} (hb : b 0) (A : fin (h.card ε ^ n).succ fin n R) :
(i₀ i₁ : fin (h.card ε ^ n).succ), i₀ i₁ (k : fin n), (abv (A i₁ k % b - A i₀ k % b)) < abv b ε

Any large enough family of vectors in R^n has a pair of elements whose remainders are close together, pointwise.

theorem absolute_value.is_admissible.exists_approx {R : Type u_1} [euclidean_domain R] {abv : absolute_value R } {ι : Type u_2} [fintype ι] {ε : } (hε : 0 < ε) {b : R} (hb : b 0) (h : abv.is_admissible) (A : fin (h.card ε ^ fintype.card ι).succ ι R) :
(i₀ i₁ : fin (h.card ε ^ fintype.card ι).succ), i₀ i₁ (k : ι), (abv (A i₁ k % b - A i₀ k % b)) < abv b ε

Any large enough family of vectors in R^ι has a pair of elements whose remainders are close together, pointwise.