Documentation

Mathlib.Topology.Maps.Basic

Specific classes of maps between topological spaces #

This file introduces the following properties of a map f : X → Y between topological spaces:

(Open and closed maps need not be continuous.)

References #

Tags #

open map, closed map, embedding, quotient map, identification map

theorem IsInducing.induced {X : Type u_1} {Y : Type u_2} [TopologicalSpace Y] (f : XY) :
@[deprecated IsInducing.induced]
theorem inducing_induced {X : Type u_1} {Y : Type u_2} [TopologicalSpace Y] (f : XY) :

Alias of IsInducing.induced.

@[deprecated IsInducing.id]
theorem inducing_id {X : Type u_1} [TopologicalSpace X] :

Alias of IsInducing.id.

theorem IsInducing.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) (hf : IsInducing f) :
@[deprecated IsInducing.comp]
theorem Inducing.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) (hf : IsInducing f) :

Alias of IsInducing.comp.

theorem IsInducing.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) :
@[deprecated IsInducing.of_comp_iff]
theorem Inducing.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) :

Alias of IsInducing.of_comp_iff.

theorem IsInducing.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsInducing (g f)) :
@[deprecated IsInducing.of_comp]
theorem inducing_of_inducing_compose {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsInducing (g f)) :

Alias of IsInducing.of_comp.

theorem isInducing_iff_nhds {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] :
IsInducing f ∀ (x : X), nhds x = Filter.comap f (nhds (f x))
@[deprecated isInducing_iff_nhds]
theorem inducing_iff_nhds {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] :
IsInducing f ∀ (x : X), nhds x = Filter.comap f (nhds (f x))

Alias of isInducing_iff_nhds.

theorem IsInducing.nhds_eq_comap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) (x : X) :
nhds x = Filter.comap f (nhds (f x))
theorem IsInducing.basis_nhds {X : Type u_1} {Y : Type u_2} {ι : Type u_4} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] {p : ιProp} {s : ιSet Y} (hf : IsInducing f) {x : X} (h_basis : (nhds (f x)).HasBasis p s) :
(nhds x).HasBasis p (Set.preimage f s)
theorem IsInducing.nhdsSet_eq_comap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) (s : Set X) :
theorem IsInducing.map_nhds_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) (x : X) :
theorem IsInducing.map_nhds_of_mem {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) (x : X) (h : Set.range f nhds (f x)) :
Filter.map f (nhds x) = nhds (f x)
theorem IsInducing.mapClusterPt_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {x : X} {l : Filter X} :
MapClusterPt (f x) l f ClusterPt x l
theorem IsInducing.image_mem_nhdsWithin {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {x : X} {s : Set X} (hs : s nhds x) :
f '' s nhdsWithin (f x) (Set.range f)
theorem IsInducing.tendsto_nhds_iff {Y : Type u_2} {Z : Type u_3} {ι : Type u_4} {g : YZ} [TopologicalSpace Y] [TopologicalSpace Z] {f : ιY} {l : Filter ι} {y : Y} (hg : IsInducing g) :
Filter.Tendsto f l (nhds y) Filter.Tendsto (g f) l (nhds (g y))
theorem IsInducing.continuousAt_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) {x : X} :
theorem IsInducing.continuous_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hg : IsInducing g) :
theorem IsInducing.continuousAt_iff' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace Y] [TopologicalSpace X] [TopologicalSpace Z] (hf : IsInducing f) {x : X} (h : Set.range f nhds (f x)) :
theorem IsInducing.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) :
theorem IsInducing.closure_eq_preimage_closure_image {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) (s : Set X) :
theorem IsInducing.isClosed_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {s : Set X} :
IsClosed s ∃ (t : Set Y), IsClosed t f ⁻¹' t = s
theorem IsInducing.isClosed_iff' {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {s : Set X} :
IsClosed s ∀ (x : X), f x closure (f '' s)x s
theorem IsInducing.isClosed_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (h : IsInducing f) (s : Set Y) (hs : IsClosed s) :
theorem IsInducing.isOpen_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {s : Set X} :
IsOpen s ∃ (t : Set Y), IsOpen t f ⁻¹' t = s
theorem IsInducing.setOf_isOpen {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) :
{s : Set X | IsOpen s} = Set.preimage f '' {t : Set Y | IsOpen t}
theorem IsInducing.dense_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace Y] [TopologicalSpace X] (hf : IsInducing f) {s : Set X} :
Dense s ∀ (x : X), f x closure (f '' s)
theorem IsInducing.of_subsingleton {X : Type u_1} {Y : Type u_2} [TopologicalSpace Y] [TopologicalSpace X] [Subsingleton X] (f : XY) :
theorem IsEmbedding.induced {X : Type u_1} {Y : Type u_2} {f : XY} [t : TopologicalSpace Y] (hf : Function.Injective f) :
@[deprecated Function.Injective.isEmbedding_induced]

Alias of IsEmbedding.induced.


Alias of IsEmbedding.induced.

theorem IsEmbedding.isInducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) :
@[deprecated IsEmbedding.isInducing]
theorem IsEmbedding.inducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) :

Alias of IsEmbedding.isInducing.

theorem IsEmbedding.mk' {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (inj : Function.Injective f) (induced : ∀ (x : X), Filter.comap f (nhds (f x)) = nhds x) :
@[deprecated IsEmbedding.mk']
theorem IsEmbedding.Embedding.mk' {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] (f : XY) (inj : Function.Injective f) (induced : ∀ (x : X), Filter.comap f (nhds (f x)) = nhds x) :

Alias of IsEmbedding.mk'.

@[deprecated IsEmbedding.id]

Alias of IsEmbedding.id.

theorem IsEmbedding.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) (hf : IsEmbedding f) :
@[deprecated IsEmbedding.comp]
theorem IsEmbedding.Embedding.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) (hf : IsEmbedding f) :

Alias of IsEmbedding.comp.

theorem IsEmbedding.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) :
@[deprecated IsEmbedding.of_comp_iff]
theorem IsEmbedding.Embedding.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) :

Alias of IsEmbedding.of_comp_iff.

theorem IsEmbedding.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsEmbedding (g f)) :
@[deprecated IsEmbedding.of_comp]
theorem IsEmbedding.embedding_of_embedding_compose {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsEmbedding (g f)) :

Alias of IsEmbedding.of_comp.

theorem IsEmbedding.of_leftInverse {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :
theorem Function.LeftInverse.isEmbedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :

Alias of IsEmbedding.of_leftInverse.

@[deprecated Function.LeftInverse.isEmbedding]
theorem Function.LeftInverse.embedding {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} {g : YX} (h : Function.LeftInverse f g) (hf : Continuous f) (hg : Continuous g) :

Alias of IsEmbedding.of_leftInverse.


Alias of IsEmbedding.of_leftInverse.

theorem IsEmbedding.map_nhds_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (x : X) :
@[deprecated IsEmbedding.map_nhds_eq]
theorem IsEmbedding.Embedding.map_nhds_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (x : X) :

Alias of IsEmbedding.map_nhds_eq.

theorem IsEmbedding.map_nhds_of_mem {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (x : X) (h : Set.range f nhds (f x)) :
Filter.map f (nhds x) = nhds (f x)
@[deprecated IsEmbedding.map_nhds_of_mem]
theorem IsEmbedding.Embedding.map_nhds_of_mem {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (x : X) (h : Set.range f nhds (f x)) :
Filter.map f (nhds x) = nhds (f x)

Alias of IsEmbedding.map_nhds_of_mem.

theorem IsEmbedding.tendsto_nhds_iff {Y : Type u_2} {Z : Type u_3} {ι : Type u_4} {g : YZ} [TopologicalSpace Y] [TopologicalSpace Z] {f : ιY} {l : Filter ι} {y : Y} (hg : IsEmbedding g) :
Filter.Tendsto f l (nhds y) Filter.Tendsto (g f) l (nhds (g y))
theorem IsEmbedding.continuous_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) :
@[deprecated IsEmbedding.continuous_iff]
theorem IsEmbedding.Embedding.continuous_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsEmbedding g) :

Alias of IsEmbedding.continuous_iff.

theorem IsEmbedding.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) :
theorem IsEmbedding.closure_eq_preimage_closure_image {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (s : Set X) :
@[deprecated IsEmbedding.closure_eq_preimage_closure_image]
theorem IsEmbedding.Embedding.closure_eq_preimage_closure_image {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (s : Set X) :

Alias of IsEmbedding.closure_eq_preimage_closure_image.

The topology induced under an inclusion f : X → Y from a discrete topological space Y is the discrete topology on X.

See also DiscreteTopology.of_continuous_injective.

@[deprecated IsEmbedding.discreteTopology]

Alias of IsEmbedding.discreteTopology.


The topology induced under an inclusion f : X → Y from a discrete topological space Y is the discrete topology on X.

See also DiscreteTopology.of_continuous_injective.

@[deprecated IsEmbedding.of_subsingleton]

Alias of IsEmbedding.of_subsingleton.

theorem isQuotientMap_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :
@[deprecated isQuotientMap_iff]
theorem quotientMap_iff {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :

Alias of isQuotientMap_iff.

@[deprecated isQuotientMap_iff_closed]
theorem quotientMap_iff_closed {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :

Alias of isQuotientMap_iff_closed.

theorem IsQuotientMap.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsQuotientMap g) (hf : IsQuotientMap f) :
theorem IsQuotientMap.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsQuotientMap (g f)) :
@[deprecated IsQuotientMap.of_comp]
theorem IsQuotientMap.of_quotientMap_compose {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : Continuous f) (hg : Continuous g) (hgf : IsQuotientMap (g f)) :

Alias of IsQuotientMap.of_comp.

theorem IsQuotientMap.of_inverse {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {g : YX} (hf : Continuous f) (hg : Continuous g) (h : Function.LeftInverse g f) :
theorem IsQuotientMap.continuous_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : IsQuotientMap f) :
theorem IsQuotientMap.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsQuotientMap f) :
theorem IsQuotientMap.isOpen_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsQuotientMap f) {s : Set Y} :
theorem IsQuotientMap.isClosed_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsQuotientMap f) {s : Set Y} :
theorem IsOpenMap.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsOpenMap g) (hf : IsOpenMap f) :
theorem IsOpenMap.isOpen_range {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) :
theorem IsOpenMap.image_mem_nhds {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) {x : X} {s : Set X} (hx : s nhds x) :
f '' s nhds (f x)
theorem IsOpenMap.range_mem_nhds {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) (x : X) :
theorem IsOpenMap.mapsTo_interior {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) {s : Set X} {t : Set Y} (h : Set.MapsTo f s t) :
theorem IsOpenMap.image_interior_subset {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) (s : Set X) :
theorem IsOpenMap.nhds_le {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) (x : X) :
nhds (f x) Filter.map f (nhds x)
theorem IsOpenMap.of_nhds_le {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : ∀ (x : X), nhds (f x) Filter.map f (nhds x)) :
theorem IsOpenMap.of_sections {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h : ∀ (x : X), ∃ (g : YX), ContinuousAt g (f x) g (f x) = x Function.RightInverse g f) :
theorem IsOpenMap.of_inverse {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {f' : YX} (h : Continuous f') (l_inv : Function.LeftInverse f f') (r_inv : Function.RightInverse f f') :
theorem IsOpenMap.isQuotientMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (open_map : IsOpenMap f) (cont : Continuous f) (surj : Function.Surjective f) :

A continuous surjective open map is a quotient map.

@[deprecated IsOpenMap.isQuotientMap]
theorem IsOpenMap.to_quotientMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (open_map : IsOpenMap f) (cont : Continuous f) (surj : Function.Surjective f) :

Alias of IsOpenMap.isQuotientMap.


A continuous surjective open map is a quotient map.

theorem IsOpenMap.preimage_interior_eq_interior_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf₁ : IsOpenMap f) (hf₂ : Continuous f) (s : Set Y) :
theorem IsOpenMap.preimage_closure_eq_closure_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) (hfc : Continuous f) (s : Set Y) :
theorem IsOpenMap.preimage_frontier_eq_frontier_preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenMap f) (hfc : Continuous f) (s : Set Y) :
theorem IsOpenMap.of_isEmpty {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [h : IsEmpty X] (f : XY) :
theorem isOpenMap_iff_nhds_le {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :
IsOpenMap f ∀ (x : X), nhds (f x) Filter.map f (nhds x)
theorem isOpenMap_iff_interior {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :
IsOpenMap f ∀ (s : Set X), f '' interior s interior (f '' s)
theorem IsInducing.isOpenMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hi : IsInducing f) (ho : IsOpen (Set.range f)) :

An inducing map with an open range is an open map.

@[deprecated IsInducing.isOpenMap]
theorem Inducing.isOpenMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hi : IsInducing f) (ho : IsOpen (Set.range f)) :

Alias of IsInducing.isOpenMap.


An inducing map with an open range is an open map.

theorem Dense.preimage {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {s : Set Y} (hs : Dense s) (hf : IsOpenMap f) :

Preimage of a dense set under an open map is dense.

theorem IsClosedMap.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsClosedMap g) (hf : IsClosedMap f) :
theorem IsClosedMap.closure_image_subset {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedMap f) (s : Set X) :
closure (f '' s) f '' closure s
theorem IsClosedMap.of_inverse {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {f' : YX} (h : Continuous f') (l_inv : Function.LeftInverse f f') (r_inv : Function.RightInverse f f') :
theorem IsClosedMap.of_nonempty {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h : ∀ (s : Set X), IsClosed ss.NonemptyIsClosed (f '' s)) :
theorem IsClosedMap.isClosed_range {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedMap f) :
@[deprecated IsClosedMap.isClosed_range]
theorem IsClosedMap.closed_range {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedMap f) :

Alias of IsClosedMap.isClosed_range.

theorem IsClosedMap.isQuotientMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Function.Surjective f) :
@[deprecated IsClosedMap.isQuotientMap]
theorem IsClosedMap.to_quotientMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hcl : IsClosedMap f) (hcont : Continuous f) (hsurj : Function.Surjective f) :

Alias of IsClosedMap.isQuotientMap.

theorem IsInducing.isClosedMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsInducing f) (h : IsClosed (Set.range f)) :
@[deprecated IsInducing.isClosedMap]
theorem Inducing.isClosedMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsInducing f) (h : IsClosed (Set.range f)) :

Alias of IsInducing.isClosedMap.

theorem isClosedMap_iff_closure_image {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :
IsClosedMap f ∀ (s : Set X), closure (f '' s) f '' closure s
theorem isClosedMap_iff_clusterPt {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] :
IsClosedMap f ∀ (s : Set X) (y : Y), MapClusterPt y (Filter.principal s) f∃ (x : X), f x = y ClusterPt x (Filter.principal s)

A map f : X → Y is closed if and only if for all sets s, any cluster point of f '' s is the image by f of some cluster point of s. If you require this for all filters instead of just principal filters, and also that f is continuous, you get the notion of proper map. See isProperMap_iff_clusterPt.

theorem IsClosedMap.closure_image_eq_of_continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (f_closed : IsClosedMap f) (f_cont : Continuous f) (s : Set X) :
closure (f '' s) = f '' closure s
theorem IsClosedMap.lift'_closure_map_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (f_closed : IsClosedMap f) (f_cont : Continuous f) (F : Filter X) :
(Filter.map f F).lift' closure = Filter.map f (F.lift' closure)
theorem IsClosedMap.mapClusterPt_iff_lift'_closure {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {F : Filter X} (f_closed : IsClosedMap f) (f_cont : Continuous f) {y : Y} :
MapClusterPt y F f (F.lift' closure Filter.principal (f ⁻¹' {y})).NeBot
theorem IsOpenEmbedding.isEmbedding {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :
theorem IsOpenEmbedding.isInducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :
@[deprecated IsOpenEmbedding.isInducing]
theorem IsOpenEmbedding.inducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :

Alias of IsOpenEmbedding.isInducing.

theorem IsOpenEmbedding.isOpenMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :
@[deprecated IsOpenEmbedding.isOpenMap]
theorem OpenEmbedding.isOpenMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :

Alias of IsOpenEmbedding.isOpenMap.

theorem IsOpenEmbedding.map_nhds_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) (x : X) :
Filter.map f (nhds x) = nhds (f x)
@[deprecated IsOpenEmbedding.map_nhds_eq]
theorem OpenEmbedding.map_nhds_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) (x : X) :
Filter.map f (nhds x) = nhds (f x)

Alias of IsOpenEmbedding.map_nhds_eq.

theorem IsOpenEmbedding.open_iff_image_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {s : Set X} :
@[deprecated IsOpenEmbedding.open_iff_image_open]
theorem OpenEmbedding.open_iff_image_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {s : Set X} :

Alias of IsOpenEmbedding.open_iff_image_open.

theorem IsOpenEmbedding.tendsto_nhds_iff {Y : Type u_2} {Z : Type u_3} {ι : Type u_4} {g : YZ} [TopologicalSpace Y] [TopologicalSpace Z] {f : ιY} {l : Filter ι} {y : Y} (hg : IsOpenEmbedding g) :
Filter.Tendsto f l (nhds y) Filter.Tendsto (g f) l (nhds (g y))
@[deprecated IsOpenEmbedding.tendsto_nhds_iff]
theorem OpenEmbedding.tendsto_nhds_iff {Y : Type u_2} {Z : Type u_3} {ι : Type u_4} {g : YZ} [TopologicalSpace Y] [TopologicalSpace Z] {f : ιY} {l : Filter ι} {y : Y} (hg : IsOpenEmbedding g) :
Filter.Tendsto f l (nhds y) Filter.Tendsto (g f) l (nhds (g y))

Alias of IsOpenEmbedding.tendsto_nhds_iff.

theorem IsOpenEmbedding.tendsto_nhds_iff' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {l : Filter Z} {x : X} :
Filter.Tendsto (g f) (nhds x) l Filter.Tendsto g (nhds (f x)) l
@[deprecated IsOpenEmbedding.tendsto_nhds_iff']
theorem OpenEmbedding.tendsto_nhds_iff' {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {l : Filter Z} {x : X} :
Filter.Tendsto (g f) (nhds x) l Filter.Tendsto g (nhds (f x)) l

Alias of IsOpenEmbedding.tendsto_nhds_iff'.

theorem IsOpenEmbedding.continuousAt_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : IsOpenEmbedding f) {x : X} :
@[deprecated IsOpenEmbedding.continuousAt_iff]
theorem OpenEmbedding.continuousAt_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hf : IsOpenEmbedding f) {x : X} :

Alias of IsOpenEmbedding.continuousAt_iff.

theorem IsOpenEmbedding.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :
@[deprecated IsOpenEmbedding.continuous]
theorem OpenEmbedding.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) :

Alias of IsOpenEmbedding.continuous.

theorem IsOpenEmbedding.open_iff_preimage_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {s : Set Y} (hs : s Set.range f) :
@[deprecated IsOpenEmbedding.open_iff_preimage_open]
theorem OpenEmbedding.open_iff_preimage_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsOpenEmbedding f) {s : Set Y} (hs : s Set.range f) :

Alias of IsOpenEmbedding.open_iff_preimage_open.

theorem IsOpenEmbedding.of_isEmbedding_isOpenMap {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h₁ : IsEmbedding f) (h₂ : IsOpenMap f) :
@[deprecated IsOpenEmbedding.of_isEmbedding_isOpenMap]
theorem isOpenEmbedding_of_embedding_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h₁ : IsEmbedding f) (h₂ : IsOpenMap f) :

Alias of IsOpenEmbedding.of_isEmbedding_isOpenMap.

@[deprecated IsOpenEmbedding.of_isEmbedding_isOpenMap]
theorem openEmbedding_of_embedding_open {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h₁ : IsEmbedding f) (h₂ : IsOpenMap f) :

Alias of IsOpenEmbedding.of_isEmbedding_isOpenMap.

A surjective embedding is an IsOpenEmbedding.

@[deprecated IsEmbedding.isOpenEmbedding_of_surjective]

Alias of IsEmbedding.isOpenEmbedding_of_surjective.


A surjective embedding is an IsOpenEmbedding.

theorem IsOpenEmbedding.of_isEmbedding {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsEmbedding f) (hsurj : Function.Surjective f) :

Alias of IsEmbedding.isOpenEmbedding_of_surjective.


A surjective embedding is an IsOpenEmbedding.

@[deprecated IsEmbedding.isOpenEmbedding_of_surjective]

Alias of IsEmbedding.isOpenEmbedding_of_surjective.


A surjective embedding is an IsOpenEmbedding.

@[deprecated isOpenEmbedding_iff_isEmbedding_isOpenMap]

Alias of isOpenEmbedding_iff_isEmbedding_isOpenMap.

@[deprecated isOpenEmbedding_iff_isEmbedding_isOpenMap]

Alias of isOpenEmbedding_iff_isEmbedding_isOpenMap.

@[deprecated isOpenEmbedding_iff_continuous_injective_open]

Alias of isOpenEmbedding_iff_continuous_injective_open.

@[deprecated isOpenEmbedding_id]

Alias of isOpenEmbedding_id.

theorem IsOpenEmbedding.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsOpenEmbedding g) (hf : IsOpenEmbedding f) :
theorem IsOpenEmbedding.isOpenMap_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsOpenEmbedding g) :
theorem IsOpenEmbedding.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (f : XY) (hg : IsOpenEmbedding g) :
theorem IsOpenEmbedding.of_comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (f : XY) (hg : IsOpenEmbedding g) (h : IsOpenEmbedding (g f)) :
theorem IsOpenEmbedding.of_isEmpty {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] [IsEmpty X] (f : XY) :
theorem IsOpenEmbedding.image_mem_nhds {X : Type u_1} {Y : Type u_2} [TopologicalSpace X] [TopologicalSpace Y] {f : XY} (hf : IsOpenEmbedding f) {s : Set X} {x : X} :
f '' s nhds (f x) s nhds x
theorem IsClosedEmbedding.isInducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) :
theorem IsClosedEmbedding.continuous {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) :
@[deprecated IsClosedEmbedding.isInducing]
theorem IsClosedEmbedding.inducing {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) :

Alias of IsClosedEmbedding.isInducing.

theorem IsClosedEmbedding.tendsto_nhds_iff {X : Type u_1} {Y : Type u_2} {ι : Type u_4} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] {g : ιX} {l : Filter ι} {x : X} (hf : IsClosedEmbedding f) :
Filter.Tendsto g l (nhds x) Filter.Tendsto (f g) l (nhds (f x))
theorem IsClosedEmbedding.closed_iff_image_closed {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) {s : Set X} :
theorem IsClosedEmbedding.closed_iff_preimage_closed {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) {s : Set Y} (hs : s Set.range f) :
@[deprecated IsClosedEmbedding.of_isEmbedding_isClosedMap]
theorem IsClosedEmbedding.of_embedding_closed {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h₁ : IsEmbedding f) (h₂ : IsClosedMap f) :

Alias of IsClosedEmbedding.of_isEmbedding_isClosedMap.

@[deprecated IsClosedEmbedding.of_isEmbedding_isClosedMap]
theorem closedEmbedding_of_embedding_closed {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (h₁ : IsEmbedding f) (h₂ : IsClosedMap f) :

Alias of IsClosedEmbedding.of_isEmbedding_isClosedMap.

@[deprecated IsClosedEmbedding.of_continuous_injective_isClosedMap]

Alias of IsClosedEmbedding.of_continuous_injective_isClosedMap.

@[deprecated isClosedEmbedding_id]

Alias of isClosedEmbedding_id.

theorem IsClosedEmbedding.comp {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsClosedEmbedding g) (hf : IsClosedEmbedding f) :
theorem IsClosedEmbedding.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsClosedEmbedding g) :
@[deprecated IsClosedEmbedding.of_comp_iff]
theorem IsClosedEmbedding.Embedding.of_comp_iff {X : Type u_1} {Y : Type u_2} {Z : Type u_3} {f : XY} {g : YZ} [TopologicalSpace X] [TopologicalSpace Y] [TopologicalSpace Z] (hg : IsClosedEmbedding g) :

Alias of IsClosedEmbedding.of_comp_iff.

theorem IsClosedEmbedding.closure_image_eq {X : Type u_1} {Y : Type u_2} {f : XY} [TopologicalSpace X] [TopologicalSpace Y] (hf : IsClosedEmbedding f) (s : Set X) :
closure (f '' s) = f '' closure s