Documentation

Mathlib.Analysis.Normed.Ring.SmoothingSeminorm

smoothingSeminorm #

In this file, we prove BGR, Proposition 1.3.2/1 : if μ is a nonarchimedean seminorm on a commutative ring R, then iInf (fun (n : PNat), (μ(x ^ (n : ℕ))) ^ (1 / (n : ℝ)))is a power-multiplicative nonarchimedean seminorm onR`.

Main Definitions #

Main Results #

References #

Tags #

smoothingSeminorm, seminorm, nonarchimedean

@[reducible, inline]
abbrev smoothingSeminormSeq {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) :

The -valued sequence sending n to (μ (x ^ n))^(1/n : ℝ).

Equations
Instances For
    theorem zero_mem_lowerBounds_smoothingSeminormSeq_range {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) :
    0 lowerBounds (Set.range fun (n : ℕ+) => μ (x ^ n) ^ (1 / n))

    0 is a lower bound of smoothingSeminormSeq.

    theorem smoothingSeminormSeq_bddBelow {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) :
    BddBelow (Set.range fun (n : ℕ+) => μ (x ^ n) ^ (1 / n))

    smoothingSeminormSeq is bounded below (by zero).

    @[reducible, inline]
    abbrev smoothingFun {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) :

    The iInf of the sequence n ↦ μ(x ^ (n : ℕ)))^(1 / (n : ℝ).

    Equations
    Instances For
      theorem tendsto_smoothingFun_of_eq_zero {R : Type u_1} [CommRing R] (μ : RingSeminorm R) {x : R} (hx : μ x = 0) :

      If μ x = 0, then smoothingFun μ x is the limit of smoothingSeminormSeq μ x.

      theorem tendsto_smoothingFun_of_ne_zero {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) {x : R} (hx : μ x 0) :

      If μ 1 ≤ 1 and μ x ≠ 0, then smoothingFun μ x is the limit of smoothingSeminormSeq μ x.

      If μ 1 ≤ 1, then smoothingFun μ x is the limit of smoothingSeminormSeq μ x as n tends to infinity.

      theorem smoothingFun_nonneg {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (x : R) :

      If μ 1 ≤ 1, then smoothingFun μ x is nonnegative.

      theorem smoothingFun_one_le {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) :

      If μ 1 ≤ 1, then smoothingFun μ 1 ≤ 1.

      theorem smoothingFun_le {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) (n : ℕ+) :
      smoothingFun μ x μ (x ^ n) ^ (1 / n)

      For any x and any positive n, smoothingFun μ x ≤ μ (x ^ (n : ℕ))^(1 / n : ℝ).

      theorem smoothingFun_le_self {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (x : R) :
      smoothingFun μ x μ x

      For all x : R, smoothingFun μ x ≤ μ x.

      theorem tendsto_smoothingFun_comp {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (x : R) {ψ : } (hψ_mono : StrictMono ψ) :
      Filter.Tendsto (fun (n : ) => μ (x ^ ψ n) ^ (1 / (ψ n))) Filter.atTop (nhds (smoothingFun μ x))
      theorem isNonarchimedean_smoothingFun {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (hna : IsNonarchimedean μ) :

      If μ 1 ≤ 1 and μ is nonarchimedean, then smoothingFun μ is nonarchimedean.

      def smoothingSeminorm {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (hna : IsNonarchimedean μ) :

      If μ 1 ≤ 1 and μ is nonarchimedean, then smoothingFun is a ring seminorm.

      Equations
      Instances For
        theorem smoothingSeminorm_map_one_le_one {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (hna : IsNonarchimedean μ) :
        (smoothingSeminorm μ hμ1 hna) 1 1

        If μ 1 ≤ 1 and μ is nonarchimedean, then smoothingSeminorm μ 1 ≤ 1.

        theorem isPowMul_smoothingFun {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) :

        If μ 1 ≤ 1 and μ is nonarchimedean, then smoothingFun μ is power-multiplicative.

        theorem smoothingFun_of_powMul {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) {x : R} (hx : ∀ (n : ), 1 nμ (x ^ n) = μ x ^ n) :
        smoothingFun μ x = μ x

        If μ 1 ≤ 1 and ∀ (1 ≤ n), μ (x ^ n) = μ x ^ n, then smoothingFun μ x = μ x.

        theorem smoothingFun_apply_of_map_mul_eq_mul {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) {x : R} (hx : ∀ (y : R), μ (x * y) = μ x * μ y) :
        smoothingFun μ x = μ x

        If μ 1 ≤ 1 and ∀ y : R, μ (x * y) = μ x * μ y, then smoothingFun μ x = μ x.

        theorem smoothingSeminorm_apply_of_map_mul_eq_mul {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (hna : IsNonarchimedean μ) {x : R} (hx : ∀ (y : R), μ (x * y) = μ x * μ y) :
        (smoothingSeminorm μ hμ1 hna) x = μ x

        If μ 1 ≤ 1, μ is nonarchimedean, and ∀ y : R, μ (x * y) = μ x * μ y, then smoothingSeminorm μ x = μ x.

        theorem smoothingFun_of_map_mul_eq_mul {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) {x : R} (hx : ∀ (y : R), μ (x * y) = μ x * μ y) (y : R) :

        If μ 1 ≤ 1, and x is multiplicative for μ, then it is multiplicative for smoothingFun.

        theorem smoothingSeminorm_of_mul {R : Type u_1} [CommRing R] (μ : RingSeminorm R) (hμ1 : μ 1 1) (hna : IsNonarchimedean μ) {x : R} (hx : ∀ (y : R), μ (x * y) = μ x * μ y) (y : R) :
        (smoothingSeminorm μ hμ1 hna) (x * y) = (smoothingSeminorm μ hμ1 hna) x * (smoothingSeminorm μ hμ1 hna) y

        If μ 1 ≤ 1, μ is nonarchimedean, and x is multiplicative for μ, then x is multiplicative for smoothingSeminorm.