Skip to content

Commit

Permalink
feat: add Continuous.image_connectedComponentIn_subset (#9983)
Browse files Browse the repository at this point in the history
and a version for homeomorphisms. From sphere-eversion; I'm just submitting things upstream.
  • Loading branch information
grunweg committed Jan 25, 2024
1 parent e7170d3 commit 47a9066
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 0 deletions.
12 changes: 12 additions & 0 deletions Mathlib/Topology/Connected/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -707,11 +707,23 @@ theorem Continuous.image_connectedComponent_subset [TopologicalSpace β] {f : α
((mem_image f (connectedComponent a) (f a)).2 ⟨a, mem_connectedComponent, rfl⟩)
#align continuous.image_connected_component_subset Continuous.image_connectedComponent_subset

theorem Continuous.image_connectedComponentIn_subset [TopologicalSpace β] {f : α → β} {s : Set α}
{a : α} (hf : Continuous f) (hx : a ∈ s) :
f '' connectedComponentIn s a ⊆ connectedComponentIn (f '' s) (f a) :=
(isPreconnected_connectedComponentIn.image _ hf.continuousOn).subset_connectedComponentIn
(mem_image_of_mem _ <| mem_connectedComponentIn hx)
(image_subset _ <| connectedComponentIn_subset _ _)

theorem Continuous.mapsTo_connectedComponent [TopologicalSpace β] {f : α → β} (h : Continuous f)
(a : α) : MapsTo f (connectedComponent a) (connectedComponent (f a)) :=
mapsTo'.2 <| h.image_connectedComponent_subset a
#align continuous.maps_to_connected_component Continuous.mapsTo_connectedComponent

theorem Continuous.mapsTo_connectedComponentIn [TopologicalSpace β] {f : α → β} {s : Set α}
(h : Continuous f) {a : α} (hx : a ∈ s) :
MapsTo f (connectedComponentIn s a) (connectedComponentIn (f '' s) (f a)) :=
mapsTo'.2 <| image_connectedComponentIn_subset h hx

theorem irreducibleComponent_subset_connectedComponent {x : α} :
irreducibleComponent x ⊆ connectedComponent x :=
isIrreducible_irreducibleComponent.isConnected.subset_connectedComponent mem_irreducibleComponent
Expand Down
7 changes: 7 additions & 0 deletions Mathlib/Topology/Homeomorph.lean
Original file line number Diff line number Diff line change
Expand Up @@ -318,6 +318,13 @@ theorem isConnected_preimage {s : Set Y} (h : X ≃ₜ Y) :
IsConnected (h ⁻¹' s) ↔ IsConnected s := by
rw [← image_symm, isConnected_image]

theorem image_connectedComponentIn {s : Set X} (h : X ≃ₜ Y) {x : X} (hx : x ∈ s) :
h '' connectedComponentIn s x = connectedComponentIn (h '' s) (h x) := by
refine (h.continuous.image_connectedComponentIn_subset hx).antisymm ?_
have := h.symm.continuous.image_connectedComponentIn_subset (mem_image_of_mem h hx)
rwa [image_subset_iff, h.preimage_symm, h.image_symm, h.preimage_image, h.symm_apply_apply]
at this

@[simp]
theorem comap_cocompact (h : X ≃ₜ Y) : comap h (cocompact Y) = cocompact X :=
(comap_cocompact_le h.continuous).antisymm <|
Expand Down

0 comments on commit 47a9066

Please sign in to comment.