Skip to content

Commit 9341c08

Browse files
committed
feat: add Normed instances for SeparationQuotient (#17452)
We already have the corresponding metric spaces structures in `MetricSpace.Basic`. This adds `Algebra`, `NormOneClass`, `Normed(Add)CommGroup`, `(Nonunital)Normed(Comm)Ring`, `NormedSpace`, and `NormedAlgebra` instances. The slightly heavy imports here can be blamed on `Topology.Algebra.SeparationQuotient` importing `LinearAlgebra.Basis.VectorSpace` for the final 50-100 lines of the file. Moving this to a new file would cut the chain, as would generalizing it to work in free modules (continuing from #17560), assuming it actually holds there.
1 parent 91bc9bf commit 9341c08

File tree

6 files changed

+98
-2
lines changed

6 files changed

+98
-2
lines changed

Mathlib/Analysis/Normed/Field/Lemmas.lean

Lines changed: 30 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -148,7 +148,7 @@ instance Pi.normedCommutativeRing {π : ι → Type*} [Fintype ι] [∀ i, Norme
148148
end NormedCommRing
149149

150150
-- see Note [lower instance priority]
151-
instance (priority := 100) semi_normed_ring_top_monoid [NonUnitalSeminormedRing α] :
151+
instance (priority := 100) NonUnitalSeminormedRing.toContinuousMul [NonUnitalSeminormedRing α] :
152152
ContinuousMul α :=
153153
⟨continuous_iff_continuousAt.2 fun x =>
154154
tendsto_iff_norm_sub_tendsto_zero.2 <| by
@@ -174,9 +174,37 @@ instance (priority := 100) semi_normed_ring_top_monoid [NonUnitalSeminormedRing
174174

175175
-- see Note [lower instance priority]
176176
/-- A seminormed ring is a topological ring. -/
177-
instance (priority := 100) semi_normed_top_ring [NonUnitalSeminormedRing α] :
177+
instance (priority := 100) NonUnitalSeminormedRing.toTopologicalRing [NonUnitalSeminormedRing α] :
178178
TopologicalRing α where
179179

180+
namespace SeparationQuotient
181+
182+
instance [NonUnitalSeminormedRing α] : NonUnitalNormedRing (SeparationQuotient α) where
183+
__ : NonUnitalRing (SeparationQuotient α) := inferInstance
184+
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
185+
norm_mul := Quotient.ind₂ norm_mul_le
186+
187+
instance [NonUnitalSeminormedCommRing α] : NonUnitalNormedCommRing (SeparationQuotient α) where
188+
__ : NonUnitalCommRing (SeparationQuotient α) := inferInstance
189+
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
190+
norm_mul := Quotient.ind₂ norm_mul_le
191+
192+
instance [SeminormedRing α] : NormedRing (SeparationQuotient α) where
193+
__ : Ring (SeparationQuotient α) := inferInstance
194+
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
195+
norm_mul := Quotient.ind₂ norm_mul_le
196+
197+
instance [SeminormedCommRing α] : NormedCommRing (SeparationQuotient α) where
198+
__ : CommRing (SeparationQuotient α) := inferInstance
199+
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
200+
norm_mul := Quotient.ind₂ norm_mul_le
201+
202+
instance [SeminormedAddCommGroup α] [One α] [NormOneClass α] :
203+
NormOneClass (SeparationQuotient α) where
204+
norm_one := norm_one (α := α)
205+
206+
end SeparationQuotient
207+
180208
section NormedDivisionRing
181209

182210
variable [NormedDivisionRing α] {a : α}

Mathlib/Analysis/Normed/Group/Basic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -770,6 +770,16 @@ theorem continuous_norm' : Continuous fun a : E => ‖a‖ := by
770770
theorem continuous_nnnorm' : Continuous fun a : E => ‖a‖₊ :=
771771
continuous_norm'.subtype_mk _
772772

773+
set_option linter.docPrime false in
774+
@[to_additive Inseparable.norm_eq_norm]
775+
theorem Inseparable.norm_eq_norm' {u v : E} (h : Inseparable u v) : ‖u‖ = ‖v‖ :=
776+
h.map continuous_norm' |>.eq
777+
778+
set_option linter.docPrime false in
779+
@[to_additive Inseparable.nnnorm_eq_nnnorm]
780+
theorem Inseparable.nnnorm_eq_nnnorm' {u v : E} (h : Inseparable u v) : ‖u‖₊ = ‖v‖₊ :=
781+
h.map continuous_nnnorm' |>.eq
782+
773783
@[to_additive]
774784
theorem mem_closure_one_iff_norm {x : E} : x ∈ closure ({1} : Set E) ↔ ‖x‖ = 0 := by
775785
rw [← closedBall_zero', mem_closedBall_one_iff, (norm_nonneg' x).le_iff_eq]

Mathlib/Analysis/Normed/Group/Uniform.lean

Lines changed: 23 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -365,6 +365,29 @@ instance (priority := 100) SeminormedCommGroup.to_uniformGroup : UniformGroup E
365365
instance (priority := 100) SeminormedCommGroup.toTopologicalGroup : TopologicalGroup E :=
366366
inferInstance
367367

368+
/-! ### SeparationQuotient -/
369+
370+
namespace SeparationQuotient
371+
372+
@[to_additive instNorm]
373+
instance instMulNorm : Norm (SeparationQuotient E) where
374+
norm := lift Norm.norm fun _ _ h => h.norm_eq_norm'
375+
376+
set_option linter.docPrime false in
377+
@[to_additive (attr := simp) norm_mk]
378+
theorem norm_mk' (p : E) : ‖mk p‖ = ‖p‖ := rfl
379+
380+
@[to_additive]
381+
instance : NormedCommGroup (SeparationQuotient E) where
382+
__ : CommGroup (SeparationQuotient E) := instCommGroup
383+
dist_eq := Quotient.ind₂ dist_eq_norm_div
384+
385+
set_option linter.docPrime false in
386+
@[to_additive (attr := simp) nnnorm_mk]
387+
theorem nnnorm_mk' (p : E) : ‖mk p‖₊ = ‖p‖₊ := rfl
388+
389+
end SeparationQuotient
390+
368391
@[to_additive]
369392
theorem cauchySeq_prod_of_eventually_eq {u v : ℕ → E} {N : ℕ} (huv : ∀ n ≥ N, u n = v n)
370393
(hv : CauchySeq fun n => ∏ k ∈ range (n + 1), v k) :

Mathlib/Analysis/Normed/Module/Basic.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -119,6 +119,9 @@ instance Pi.normedSpace {ι : Type*} {E : ι → Type*} [Fintype ι] [∀ i, Sem
119119
NNReal.mul_finset_sup]
120120
exact Finset.sup_mono_fun fun _ _ => norm_smul_le a _
121121

122+
instance SeparationQuotient.instNormedSpace : NormedSpace 𝕜 (SeparationQuotient E) where
123+
norm_smul_le := norm_smul_le
124+
122125
instance MulOpposite.instNormedSpace : NormedSpace 𝕜 Eᵐᵒᵖ where
123126
norm_smul_le _ x := norm_smul_le _ x.unop
124127

@@ -348,6 +351,10 @@ instance Pi.normedAlgebra {ι : Type*} {E : ι → Type*} [Fintype ι] [∀ i, S
348351

349352
variable [SeminormedRing E] [NormedAlgebra 𝕜 E]
350353

354+
instance SeparationQuotient.instNormedAlgebra : NormedAlgebra 𝕜 (SeparationQuotient E) where
355+
__ : NormedSpace 𝕜 (SeparationQuotient E) := inferInstance
356+
__ : Algebra 𝕜 (SeparationQuotient E) := inferInstance
357+
351358
instance MulOpposite.instNormedAlgebra {E : Type*} [SeminormedRing E] [NormedAlgebra 𝕜 E] :
352359
NormedAlgebra 𝕜 Eᵐᵒᵖ where
353360
__ := instAlgebra

Mathlib/Topology/Algebra/SeparationQuotient.lean

Lines changed: 21 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -333,6 +333,12 @@ instance instCommRing [CommRing R] [TopologicalRing R] :
333333
surjective_mk.commRing mk mk_zero mk_one mk_add mk_mul mk_neg mk_sub mk_smul mk_smul mk_pow
334334
mk_natCast mk_intCast
335335

336+
/-- `SeparationQuotient.mk` as a `RingHom`. -/
337+
@[simps]
338+
def mkRingHom [NonAssocSemiring R] [TopologicalSemiring R] : R →+* SeparationQuotient R where
339+
toFun := mk
340+
map_one' := mk_one; map_zero' := mk_zero; map_add' := mk_add; map_mul' := mk_mul
341+
336342
end Ring
337343

338344
section DistribSMul
@@ -375,6 +381,21 @@ def mkCLM : M →L[R] SeparationQuotient M where
375381

376382
end Module
377383

384+
section Algebra
385+
variable {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A]
386+
[TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A]
387+
388+
instance instAlgebra : Algebra R (SeparationQuotient A) where
389+
toRingHom := mkRingHom.comp (algebraMap R A)
390+
commutes' r := Quotient.ind fun a => congrArg _ <| Algebra.commutes r a
391+
smul_def' r := Quotient.ind fun a => congrArg _ <| Algebra.smul_def r a
392+
393+
@[simp]
394+
theorem mk_algebraMap (r : R) : mk (algebraMap R A r) = algebraMap R (SeparationQuotient A) r :=
395+
rfl
396+
397+
end Algebra
398+
378399
section VectorSpace
379400

380401
variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E]

Mathlib/Topology/MetricSpace/Algebra.lean

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Heather Macbeth
66
import Mathlib.Topology.Algebra.MulAction
77
import Mathlib.Topology.Algebra.UniformMulAction
88
import Mathlib.Topology.MetricSpace.Lipschitz
9+
import Mathlib.Topology.Algebra.SeparationQuotient
910

1011
/-!
1112
# Compatibility of algebraic operations with metric space structures
@@ -213,5 +214,11 @@ instance Prod.instBoundedSMul {α β γ : Type*} [PseudoMetricSpace α] [PseudoM
213214
max_le ((dist_pair_smul _ _ _).trans <| mul_le_mul_of_nonneg_left (le_max_left _ _) dist_nonneg)
214215
((dist_pair_smul _ _ _).trans <| mul_le_mul_of_nonneg_left (le_max_right _ _) dist_nonneg)
215216

217+
instance {α β : Type*}
218+
[PseudoMetricSpace α] [PseudoMetricSpace β] [Zero α] [Zero β] [SMul α β] [BoundedSMul α β] :
219+
BoundedSMul α (SeparationQuotient β) where
220+
dist_smul_pair' _ := Quotient.ind₂ <| dist_smul_pair _
221+
dist_pair_smul' _ _ := Quotient.ind <| dist_pair_smul _ _
222+
216223
-- We don't have the `SMul α γ → SMul β δ → SMul (α × β) (γ × δ)` instance, but if we did, then
217224
-- `BoundedSMul α γ → BoundedSMul β δ → BoundedSMul (α × β) (γ × δ)` would hold

0 commit comments

Comments
 (0)