Gauss norm for power series #
This file defines the Gauss norm for power series. Given a power series f
in R⟦X⟧
, a function
v : R → ℝ
and a real number c
, the Gauss norm is defined as the supremum of the set of all
values of v (f.coeff R i) * c ^ i
for all i : ℕ
.
In case f
is a polynomial, v
is a non-negative function with v 0 = 0
and c ≥ 0
,
f.gaussNorm v c
reduces to the Gauss norm defined in RingTheory/Polynomial/GaussNorm
, see
Polynomial.gaussNorm_coe_powerSeries
.
Main Definitions and Results #
PowerSeries.gaussNorm
is the supremum of the set of all values ofv (f.coeff R i) * c ^ i
for alli : ℕ
, wheref
is a power series inR⟦X⟧
,v : R → ℝ
is a function andc
is a real number.PowerSeries.gaussNorm_nonneg
: ifv
is a non-negative function, then the Gauss norm is non-negative.PowerSeries.gaussNorm_eq_zero_iff
: ifv
is a non-negative function andv x = 0 ↔ x = 0
for allx : R
andc
is positive, then the Gauss norm is zero if and only if the power series is zero.
noncomputable def
PowerSeries.gaussNorm
{R : Type u_1}
{F : Type u_2}
[Semiring R]
[FunLike F R ℝ]
(v : F)
(c : ℝ)
(f : PowerSeries R)
:
Given a power series f
in R⟦X⟧
, a function v : R → ℝ
and a real number c
, the Gauss norm
is defined as the supremum of the set of all values of v (f.coeff R i) * c ^ i
for all i : ℕ
.
Equations
- PowerSeries.gaussNorm v c f = ⨆ (i : ℕ), v ((PowerSeries.coeff R i) f) * c ^ i
Instances For
theorem
PowerSeries.gaussNorm_nonneg
{R : Type u_1}
{F : Type u_2}
[Semiring R]
[FunLike F R ℝ]
(v : F)
(c : ℝ)
(f : PowerSeries R)
[NonnegHomClass F R ℝ]
:
@[simp]
theorem
PowerSeries.gaussNorm_eq_zero_iff
{R : Type u_1}
{F : Type u_2}
[Semiring R]
[FunLike F R ℝ]
[ZeroHomClass F R ℝ]
[NonnegHomClass F R ℝ]
{v : F}
(h_eq_zero : ∀ (x : R), v x = 0 → x = 0)
{f : PowerSeries R}
{c : ℝ}
(hc : 0 < c)
(hbd : BddAbove (Set.range fun (i : ℕ) => v ((coeff R i) f) * c ^ i))
: