![]() |
Metamath
Proof Explorer Theorem List (p. 217 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | regr1 21601 | A regular space is R1, which means that any two topologically distinct points can be separated by neighborhoods. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Reg → (KQ‘𝐽) ∈ Haus) | ||
Theorem | kqreg 21602 | The Kolmogorov quotient of a regular space is regular. By regr1 21601 it is also Hausdorff, so we can also say that a space is regular iff the Kolmogorov quotient is regular Hausdorff (T3). (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Reg ↔ (KQ‘𝐽) ∈ Reg) | ||
Theorem | kqnrm 21603 | The Kolmogorov quotient of a normal space is normal. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Nrm ↔ (KQ‘𝐽) ∈ Nrm) | ||
Syntax | chmeo 21604 | Extend class notation with the class of all homeomorphisms. |
class Homeo | ||
Syntax | chmph 21605 | Extend class notation with the relation "is homeomorphic to.". |
class ≃ | ||
Definition | df-hmeo 21606* | Function returning all the homeomorphisms from topology 𝑗 to topology 𝑘. (Contributed by FL, 14-Feb-2007.) |
⊢ Homeo = (𝑗 ∈ Top, 𝑘 ∈ Top ↦ {𝑓 ∈ (𝑗 Cn 𝑘) ∣ ◡𝑓 ∈ (𝑘 Cn 𝑗)}) | ||
Definition | df-hmph 21607 | Definition of the relation 𝑥 is homeomorphic to 𝑦. (Contributed by FL, 14-Feb-2007.) |
⊢ ≃ = (◡Homeo “ (V ∖ 1𝑜)) | ||
Theorem | hmeofn 21608 | The set of homeomorphisms is a function on topologies. (Contributed by Mario Carneiro, 23-Aug-2015.) |
⊢ Homeo Fn (Top × Top) | ||
Theorem | hmeofval 21609* | The set of all the homeomorphisms between two topologies. (Contributed by FL, 14-Feb-2007.) (Revised by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐽Homeo𝐾) = {𝑓 ∈ (𝐽 Cn 𝐾) ∣ ◡𝑓 ∈ (𝐾 Cn 𝐽)} | ||
Theorem | ishmeo 21610 | The predicate F is a homeomorphism between topology 𝐽 and topology 𝐾. Proposition of [BourbakiTop1] p. I.2. (Contributed by FL, 14-Feb-2007.) (Revised by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) ↔ (𝐹 ∈ (𝐽 Cn 𝐾) ∧ ◡𝐹 ∈ (𝐾 Cn 𝐽))) | ||
Theorem | hmeocn 21611 | A homeomorphism is continuous. (Contributed by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) → 𝐹 ∈ (𝐽 Cn 𝐾)) | ||
Theorem | hmeocnvcn 21612 | The converse of a homeomorphism is continuous. (Contributed by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) → ◡𝐹 ∈ (𝐾 Cn 𝐽)) | ||
Theorem | hmeocnv 21613 | The converse of a homeomorphism is a homeomorphism. (Contributed by FL, 5-Mar-2007.) (Revised by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) → ◡𝐹 ∈ (𝐾Homeo𝐽)) | ||
Theorem | hmeof1o2 21614 | A homeomorphism is a 1-1-onto mapping. (Contributed by Mario Carneiro, 22-Aug-2015.) |
⊢ ((𝐽 ∈ (TopOn‘𝑋) ∧ 𝐾 ∈ (TopOn‘𝑌) ∧ 𝐹 ∈ (𝐽Homeo𝐾)) → 𝐹:𝑋–1-1-onto→𝑌) | ||
Theorem | hmeof1o 21615 | A homeomorphism is a 1-1-onto mapping. (Contributed by FL, 5-Mar-2007.) (Revised by Mario Carneiro, 30-May-2014.) |
⊢ 𝑋 = ∪ 𝐽 & ⊢ 𝑌 = ∪ 𝐾 ⇒ ⊢ (𝐹 ∈ (𝐽Homeo𝐾) → 𝐹:𝑋–1-1-onto→𝑌) | ||
Theorem | hmeoima 21616 | The image of an open set by a homeomorphism is an open set. (Contributed by FL, 5-Mar-2007.) (Revised by Mario Carneiro, 22-Aug-2015.) |
⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐴 ∈ 𝐽) → (𝐹 “ 𝐴) ∈ 𝐾) | ||
Theorem | hmeoopn 21617 | Homeomorphisms preserve openness. (Contributed by Jeff Madsen, 2-Sep-2009.) (Proof shortened by Mario Carneiro, 25-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐴 ⊆ 𝑋) → (𝐴 ∈ 𝐽 ↔ (𝐹 “ 𝐴) ∈ 𝐾)) | ||
Theorem | hmeocld 21618 | Homeomorphisms preserve closedness. (Contributed by Jeff Madsen, 2-Sep-2009.) (Proof shortened by Mario Carneiro, 25-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐴 ⊆ 𝑋) → (𝐴 ∈ (Clsd‘𝐽) ↔ (𝐹 “ 𝐴) ∈ (Clsd‘𝐾))) | ||
Theorem | hmeocls 21619 | Homeomorphisms preserve closures. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐴 ⊆ 𝑋) → ((cls‘𝐾)‘(𝐹 “ 𝐴)) = (𝐹 “ ((cls‘𝐽)‘𝐴))) | ||
Theorem | hmeontr 21620 | Homeomorphisms preserve interiors. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐴 ⊆ 𝑋) → ((int‘𝐾)‘(𝐹 “ 𝐴)) = (𝐹 “ ((int‘𝐽)‘𝐴))) | ||
Theorem | hmeoimaf1o 21621* | The function mapping open sets to their images under a homeomorphism is a bijection of topologies. (Contributed by Mario Carneiro, 10-Sep-2015.) |
⊢ 𝐺 = (𝑥 ∈ 𝐽 ↦ (𝐹 “ 𝑥)) ⇒ ⊢ (𝐹 ∈ (𝐽Homeo𝐾) → 𝐺:𝐽–1-1-onto→𝐾) | ||
Theorem | hmeores 21622 | The restriction of a homeomorphism is a homeomorphism. (Contributed by Mario Carneiro, 14-Sep-2014.) (Proof shortened by Mario Carneiro, 22-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝑌 ⊆ 𝑋) → (𝐹 ↾ 𝑌) ∈ ((𝐽 ↾t 𝑌)Homeo(𝐾 ↾t (𝐹 “ 𝑌)))) | ||
Theorem | hmeoco 21623 | The composite of two homeomorphisms is a homeomorphism. (Contributed by FL, 9-Mar-2007.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ ((𝐹 ∈ (𝐽Homeo𝐾) ∧ 𝐺 ∈ (𝐾Homeo𝐿)) → (𝐺 ∘ 𝐹) ∈ (𝐽Homeo𝐿)) | ||
Theorem | idhmeo 21624 | The identity function is a homeomorphism. (Contributed by FL, 14-Feb-2007.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐽 ∈ (TopOn‘𝑋) → ( I ↾ 𝑋) ∈ (𝐽Homeo𝐽)) | ||
Theorem | hmeocnvb 21625 | The converse of a homeomorphism is a homeomorphism. (Contributed by FL, 5-Mar-2007.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ (Rel 𝐹 → (◡𝐹 ∈ (𝐽Homeo𝐾) ↔ 𝐹 ∈ (𝐾Homeo𝐽))) | ||
Theorem | hmeoqtop 21626 | A homeomorphism is a quotient map. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) → 𝐾 = (𝐽 qTop 𝐹)) | ||
Theorem | hmph 21627 | Express the predicate 𝐽 is homeomorphic to 𝐾. (Contributed by FL, 14-Feb-2007.) (Revised by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 ↔ (𝐽Homeo𝐾) ≠ ∅) | ||
Theorem | hmphi 21628 | If there is a homeomorphism between spaces, then the spaces are homeomorphic. (Contributed by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐹 ∈ (𝐽Homeo𝐾) → 𝐽 ≃ 𝐾) | ||
Theorem | hmphtop 21629 | Reverse closure for the homeomorphic predicate. (Contributed by Mario Carneiro, 22-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Top ∧ 𝐾 ∈ Top)) | ||
Theorem | hmphtop1 21630 | The relation "being homeomorphic to" implies the operands are topologies. (Contributed by FL, 23-Mar-2007.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → 𝐽 ∈ Top) | ||
Theorem | hmphtop2 21631 | The relation "being homeomorphic to" implies the operands are topologies. (Contributed by FL, 23-Mar-2007.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → 𝐾 ∈ Top) | ||
Theorem | hmphref 21632 | "Is homeomorphic to" is reflexive. (Contributed by FL, 25-Feb-2007.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐽 ∈ Top → 𝐽 ≃ 𝐽) | ||
Theorem | hmphsym 21633 | "Is homeomorphic to" is symmetric. (Contributed by FL, 8-Mar-2007.) (Proof shortened by Mario Carneiro, 30-May-2014.) |
⊢ (𝐽 ≃ 𝐾 → 𝐾 ≃ 𝐽) | ||
Theorem | hmphtr 21634 | "Is homeomorphic to" is transitive. (Contributed by FL, 9-Mar-2007.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ ((𝐽 ≃ 𝐾 ∧ 𝐾 ≃ 𝐿) → 𝐽 ≃ 𝐿) | ||
Theorem | hmpher 21635 | "Is homeomorphic to" is an equivalence relation. (Contributed by FL, 22-Mar-2007.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ ≃ Er Top | ||
Theorem | hmphen 21636 | Homeomorphisms preserve the cardinality of the topologies. (Contributed by FL, 1-Jun-2008.) (Revised by Mario Carneiro, 10-Sep-2015.) |
⊢ (𝐽 ≃ 𝐾 → 𝐽 ≈ 𝐾) | ||
Theorem | hmphsymb 21637 | "Is homeomorphic to" is symmetric. (Contributed by FL, 22-Feb-2007.) |
⊢ (𝐽 ≃ 𝐾 ↔ 𝐾 ≃ 𝐽) | ||
Theorem | haushmphlem 21638* | Lemma for haushmph 21643 and similar theorems. If the topological property 𝐴 is preserved under injective preimages, then property 𝐴 is preserved under homeomorphisms. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ 𝐴 → 𝐽 ∈ Top) & ⊢ ((𝐽 ∈ 𝐴 ∧ 𝑓:∪ 𝐾–1-1→∪ 𝐽 ∧ 𝑓 ∈ (𝐾 Cn 𝐽)) → 𝐾 ∈ 𝐴) ⇒ ⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ 𝐴 → 𝐾 ∈ 𝐴)) | ||
Theorem | cmphmph 21639 | Compactness is a topological property-that is, for any two homeomorphic topologies, either both are compact or neither is. (Contributed by Jeff Hankins, 30-Jun-2009.) (Revised by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Comp → 𝐾 ∈ Comp)) | ||
Theorem | connhmph 21640 | Connectedness is a topological property. (Contributed by Jeff Hankins, 3-Jul-2009.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Conn → 𝐾 ∈ Conn)) | ||
Theorem | t0hmph 21641 | T0 is a topological property. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Kol2 → 𝐾 ∈ Kol2)) | ||
Theorem | t1hmph 21642 | T1 is a topological property. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Fre → 𝐾 ∈ Fre)) | ||
Theorem | haushmph 21643 | Hausdorff-ness is a topological property. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Haus → 𝐾 ∈ Haus)) | ||
Theorem | reghmph 21644 | Regularity is a topological property. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Reg → 𝐾 ∈ Reg)) | ||
Theorem | nrmhmph 21645 | Normality is a topological property. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ≃ 𝐾 → (𝐽 ∈ Nrm → 𝐾 ∈ Nrm)) | ||
Theorem | hmph0 21646 | A topology homeomorphic to the empty set is empty. (Contributed by FL, 18-Aug-2008.) (Revised by Mario Carneiro, 10-Sep-2015.) |
⊢ (𝐽 ≃ {∅} ↔ 𝐽 = {∅}) | ||
Theorem | hmphdis 21647 | Homeomorphisms preserve topological discretion. (Contributed by Mario Carneiro, 10-Sep-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ (𝐽 ≃ 𝒫 𝐴 → 𝐽 = 𝒫 𝑋) | ||
Theorem | hmphindis 21648 | Homeomorphisms preserve topological indiscretion. (Contributed by FL, 18-Aug-2008.) (Revised by Mario Carneiro, 10-Sep-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ (𝐽 ≃ {∅, 𝐴} → 𝐽 = {∅, 𝑋}) | ||
Theorem | indishmph 21649 | Equinumerous sets equipped with their indiscrete topologies are homeomorphic (which means in that particular case that a segment is homeomorphic to a circle contrary to what Wikipedia claims). (Contributed by FL, 17-Aug-2008.) (Proof shortened by Mario Carneiro, 10-Sep-2015.) |
⊢ (𝐴 ≈ 𝐵 → {∅, 𝐴} ≃ {∅, 𝐵}) | ||
Theorem | hmphen2 21650 | Homeomorphisms preserve the cardinality of the underlying sets. (Contributed by FL, 17-Aug-2008.) (Revised by Mario Carneiro, 10-Sep-2015.) |
⊢ 𝑋 = ∪ 𝐽 & ⊢ 𝑌 = ∪ 𝐾 ⇒ ⊢ (𝐽 ≃ 𝐾 → 𝑋 ≈ 𝑌) | ||
Theorem | cmphaushmeo 21651 | A continuous bijection from a compact space to a Hausdorff space is a homeomorphism. (Contributed by Mario Carneiro, 17-Feb-2015.) |
⊢ 𝑋 = ∪ 𝐽 & ⊢ 𝑌 = ∪ 𝐾 ⇒ ⊢ ((𝐽 ∈ Comp ∧ 𝐾 ∈ Haus ∧ 𝐹 ∈ (𝐽 Cn 𝐾)) → (𝐹 ∈ (𝐽Homeo𝐾) ↔ 𝐹:𝑋–1-1-onto→𝑌)) | ||
Theorem | ordthmeolem 21652 | Lemma for ordthmeo 21653. (Contributed by Mario Carneiro, 9-Sep-2015.) |
⊢ 𝑋 = dom 𝑅 & ⊢ 𝑌 = dom 𝑆 ⇒ ⊢ ((𝑅 ∈ 𝑉 ∧ 𝑆 ∈ 𝑊 ∧ 𝐹 Isom 𝑅, 𝑆 (𝑋, 𝑌)) → 𝐹 ∈ ((ordTop‘𝑅) Cn (ordTop‘𝑆))) | ||
Theorem | ordthmeo 21653 | An order isomorphism is a homeomorphism on the respective order topologies. (Contributed by Mario Carneiro, 9-Sep-2015.) |
⊢ 𝑋 = dom 𝑅 & ⊢ 𝑌 = dom 𝑆 ⇒ ⊢ ((𝑅 ∈ 𝑉 ∧ 𝑆 ∈ 𝑊 ∧ 𝐹 Isom 𝑅, 𝑆 (𝑋, 𝑌)) → 𝐹 ∈ ((ordTop‘𝑅)Homeo(ordTop‘𝑆))) | ||
Theorem | txhmeo 21654* | Lift a pair of homeomorphisms on the factors to a homeomorphism of product topologies. (Contributed by Mario Carneiro, 2-Sep-2015.) |
⊢ 𝑋 = ∪ 𝐽 & ⊢ 𝑌 = ∪ 𝐾 & ⊢ (𝜑 → 𝐹 ∈ (𝐽Homeo𝐿)) & ⊢ (𝜑 → 𝐺 ∈ (𝐾Homeo𝑀)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ 〈(𝐹‘𝑥), (𝐺‘𝑦)〉) ∈ ((𝐽 ×t 𝐾)Homeo(𝐿 ×t 𝑀))) | ||
Theorem | txswaphmeolem 21655* | Show inverse for the "swap components" operation on a Cartesian product. (Contributed by Mario Carneiro, 21-Mar-2015.) |
⊢ ((𝑦 ∈ 𝑌, 𝑥 ∈ 𝑋 ↦ 〈𝑥, 𝑦〉) ∘ (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ 〈𝑦, 𝑥〉)) = ( I ↾ (𝑋 × 𝑌)) | ||
Theorem | txswaphmeo 21656* | There is a homeomorphism from 𝑋 × 𝑌 to 𝑌 × 𝑋. (Contributed by Mario Carneiro, 21-Mar-2015.) |
⊢ ((𝐽 ∈ (TopOn‘𝑋) ∧ 𝐾 ∈ (TopOn‘𝑌)) → (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ 〈𝑦, 𝑥〉) ∈ ((𝐽 ×t 𝐾)Homeo(𝐾 ×t 𝐽))) | ||
Theorem | pt1hmeo 21657* | The canonical homeomorphism from a topological product on a singleton to the topology of the factor. (Contributed by Mario Carneiro, 3-Feb-2015.) (Proof shortened by AV, 18-Apr-2021.) |
⊢ 𝐾 = (∏t‘{〈𝐴, 𝐽〉}) & ⊢ (𝜑 → 𝐴 ∈ 𝑉) & ⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) ⇒ ⊢ (𝜑 → (𝑥 ∈ 𝑋 ↦ {〈𝐴, 𝑥〉}) ∈ (𝐽Homeo𝐾)) | ||
Theorem | ptuncnv 21658* | Exhibit the converse function of the map 𝐺 which joins two product topologies on disjoint index sets. (Contributed by Mario Carneiro, 8-Feb-2015.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐾 & ⊢ 𝑌 = ∪ 𝐿 & ⊢ 𝐽 = (∏t‘𝐹) & ⊢ 𝐾 = (∏t‘(𝐹 ↾ 𝐴)) & ⊢ 𝐿 = (∏t‘(𝐹 ↾ 𝐵)) & ⊢ 𝐺 = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ (𝑥 ∪ 𝑦)) & ⊢ (𝜑 → 𝐶 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝐶⟶Top) & ⊢ (𝜑 → 𝐶 = (𝐴 ∪ 𝐵)) & ⊢ (𝜑 → (𝐴 ∩ 𝐵) = ∅) ⇒ ⊢ (𝜑 → ◡𝐺 = (𝑧 ∈ ∪ 𝐽 ↦ 〈(𝑧 ↾ 𝐴), (𝑧 ↾ 𝐵)〉)) | ||
Theorem | ptunhmeo 21659* | Define a homeomorphism from a binary product of indexed product topologies to an indexed product topology on the union of the index sets. This is the topological analogue of (𝐴↑𝐵) · (𝐴↑𝐶) = 𝐴↑(𝐵 + 𝐶). (Contributed by Mario Carneiro, 8-Feb-2015.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐾 & ⊢ 𝑌 = ∪ 𝐿 & ⊢ 𝐽 = (∏t‘𝐹) & ⊢ 𝐾 = (∏t‘(𝐹 ↾ 𝐴)) & ⊢ 𝐿 = (∏t‘(𝐹 ↾ 𝐵)) & ⊢ 𝐺 = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ (𝑥 ∪ 𝑦)) & ⊢ (𝜑 → 𝐶 ∈ 𝑉) & ⊢ (𝜑 → 𝐹:𝐶⟶Top) & ⊢ (𝜑 → 𝐶 = (𝐴 ∪ 𝐵)) & ⊢ (𝜑 → (𝐴 ∩ 𝐵) = ∅) ⇒ ⊢ (𝜑 → 𝐺 ∈ ((𝐾 ×t 𝐿)Homeo𝐽)) | ||
Theorem | xpstopnlem1 21660* | The function 𝐹 used in xpsval 16279 is a homeomorphism from the binary product topology to the indexed product topology. (Contributed by Mario Carneiro, 2-Sep-2015.) |
⊢ 𝐹 = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ ◡({𝑥} +𝑐 {𝑦})) & ⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐾 ∈ (TopOn‘𝑌)) ⇒ ⊢ (𝜑 → 𝐹 ∈ ((𝐽 ×t 𝐾)Homeo(∏t‘◡({𝐽} +𝑐 {𝐾})))) | ||
Theorem | xpstps 21661 | A binary product of topologies is a topological space. (Contributed by Mario Carneiro, 27-Aug-2015.) |
⊢ 𝑇 = (𝑅 ×s 𝑆) ⇒ ⊢ ((𝑅 ∈ TopSp ∧ 𝑆 ∈ TopSp) → 𝑇 ∈ TopSp) | ||
Theorem | xpstopnlem2 21662* | Lemma for xpstopn 21663. (Contributed by Mario Carneiro, 27-Aug-2015.) |
⊢ 𝑇 = (𝑅 ×s 𝑆) & ⊢ 𝐽 = (TopOpen‘𝑅) & ⊢ 𝐾 = (TopOpen‘𝑆) & ⊢ 𝑂 = (TopOpen‘𝑇) & ⊢ 𝑋 = (Base‘𝑅) & ⊢ 𝑌 = (Base‘𝑆) & ⊢ 𝐹 = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ ◡({𝑥} +𝑐 {𝑦})) ⇒ ⊢ ((𝑅 ∈ TopSp ∧ 𝑆 ∈ TopSp) → 𝑂 = (𝐽 ×t 𝐾)) | ||
Theorem | xpstopn 21663 | The topology on a binary product of topological spaces, as we have defined it (transferring the indexed product topology on functions on {∅, 1𝑜} to (𝑋 × 𝑌) by the canonical bijection), coincides with the usual topological product (generated by a base of rectangles). (Contributed by Mario Carneiro, 27-Aug-2015.) |
⊢ 𝑇 = (𝑅 ×s 𝑆) & ⊢ 𝐽 = (TopOpen‘𝑅) & ⊢ 𝐾 = (TopOpen‘𝑆) & ⊢ 𝑂 = (TopOpen‘𝑇) ⇒ ⊢ ((𝑅 ∈ TopSp ∧ 𝑆 ∈ TopSp) → 𝑂 = (𝐽 ×t 𝐾)) | ||
Theorem | ptcmpfi 21664 | A topological product of finitely many compact spaces is compact. This weak version of Tychonoff's theorem does not require the axiom of choice. (Contributed by Mario Carneiro, 8-Feb-2015.) |
⊢ ((𝐴 ∈ Fin ∧ 𝐹:𝐴⟶Comp) → (∏t‘𝐹) ∈ Comp) | ||
Theorem | xkocnv 21665* | The inverse of the "currying" function 𝐹 is the uncurrying function. (Contributed by Mario Carneiro, 13-Apr-2015.) |
⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐾 ∈ (TopOn‘𝑌)) & ⊢ 𝐹 = (𝑓 ∈ ((𝐽 ×t 𝐾) Cn 𝐿) ↦ (𝑥 ∈ 𝑋 ↦ (𝑦 ∈ 𝑌 ↦ (𝑥𝑓𝑦)))) & ⊢ (𝜑 → 𝐽 ∈ 𝑛-Locally Comp) & ⊢ (𝜑 → 𝐾 ∈ 𝑛-Locally Comp) & ⊢ (𝜑 → 𝐿 ∈ Top) ⇒ ⊢ (𝜑 → ◡𝐹 = (𝑔 ∈ (𝐽 Cn (𝐿 ^ko 𝐾)) ↦ (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑌 ↦ ((𝑔‘𝑥)‘𝑦)))) | ||
Theorem | xkohmeo 21666* | The Exponential Law for topological spaces. The "currying" function 𝐹 is a homeomorphism on function spaces when 𝐽 and 𝐾 are exponentiable spaces (by xkococn 21511, it is sufficient to assume that 𝐽, 𝐾 are locally compact to ensure exponentiability). (Contributed by Mario Carneiro, 13-Apr-2015.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐾 ∈ (TopOn‘𝑌)) & ⊢ 𝐹 = (𝑓 ∈ ((𝐽 ×t 𝐾) Cn 𝐿) ↦ (𝑥 ∈ 𝑋 ↦ (𝑦 ∈ 𝑌 ↦ (𝑥𝑓𝑦)))) & ⊢ (𝜑 → 𝐽 ∈ 𝑛-Locally Comp) & ⊢ (𝜑 → 𝐾 ∈ 𝑛-Locally Comp) & ⊢ (𝜑 → 𝐿 ∈ Top) ⇒ ⊢ (𝜑 → 𝐹 ∈ ((𝐿 ^ko (𝐽 ×t 𝐾))Homeo((𝐿 ^ko 𝐾) ^ko 𝐽))) | ||
Theorem | qtopf1 21667 | If a quotient map is injective, then it is a homeomorphism. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐹:𝑋–1-1→𝑌) ⇒ ⊢ (𝜑 → 𝐹 ∈ (𝐽Homeo(𝐽 qTop 𝐹))) | ||
Theorem | qtophmeo 21668* | If two functions on a base topology 𝐽 make the same identifications in order to create quotient spaces 𝐽 qTop 𝐹 and 𝐽 qTop 𝐺, then not only are 𝐽 qTop 𝐹 and 𝐽 qTop 𝐺 homeomorphic, but there is a unique homeomorphism that makes the diagram commute. (Contributed by Mario Carneiro, 24-Mar-2015.) (Proof shortened by Mario Carneiro, 23-Aug-2015.) |
⊢ (𝜑 → 𝐽 ∈ (TopOn‘𝑋)) & ⊢ (𝜑 → 𝐹:𝑋–onto→𝑌) & ⊢ (𝜑 → 𝐺:𝑋–onto→𝑌) & ⊢ ((𝜑 ∧ (𝑥 ∈ 𝑋 ∧ 𝑦 ∈ 𝑋)) → ((𝐹‘𝑥) = (𝐹‘𝑦) ↔ (𝐺‘𝑥) = (𝐺‘𝑦))) ⇒ ⊢ (𝜑 → ∃!𝑓 ∈ ((𝐽 qTop 𝐹)Homeo(𝐽 qTop 𝐺))𝐺 = (𝑓 ∘ 𝐹)) | ||
Theorem | t0kq 21669* | A topological space is T0 iff the quotient map is a homeomorphism onto the space's Kolmogorov quotient. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ 𝐹 = (𝑥 ∈ 𝑋 ↦ {𝑦 ∈ 𝐽 ∣ 𝑥 ∈ 𝑦}) ⇒ ⊢ (𝐽 ∈ (TopOn‘𝑋) → (𝐽 ∈ Kol2 ↔ 𝐹 ∈ (𝐽Homeo(KQ‘𝐽)))) | ||
Theorem | kqhmph 21670 | A topological space is T0 iff it is homeomorphic to its Kolmogorov quotient. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Kol2 ↔ 𝐽 ≃ (KQ‘𝐽)) | ||
Theorem | ist1-5lem 21671 | Lemma for ist1-5 21673 and similar theorems. If 𝐴 is a topological property which implies T0, such as T1 or T2, the property can be "decomposed" into T0 and a non-T0 version of property 𝐴 (which is defined as stating that the Kolmogorov quotient of the space has property 𝐴). For example, if 𝐴 is T1, then the theorem states that a space is T1 iff it is T0 and its Kolmogorov quotient is T1 (we call this property R0). (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ 𝐴 → 𝐽 ∈ Kol2) & ⊢ (𝐽 ≃ (KQ‘𝐽) → (𝐽 ∈ 𝐴 → (KQ‘𝐽) ∈ 𝐴)) & ⊢ ((KQ‘𝐽) ≃ 𝐽 → ((KQ‘𝐽) ∈ 𝐴 → 𝐽 ∈ 𝐴)) ⇒ ⊢ (𝐽 ∈ 𝐴 ↔ (𝐽 ∈ Kol2 ∧ (KQ‘𝐽) ∈ 𝐴)) | ||
Theorem | t1r0 21672 | A T1 space is R0. That is, the Kolmogorov quotient of a T1 space is also T1 (because they are homeomorphic). (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Fre → (KQ‘𝐽) ∈ Fre) | ||
Theorem | ist1-5 21673 | A topological space is T1 iff it is both T0 and R0. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Fre ↔ (𝐽 ∈ Kol2 ∧ (KQ‘𝐽) ∈ Fre)) | ||
Theorem | ishaus3 21674 | A topological space is Hausdorff iff it is both T0 and R1 (where R1 means that any two topologically distinct points are separated by neighborhoods). (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ (𝐽 ∈ Haus ↔ (𝐽 ∈ Kol2 ∧ (KQ‘𝐽) ∈ Haus)) | ||
Theorem | nrmreg 21675 | A normal T1 space is regular Hausdorff. In other words, a T4 space is T3 . One can get away with slightly weaker assumptions; see nrmr0reg 21600. (Contributed by Mario Carneiro, 25-Aug-2015.) |
⊢ ((𝐽 ∈ Nrm ∧ 𝐽 ∈ Fre) → 𝐽 ∈ Reg) | ||
Theorem | reghaus 21676 | A regular T0 space is Hausdorff. In other words, a T3 space is T2 . A regular Hausdorff or T0 space is also known as a T3 space. (Contributed by Mario Carneiro, 24-Aug-2015.) |
⊢ (𝐽 ∈ Reg → (𝐽 ∈ Haus ↔ 𝐽 ∈ Kol2)) | ||
Theorem | nrmhaus 21677 | A T1 normal space is Hausdorff. A Hausdorff or T1 normal space is also known as a T4 space. (Contributed by Mario Carneiro, 24-Aug-2015.) |
⊢ (𝐽 ∈ Nrm → (𝐽 ∈ Haus ↔ 𝐽 ∈ Fre)) | ||
Theorem | elmptrab 21678* | Membership in a one-parameter class of sets. (Contributed by Stefan O'Rear, 28-Jul-2015.) |
⊢ 𝐹 = (𝑥 ∈ 𝐷 ↦ {𝑦 ∈ 𝐵 ∣ 𝜑}) & ⊢ ((𝑥 = 𝑋 ∧ 𝑦 = 𝑌) → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑋 → 𝐵 = 𝐶) & ⊢ (𝑥 ∈ 𝐷 → 𝐵 ∈ 𝑉) ⇒ ⊢ (𝑌 ∈ (𝐹‘𝑋) ↔ (𝑋 ∈ 𝐷 ∧ 𝑌 ∈ 𝐶 ∧ 𝜓)) | ||
Theorem | elmptrab2 21679* | Membership in a one-parameter class of sets, indexed by arbitrary base sets. (Contributed by Stefan O'Rear, 28-Jul-2015.) (Revised by AV, 26-Mar-2021.) |
⊢ 𝐹 = (𝑥 ∈ V ↦ {𝑦 ∈ 𝐵 ∣ 𝜑}) & ⊢ ((𝑥 = 𝑋 ∧ 𝑦 = 𝑌) → (𝜑 ↔ 𝜓)) & ⊢ (𝑥 = 𝑋 → 𝐵 = 𝐶) & ⊢ 𝐵 ∈ V & ⊢ (𝑌 ∈ 𝐶 → 𝑋 ∈ 𝑊) ⇒ ⊢ (𝑌 ∈ (𝐹‘𝑋) ↔ (𝑌 ∈ 𝐶 ∧ 𝜓)) | ||
Theorem | isfbas 21680* | The predicate "𝐹 is a filter base." Note that some authors require filter bases to be closed under pairwise intersections, but that is not necessary under our definition. One advantage of this definition is that tails in a directed set form a filter base under our meaning. (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐵 ∈ 𝐴 → (𝐹 ∈ (fBas‘𝐵) ↔ (𝐹 ⊆ 𝒫 𝐵 ∧ (𝐹 ≠ ∅ ∧ ∅ ∉ 𝐹 ∧ ∀𝑥 ∈ 𝐹 ∀𝑦 ∈ 𝐹 (𝐹 ∩ 𝒫 (𝑥 ∩ 𝑦)) ≠ ∅)))) | ||
Theorem | fbasne0 21681 | There are no empty filter bases. (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐹 ∈ (fBas‘𝐵) → 𝐹 ≠ ∅) | ||
Theorem | 0nelfb 21682 | No filter base contains the empty set. (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐹 ∈ (fBas‘𝐵) → ¬ ∅ ∈ 𝐹) | ||
Theorem | fbsspw 21683 | A filter base on a set is a subset of the power set. (Contributed by Stefan O'Rear, 28-Jul-2015.) |
⊢ (𝐹 ∈ (fBas‘𝐵) → 𝐹 ⊆ 𝒫 𝐵) | ||
Theorem | fbelss 21684 | An element of the filter base is a subset of the base set. (Contributed by Stefan O'Rear, 28-Jul-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝐵) ∧ 𝑋 ∈ 𝐹) → 𝑋 ⊆ 𝐵) | ||
Theorem | fbdmn0 21685 | The domain of a filter base is nonempty. (Contributed by Mario Carneiro, 28-Nov-2013.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ (𝐹 ∈ (fBas‘𝐵) → 𝐵 ≠ ∅) | ||
Theorem | isfbas2 21686* | The predicate "𝐹 is a filter base." (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ (𝐵 ∈ 𝐴 → (𝐹 ∈ (fBas‘𝐵) ↔ (𝐹 ⊆ 𝒫 𝐵 ∧ (𝐹 ≠ ∅ ∧ ∅ ∉ 𝐹 ∧ ∀𝑥 ∈ 𝐹 ∀𝑦 ∈ 𝐹 ∃𝑧 ∈ 𝐹 𝑧 ⊆ (𝑥 ∩ 𝑦))))) | ||
Theorem | fbasssin 21687* | A filter base contains subsets of its pairwise intersections. (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Jeff Hankins, 1-Dec-2010.) |
⊢ ((𝐹 ∈ (fBas‘𝑋) ∧ 𝐴 ∈ 𝐹 ∧ 𝐵 ∈ 𝐹) → ∃𝑥 ∈ 𝐹 𝑥 ⊆ (𝐴 ∩ 𝐵)) | ||
Theorem | fbssfi 21688* | A filter base contains subsets of its finite intersections. (Contributed by Mario Carneiro, 26-Nov-2013.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝑋) ∧ 𝐴 ∈ (fi‘𝐹)) → ∃𝑥 ∈ 𝐹 𝑥 ⊆ 𝐴) | ||
Theorem | fbssint 21689* | A filter base contains subsets of its finite intersections. (Contributed by Jeff Hankins, 1-Sep-2009.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝐵) ∧ 𝐴 ⊆ 𝐹 ∧ 𝐴 ∈ Fin) → ∃𝑥 ∈ 𝐹 𝑥 ⊆ ∩ 𝐴) | ||
Theorem | fbncp 21690 | A filter base does not contain complements of its elements. (Contributed by Mario Carneiro, 26-Nov-2013.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝑋) ∧ 𝐴 ∈ 𝐹) → ¬ (𝐵 ∖ 𝐴) ∈ 𝐹) | ||
Theorem | fbun 21691* | A necessary and sufficient condition for the union of two filter bases to also be a filter base. (Contributed by Mario Carneiro, 28-Nov-2013.) (Revised by Stefan O'Rear, 2-Aug-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝑋) ∧ 𝐺 ∈ (fBas‘𝑋)) → ((𝐹 ∪ 𝐺) ∈ (fBas‘𝑋) ↔ ∀𝑥 ∈ 𝐹 ∀𝑦 ∈ 𝐺 ∃𝑧 ∈ (𝐹 ∪ 𝐺)𝑧 ⊆ (𝑥 ∩ 𝑦))) | ||
Theorem | fbfinnfr 21692 | No filter base containing a finite element is free. (Contributed by Jeff Hankins, 5-Dec-2009.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝐵) ∧ 𝑆 ∈ 𝐹 ∧ 𝑆 ∈ Fin) → ∩ 𝐹 ≠ ∅) | ||
Theorem | opnfbas 21693* | The collection of open supersets of a nonempty set in a topology is a neighborhoods of the set, one of the motivations for the filter concept. (Contributed by Jeff Hankins, 2-Sep-2009.) (Revised by Mario Carneiro, 7-Aug-2015.) |
⊢ 𝑋 = ∪ 𝐽 ⇒ ⊢ ((𝐽 ∈ Top ∧ 𝑆 ⊆ 𝑋 ∧ 𝑆 ≠ ∅) → {𝑥 ∈ 𝐽 ∣ 𝑆 ⊆ 𝑥} ∈ (fBas‘𝑋)) | ||
Theorem | trfbas2 21694 | Conditions for the trace of a filter base 𝐹 to be a filter base. (Contributed by Mario Carneiro, 13-Oct-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝑌) ∧ 𝐴 ⊆ 𝑌) → ((𝐹 ↾t 𝐴) ∈ (fBas‘𝐴) ↔ ¬ ∅ ∈ (𝐹 ↾t 𝐴))) | ||
Theorem | trfbas 21695* | Conditions for the trace of a filter base 𝐹 to be a filter base. (Contributed by Mario Carneiro, 13-Oct-2015.) |
⊢ ((𝐹 ∈ (fBas‘𝑌) ∧ 𝐴 ⊆ 𝑌) → ((𝐹 ↾t 𝐴) ∈ (fBas‘𝐴) ↔ ∀𝑣 ∈ 𝐹 (𝑣 ∩ 𝐴) ≠ ∅)) | ||
Syntax | cfil 21696 | Extend class notation with the set of filters on a set. |
class Fil | ||
Definition | df-fil 21697* | The set of filters on a set. Definition 1 (axioms FI, FIIa, FIIb, FIII) of [BourbakiTop1] p. I.36. Filters are used to define the concept of limit in the general case. They are a generalization of the idea of neighborhoods. Suppose you are in ℝ. With neighborhoods you can express the idea of a variable that tends to a specific number but you can't express the idea of a variable that tends to infinity. Filters relax the "axioms" of neighborhoods and then succeed in expressing the idea of something that tends to infinity. Filters were invented by Cartan in 1937 and made famous by Bourbaki in his treatise. A notion similar to the notion of filter is the concept of net invented by Moore and Smith in 1922. (Contributed by FL, 20-Jul-2007.) (Revised by Stefan O'Rear, 28-Jul-2015.) |
⊢ Fil = (𝑧 ∈ V ↦ {𝑓 ∈ (fBas‘𝑧) ∣ ∀𝑥 ∈ 𝒫 𝑧((𝑓 ∩ 𝒫 𝑥) ≠ ∅ → 𝑥 ∈ 𝑓)}) | ||
Theorem | isfil 21698* | The predicate "is a filter." (Contributed by FL, 20-Jul-2007.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐹 ∈ (Fil‘𝑋) ↔ (𝐹 ∈ (fBas‘𝑋) ∧ ∀𝑥 ∈ 𝒫 𝑋((𝐹 ∩ 𝒫 𝑥) ≠ ∅ → 𝑥 ∈ 𝐹))) | ||
Theorem | filfbas 21699 | A filter is a filter base. (Contributed by Jeff Hankins, 2-Sep-2009.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐹 ∈ (Fil‘𝑋) → 𝐹 ∈ (fBas‘𝑋)) | ||
Theorem | 0nelfil 21700 | The empty set doesn't belong to a filter. (Contributed by FL, 20-Jul-2007.) (Revised by Mario Carneiro, 28-Jul-2015.) |
⊢ (𝐹 ∈ (Fil‘𝑋) → ¬ ∅ ∈ 𝐹) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |