Skip to content

Commit

Permalink
feat: add some consequences of Tychonoff's theorem (#10161)
Browse files Browse the repository at this point in the history
Preliminaries for #6844 
The way I prove Ascoli's theorem is by using equicontinuity to reduce it to the case of the product topology, where one can use Tychonoff, hence these variations.
  • Loading branch information
ADedecker committed Feb 2, 2024
1 parent 41ab332 commit b75513b
Show file tree
Hide file tree
Showing 2 changed files with 28 additions and 1 deletion.
17 changes: 16 additions & 1 deletion Mathlib/Topology/Compactness/Compact.lean
Expand Up @@ -25,8 +25,8 @@ We define the following properties for sets in a topological space:
* `isCompact_univ_pi`: **Tychonov's theorem** - an arbitrary product of compact sets
is compact.
-/
open Set Filter Topology TopologicalSpace Classical

open Set Filter Topology TopologicalSpace Classical Function

universe u v

Expand Down Expand Up @@ -1084,6 +1084,21 @@ instance Function.compactSpace [CompactSpace Y] : CompactSpace (ι → Y) :=
Pi.compactSpace
#align function.compact_space Function.compactSpace

lemma Pi.isCompact_iff_of_isClosed {s : Set (Π i, X i)} (hs : IsClosed s) :
IsCompact s ↔ ∀ i, IsCompact (eval i '' s) := by
constructor <;> intro H
· exact fun i ↦ H.image <| continuous_apply i
· exact IsCompact.of_isClosed_subset (isCompact_univ_pi H) hs (subset_pi_eval_image univ s)

protected lemma Pi.exists_compact_superset_iff {s : Set (Π i, X i)} :
(∃ K, IsCompact K ∧ s ⊆ K) ↔ ∀ i, ∃ Ki, IsCompact Ki ∧ s ⊆ eval i ⁻¹' Ki := by
constructor
· intro ⟨K, hK, hsK⟩ i
exact ⟨eval i '' K, hK.image <| continuous_apply i, hsK.trans <| K.subset_preimage_image _⟩
· intro H
choose K hK hsK using H
exact ⟨pi univ K, isCompact_univ_pi hK, fun _ hx i _ ↦ hsK i hx⟩

/-- **Tychonoff's theorem** formulated in terms of filters: `Filter.cocompact` on an indexed product
type `Π d, X d` the `Filter.coprodᵢ` of filters `Filter.cocompact` on `X d`. -/
theorem Filter.coprodᵢ_cocompact {X : ι → Type*} [∀ d, TopologicalSpace (X d)] :
Expand Down
12 changes: 12 additions & 0 deletions Mathlib/Topology/Separation.lean
Expand Up @@ -1675,6 +1675,18 @@ theorem IsCompact.preimage_continuous [CompactSpace X] [T2Space Y] {f : X → Y}
(hs : IsCompact s) (hf : Continuous f) : IsCompact (f ⁻¹' s) :=
(hs.isClosed.preimage hf).isCompact

lemma Pi.isCompact_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)]
[∀ i, T2Space (π i)] {s : Set (Π i, π i)} :
IsCompact s ↔ IsClosed s ∧ ∀ i, IsCompact (eval i '' s):= by
constructor <;> intro H
· exact ⟨H.isClosed, fun i ↦ H.image <| continuous_apply i⟩
· exact IsCompact.of_isClosed_subset (isCompact_univ_pi H.2) H.1 (subset_pi_eval_image univ s)

lemma Pi.isCompact_closure_iff {ι : Type*} {π : ι → Type*} [∀ i, TopologicalSpace (π i)]
[∀ i, T2Space (π i)] {s : Set (Π i, π i)} :
IsCompact (closure s) ↔ ∀ i, IsCompact (closure <| eval i '' s) := by
simp_rw [← exists_isCompact_superset_iff, Pi.exists_compact_superset_iff, image_subset_iff]

/-- If `V : ι → Set X` is a decreasing family of compact sets then any neighborhood of
`⋂ i, V i` contains some `V i`. This is a version of `exists_subset_nhds_of_isCompact'` where we
don't need to assume each `V i` closed because it follows from compactness since `X` is
Expand Down

0 comments on commit b75513b

Please sign in to comment.