Skip to content

Commit

Permalink
feat: add Homeomorph.subtype for lifting homeomorphisms to subtypes (
Browse files Browse the repository at this point in the history
…#9959)

This extends `Equiv.subtypeEquiv`, which promotes `e : α ≃ β` to `e.subtypeEquiv _ : {a : α // p a} ≃ {b : β // q b}`, to homeomorphisms.

We also add a missing lemma linking `Equiv.subtypeEquiv` to `Subtype.map`, and update the definition to use `Subtype.map` also.
  • Loading branch information
j-loreaux committed Jan 25, 2024
1 parent 0c34368 commit 8c661e5
Show file tree
Hide file tree
Showing 2 changed files with 24 additions and 0 deletions.
4 changes: 4 additions & 0 deletions Mathlib/Logic/Equiv/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1155,6 +1155,10 @@ def subtypeEquiv {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a,
right_inv b := Subtype.ext <| by simp
#align equiv.subtype_equiv Equiv.subtypeEquiv

lemma coe_subtypeEquiv_eq_map {X Y : Type*} {p : X → Prop} {q : Y → Prop} (e : X ≃ Y)
(h : ∀ x, p x ↔ q (e x)) : ⇑(e.subtypeEquiv h) = Subtype.map e (h · |>.mp) :=
rfl

@[simp]
theorem subtypeEquiv_refl {p : α → Prop} (h : ∀ a, p a ↔ p (Equiv.refl _ a) := fun a => Iff.rfl) :
(Equiv.refl α).subtypeEquiv h = Equiv.refl { a : α // p a } := by
Expand Down
20 changes: 20 additions & 0 deletions Mathlib/Topology/Homeomorph.lean
Original file line number Diff line number Diff line change
Expand Up @@ -518,6 +518,26 @@ theorem comp_isOpenMap_iff' (h : X ≃ₜ Y) {f : Y → Z} : IsOpenMap (f ∘ h)
exact hf.comp h.symm.isOpenMap
#align homeomorph.comp_is_open_map_iff' Homeomorph.comp_isOpenMap_iff'

/-- A homeomorphism `h : X ≃ₜ Y` lifts to a homeomorphism between subtypes corresponding to
predicates `p : X → Prop` and `q : Y → Prop` so long as `p = q ∘ h`. -/
@[simps!]
def subtype {p : X → Prop} {q : Y → Prop} (h : X ≃ₜ Y) (h_iff : ∀ x, p x ↔ q (h x)) :
{x // p x} ≃ₜ {y // q y} where
continuous_toFun := by simpa [Equiv.coe_subtypeEquiv_eq_map] using h.continuous.subtype_map _
continuous_invFun := by simpa [Equiv.coe_subtypeEquiv_eq_map] using
h.symm.continuous.subtype_map _
__ := h.subtypeEquiv h_iff

@[simp]
lemma subtype_toEquiv {p : X → Prop} {q : Y → Prop} (h : X ≃ₜ Y) (h_iff : ∀ x, p x ↔ q (h x)) :
(h.subtype h_iff).toEquiv = h.toEquiv.subtypeEquiv h_iff :=
rfl

/-- A homeomorphism `h : X ≃ₜ Y` lifts to a homeomorphism between sets `s : Set X` and `t : Set Y`
whenever `h` maps `s` onto `t`. -/
abbrev sets {s : Set X} {t : Set Y} (h : X ≃ₜ Y) (h_eq : s = h ⁻¹' t) : s ≃ₜ t :=
h.subtype <| Set.ext_iff.mp h_eq

/-- If two sets are equal, then they are homeomorphic. -/
def setCongr {s t : Set X} (h : s = t) : s ≃ₜ t where
continuous_toFun := continuous_inclusion h.subset
Expand Down

0 comments on commit 8c661e5

Please sign in to comment.