@@ -12,6 +12,7 @@ import algebra.algebra.basic
1212import linear_algebra.projection
1313import linear_algebra.pi
1414import 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
167168end submodule
168169
169170section 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
175179lemma 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 :=
239243is_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+
241263end closure
242264
243265/-- Continuous linear maps between modules. We only put the type classes that are necessary for the
0 commit comments