Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 531db2e

Browse files
committed
feat(topology/algebra/uniform_group): add instance topological_group_is_uniform_of_compact_space (#16027)
Also update doc string for `topological_group.to_uniform_space`. cc @ADedecker
1 parent 07d11ba commit 531db2e

File tree

9 files changed

+35
-15
lines changed

9 files changed

+35
-15
lines changed

src/analysis/normed_space/compact_operator.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -323,7 +323,7 @@ variables {𝕜₁ 𝕜₂ : Type*} [nontrivially_normed_field 𝕜₁] [nontriv
323323
(hf : is_compact_operator f) : continuous f :=
324324
begin
325325
letI : uniform_space M₂ := topological_add_group.to_uniform_space _,
326-
haveI : uniform_add_group M₂ := topological_add_group_is_uniform,
326+
haveI : uniform_add_group M₂ := topological_add_comm_group_is_uniform,
327327
-- Since `f` is linear, we only need to show that it is continuous at zero.
328328
-- Let `U` be a neighborhood of `0` in `M₂`.
329329
refine continuous_of_continuous_at_zero f (λ U hU, _),

src/topology/algebra/group.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -320,7 +320,7 @@ continuous.
320320
321321
When you declare an instance that does not already have a `uniform_space` instance,
322322
you should also provide an instance of `uniform_space` and `uniform_group` using
323-
`topological_group.to_uniform_space` and `topological_group_is_uniform`. -/
323+
`topological_group.to_uniform_space` and `topological_comm_group_is_uniform`. -/
324324
@[to_additive]
325325
class topological_group (G : Type*) [topological_space G] [group G]
326326
extends has_continuous_mul G, has_continuous_inv G : Prop

src/topology/algebra/infinite_sum.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1322,7 +1322,7 @@ lemma summable.vanishing (hf : summable f) ⦃e : set G⦄ (he : e ∈ 𝓝 (0 :
13221322
∃ s : finset α, ∀ t, disjoint t s → ∑ k in t, f k ∈ e :=
13231323
begin
13241324
letI : uniform_space G := topological_add_group.to_uniform_space G,
1325-
letI : uniform_add_group G := topological_add_group_is_uniform,
1325+
letI : uniform_add_group G := topological_add_comm_group_is_uniform,
13261326
rcases hf with ⟨y, hy⟩,
13271327
exact cauchy_seq_finset_iff_vanishing.1 hy.cauchy_seq e he
13281328
end

src/topology/algebra/module/finite_dimension.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -205,7 +205,7 @@ private lemma continuous_equiv_fun_basis_aux [ht2 : t2_space E] {ι : Type v} [f
205205
(ξ : basis ι 𝕜 E) : continuous ξ.equiv_fun :=
206206
begin
207207
letI : uniform_space E := topological_add_group.to_uniform_space E,
208-
letI : uniform_add_group E := topological_add_group_is_uniform,
208+
letI : uniform_add_group E := topological_add_comm_group_is_uniform,
209209
letI : separated_space E := separated_iff_t2.mpr ht2,
210210
unfreezingI { induction hn : fintype.card ι with n IH generalizing ι E },
211211
{ rw fintype.card_eq_zero_iff at hn,

src/topology/algebra/nonarchimedean/adic_topology.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -219,7 +219,7 @@ variables (R) [with_ideal R]
219219
topological_add_group.to_uniform_space R
220220

221221
@[priority 100] instance : uniform_add_group R :=
222-
topological_add_group_is_uniform
222+
topological_add_comm_group_is_uniform
223223

224224
/-- The adic topology on a `R` module coming from the ideal `with_ideal.I`.
225225
This cannot be an instance because `R` cannot be inferred from `M`. -/

src/topology/algebra/uniform_filter_basis.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -32,7 +32,7 @@ protected def uniform_space : uniform_space G :=
3232
/-- The uniform space structure associated to an abelian group filter basis via the associated
3333
topological abelian group structure is compatible with its group structure. -/
3434
protected lemma uniform_add_group : @uniform_add_group G B.uniform_space _:=
35-
@topological_add_group_is_uniform G _ B.topology B.is_topological_add_group
35+
@topological_add_comm_group_is_uniform G _ B.topology B.is_topological_add_group
3636

3737
lemma cauchy_iff {F : filter G} :
3838
@cauchy G B.uniform_space F ↔ F.ne_bot ∧ ∀ U ∈ B, ∃ M ∈ F, ∀ x y ∈ M, y - x ∈ U :=

src/topology/algebra/uniform_group.lean

Lines changed: 26 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl
66
import topology.uniform_space.uniform_convergence
77
import topology.uniform_space.uniform_embedding
88
import topology.uniform_space.complete_separated
9+
import topology.uniform_space.compact_separated
910
import topology.algebra.group
1011
import tactic.abel
1112

@@ -22,8 +23,8 @@ group naturally induces a uniform structure.
2223
2324
## Main results
2425
25-
* `topological_add_group.to_uniform_space` and `topological_add_group_is_uniform` can be used to
26-
construct a canonical uniformity for a topological add group.
26+
* `topological_add_group.to_uniform_space` and `topological_add_comm_group_is_uniform` can be used
27+
to construct a canonical uniformity for a topological add group.
2728
2829
* extension of ℤ-bilinear maps to complete groups (useful for ring completions)
2930
-/
@@ -347,8 +348,19 @@ section topological_group
347348
open filter
348349
variables (G : Type*) [group G] [topological_space G] [topological_group G]
349350

350-
/-- The right uniformity on a topological group. -/
351-
@[to_additive "The right uniformity on a topological additive group"]
351+
/-- The right uniformity on a topological group (as opposed to the left uniformity).
352+
353+
Warning: in general the right and left uniformities do not coincide and so one does not obtain a
354+
`uniform_group` structure. Two important special cases where they _do_ coincide are for
355+
commutative groups (see `topological_comm_group_is_uniform`) and for compact Hausdorff groups (see
356+
`topological_group_is_uniform_of_compact_space`). -/
357+
@[to_additive "The right uniformity on a topological additive group (as opposed to the left
358+
uniformity).
359+
360+
Warning: in general the right and left uniformities do not coincide and so one does not obtain a
361+
`uniform_add_group` structure. Two important special cases where they _do_ coincide are for
362+
commutative additive groups (see `topological_add_comm_group_is_uniform`) and for compact Hausdorff
363+
additive groups (see `topological_add_comm_group_is_uniform_of_compact_space`)."]
352364
def topological_group.to_uniform_space : uniform_space G :=
353365
{ uniformity := comap (λp:G×G, p.2 / p.1) (𝓝 1),
354366
refl :=
@@ -400,6 +412,14 @@ local attribute [instance] topological_group.to_uniform_space
400412
@[to_additive] lemma uniformity_eq_comap_nhds_one' :
401413
𝓤 G = comap (λp:G×G, p.2 / p.1) (𝓝 (1 : G)) := rfl
402414

415+
@[to_additive] lemma topological_group_is_uniform_of_compact_space
416+
[compact_space G] [t2_space G] : uniform_group G :=
417+
begin
418+
haveI : separated_space G := separated_iff_t2.mpr (by apply_instance),
419+
apply compact_space.uniform_continuous_of_continuous,
420+
exact continuous_div',
421+
end
422+
403423
variables {G}
404424

405425
@[to_additive] lemma topological_group.tendsto_uniformly_iff
@@ -442,7 +462,7 @@ section
442462
local attribute [instance] topological_group.to_uniform_space
443463

444464
variable {G}
445-
@[to_additive] lemma topological_group_is_uniform : uniform_group G :=
465+
@[to_additive] lemma topological_comm_group_is_uniform : uniform_group G :=
446466
have tendsto
447467
((λp:(G×G), p.1 / p.2) ∘ (λp:(G×G)×(G×G), (p.1.2 / p.1.1, p.2.2 / p.2.1)))
448468
(comap (λp:(G×G)×(G×G), (p.1.2 / p.1.1, p.2.2 / p.2.1)) ((𝓝 1).prod (𝓝 1)))
@@ -460,7 +480,7 @@ open set
460480
@[to_additive] lemma topological_group.t2_space_iff_one_closed :
461481
t2_space G ↔ is_closed ({1} : set G) :=
462482
begin
463-
haveI : uniform_group G := topological_group_is_uniform,
483+
haveI : uniform_group G := topological_comm_group_is_uniform,
464484
rw [← separated_iff_t2, separated_space_iff, ← closure_eq_iff_is_closed],
465485
split; intro h,
466486
{ apply subset.antisymm,

src/topology/algebra/valuation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -102,7 +102,7 @@ structure. -/
102102
def mk' (v : valuation R Γ₀) : valued R Γ₀ :=
103103
{ v := v,
104104
to_uniform_space := @topological_add_group.to_uniform_space R _ v.subgroups_basis.topology _,
105-
to_uniform_add_group := @topological_add_group_is_uniform _ _ v.subgroups_basis.topology _,
105+
to_uniform_add_group := @topological_add_comm_group_is_uniform _ _ v.subgroups_basis.topology _,
106106
is_topological_valuation :=
107107
begin
108108
letI := @topological_add_group.to_uniform_space R _ v.subgroups_basis.topology _,

src/topology/continuous_function/algebra.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -287,7 +287,7 @@ coe_injective.comm_group _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow
287287
[topological_space β] [comm_group β] [topological_group β] : topological_group C(α, β) :=
288288
{ continuous_mul := by
289289
{ letI : uniform_space β := topological_group.to_uniform_space β,
290-
have : uniform_group β := topological_group_is_uniform,
290+
have : uniform_group β := topological_comm_group_is_uniform,
291291
rw continuous_iff_continuous_at,
292292
rintros ⟨f, g⟩,
293293
rw [continuous_at, tendsto_iff_forall_compact_tendsto_uniformly_on, nhds_prod_eq],
@@ -296,7 +296,7 @@ coe_injective.comm_group _ coe_one coe_mul coe_inv coe_div coe_pow coe_zpow
296296
(tendsto_iff_forall_compact_tendsto_uniformly_on.mp filter.tendsto_id K hK)), },
297297
continuous_inv := by
298298
{ letI : uniform_space β := topological_group.to_uniform_space β,
299-
have : uniform_group β := topological_group_is_uniform,
299+
have : uniform_group β := topological_comm_group_is_uniform,
300300
rw continuous_iff_continuous_at,
301301
intro f,
302302
rw [continuous_at, tendsto_iff_forall_compact_tendsto_uniformly_on],

0 commit comments

Comments
 (0)