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 on
R`.
Main Definitions #
smoothingSeminormSeq
: theℝ
-valued sequence sendingn
to((f( (x ^ n)) ^ (1 / n : ℝ)
.smoothingFun
: the iInf of the sequencen ↦ f(x ^ (n : ℕ))) ^ (1 / (n : ℝ)
.smoothingSeminorm
: ifμ 1 ≤ 1
andμ
is nonarchimedean, thensmoothingFun
is a ring seminorm.
Main Results #
tendsto_smoothingFun_of_map_one_le_one
: ifμ 1 ≤ 1
, thensmoothingFun μ x
is the limit ofsmoothingSeminormSeq μ x
asn
tends to infinity.isNonarchimedean_smoothingFun
: ifμ 1 ≤ 1
andμ
is nonarchimedean, thensmoothingFun
is nonarchimedean.isPowMul_smoothingFun
: ifμ 1 ≤ 1
andμ
is nonarchimedean, thensmoothingFun μ
is power-multiplicative.
References #
Tags #
smoothingSeminorm, seminorm, nonarchimedean
The ℝ
-valued sequence sending n
to (μ (x ^ n))^(1/n : ℝ)
.
Equations
- smoothingSeminormSeq μ x n = μ (x ^ n) ^ (1 / ↑n)
Instances For
0
is a lower bound of smoothingSeminormSeq
.
smoothingSeminormSeq
is bounded below (by zero).
The iInf of the sequence n ↦ μ(x ^ (n : ℕ)))^(1 / (n : ℝ)
.
Instances For
If μ x = 0
, then smoothingFun μ x
is the limit of smoothingSeminormSeq μ x
.
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.
If μ 1 ≤ 1
, then smoothingFun μ x
is nonnegative.
If μ 1 ≤ 1
, then smoothingFun μ 1 ≤ 1
.
For any x
and any positive n
, smoothingFun μ x ≤ μ (x ^ (n : ℕ))^(1 / n : ℝ)
.
For all x : R
, smoothingFun μ x ≤ μ x
.
If μ 1 ≤ 1
and μ
is nonarchimedean, then smoothingFun μ
is nonarchimedean.
If μ 1 ≤ 1
and μ
is nonarchimedean, then smoothingFun
is a ring seminorm.
Equations
- smoothingSeminorm μ hμ1 hna = { toFun := smoothingFun μ, map_zero' := ⋯, add_le' := ⋯, neg' := ⋯, mul_le' := ⋯ }
Instances For
If μ 1 ≤ 1
and μ
is nonarchimedean, then smoothingSeminorm μ 1 ≤ 1
.
If μ 1 ≤ 1
and μ
is nonarchimedean, then smoothingFun μ
is
power-multiplicative.
If μ 1 ≤ 1
and ∀ (1 ≤ n), μ (x ^ n) = μ x ^ n
, then smoothingFun μ x = μ x
.
If μ 1 ≤ 1
and ∀ y : R, μ (x * y) = μ x * μ y
, then smoothingFun μ x = μ x
.
If μ 1 ≤ 1
, μ
is nonarchimedean, and ∀ y : R, μ (x * y) = μ x * μ y
, then
smoothingSeminorm μ x = μ x
.
If μ 1 ≤ 1
, and x
is multiplicative for μ
, then it is multiplicative for
smoothingFun
.
If μ 1 ≤ 1
, μ
is nonarchimedean, and x
is multiplicative for μ
, then x
is
multiplicative for smoothingSeminorm
.