Ayberk Tosun, 19 August 2023
The module contains the definitions of
(1) a compact open of a locale, and
(2) a compact locale.
These used to live in the `CompactRegular` module which is now deprecated and
will be broken down into smaller modules.
\begin{code}[hide]
{-# OPTIONS --safe --without-K #-}
open import MLTT.Spartan hiding (J)
open import UF.Base
open import UF.Classifiers
open import UF.FunExt
open import UF.PropTrunc
open import UF.Sets
open import UF.Subsingletons
open import UF.Subsingletons-FunExt
open import UF.Subsingletons-Properties
open import UF.SubtypeClassifier
module Locales.Compactness.Definition
(pt : propositional-truncations-exist)
(fe : Fun-Ext)
where
open import Fin.Kuratowski pt
open import Fin.Type
open import Locales.Frame pt fe
open import Locales.WayBelowRelation.Definition pt fe
open import MLTT.List using (member; []; _β·_; List; in-head; in-tail; length)
open import Slice.Family
open import Taboos.FiniteSubsetTaboo pt fe
open import UF.Equiv hiding (_β )
open import UF.EquivalenceExamples
open import UF.ImageAndSurjection pt
open import UF.Logic
open import UF.Powerset-Fin pt
open import UF.Powerset-MultiUniverse
open import UF.Sets-Properties
open PropositionalTruncation pt
open AllCombinators pt fe
open Locale
\end{code}
An open `U : πͺ(X)` of a locale `X` is called compact if it is way below itself.
\begin{code}
is-compact-open : (X : Locale π€ π₯ π¦) β β¨ πͺ X β© β Ξ© (π€ β π₯ β π¦ βΊ)
is-compact-open X U = U βͺ[ πͺ X ] U
\end{code}
A locale `X` is called compact if its top element `π` is compact.
\begin{code}
is-compact : Locale π€ π₯ π¦ β Ξ© (π€ β π₯ β π¦ βΊ)
is-compact X = is-compact-open X π[ πͺ X ]
\end{code}
We also define the type `π¦ X` expressing the type of compact opens of a locale
`X`.
\begin{code}
π¦ : Locale π€ π₯ π¦ β π€ β π₯ β π¦ βΊ Μ
π¦ X = Ξ£ U κ β¨ πͺ X β© , is-compact-open X U holds
π¦-is-set : (X : Locale π€ π₯ π¦) β is-set (π¦ X)
π¦-is-set X {(Kβ , ΞΊβ)} {(Kβ , ΞΊβ)} =
Ξ£-is-set
carrier-of-[ poset-of (πͺ X) ]-is-set
(Ξ» U β props-are-sets (holds-is-prop (is-compact-open X U)))
to-π¦-οΌ : (X : Locale π€ π₯ π¦) {Kβ Kβ : β¨ πͺ X β©}
β (ΞΊβ : is-compact-open X Kβ holds)
β (ΞΊβ : is-compact-open X Kβ holds)
β Kβ οΌ Kβ
β (Kβ , ΞΊβ) οΌ (Kβ , ΞΊβ)
to-π¦-οΌ X ΞΊβ ΞΊβ = to-subtype-οΌ (holds-is-prop β is-compact-open X)
\end{code}
Using this, we could define a family giving the compact opens of a locale `X`:
\begin{code}
β¬-compact : (X : Locale π€ π₯ π¦) β Fam (π€ β π₯ β π¦ βΊ) β¨ πͺ X β©
β¬-compact X = π¦ X , prβ
\end{code}
but the index of this family lives in `π€ β π₯ β π¦ βΊ`. This is to say that, if one
starts with a large and locally small locale, the resulting family would live in
`π€ βΊ` which means it would be *too big*.
\begin{code}
β¬-compactβ : (X : Locale (π€ βΊ) π€ π€) β Fam (π€ βΊ) β¨ πͺ X β©
β¬-compactβ = β¬-compact
\end{code}
\section{Properties of compactness}
\begin{code}
π-is-compact : (X : Locale π€ π₯ π¦) β is-compact-open X π[ πͺ X ] holds
π-is-compact X S (β£iβ£ , _) p = β₯β₯-rec β-is-prop β β£iβ£
where
β : index S β β i κ index S , (π[ πͺ X ] β€[ poset-of (πͺ X) ] S [ i ]) holds
β i = β£ i , π-is-bottom (πͺ X) (S [ i ]) β£
πβ[_] : (X : Locale π€ π₯ π¦) β π¦ X
πβ[_] X = π[ πͺ X ] , π-is-compact X
\end{code}
The binary join of two compact opens is compact.
\begin{code}
compact-opens-are-closed-under-β¨ : (X : Locale π€ π₯ π¦) (Kβ Kβ : β¨ πͺ X β©)
β is-compact-open X Kβ holds
β is-compact-open X Kβ holds
β is-compact-open X (Kβ β¨[ πͺ X ] Kβ) holds
compact-opens-are-closed-under-β¨ X U V ΞΊβ ΞΊβ S Ξ΄ p =
β₯β₯-recβ β-is-prop β (ΞΊβ S Ξ΄ Ο) (ΞΊβ S Ξ΄ Ο)
where
open PosetNotation (poset-of (πͺ X)) using (_β€_)
open PosetReasoning (poset-of (πͺ X))
β : Ξ£ iβ κ index S , (U β€[ poset-of (πͺ X) ] S [ iβ ]) holds
β Ξ£ iβ κ index S , (V β€[ poset-of (πͺ X) ] S [ iβ ]) holds
β β iβ κ index S , ((U β¨[ πͺ X ] V) β€ S [ iβ ]) holds
β (iβ , pβ) (iβ , pβ) = β₯β₯-rec β-is-prop β‘ (prβ Ξ΄ iβ iβ)
where
β‘ : Ξ£ iβ κ index S , (S [ iβ ] β€ S [ iβ ]) holds
Γ (S [ iβ ] β€ S [ iβ ]) holds
β β iβ κ index S , ((U β¨[ πͺ X ] V) β€ S [ iβ ]) holds
β‘ (iβ , q , r) = β£ iβ , β¨[ πͺ X ]-least β β£ β£
where
β : (U β€[ poset-of (πͺ X) ] (S [ iβ ])) holds
β = U β€β¨ pβ β© S [ iβ ] β€β¨ q β© S [ iβ ] β
β£ : (V β€[ poset-of (πͺ X) ] (S [ iβ ])) holds
β£ = V β€β¨ pβ β© S [ iβ ] β€β¨ r β© S [ iβ ] β
Ο : (U β€ (β[ πͺ X ] S)) holds
Ο = U β€β¨ β¨[ πͺ X ]-upperβ U V β© U β¨[ πͺ X ] V β€β¨ p β© β[ πͺ X ] S β
Ο : (V β€[ poset-of (πͺ X) ] (β[ πͺ X ] S)) holds
Ο = V β€β¨ β¨[ πͺ X ]-upperβ U V β© U β¨[ πͺ X ] V β€β¨ p β© β[ πͺ X ] S β
\end{code}