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

Commit e05f130

Browse files
committed
feat(topology/algebra/module/basic): a hyperplane is either closed or dense (#16782)
1 parent 4a4d740 commit e05f130

File tree

2 files changed

+36
-1
lines changed

2 files changed

+36
-1
lines changed

src/topology/algebra/module/basic.lean

Lines changed: 23 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ import algebra.algebra.basic
1212
import linear_algebra.projection
1313
import linear_algebra.pi
1414
import linear_algebra.determinant
15+
import ring_theory.simple_module
1516

1617
/-!
1718
# Theory of topological modules and continuous linear maps.
@@ -167,10 +168,13 @@ S.to_add_subgroup.topological_add_group
167168
end submodule
168169

169170
section closure
170-
variables {R : Type u} {M : Type v}
171+
variables {R R' : Type u} {M M' : Type v}
171172
[semiring R] [topological_space R]
173+
[ring R'] [topological_space R']
172174
[topological_space M] [add_comm_monoid M]
175+
[topological_space M'] [add_comm_group M']
173176
[module R M] [has_continuous_smul R M]
177+
[module R' M'] [has_continuous_smul R' M']
174178

175179
lemma submodule.closure_smul_self_subset (s : submodule R M) :
176180
(λ p : R × M, p.1 • p.2) '' (set.univ ×ˢ closure s) ⊆ closure s :=
@@ -238,6 +242,24 @@ instance {M' : Type*} [add_comm_monoid M'] [module R M'] [uniform_space M']
238242
complete_space U.topological_closure :=
239243
is_closed_closure.complete_space_coe
240244

245+
/-- A maximal proper subspace of a topological module (i.e a `submodule` satisfying `is_coatom`)
246+
is either closed or dense. -/
247+
lemma submodule.is_closed_or_dense_of_is_coatom (s : submodule R M) (hs : is_coatom s) :
248+
is_closed (s : set M) ∨ dense (s : set M) :=
249+
(hs.le_iff.mp s.submodule_topological_closure).swap.imp (is_closed_of_closure_subset ∘ eq.le)
250+
submodule.dense_iff_topological_closure_eq_top.mpr
251+
252+
lemma linear_map.is_closed_or_dense_ker [has_continuous_add M'] [is_simple_module R' R']
253+
(l : M' →ₗ[R'] R') :
254+
is_closed (l.ker : set M') ∨ dense (l.ker : set M') :=
255+
begin
256+
rcases l.surjective_or_eq_zero with (hl|rfl),
257+
{ refine l.ker.is_closed_or_dense_of_is_coatom (linear_map.is_coatom_ker_of_surjective hl) },
258+
{ rw linear_map.ker_zero,
259+
left,
260+
exact is_closed_univ },
261+
end
262+
241263
end closure
242264

243265
/-- Continuous linear maps between modules. We only put the type classes that are necessary for the

src/topology/algebra/module/finite_dimension.lean

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -197,6 +197,19 @@ lemma linear_map.continuous_iff_is_closed_ker (l : E →ₗ[𝕜] 𝕜) :
197197
continuous l ↔ is_closed (l.ker : set E) :=
198198
⟨λ h, is_closed_singleton.preimage h, l.continuous_of_is_closed_ker⟩
199199

200+
/-- Over a nontrivially normed field, any linear form which is nonzero on a nonempty open set is
201+
automatically continuous. -/
202+
lemma linear_map.continuous_of_nonzero_on_open (l : E →ₗ[𝕜] 𝕜) (s : set E) (hs₁ : is_open s)
203+
(hs₂ : s.nonempty) (hs₃ : ∀ x ∈ s, l x ≠ 0) : continuous l :=
204+
begin
205+
refine l.continuous_of_is_closed_ker (l.is_closed_or_dense_ker.resolve_right $ λ hl, _),
206+
rcases hs₂ with ⟨x, hx⟩,
207+
have : x ∈ interior (l.ker : set E)ᶜ,
208+
{ rw mem_interior_iff_mem_nhds,
209+
exact mem_of_superset (hs₁.mem_nhds hx) hs₃ },
210+
rwa hl.interior_compl at this
211+
end
212+
200213
variables [complete_space 𝕜]
201214

202215
/-- This version imposes `ι` and `E` to live in the same universe, so you should instead use

0 commit comments

Comments
 (0)