Skip to content

Commit

Permalink
feat: add Normed instances for SeparationQuotient (#17452)
Browse files Browse the repository at this point in the history
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.
  • Loading branch information
eric-wieser committed Oct 8, 2024
1 parent 91bc9bf commit 9341c08
Show file tree
Hide file tree
Showing 6 changed files with 98 additions and 2 deletions.
32 changes: 30 additions & 2 deletions Mathlib/Analysis/Normed/Field/Lemmas.lean
Original file line number Diff line number Diff line change
Expand Up @@ -148,7 +148,7 @@ instance Pi.normedCommutativeRing {π : ι → Type*} [Fintype ι] [∀ i, Norme
end NormedCommRing

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

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

namespace SeparationQuotient

instance [NonUnitalSeminormedRing α] : NonUnitalNormedRing (SeparationQuotient α) where
__ : NonUnitalRing (SeparationQuotient α) := inferInstance
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
norm_mul := Quotient.ind₂ norm_mul_le

instance [NonUnitalSeminormedCommRing α] : NonUnitalNormedCommRing (SeparationQuotient α) where
__ : NonUnitalCommRing (SeparationQuotient α) := inferInstance
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
norm_mul := Quotient.ind₂ norm_mul_le

instance [SeminormedRing α] : NormedRing (SeparationQuotient α) where
__ : Ring (SeparationQuotient α) := inferInstance
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
norm_mul := Quotient.ind₂ norm_mul_le

instance [SeminormedCommRing α] : NormedCommRing (SeparationQuotient α) where
__ : CommRing (SeparationQuotient α) := inferInstance
__ : NormedAddCommGroup (SeparationQuotient α) := inferInstance
norm_mul := Quotient.ind₂ norm_mul_le

instance [SeminormedAddCommGroup α] [One α] [NormOneClass α] :
NormOneClass (SeparationQuotient α) where
norm_one := norm_one (α := α)

end SeparationQuotient

section NormedDivisionRing

variable [NormedDivisionRing α] {a : α}
Expand Down
10 changes: 10 additions & 0 deletions Mathlib/Analysis/Normed/Group/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -770,6 +770,16 @@ theorem continuous_norm' : Continuous fun a : E => ‖a‖ := by
theorem continuous_nnnorm' : Continuous fun a : E => ‖a‖₊ :=
continuous_norm'.subtype_mk _

set_option linter.docPrime false in
@[to_additive Inseparable.norm_eq_norm]
theorem Inseparable.norm_eq_norm' {u v : E} (h : Inseparable u v) : ‖u‖ = ‖v‖ :=
h.map continuous_norm' |>.eq

set_option linter.docPrime false in
@[to_additive Inseparable.nnnorm_eq_nnnorm]
theorem Inseparable.nnnorm_eq_nnnorm' {u v : E} (h : Inseparable u v) : ‖u‖₊ = ‖v‖₊ :=
h.map continuous_nnnorm' |>.eq

@[to_additive]
theorem mem_closure_one_iff_norm {x : E} : x ∈ closure ({1} : Set E) ↔ ‖x‖ = 0 := by
rw [← closedBall_zero', mem_closedBall_one_iff, (norm_nonneg' x).le_iff_eq]
Expand Down
23 changes: 23 additions & 0 deletions Mathlib/Analysis/Normed/Group/Uniform.lean
Original file line number Diff line number Diff line change
Expand Up @@ -365,6 +365,29 @@ instance (priority := 100) SeminormedCommGroup.to_uniformGroup : UniformGroup E
instance (priority := 100) SeminormedCommGroup.toTopologicalGroup : TopologicalGroup E :=
inferInstance

/-! ### SeparationQuotient -/

namespace SeparationQuotient

@[to_additive instNorm]
instance instMulNorm : Norm (SeparationQuotient E) where
norm := lift Norm.norm fun _ _ h => h.norm_eq_norm'

set_option linter.docPrime false in
@[to_additive (attr := simp) norm_mk]
theorem norm_mk' (p : E) : ‖mk p‖ = ‖p‖ := rfl

@[to_additive]
instance : NormedCommGroup (SeparationQuotient E) where
__ : CommGroup (SeparationQuotient E) := instCommGroup
dist_eq := Quotient.ind₂ dist_eq_norm_div

set_option linter.docPrime false in
@[to_additive (attr := simp) nnnorm_mk]
theorem nnnorm_mk' (p : E) : ‖mk p‖₊ = ‖p‖₊ := rfl

end SeparationQuotient

@[to_additive]
theorem cauchySeq_prod_of_eventually_eq {u v : ℕ → E} {N : ℕ} (huv : ∀ n ≥ N, u n = v n)
(hv : CauchySeq fun n => ∏ k ∈ range (n + 1), v k) :
Expand Down
7 changes: 7 additions & 0 deletions Mathlib/Analysis/Normed/Module/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -119,6 +119,9 @@ instance Pi.normedSpace {ι : Type*} {E : ι → Type*} [Fintype ι] [∀ i, Sem
NNReal.mul_finset_sup]
exact Finset.sup_mono_fun fun _ _ => norm_smul_le a _

instance SeparationQuotient.instNormedSpace : NormedSpace 𝕜 (SeparationQuotient E) where
norm_smul_le := norm_smul_le

instance MulOpposite.instNormedSpace : NormedSpace 𝕜 Eᵐᵒᵖ where
norm_smul_le _ x := norm_smul_le _ x.unop

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

variable [SeminormedRing E] [NormedAlgebra 𝕜 E]

instance SeparationQuotient.instNormedAlgebra : NormedAlgebra 𝕜 (SeparationQuotient E) where
__ : NormedSpace 𝕜 (SeparationQuotient E) := inferInstance
__ : Algebra 𝕜 (SeparationQuotient E) := inferInstance

instance MulOpposite.instNormedAlgebra {E : Type*} [SeminormedRing E] [NormedAlgebra 𝕜 E] :
NormedAlgebra 𝕜 Eᵐᵒᵖ where
__ := instAlgebra
Expand Down
21 changes: 21 additions & 0 deletions Mathlib/Topology/Algebra/SeparationQuotient.lean
Original file line number Diff line number Diff line change
Expand Up @@ -333,6 +333,12 @@ instance instCommRing [CommRing R] [TopologicalRing R] :
surjective_mk.commRing mk mk_zero mk_one mk_add mk_mul mk_neg mk_sub mk_smul mk_smul mk_pow
mk_natCast mk_intCast

/-- `SeparationQuotient.mk` as a `RingHom`. -/
@[simps]
def mkRingHom [NonAssocSemiring R] [TopologicalSemiring R] : R →+* SeparationQuotient R where
toFun := mk
map_one' := mk_one; map_zero' := mk_zero; map_add' := mk_add; map_mul' := mk_mul

end Ring

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

end Module

section Algebra
variable {R A : Type*} [CommSemiring R] [Semiring A] [Algebra R A]
[TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A]

instance instAlgebra : Algebra R (SeparationQuotient A) where
toRingHom := mkRingHom.comp (algebraMap R A)
commutes' r := Quotient.ind fun a => congrArg _ <| Algebra.commutes r a
smul_def' r := Quotient.ind fun a => congrArg _ <| Algebra.smul_def r a

@[simp]
theorem mk_algebraMap (r : R) : mk (algebraMap R A r) = algebraMap R (SeparationQuotient A) r :=
rfl

end Algebra

section VectorSpace

variable (K E : Type*) [DivisionRing K] [AddCommGroup E] [Module K E]
Expand Down
7 changes: 7 additions & 0 deletions Mathlib/Topology/MetricSpace/Algebra.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Heather Macbeth
import Mathlib.Topology.Algebra.MulAction
import Mathlib.Topology.Algebra.UniformMulAction
import Mathlib.Topology.MetricSpace.Lipschitz
import Mathlib.Topology.Algebra.SeparationQuotient

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

instance {α β : Type*}
[PseudoMetricSpace α] [PseudoMetricSpace β] [Zero α] [Zero β] [SMul α β] [BoundedSMul α β] :
BoundedSMul α (SeparationQuotient β) where
dist_smul_pair' _ := Quotient.ind₂ <| dist_smul_pair _
dist_pair_smul' _ _ := Quotient.ind <| dist_pair_smul _ _

-- We don't have the `SMul α γ → SMul β δ → SMul (α × β) (γ × δ)` instance, but if we did, then
-- `BoundedSMul α γ → BoundedSMul β δ → BoundedSMul (α × β) (γ × δ)` would hold

0 comments on commit 9341c08

Please sign in to comment.