Skip to content

Commit

Permalink
feat: add PartialHomeomorph.extend_target' (#9977)
Browse files Browse the repository at this point in the history
Inspired by sphere-eversion; similar to `PartialEquiv.image_source_eq_target`.
  • Loading branch information
grunweg committed Jan 26, 2024
1 parent 0500719 commit 19ab970
Showing 1 changed file with 3 additions and 0 deletions.
3 changes: 3 additions & 0 deletions Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean
Expand Up @@ -980,6 +980,9 @@ theorem extend_target : (f.extend I).target = I.symm ⁻¹' f.target ∩ range I
simp_rw [extend, PartialEquiv.trans_target, I.target_eq, I.toPartialEquiv_coe_symm, inter_comm]
#align local_homeomorph.extend_target PartialHomeomorph.extend_target

theorem extend_target' : (f.extend I).target = I '' f.target := by
rw [extend, PartialEquiv.trans_target'', I.source_eq, univ_inter, I.toPartialEquiv_coe]

lemma isOpen_extend_target [I.Boundaryless] : IsOpen (f.extend I).target := by
rw [extend_target, I.range_eq_univ, inter_univ]
exact I.continuous_symm.isOpen_preimage _ f.open_target
Expand Down

0 comments on commit 19ab970

Please sign in to comment.