Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat(RingTheory.MvPowerSeries.Topology): define the product topology on mv power series #14866

Open
wants to merge 82 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from 74 commits
Commits
Show all changes
82 commits
Select commit Hold shift + click to select a range
7cb1908
coefficients of powers of MvPowerSeries
AntoineChambert-Loir Jul 18, 2024
bdef1b0
initial commit
AntoineChambert-Loir Jul 18, 2024
06c9528
update references.bib
AntoineChambert-Loir Jul 18, 2024
543b97f
delete one blank line
AntoineChambert-Loir Jul 18, 2024
c2c4976
Merge branch 'ACLMIdFF/MvPowerSeries_coeffpow' into ACLMIdFF/MvPowerS…
AntoineChambert-Loir Jul 18, 2024
3d5a7d0
defines the topology (removes homogeneous components)
AntoineChambert-Loir Jul 18, 2024
2649640
correct def to theorem
AntoineChambert-Loir Jul 18, 2024
8345665
update
AntoineChambert-Loir Jul 18, 2024
805f1d5
mention non-relation with adic topology
AntoineChambert-Loir Jul 18, 2024
54bbcea
correct centered dots
AntoineChambert-Loir Jul 18, 2024
8d001f7
update Mathlib.lean
AntoineChambert-Loir Jul 18, 2024
d507e68
make lint happy (try to)
AntoineChambert-Loir Jul 18, 2024
1191b60
update
AntoineChambert-Loir Jul 18, 2024
cb8c681
move weights to other file
AntoineChambert-Loir Jul 18, 2024
5820cc0
homogeneize
AntoineChambert-Loir Jul 18, 2024
6e26559
adjust for degrees
AntoineChambert-Loir Jul 18, 2024
aae13ce
add le_degree
AntoineChambert-Loir Jul 18, 2024
df1d019
update mathlib.lean
AntoineChambert-Loir Jul 18, 2024
b2bbb36
try to make lint happy
AntoineChambert-Loir Jul 18, 2024
f649884
adjust
AntoineChambert-Loir Jul 18, 2024
4831b9b
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
bdfe910
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
bafb2d7
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
d744b9a
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
bc296a4
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
93df652
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
da2b237
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 18, 2024
c6886c7
apply (almost all of) Filippo's comments
AntoineChambert-Loir Jul 18, 2024
a2049f3
adjust
AntoineChambert-Loir Jul 18, 2024
e7c89bd
adjust and simplify
AntoineChambert-Loir Jul 18, 2024
165e928
make a def a theorem
AntoineChambert-Loir Jul 18, 2024
3a0ec1e
add a docstring
AntoineChambert-Loir Jul 18, 2024
4ecd44c
adjust to import - linter
AntoineChambert-Loir Jul 18, 2024
8dea296
adjust scoped instances following linter
AntoineChambert-Loir Jul 18, 2024
e0eb0fb
adjust import
AntoineChambert-Loir Jul 18, 2024
63924ca
adjust import
AntoineChambert-Loir Jul 18, 2024
72f8dce
Merge branch 'master' into ACLMIdFF/degree
AntoineChambert-Loir Jul 19, 2024
521492a
remove all #align and #align_import
AntoineChambert-Loir Jul 19, 2024
3054cbb
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
ab06f0f
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
ecade7c
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
c82153a
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
94ea8bf
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
c3e9092
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
54dafe5
Update Mathlib/RingTheory/MvPowerSeries/Basic.lean
AntoineChambert-Loir Jul 19, 2024
6758185
adjust
AntoineChambert-Loir Jul 19, 2024
9123cf2
linter 100
AntoineChambert-Loir Jul 19, 2024
f1f83c1
Merge branch 'ACLMIdFF/degree' into ACLMIdFF/MvPowerSeries_coeffpow
AntoineChambert-Loir Jul 20, 2024
761bbcf
adjust
AntoineChambert-Loir Jul 20, 2024
0d1c9d6
make some arguments implicit
AntoineChambert-Loir Jul 20, 2024
b9da976
Merge branch 'ACLMIdFF/degree' into ACLMIdFF/MvPowerSeries_coeffpow
AntoineChambert-Loir Jul 20, 2024
c96c50c
add docstring
AntoineChambert-Loir Jul 20, 2024
0492feb
linter 100 / adjust
AntoineChambert-Loir Jul 20, 2024
c2d9b6f
Merge branch 'master' into ACLMIdFF/MvPowerSeries_coeffpow
AntoineChambert-Loir Jul 21, 2024
48d2f35
change refine' to apply
AntoineChambert-Loir Jul 21, 2024
13cb3b2
Merge branch 'ACLMIdFF/MvPowerSeries_coeffpow' into ACLMIdFF/MvPowerS…
AntoineChambert-Loir Jul 21, 2024
d167eb3
adjust
AntoineChambert-Loir Jul 21, 2024
8d7a099
better docstring
AntoineChambert-Loir Jul 21, 2024
c21f5dd
better preamble
AntoineChambert-Loir Jul 21, 2024
f46bac1
linter 100
AntoineChambert-Loir Jul 21, 2024
ddd42b9
change α to R
AntoineChambert-Loir Jul 23, 2024
3a3dbf4
answer JC's comments + some lemmas hold more generally
AntoineChambert-Loir Jul 23, 2024
082d250
linter 100
AntoineChambert-Loir Jul 23, 2024
19944af
generalize
AntoineChambert-Loir Jul 23, 2024
9bf21d8
Merge branch 'master' into ACLMIdFF/MvPowerSeries_Topology
AntoineChambert-Loir Jul 29, 2024
d03c51f
rename Topology to PiTopology
AntoineChambert-Loir Jul 29, 2024
62170fe
mv stuff to BigOperators and change apply_eq_coeff
AntoineChambert-Loir Jul 29, 2024
0493cf1
never consider power series as functions
AntoineChambert-Loir Jul 29, 2024
379fbca
cleanup, docstrings
AntoineChambert-Loir Jul 29, 2024
d0cf16b
add PowerSeries/PiTopology
AntoineChambert-Loir Jul 29, 2024
e6e52e8
correct class
AntoineChambert-Loir Jul 29, 2024
71ed0cc
use Pi instances, reorder file
AntoineChambert-Loir Sep 6, 2024
b369662
fix namespaces
mariainesdff Sep 13, 2024
5d19e4d
expand comments
mariainesdff Sep 13, 2024
c6faa8a
Update Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
13fa966
Update Mathlib/RingTheory/PowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
3546540
Update Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
596670e
Update Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
a817bc3
Update Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
8212c6d
Update Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
AntoineChambert-Loir Sep 20, 2024
054ea42
remove variable declaration and DecidableEq assumption
AntoineChambert-Loir Sep 20, 2024
57fce93
add classical
AntoineChambert-Loir Sep 20, 2024
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
2 changes: 2 additions & 0 deletions Mathlib.lean
Original file line number Diff line number Diff line change
Expand Up @@ -3822,6 +3822,7 @@ import Mathlib.RingTheory.MvPolynomial.WeightedHomogeneous
import Mathlib.RingTheory.MvPowerSeries.Basic
import Mathlib.RingTheory.MvPowerSeries.Inverse
import Mathlib.RingTheory.MvPowerSeries.NoZeroDivisors
import Mathlib.RingTheory.MvPowerSeries.PiTopology
import Mathlib.RingTheory.MvPowerSeries.Trunc
import Mathlib.RingTheory.Nakayama
import Mathlib.RingTheory.Nilpotent.Basic
Expand Down Expand Up @@ -3873,6 +3874,7 @@ import Mathlib.RingTheory.PowerSeries.Basic
import Mathlib.RingTheory.PowerSeries.Derivative
import Mathlib.RingTheory.PowerSeries.Inverse
import Mathlib.RingTheory.PowerSeries.Order
import Mathlib.RingTheory.PowerSeries.PiTopology
import Mathlib.RingTheory.PowerSeries.Trunc
import Mathlib.RingTheory.PowerSeries.WellKnown
import Mathlib.RingTheory.Presentation
Expand Down
8 changes: 8 additions & 0 deletions Mathlib/Algebra/BigOperators/Ring.lean
Original file line number Diff line number Diff line change
Expand Up @@ -176,6 +176,14 @@ theorem prod_add (f g : ι → α) (s : Finset ι) :
simp only [mem_filter, mem_sdiff, not_and, not_exists, and_congr_right_iff]
tauto)

theorem prod_one_add {f : ι → α} (s : Finset ι) :
(s.prod fun i => 1 + f i) = s.powerset.sum fun t => t.prod f := by
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
(s.prod fun i => 1 + f i) = s.powerset.sum fun t => t.prod f := by
∏ i ∈ s, (1 + f i) = ∑ t ∈ s.powerset, t.prod f := by

simp_rw [add_comm, Finset.prod_add]
congr
ext t
convert mul_one (Finset.prod t fun a => f a)
exact Finset.prod_eq_one (fun i _ => rfl)

end DecidableEq

/-- `∏ i, (f i + g i) = (∏ i, f i) + ∑ i, g i * (∏ j < i, f j + g j) * (∏ j > i, f j)`. -/
Expand Down
259 changes: 259 additions & 0 deletions Mathlib/RingTheory/MvPowerSeries/PiTopology.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,259 @@
/-
Copyright (c) 2024 Antoine Chambert-Loir, María Inés de Frutos Fernández. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Antoine Chambert-Loir, María Inés de Frutos Fernández
-/

import Mathlib.RingTheory.Nilpotent.Defs
import Mathlib.RingTheory.MvPowerSeries.Basic
import Mathlib.Topology.Algebra.InfiniteSum.Constructions
import Mathlib.Topology.Algebra.Ring.Basic
import Mathlib.Topology.Algebra.UniformGroup
import Mathlib.Topology.UniformSpace.Pi

/-! # Product topology on mv power series

Let `R` be with `Semiring R` and `TopologicalSpace R`
In this file we define the topology on `MvPowerSeries σ R`
that corresponds to the simple convergence on its coefficients.
It is the coarsest topology for which all coefficients maps are continuous.

When `R` has `UniformSpace R`, we define the corresponding uniform structure.

This topology can be included by writing `open scoped MvPowerSeries.WithPiTopology`.

When the type of coefficients has the discrete topology,
it corresponds to the topology defined by [bourbaki1981], chapter 4, §4, n°2.

It is *not* the adic topology in general.

- `MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_nilpotent`,
`MvPowerSeries.WithPiTopology.tendsto_pow_zero_of_constantCoeff_zero`: if the constant coefficient
of `f` is nilpotent, or vanishes, then the powers of `f` converge to zero.

- `MvPowerSeries.WithPiTopology.tendsto_pow_of_constantCoeff_nilpotent_iff` : the powers of `f`
converge to zero iff the constant coefficient of `f` is nilpotent.

- `MvPowerSeries.WithPiTopology.hasSum_of_monomials_self` : viewed as an infinite sum, a power
series coverges to itself.

TODO: add the similar result for the series of homogeneous components.

## Instances

- If `R` is a topological (semi)ring, then so is `MvPowerSeries σ R`.

- If the topology of `R` is T0 or T2, then so is that of `MvPowerSeries σ R`.

- If `R` is a `UniformAddGroup`, then so is `MvPowerSeries σ R`.

- If `R` is complete, then so is `MvPowerSeries σ R`.
AntoineChambert-Loir marked this conversation as resolved.
Show resolved Hide resolved

-/


theorem MvPowerSeries.apply_eq_coeff {σ R : Type _} [Semiring R] (f : MvPowerSeries σ R)
(d : σ →₀ ℕ) : f d = MvPowerSeries.coeff R d f :=
rfl
Comment on lines +52 to +54
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Should this go elsewhere?


namespace MvPowerSeries

open Function

variable {σ R : Type*}

namespace WithPiTopology

section Topology

variable [TopologicalSpace R]

variable (R) in
/-- The pointwise topology on MvPowerSeries -/
scoped instance : TopologicalSpace (MvPowerSeries σ R) :=
Pi.topologicalSpace

/-- MvPowerSeries on a T0Space form a T0Space -/
@[scoped instance]
theorem instT0Space [T0Space R] : T0Space (MvPowerSeries σ R) := Pi.instT0Space

/-- MvPowerSeries on a T2Space form a T2Space -/
@[scoped instance]
theorem instT2Space [T2Space R] : T2Space (MvPowerSeries σ R) := Pi.t2Space

variable (R) in
/-- coeff are continuous -/
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
/-- coeff are continuous -/
/-- `MvPowerSeries.coeff` is continuous. -/

theorem continuous_coeff [Semiring R] (d : σ →₀ ℕ) : Continuous (MvPowerSeries.coeff R d) :=
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
theorem continuous_coeff [Semiring R] (d : σ →₀ ℕ) : Continuous (MvPowerSeries.coeff R d) :=
@[fun_prop]
theorem continuous_coeff [Semiring R] (d : σ →₀ ℕ) : Continuous (MvPowerSeries.coeff R d) :=

continuous_pi_iff.mp continuous_id d

variable (R) in
/-- constant_coeff is continuous -/
theorem continuous_constantCoeff [Semiring R] : Continuous (constantCoeff σ R) :=
continuous_coeff R 0

/-- A family of power series converges iff it converges coefficientwise -/
theorem tendsto_iff_coeff_tendsto [Semiring R] {ι : Type*}
(f : ι → MvPowerSeries σ R) (u : Filter ι) (g : MvPowerSeries σ R) :
Filter.Tendsto f u (nhds g) ↔
∀ d : σ →₀ ℕ, Filter.Tendsto (fun i => coeff R d (f i)) u (nhds (coeff R d g)) := by
rw [nhds_pi, Filter.tendsto_pi]
exact forall_congr' (fun d => Iff.rfl)

variable (σ R)

/-- The semiring topology on MvPowerSeries of a topological semiring -/
@[scoped instance]
theorem instTopologicalSemiring [Semiring R] [TopologicalSemiring R] :
TopologicalSemiring (MvPowerSeries σ R) where
continuous_add := continuous_pi fun d => Continuous.comp continuous_add
(Continuous.prod_mk (Continuous.fst' (continuous_coeff R d))
(Continuous.snd' (continuous_coeff R d)))
continuous_mul := continuous_pi fun _ => continuous_finset_sum _ (fun i _ => Continuous.comp
continuous_mul (Continuous.prod_mk (Continuous.fst' (continuous_coeff R i.fst))
(Continuous.snd' (continuous_coeff R i.snd))))
Comment on lines +103 to +110
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This can probably be shortened using dot-notation

Comment on lines +105 to +110
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Could you try using fun_prop here?


/-- The ring topology on MvPowerSeries of a topological ring -/
@[scoped instance]
theorem instTopologicalRing (R : Type*) [TopologicalSpace R] [Ring R] [TopologicalRing R] :
TopologicalRing (MvPowerSeries σ R) :=
{ instTopologicalSemiring σ R with
continuous_neg := continuous_pi fun d ↦ Continuous.comp continuous_neg
(continuous_coeff R d) }

variable {σ R}

variable [DecidableEq σ] [TopologicalSpace R]

theorem continuous_C [Ring R] [TopologicalRing R] :
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
theorem continuous_C [Ring R] [TopologicalRing R] :
@[fun_prop]
theorem continuous_C [Ring R] [TopologicalRing R] :

Continuous (C σ R) := by
apply continuous_of_continuousAt_zero
rw [continuousAt_pi]
intro d
change ContinuousAt (fun y => coeff R d ((C σ R) y)) 0
by_cases hd : d = 0
· convert continuousAt_id
rw [hd, coeff_zero_C, id_eq]
· convert continuousAt_const
rw [coeff_C, if_neg hd]

theorem variables_tendsto_zero [Semiring R] :
Filter.Tendsto (fun s : σ => (X s : MvPowerSeries σ R)) Filter.cofinite (nhds 0) := by
AntoineChambert-Loir marked this conversation as resolved.
Show resolved Hide resolved
classical
rw [tendsto_pi_nhds]
intro d s hs
rw [Filter.mem_map, Filter.mem_cofinite, ← Set.preimage_compl]
by_cases h : ∃ i, d = Finsupp.single i 1
· obtain ⟨i, rfl⟩ := h
apply Set.Finite.subset (Set.finite_singleton i)
intro x
simp only [OfNat.ofNat, Zero.zero] at hs
rw [Set.mem_preimage, Set.mem_compl_iff, Set.mem_singleton_iff, not_imp_comm]
intro hx
convert mem_of_mem_nhds hs
rw [apply_eq_coeff (X x) (Finsupp.single i 1), coeff_X, if_neg]
rfl
· simp only [Finsupp.single_eq_single_iff, Ne.symm hx, and_true, one_ne_zero, and_self,
or_self, not_false_eq_true]
· convert Set.finite_empty
rw [Set.eq_empty_iff_forall_not_mem]
intro x
rw [Set.mem_preimage, Set.not_mem_compl_iff]
convert mem_of_mem_nhds hs using 1
rw [apply_eq_coeff (X x) d, coeff_X, if_neg]
rfl
· intro h'
apply h
exact ⟨x, h'⟩
AntoineChambert-Loir marked this conversation as resolved.
Show resolved Hide resolved

theorem tendsto_pow_zero_of_constantCoeff_nilpotent [CommSemiring R]
{f} (hf : IsNilpotent (constantCoeff σ R f)) :
Filter.Tendsto (fun n : ℕ => f ^ n) Filter.atTop (nhds 0) := by
classical
obtain ⟨m, hm⟩ := hf
simp_rw [tendsto_iff_coeff_tendsto, coeff_zero]
exact fun d => tendsto_atTop_of_eventually_const fun n hn =>
AntoineChambert-Loir marked this conversation as resolved.
Show resolved Hide resolved
coeff_eq_zero_of_constantCoeff_nilpotent hm hn

theorem tendsto_pow_zero_of_constantCoeff_zero [CommSemiring R]
{f} (hf : constantCoeff σ R f = 0) :
Filter.Tendsto (fun n : ℕ => f ^ n) Filter.atTop (nhds 0) := by
apply tendsto_pow_zero_of_constantCoeff_nilpotent
rw [hf]
exact IsNilpotent.zero

/-- The powers of a `MvPowerSeries` converge to 0 iff its constant coefficient is nilpotent.
N. Bourbaki, *Algebra II*, [bourbaki1981] (chap. 4, §4, n°2, corollaire de la prop. 3) -/
theorem tendsto_pow_of_constantCoeff_nilpotent_iff [CommRing R] [DiscreteTopology R] (f) :
Filter.Tendsto (fun n : ℕ => f ^ n) Filter.atTop (nhds 0) ↔
IsNilpotent (constantCoeff σ R f) := by
refine ⟨?_, tendsto_pow_zero_of_constantCoeff_nilpotent⟩
intro h
suffices Filter.Tendsto (fun n : ℕ => constantCoeff σ R (f ^ n)) Filter.atTop (nhds 0) by
simp only [Filter.tendsto_def] at this
specialize this {0} _
suffices ∀ x : R, {x} ∈ nhds x by exact this 0
AntoineChambert-Loir marked this conversation as resolved.
Show resolved Hide resolved
rw [← discreteTopology_iff_singleton_mem_nhds]; infer_instance
simp only [map_pow, Filter.mem_atTop_sets, ge_iff_le, Set.mem_preimage,
Set.mem_singleton_iff] at this
obtain ⟨m, hm⟩ := this
use m
apply hm m (le_refl m)
simp only [← @comp_apply _ R ℕ, ← Filter.tendsto_map'_iff]
simp only [Filter.Tendsto, Filter.map_le_iff_le_comap] at h ⊢
refine le_trans h (Filter.comap_mono ?_)
rw [← Filter.map_le_iff_le_comap]
exact Continuous.continuousAt (continuous_constantCoeff R)

variable [Semiring R]

/-- A power series is the sum (in the sense of summable families) of its monomials -/
theorem hasSum_of_monomials_self (f : MvPowerSeries σ R) :
HasSum (fun d : σ →₀ ℕ => monomial R d (coeff R d f)) f := by
rw [Pi.hasSum]
intro d
convert hasSum_single d ?_ using 1
exact (coeff_monomial_same d _).symm
exact fun d' h ↦ coeff_monomial_ne (Ne.symm h) _

/-- If the coefficient space is T2, then the power series is `tsum` of its monomials -/
theorem as_tsum [T2Space R] (f : MvPowerSeries σ R) :
f = tsum fun d : σ →₀ ℕ => monomial R d (coeff R d f) :=
(HasSum.tsum_eq (hasSum_of_monomials_self _)).symm

end Topology

section Uniformity

variable [UniformSpace R]

variable (σ R) in
/-- The componentwise uniformity on MvPowerSeries -/
scoped instance : UniformSpace (MvPowerSeries σ R) :=
Pi.uniformSpace fun _ : σ →₀ ℕ => R

variable (R)

/-- Coefficients are uniformly continuous -/
theorem uniformContinuous_coeff [Semiring R] (d : σ →₀ ℕ) :
UniformContinuous fun f : MvPowerSeries σ R => coeff R d f :=
uniformContinuous_pi.mp uniformContinuous_id d

variable [Ring R]

variable (σ)

/-- The `UniformAddGroup` structure on `MvPowerSeries` of a `UniformAddGroup` -/
@[scoped instance]
theorem instUniformAddGroup [UniformAddGroup R] :
UniformAddGroup (MvPowerSeries σ R) := Pi.instUniformAddGroup

/-- Completeness of the uniform structure on MvPowerSeries -/
@[scoped instance]
theorem instCompleteSpace [CompleteSpace R] :
CompleteSpace (MvPowerSeries σ R) := Pi.complete _

end Uniformity

end WithPiTopology

end MvPowerSeries
Loading
Loading