Computational realization of topological spaces (experimental) #
This file provides infrastructure to compute with topological spaces.
Main declarations #
Ctop
: Realization of a topology basis.Ctop.Realizer
: Realization of a topological space.Ctop
that generates the given topology.LocallyFinite.Realizer
: Realization of the local finiteness of an indexed family of sets.Compact.Realizer
: Realization of the compactness of a set.
theorem
Ctop.coe_mk
{α : Type u_1}
{σ : Type u_3}
(f : σ → Set α)
(T : α → σ)
(h₁ : ∀ (x : α), x ∈ f (T x))
(I : (a b : σ) → (x : α) → x ∈ f a ∩ f b → σ)
(h₂ : ∀ (a b : σ) (x : α) (h : x ∈ f a ∩ f b), x ∈ f (I a b x h))
(h₃ : ∀ (a b : σ) (x : α) (h : x ∈ f a ∩ f b), f (I a b x h) ⊆ f a ∩ f b)
(a : σ)
:
A Ctop
realizes the topological space it generates.
Equations
- F.toRealizer = { σ := σ, F := F, eq := ⋯ }
Instances For
Equations
- instInhabitedRealizer F = { default := F.toRealizer }
The topological space realizer made of the open sets.
Equations
- One or more equations did not get rendered due to their size.
Instances For
def
Ctop.Realizer.ofEquiv
{α : Type u_1}
{τ : Type u_4}
[TopologicalSpace α]
(F : Realizer α)
(E : F.σ ≃ τ)
:
Realizer α
Replace the representation type of a Ctop
realizer.
Equations
- F.ofEquiv E = { σ := τ, F := Ctop.ofEquiv E F.F, eq := ⋯ }
Instances For
A realizer of the neighborhood of a point.
Equations
- One or more equations did not get rendered due to their size.
Instances For
structure
LocallyFinite.Realizer
{α : Type u_1}
{β : Type u_2}
[TopologicalSpace α]
(F : Ctop.Realizer α)
(f : β → Set α)
:
Type (max (max u_1 u_2) u_5)
A LocallyFinite.Realizer F f
is a realization that f
is locally finite, namely it is a
choice of open sets from the basis of F
such that they intersect only finitely many of the values
of f
.
Instances For
theorem
LocallyFinite.Realizer.to_locallyFinite
{α : Type u_1}
{β : Type u_2}
[TopologicalSpace α]
{F : Ctop.Realizer α}
{f : β → Set α}
(R : Realizer F f)
:
theorem
locallyFinite_iff_exists_realizer
{α : Type u_1}
{β : Type u_2}
[TopologicalSpace α]
(F : Ctop.Realizer α)
{f : β → Set α}
:
instance
instNonemptyRealizerOfFinite
{α : Type u_1}
{β : Type u_2}
[TopologicalSpace α]
[Finite β]
(F : Ctop.Realizer α)
(f : β → Set α)
:
def
Compact.Realizer
{α : Type u_1}
[TopologicalSpace α]
(s : Set α)
:
Type (max (max (max u_1 (u_5 + 1)) u_5) u_1)
A Compact.Realizer s
is a realization that s
is compact, namely it is a
choice of finite open covers for each set family covering s
.