Martin Escardo, January 2018 Two weak notions of compactness: β-compactness and Ξ -compactness. See the module CompactTypes for the strong notion. \begin{code} {-# OPTIONS --without-K --exact-split --safe --auto-inline #-} open import MLTT.Spartan open import TypeTopology.CompactTypes open import TypeTopology.TotallySeparated open import CoNaturals.GenericConvergentSequence open import MLTT.Two-Properties open import MLTT.Plus-Properties open import TypeTopology.DisconnectedTypes open import TypeTopology.DiscreteAndSeparated open import Taboos.WLPO open import Notation.Order open import UF.Base open import UF.Subsingletons open import UF.Subsingletons-FunExt open import UF.FunExt open import UF.PropTrunc open import UF.Retracts open import UF.Retracts-FunExt open import UF.Equiv open import UF.Miscelanea module TypeTopology.WeaklyCompactTypes (fe : FunExt) (pt : propositional-truncations-exist) where open PropositionalTruncation pt open import NotionsOfDecidability.Decidable open import NotionsOfDecidability.Complemented β-compact : π€ Μ β π€ Μ β-compact X = (p : X β π) β decidable (β x κ X , p x οΌ β) β-compactness-is-prop : {X : π€ Μ } β is-prop (β-compact X) β-compactness-is-prop {π€} {X} = Ξ -is-prop (fe π€ π€) (Ξ» _ β decidability-of-prop-is-prop (fe π€ π€β) β₯β₯-is-prop) β-compactness-gives-Markov : {X : π€ Μ } β β-compact X β (p : X β π) β ¬¬ (β x κ X , p x οΌ β) β β x κ X , p x οΌ β β-compactness-gives-Markov {π€} {X} c p Ο = g (c p) where g : decidable (β x κ X , p x οΌ β) β β x κ X , p x οΌ β g (inl e) = e g (inr u) = π-elim (Ο u) \end{code} The relation of β-compactness with compactness is the same as that of LPO with WLPO. \begin{code} Ξ -compact : π€ Μ β π€ Μ Ξ -compact X = (p : X β π) β decidable ((x : X) β p x οΌ β) Ξ -compactness-is-prop : {X : π€ Μ } β is-prop (Ξ -compact X) Ξ -compactness-is-prop {π€} = Ξ -is-prop (fe π€ π€) (Ξ» _ β decidability-of-prop-is-prop (fe π€ π€β) (Ξ -is-prop (fe π€ π€β) Ξ» _ β π-is-set)) β-compact-gives-Ξ -compact : {X : π€ Μ } β β-compact X β Ξ -compact X β-compact-gives-Ξ -compact {π€} {X} c p = f (c p) where f : decidable (β x κ X , p x οΌ β) β decidable (Ξ x κ X , p x οΌ β) f (inl s) = inr (Ξ» Ξ± β β₯β₯-rec π-is-prop (g Ξ±) s) where g : ((x : X) β p x οΌ β) β Β¬ (Ξ£ x κ X , p x οΌ β) g Ξ± (x , r) = zero-is-not-one (r β»ΒΉ β Ξ± x) f (inr u) = inl (not-existsβ-implies-forallβ p u) empty-types-are-β-compact : {X : π€ Μ } β is-empty X β β-compact X empty-types-are-β-compact u p = inr (β₯β₯-rec π-is-prop Ξ» Ο β u (prβ Ο)) empty-types-are-Ξ -compact : {X : π€ Μ } β is-empty X β Ξ -compact X empty-types-are-Ξ -compact u p = inl (Ξ» x β π-elim (u x)) \end{code} The β-compactness, and hence Ξ -compactness, of compact sets (and hence of ββ, for example): \begin{code} compact-types-are-β-compact : {X : π€ Μ } β compact X β β-compact X compact-types-are-β-compact {π€} {X} Ο p = g (Ο p) where g : ((Ξ£ x κ X , p x οΌ β) + ((x : X) β p x οΌ β)) β decidable (β x κ X , p x οΌ β) g (inl (x , r)) = inl β£ x , r β£ g (inr Ξ±) = inr (forallβ-implies-not-existsβ p Ξ±) β₯Compactβ₯-types-are-β-compact : {X : π€ Μ } β β₯ Compact X β₯ β β-compact X β₯Compactβ₯-types-are-β-compact {π€} {X} = β₯β₯-rec β-compactness-is-prop (compact-types-are-β-compact β Compact-gives-compact) \end{code} But notice that the Ξ -compactness of β is WLPO and its β-compactness is amounts to LPO. The Ξ -compactness of X is equivalent to the isolatedness of the boolean predicate Ξ» x β β: \begin{code} Ξ -compact' : π€ Μ β π€ Μ Ξ -compact' X = is-isolated' (Ξ» (x : X) β β) Ξ -compactness'-is-prop : {X : π€ Μ } β is-prop (Ξ -compact' X) Ξ -compactness'-is-prop {π€} = being-isolated'-is-prop fe (Ξ» x β β) Ξ -compact'-gives-Ξ -compact : {X : π€ Μ } β Ξ -compact' X β Ξ -compact X Ξ -compact'-gives-Ξ -compact {π€} {X} c' p = g (c' p) where g : decidable (p οΌ Ξ» x β β) β decidable ((x : X) β p x οΌ β) g (inl r) = inl (happly r) g (inr u) = inr (contrapositive (dfunext (fe π€ π€β)) u) Ξ -compact-gives-Ξ -compact' : {X : π€ Μ } β Ξ -compact X β Ξ -compact' X Ξ -compact-gives-Ξ -compact' {π€} {X} c p = g (c p) where g : decidable ((x : X) β p x οΌ β) β decidable (p οΌ Ξ» x β β) g (inl Ξ±) = inl (dfunext (fe π€ π€β) Ξ±) g (inr u) = inr (contrapositive happly u) \end{code} In classical topology, the Tychonoff Theorem gives that compact to the power discrete is compact (where we read the function type X β Y as "Y to the power X", with Y the base and X the exponent, and call it an exponential). Here we don't have the Tychonoff Theorem (in the absence of anti-classical intuitionistic assumptions). It is less well-known that in classical topology we also have that discrete to the power compact is discrete. This we do have here, without the need of any assumption: \begin{code} discrete-to-the-power-Ξ -compact-is-discrete : {X : π€ Μ } {Y : π₯ Μ } β Ξ -compact X β is-discrete Y β is-discrete (X β Y) discrete-to-the-power-Ξ -compact-is-discrete {π€} {π₯} {X} {Y} c d f g = Ξ΄ where p : X β π p = prβ (co-characteristic-function (Ξ» x β d (f x) (g x))) r : (x : X) β (p x οΌ β β Β¬ (f x οΌ g x)) Γ (p x οΌ β β f x οΌ g x) r = prβ (co-characteristic-function Ξ» x β d (f x) (g x)) Ο : ((x : X) β p x οΌ β) β f οΌ g Ο Ξ± = (dfunext (fe π€ π₯) (Ξ» x β prβ (r x) (Ξ± x))) Ξ³ : f οΌ g β (x : X) β p x οΌ β Ξ³ t x = different-from-β-equal-β (Ξ» u β prβ (r x) u (happly t x)) h : decidable ((x : X) β p x οΌ β) β decidable (f οΌ g) h (inl Ξ±) = inl (Ο Ξ±) h (inr u) = inr (contrapositive Ξ³ u) Ξ΄ : decidable (f οΌ g) Ξ΄ = h (c p) \end{code} If an exponential with discrete base is discrete, then its exponent is compact, provided the base has at least two points. First, to decide Ξ (p : X β π), p x οΌ 1, decide p οΌ Ξ» x β β: \begin{code} power-of-two-discrete-gives-compact-exponent : {X : π€ Μ } β is-discrete (X β π) β Ξ -compact X power-of-two-discrete-gives-compact-exponent d = Ξ -compact'-gives-Ξ -compact (Ξ» p β d p (Ξ» x β β)) discrete-power-of-disconnected-gives-compact-exponent : {X : π€ Μ } {Y : π₯ Μ } β disconnected Y β is-discrete (X β Y) β Ξ -compact X discrete-power-of-disconnected-gives-compact-exponent {π€} {π₯} {X} {Y} Ο d = Ξ³ where a : retract (X β π) of (X β Y) a = retract-contravariance (fe π€ π€β) Ο b : is-discrete (X β π) b = retract-is-discrete a d Ξ³ : Ξ -compact X Ξ³ = power-of-two-discrete-gives-compact-exponent b discrete-power-of-non-trivial-discrete-gives-compact-exponent' : {X : π€ Μ } {Y : π₯ Μ } β (Ξ£ yβ κ Y , Ξ£ yβ κ Y , yβ β yβ) β is-discrete Y β is-discrete (X β Y) β Ξ -compact X discrete-power-of-non-trivial-discrete-gives-compact-exponent' w d = discrete-power-of-disconnected-gives-compact-exponent (discrete-type-with-two-different-points-gives-disconnected w d) \end{code} So, in summary, if Y is a non-trivial discrete type, then X is Ξ -compact iff (X β Y) is discrete. Compactness of images: \begin{code} open import UF.ImageAndSurjection pt surjection-β-compact : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-surjection f β β-compact X β β-compact Y surjection-β-compact {π€} {π₯} {X} {Y} f su c q = g (c (q β f)) where h : (Ξ£ x κ X , q (f x) οΌ β) β Ξ£ y κ Y , q y οΌ β h (x , r) = (f x , r) l : (y : Y) β q y οΌ β β (Ξ£ x κ X , f x οΌ y) β Ξ£ x κ X , q (f x) οΌ β l y r (x , s) = (x , (ap q s β r)) k : (Ξ£ y κ Y , q y οΌ β) β β x κ X , q (f x) οΌ β k (y , r) = β₯β₯-functor (l y r) (su y) g : decidable (β x κ X , q (f x) οΌ β)Β β decidable (β y κ Y , q y οΌ β) g (inl s) = inl (β₯β₯-functor h s) g (inr u) = inr (contrapositive (β₯β₯-rec β₯β₯-is-prop k) u) image-β-compact : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β β-compact X β β-compact (image f) image-β-compact f = surjection-β-compact (corestriction f) (corestrictions-are-surjections f) surjection-Ξ -compact : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β is-surjection f β Ξ -compact X β Ξ -compact Y surjection-Ξ -compact {π€} {π₯} {X} {Y} f su c q = g (c (q β f)) where g : decidable ((x : X) β q (f x) οΌ β) β decidable ((x : Y) β q x οΌ β) g (inl s) = inl (surjection-induction f su (Ξ» y β q y οΌ β) (Ξ» _ β π-is-set) s) g (inr u) = inr (contrapositive (Ξ» Ο x β Ο (f x)) u) retract-β-compact : {X : π€ Μ } {Y : π₯ Μ } β retract Y of X β β-compact X β β-compact Y retract-β-compact (f , hass) = surjection-β-compact f (retractions-are-surjections f hass) retract-β-compact' : {X : π€ Μ } {Y : π₯ Μ } β β₯ retract Y of X β₯ β β-compact X β β-compact Y retract-β-compact' t c = β₯β₯-rec β-compactness-is-prop (Ξ» r β retract-β-compact r c) t image-Ξ -compact : {X : π€ Μ } {Y : π₯ Μ } (f : X β Y) β Ξ -compact X β Ξ -compact (image f) image-Ξ -compact f = surjection-Ξ -compact (corestriction f) (corestrictions-are-surjections f) retract-Ξ -compact : {X : π€ Μ } {Y : π₯ Μ } β retract Y of X β Ξ -compact X β Ξ -compact Y retract-Ξ -compact (f , hass) = surjection-Ξ -compact f (retractions-are-surjections f hass) retract-Ξ -compact' : {X : π€ Μ } {Y : π₯ Μ } β β₯ retract Y of X β₯ β Ξ -compact X β Ξ -compact Y retract-Ξ -compact' t c = β₯β₯-rec Ξ -compactness-is-prop (Ξ» r β retract-Ξ -compact r c) t Ξ -compact-exponential-with-pointed-domain-has-Ξ -compact-domain : {X : π€ Μ } {Y : π₯ Μ } β X β Ξ -compact (X β Y) β Ξ -compact Y Ξ -compact-exponential-with-pointed-domain-has-Ξ -compact-domain x = retract-Ξ -compact (codomain-is-retract-of-function-space-with-pointed-domain x) \end{code} A main reason to consider the notion of total separatedness is that the totally separated reflection π X of X has the same supply of boolean predicates as X, and hence X is β-compact (respectively Ξ -compact) iff π X is, as we show now. \begin{code} module _ (X : π€ Μ ) where open TotallySeparatedReflection fe pt private extension : (X β π) β (π X β π) extension p = prβ (prβ (totally-separated-reflection π-is-totally-separated p)) extension-property : (p : X β π) (x : X) β extension p (Ξ· x) οΌ p x extension-property p = happly (prβ (prβ (totally-separated-reflection π-is-totally-separated p))) β-compact-gives-β-compact-π : β-compact X β β-compact (π X) β-compact-gives-β-compact-π = surjection-β-compact Ξ· Ξ·-is-surjection β-compact-π-gives-β-compact : β-compact (π X) β β-compact X β-compact-π-gives-β-compact c p = h (c (extension p)) where f : (Ξ£ x' κ π X , extension p x' οΌ β) β β x κ X , p x οΌ β f (x' , r) = β₯β₯-functor f' (Ξ·-is-surjection x') where f' : (Ξ£ x κ X , Ξ· x οΌ x') β Ξ£ x κ X , p x οΌ β f' (x , s) = x , ((extension-property p x) β»ΒΉ β ap (extension p) s β r) g : (Ξ£ x κ X , p x οΌ β) β Ξ£ x' κ π X , extension p x' οΌ β g (x , r) = Ξ· x , (extension-property p x β r) h : decidable (β x' κ π X , extension p x' οΌ β) β decidable (β x κ X , p x οΌ β) h (inl x) = inl (β₯β₯-rec β₯β₯-is-prop f x) h (inr u) = inr (contrapositive (β₯β₯-functor g) u) Ξ -compact-gives-Ξ -compact-π : Ξ -compact X β Ξ -compact (π X) Ξ -compact-gives-Ξ -compact-π = surjection-Ξ -compact Ξ· (Ξ·-is-surjection) Ξ -compact-π-gives-Ξ -compact : Ξ -compact (π X) β Ξ -compact X Ξ -compact-π-gives-Ξ -compact c p = h (c (extension p)) where f : ((x' : π X) β extension p x' οΌ β) β ((x : X) β p x οΌ β) f Ξ± x = (extension-property p x)β»ΒΉ β Ξ± (Ξ· x) g : (Ξ± : (x : X) β p x οΌ β) β ((x' : π X) β extension p x' οΌ β) g Ξ± = Ξ·-induction (Ξ» x' β extension p x' οΌ β) (Ξ» _ β π-is-set) g' where g' : (x : X) β extension p (Ξ· x) οΌ β g' x = extension-property p x β Ξ± x h : decidable ((x' : π X) β extension p x' οΌ β) β decidable ((x : X) β p x οΌ β) h (inl Ξ±) = inl (f Ξ±) h (inr u) = inr (contrapositive g u) \end{code} If X is totally separated, and (X β π) is compact, then X is discrete. More generally, if π is a retract of Y and (X β Y) is compact, then X is discrete if it is totally separated. This is a new result as far as I know. I didn't know it before 12th January 2018. The following proof works as follows. For any given x,y:X, define q:(Xβπ)βπ such that q(p)=1 β p(x) = p(y), which is possible because π has decidable equality (it is discrete). By the Ξ -compactness of Xβπ, the condition (p:Xβπ) β q(p)=1 is decidable, which amounts to saying that (p:Xβπ) β p (x)=p (y) is decidable. But because X is totally separated, the latter is equivalent to x=y, which shows that X is discrete. \begin{code} tscd : {X : π€ Μ } β is-totally-separated X β Ξ -compact (X β π) β is-discrete X tscd {π€} {X} ts c x y = g (a s) where q : (X β π) β π q = prβ (co-characteristic-function (Ξ» p β π-is-discrete (p x) (p y))) r : (p : X β π) β (q p οΌ β β p x β p y) Γ (q p οΌ β β p x οΌ p y) r = prβ (co-characteristic-function (Ξ» p β π-is-discrete (p x) (p y))) s : decidable ((p : X β π) β q p οΌ β) s = c q b : (p : X β π) β p x οΌ p y β q p οΌ β b p u = different-from-β-equal-β (Ξ» v β prβ (r p) v u) a : decidable ((p : X β π) β q p οΌ β) β decidable ((p : X β π) β p x οΌ p y) a (inl f) = inl (Ξ» p β prβ (r p) (f p)) a (inr Ο) = inr h where h : Β¬ ((p : X β π) β p x οΌ p y) h Ξ± = Ο (Ξ» p β b p (Ξ± p)) g : decidable ((p : X β π) β p x οΌ p y) β decidable (x οΌ y) g (inl Ξ±) = inl (ts Ξ±) g (inr u) = inr (contrapositive (Ξ» e p β ap p e) u) \end{code} We are interested in the following two generalizations, which arise as corollaries: \begin{code} tscdβ : {X : π€ Μ } {Y : π₯ Μ } β is-totally-separated X β disconnected Y β Ξ -compact (X β Y) β is-discrete X tscdβ {π€} {π₯} {X} {Y} ts r c = tscd ts (retract-Ξ -compact (retract-contravariance (fe π€ π€β) r) c) open TotallySeparatedReflection fe pt tscdβ : {X : π€ Μ } {Y : π₯ Μ } β disconnected Y β Ξ -compact (X β Y) β is-discrete (π X) tscdβ {π€} {π₯} {X} {Y} r c = f where z : retract (X β π) of (X β Y) z = retract-contravariance (fe π€ π€β) r a : (π X β π) β (X β π) a = totally-separated-reflection'' π-is-totally-separated b : retract (π X β π) of (X β π) b = β-gives-β a d : retract (π X β π) of (X β Y) d = retracts-compose z b e : Ξ -compact (π X β π) e = retract-Ξ -compact d c f : is-discrete (π X) f = tscd Ο e \end{code} In topological models, Ξ -compactness is the same as topological compactess in the presence of total separatedness, at least for some spaces, including the Kleene-Kreisel spaces, which model the simple types (see the module SimpleTypes). Hence, for example, the topological space (βββπ) is not Ξ -compact because it is countably discrete, as it is a theorem of topology that discrete to the power compact is again discrete, which is compact iff it is finite. This argument is both classical and external. But here we have that the type (βββπ) is "not" Ξ -compact, internally and constructively. \begin{code} [βββπ]-compact-implies-WLPO : Ξ -compact (ββ β π) β WLPO [βββπ]-compact-implies-WLPO c = ββ-discrete-gives-WLPO (tscd (ββ-is-totally-separated (fe π€β π€β)) c) \end{code} Closure of compactness under sums (and hence binary products): \begin{code} Ξ -compact-closed-under-Ξ£ : {X : π€ Μ } {Y : X β π₯ Μ } β Ξ -compact X β ((x : X) β Ξ -compact (Y x)) β Ξ -compact (Ξ£ Y) Ξ -compact-closed-under-Ξ£ {π€} {π₯} {X} {Y} c d p = g e where f : β x β decidable (β y β p (x , y) οΌ β) f x = d x (Ξ» y β p (x , y)) q : X β π q = prβ (co-characteristic-function f) qβ : (x : X) β q x οΌ β β Β¬ ((y : Y x) β p (x , y) οΌ β) qβ x = prβ (prβ (co-characteristic-function f) x) qβ : (x : X) β q x οΌ β β (y : Y x) β p (x , y) οΌ β qβ x = prβ (prβ (co-characteristic-function f) x) e : decidable (β x β q x οΌ β) e = c q g : decidable (β x β q x οΌ β) β decidable (β Ο β p Ο οΌ β) g (inl Ξ±) = inl h where h : (Ο : Ξ£ Y) β p Ο οΌ β h (x , y) = qβ x (Ξ± x) y g (inr u) = inr (contrapositive h u) where h : ((Ο : Ξ£ Y) β p Ο οΌ β) β (x : X) β q x οΌ β h Ξ² x = different-from-β-equal-β (Ξ» r β qβ x r (Ξ» y β Ξ² (x , y))) \end{code} TODO. Consider also other possible closure properties, and β-compactness. We now turn to propositions. A proposition is β-compact iff it is decidable. Regarding the compactness of propositions, we have partial information for the moment. \begin{code} β-compact-propositions-are-decidable : (X : π€ Μ ) β is-prop X β β-compact X β decidable X β-compact-propositions-are-decidable X isp c = f a where a : decidable β₯ X Γ (β οΌ β) β₯ a = c (Ξ» x β β) f : decidable β₯ X Γ (β οΌ β) β₯ β decidable X f (inl s) = inl (β₯β₯-rec isp prβ s) f (inr u) = inr (Ξ» x β u β£ x , refl β£) β-compact-types-have-decidable-support : {X : π€ Μ } β β-compact X β decidable β₯ X β₯ β-compact-types-have-decidable-support {π€} {X} c = β-compact-propositions-are-decidable β₯ X β₯ β₯β₯-is-prop (surjection-β-compact β£_β£ pt-is-surjection c) β-compact-non-empty-types-are-inhabited : {X : π€ Μ } β β-compact X β ¬¬ X β β₯ X β₯ β-compact-non-empty-types-are-inhabited {π€} {X} c Ο = g (β-compact-types-have-decidable-support c) where g : decidable β₯ X β₯ β β₯ X β₯ g (inl s) = s g (inr u) = π-elim (Ο (Ξ» x β u β£ x β£)) decidable-propositions-are-β-compact : (X : π€ Μ ) β is-prop X β decidable X β β-compact X decidable-propositions-are-β-compact X isp d p = g d where g : decidable X β decidable (β x κ X , p x οΌ β) g (inl x) = π-equality-cases b c where b : p x οΌ β β decidable (β x κ X , p x οΌ β) b r = inl β£ x , r β£ c : p x οΌ β β decidable (β x κ X , p x οΌ β) c r = inr (β₯β₯-rec (π-is-prop) f) where f : Β¬ (Ξ£ y κ X , p y οΌ β) f (y , q) = zero-is-not-one (transport (Ξ» - β p - οΌ β) (isp y x) q β»ΒΉ β r) g (inr u) = inr (β₯β₯-rec π-is-prop (Ξ» Ο β u (prβ Ο))) negations-of-Ξ -compact-propositions-are-decidable : (X : π€ Μ ) β is-prop X β Ξ -compact X β decidable (Β¬ X) negations-of-Ξ -compact-propositions-are-decidable X isp c = f a where a : decidable (X β β οΌ β) a = c (Ξ» x β β) f : decidable (X β β οΌ β) β decidable (Β¬ X) f (inl u) = inl (zero-is-not-one β u) f (inr Ο) = inr (Ξ» u β Ο (Ξ» x β π-elim (u x))) negations-of-propositions-whose-decidability-is-Ξ -compact-are-decidable : (X : π€ Μ ) β is-prop X β Ξ -compact (decidable X) β decidable (Β¬ X) negations-of-propositions-whose-decidability-is-Ξ -compact-are-decidable X isp c = Cases a l m where p : X + Β¬ X β π p (inl x) = β p (inr u) = β a : decidable ((z : X + Β¬ X) β p z οΌ β) a = c p l : ((z : X + Β¬ X) β p z οΌ β) β Β¬ X + ¬¬ X l Ξ± = inl (Ξ» x β π-elim (zero-is-not-one (Ξ± (inl x)))) Ξ± : (u : X β π) (z : X + Β¬ X) β p z οΌ β Ξ± u (inl x) = π-elim (u x) Ξ± u (inr v) = refl m : Β¬ ((z : X + Β¬ X) β p z οΌ β) β Β¬ X + ¬¬ X m Ο = inr (Ξ» u β Ο (Ξ± u)) \end{code} 8th Feb 2018: A pointed detachable subset of any type is a retract. Hence any detachable (pointed or not) subset of a β-compact type is compact. The first construction should probably go to another module. \begin{code} detachable-subset-retract : {X : π€ Μ } {A : X β π} β (Ξ£ x κ X , A x οΌ β) β retract (Ξ£ x κ X , A x οΌ β) of X detachable-subset-retract {π€} {X} {A} (xβ , eβ) = r , prβ , rs where r : X β Ξ£ x κ X , A x οΌ β r x = π-equality-cases (Ξ» (e : A x οΌ β) β (x , e)) (Ξ» (e : A x οΌ β) β (xβ , eβ)) rs : (Ο : Ξ£ x κ X , A x οΌ β) β r (prβ Ο) οΌ Ο rs (x , e) = w where s : (b : π) β b οΌ β β π-equality-cases (Ξ» (_ : b οΌ β) β (x , e)) (Ξ» (_ : b οΌ β) β (xβ , eβ)) οΌ (x , e) s β refl = refl s β r = π-elim (one-is-not-zero r) t : π-equality-cases (Ξ» (_ : A x οΌ β) β x , e) (Ξ» (_ : A x οΌ β) β xβ , eβ) οΌ (x , e) t = s (A x) e u : (Ξ» e' β x , e') οΌ (Ξ» _ β x , e) u = dfunext (fe π€β π€) Ξ» e' β ap (Ξ» - β (x , -)) (π-is-set e' e) v : r x οΌ π-equality-cases (Ξ» (_ : A x οΌ β) β x , e) (Ξ» (_ : A x οΌ β) β xβ , eβ) v = ap (Ξ» - β π-equality-cases - (Ξ» (_ : A x οΌ β) β xβ , eβ)) u w : r x οΌ x , e w = v β t \end{code} Notice that in the above lemma we need to assume that the detachable set is pointed. But its use below doesn't, because β-compactness allows us to decide inhabitedness, and β-compactness is a proposition. \begin{code} detachable-subset-β-compact : {X : π€ Μ } (A : X β π) β β-compact X β β-compact (Ξ£ x κ X , A x οΌ β) detachable-subset-β-compact {π€} {X} A c = g (c A) where g : decidable (β x κ X , A x οΌ β) β β-compact (Ξ£ x κ X , A (x) οΌ β) g (inl e) = retract-β-compact' (β₯β₯-functor detachable-subset-retract e) c g (inr u) = empty-types-are-β-compact (contrapositive β£_β£ u) \end{code} For the compact case, the retraction method to prove the last theorem is not available, but the conclusion holds, with some of the same ingredients (and with a longer proof (is there a shorter one?)). \begin{code} detachable-subset-Ξ -compact : {X : π€ Μ } (A : X β π) β Ξ -compact X β Ξ -compact (Ξ£ x κ X , A x οΌ β) detachable-subset-Ξ -compact {π€} {X} A c q = g (c p) where pβ : (x : X) β A x οΌ β β π pβ x e = β pβ : (x : X) β A x οΌ β β π pβ x e = q (x , e) p : X β π p x = π-equality-cases (pβ x) (pβ x) p-specβ : (x : X) β A x οΌ β β p x οΌ β p-specβ x e = s (A x) e (pβ x) where s : (b : π) β b οΌ β β (fβ : b οΌ β β π) β π-equality-cases (Ξ» (_ : b οΌ β) β β) fβ οΌ β s β refl = Ξ» fβ β refl s β r = π-elim (one-is-not-zero r) p-specβ : (x : X) (e : A x οΌ β) β p x οΌ q (x , e) p-specβ x e = u β t where y : A x οΌ β β π y _ = q (x , e) r : pβ x οΌ y r = dfunext (fe π€β π€β) (Ξ» e' β ap (pβ x) (π-is-set e' e)) s : (b : π) β b οΌ β β π-equality-cases (Ξ» (_ : b οΌ β) β β) (Ξ» (_ : b οΌ β) β q (x , e)) οΌ q (x , e) s β r = π-elim (zero-is-not-one r) s β refl = refl t : π-equality-cases (pβ x) y οΌ q (x , e) t = s (A x) e u : p x οΌ π-equality-cases (pβ x) y u = ap (π-equality-cases (pβ x)) r g : decidable ((x : X) β p x οΌ β) β decidable ((Ο : Ξ£ x κ X , A x οΌ β) β q Ο οΌ β) g (inl Ξ±) = inl h where h : (Ο : Ξ£ x κ X , A x οΌ β) β q Ο οΌ β h (x , e) = (p-specβ x e) β»ΒΉ β Ξ± x g (inr u) = inr (contrapositive h u) where h : ((Ο : Ξ£ x κ X , A x οΌ β) β q Ο οΌ β) β (x : X) β p x οΌ β h Ξ² x = π-equality-cases (p-specβ x) (Ξ» e β p-specβ x e β Ξ² (x , e)) \end{code} 20 Jan 2018. We now consider a truncated version of pointed compactness (see the module CompactTypes). \begin{code} β-compactβ : π€ Μ β π€ Μ β-compactβ X = (p : X β π) β β xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β) β-compactnessβ-is-prop : {X : π€ Μ } β is-prop (β-compactβ X) β-compactnessβ-is-prop {π€} = Ξ -is-prop (fe π€ π€) (Ξ» _ β β₯β₯-is-prop) \end{code} Notice that, in view of the above results, inhabitedness can be replaced by non-emptiness in the following results: \begin{code} β-compactβ-gives-inhabited-and-compact : {X : π€ Μ } β β-compactβ X β β₯ X β₯ Γ β-compact X β-compactβ-gives-inhabited-and-compact {π€} {X} c = Ξ³ where gβ : β₯ Ξ£ (Ξ» xβ β β οΌ β β (x : X) β β οΌ β) β₯ gβ = c (Ξ» x β β) gβ : (p : X β π) β (Ξ£ xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β)) β decidable (β x κ X , p x οΌ β) gβ p (xβ , Ο) = h (π-is-discrete (p xβ) β) where h : decidable (p xβ οΌ β) β decidable (β x κ X , p x οΌ β) h (inl r) = inr (β₯β₯-rec π-is-prop f) where f : Β¬ (Ξ£ x κ X , p x οΌ β) f (x , s) = zero-is-not-one (s β»ΒΉ β Ο r x) h (inr u) = inl β£ xβ , (different-from-β-equal-β u) β£ Ξ³ : β₯ X β₯ Γ β-compact X Ξ³ = β₯β₯-functor prβ gβ , (Ξ» p β β₯β₯-rec (decidability-of-prop-is-prop (fe π€ π€β) β₯β₯-is-prop) (gβ p) (c p)) inhabited-and-compact-gives-β-compactβ : {X : π€ Μ } β β₯ X β₯ Γ β-compact X β β-compactβ X inhabited-and-compact-gives-β-compactβ {π€} {X} (t , c) p = Ξ³ where f : X β β xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β) f xβ = g (π-is-discrete (p xβ) β) (c p) where g : decidable (p xβ οΌ β) β decidable (β x κ X , p x οΌ β) β β xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β) g (inl r) _ = β£ xβ , (Ξ» s _ β π-elim (zero-is-not-one (r β»ΒΉ β s))) β£ g (inr _) (inl t) = β₯β₯-functor h t where h : (Ξ£ x κ X , p x οΌ β) β Ξ£ xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β) h (x , r) = x , Ξ» s _ β π-elim (zero-is-not-one (r β»ΒΉ β s)) g (inr _) (inr v) = β£ xβ , (Ξ» _ β not-existsβ-implies-forallβ p v) β£ Ξ³ : β xβ κ X , (p xβ οΌ β β (x : X) β p x οΌ β) Ξ³ = β₯β₯-rec β₯β₯-is-prop f t \end{code} This characterizes the β-compactβ types as those that are β-compact and inhabited. We can also characterize the β-compact types as those that are β-compactβ or empty: \begin{code} being-β-compactβ-and-empty-is-prop : {X : π€ Μ } β is-prop (β-compactβ X + is-empty X) being-β-compactβ-and-empty-is-prop {π€} {X} = sum-of-contradictory-props β-compactnessβ-is-prop (Ξ -is-prop (fe π€ π€β) (Ξ» _ β π-is-prop)) (Ξ» c u β β₯β₯-rec π-is-prop (contrapositive prβ u) (c (Ξ» _ β β))) β-compactβ-or-empty-gives-β-compact : {X : π€ Μ } β β-compactβ X + is-empty X β β-compact X β-compactβ-or-empty-gives-β-compact (inl c) = prβ (β-compactβ-gives-inhabited-and-compact c) β-compactβ-or-empty-gives-β-compact (inr u) = empty-types-are-β-compact u β-compact-gives-β-compactβ-or-empty : {X : π€ Μ } β β-compact X β β-compactβ X + is-empty X β-compact-gives-β-compactβ-or-empty {π€} {X} c = g where h : decidable (β x κ X , β οΌ β) β β-compactβ X + is-empty X h (inl t) = inl (inhabited-and-compact-gives-β-compactβ (β₯β₯-functor prβ t , c)) h (inr u) = inr (contrapositive (Ξ» x β β£ x , refl β£) u) g : β-compactβ X + is-empty X g = h (c (Ξ» _ β β)) \end{code} 8 Feb 2018: A type X is Ξ -compact iff every map X β π has an infimum: \begin{code} _has-inf_ : {X : π€ Μ } β (X β π) β π β π€ Μ p has-inf n = (β x β n β€ p x) Γ (β (m : π) β (β x β m β€ p x) β m β€ n) having-inf-is-prop : {X : π€ Μ } (p : X β π) (n : π) β is-prop (p has-inf n) having-inf-is-prop {π€} {X} p n (f , g) (f' , g') = to-Γ-οΌ r s where r : f οΌ f' r = dfunext (fe π€ π€β) (Ξ» x β β€β-is-prop-valued (f x) (f' x)) s : g οΌ g' s = dfunext (fe π€β π€) (Ξ» m β dfunext (fe π€ π€β) (Ξ» Ο β β€β-is-prop-valued (g m Ο) (g' m Ο))) at-most-one-inf : {X : π€ Μ } (p : X β π) β is-prop (Ξ£ n κ π , p has-inf n) at-most-one-inf p (n , f , g) (n' , f' , g') = to-Ξ£-οΌ (β€β-anti (g' n f) (g n' f') , having-inf-is-prop p n' _ _) has-infs : π€ Μ β π€ Μ has-infs X = β (p : X β π) β Ξ£ n κ π , p has-inf n having-infs-is-prop : {X : π€ Μ } β is-prop (has-infs X) having-infs-is-prop {π€} {X} = Ξ -is-prop (fe π€ π€) at-most-one-inf Ξ -compact-has-infs : {X : π€ Μ } β Ξ -compact X β has-infs X Ξ -compact-has-infs c p = g (c p) where g : decidable (β x β p x οΌ β) β Ξ£ n κ π , p has-inf n g (inl Ξ±) = β , (Ξ» x β transportβ»ΒΉ (β β€β_) (Ξ± x) (β€β-refl {β})) , Ξ» m Ο β β-top g (inr u) = β , (Ξ» _ β β-bottom {β}) , h where h : (m : π) β (β x β m β€ p x) β m β€ β h m Ο = β€β-criterion f where f : m οΌ β β β οΌ β f r = π-elim (u Ξ±) where Ξ± : β x β p x οΌ β Ξ± x = β-maximal (transport (_β€ p x) r (Ο x)) has-infs-Ξ -compact : {X : π€ Μ } β has-infs X β Ξ -compact X has-infs-Ξ -compact h p = f (h p) where f : (Ξ£ n κ π , p has-inf n) β decidable (β x β p x οΌ β) f (β , _ , l) = inr u where u : Β¬ β x β p x οΌ β u Ξ± = l β (Ξ» x β β€β-criterion (Ξ» _ β Ξ± x)) f (β , g , _) = inl Ξ± where Ξ± : β x β p x οΌ β Ξ± x = β-maximal (g x) \end{code} TODO. Show equivalence with existence of suprema. Is there a similar characterization of β-compactness? Implicit application of type-theoretical choice: \begin{code} inf : {X : π€ Μ } β Ξ -compact X β (X β π) β π inf c p = prβ (Ξ -compact-has-infs c p) inf-property : {X : π€ Μ } β (c : Ξ -compact X) (p : X β π) β p has-inf (inf c p) inf-property c p = prβ (Ξ -compact-has-infs c p) infβ : {X : π€ Μ } (c : Ξ -compact X) {p : X β π} β inf c p οΌ β β β x β p x οΌ β infβ c {p} r x = β€β-criterion-converse (prβ (inf-property c p) x) r infβ-converse : {X : π€ Μ } (c : Ξ -compact X) {p : X β π} β (β x β p x οΌ β) β inf c p οΌ β infβ-converse c {p} Ξ± = β-maximal (h g) where h : (β x β β β€ p x) β β β€ inf c p h = prβ (inf-property c p) β g : β x β β β€ p x g x = β-maximal-converse (Ξ± x) \end{code} 21 Feb 2018. It is well known that infima and suprema are characterized as adjoints. TODO. Link the above development with the following (easy). In synthetic topology with the dominance π, a type is called π-compact if the map Ξ : π β (X β π) has a right adjoint A : (X β π) β π, with respect to the natural ordering of π and the pointwise order of the function type (X β π), and π-overt if it has a left-adjoint E : (X β π) β π. Ξ is the usual combinator (written Kappa rather than Kay here): \begin{code} Ξ : {X : π€ Μ } {Y : π₯ Μ } β Y β (X β Y) Ξ y x = y \end{code} The pointwise order on boolean predicates: \begin{code} _β€Μ_ : {X : π€ Μ } β (X β π) β (X β π) β π€ Μ p β€Μ q = β x β p x β€ q x \end{code} We define adjunctions in the two special cases where one of the sides is Ξ with Y=π, for simplicity, rather than in full generality: \begin{code} Ξβ£ : {X : π€ Μ } β ((X β π) β π) β π€ Μ Ξβ£ A = (n : π) (p : _ β π) β Ξ n β€Μ p β n β€ A p _β£Ξ : {X : π€ Μ } β ((X β π) β π) β π€ Μ E β£Ξ = (n : π) (p : _ β π) β E p β€ n β p β€Μ Ξ n \end{code} TODO: The types Ξβ£ A and E β£Ξ are propositions, and so are the types Ξ£ A κ ((X β π) β π) , Ξβ£ A (compactness) and Ξ£ E κ (X β π) β π) , E β£Ξ (overtness). Right adjoints to Ξ are characterized as follows: \begin{code} Ξβ£-charac : {X : π€ Μ } β (A : (X β π) β π) β Ξβ£ A β ((p : X β π) β A p οΌ β β p οΌ (Ξ» x β β)) Ξβ£-charac {π€} {X} A = f , g where f : Ξβ£ A β (p : X β π) β A p οΌ β β p οΌ (Ξ» x β β) f Ο p = fβ , fβ where fβ : A p οΌ β β p οΌ (Ξ» x β β) fβ r = dfunext (fe π€ π€β) lβ where lβ : β β€ A p β Ξ β β€Μ p lβ = prβ (Ο β p) lβ : Ξ β β€Μ p lβ = lβ (β-maximal-converse r) lβ : (x : X) β β β€ p x lβ = lβ lβ : (x : X) β p x οΌ β lβ x = β€β-criterion-converse (lβ x) refl fβ : p οΌ (Ξ» x β β) β A p οΌ β fβ s = β€β-criterion-converse lβ refl where lβ : (x : X) β p x οΌ β lβ = happly s lβ : (x : X) β β β€ p x lβ x = β-maximal-converse (lβ x) lβ : Ξ β β€Μ p lβ = lβ lβ : β β€ A p lβ = prβ (Ο β p) lβ g : ((p : X β π) β A p οΌ β β p οΌ (Ξ» x β β)) β Ξβ£ A g Ξ³ n p = (gβ n refl , gβ n refl) where gβ : β m β m οΌ n β Ξ m β€Μ p β m β€ A p gβ β r l = β-bottom {β} gβ β refl l = β-maximal-converse (prβ (Ξ³ p) lβ) where lβ : (x : X) β p x οΌ β lβ x = β-maximal (l x) lβ : p οΌ (Ξ» x β β) lβ = dfunext (fe π€ π€β) lβ gβ : β m β m οΌ n β m β€ A p β Ξ m β€Μ p gβ β r l x = β-bottom {β} gβ β refl l x = β-maximal-converse (lβ x) where lβ : p οΌ (Ξ» x β β) lβ = prβ (Ξ³ p) (β-maximal l) lβ : (x : X) β p x οΌ β lβ = happly lβ \end{code} Using this as a lemma, we see that a type is Ξ -compact in the sense we defined iff it is compact in the usual sense of synthetic topology for the dominance π. \begin{code} Ξ -compact-iff-Ξ-has-right-adjoint : {X : π€ Μ } β Ξ -compact X β (Ξ£ A κ ((X β π) β π), Ξβ£ A) Ξ -compact-iff-Ξ-has-right-adjoint {π€} {X} = (f , g) where f : Ξ -compact X β Ξ£ A κ ((X β π) β π), Ξβ£ A f c = (A , prβ (Ξβ£-charac A) lβ) where c' : (p : X β π) β decidable (p οΌ (Ξ» x β β)) c' = Ξ -compact-gives-Ξ -compact' c lβ : (p : X β π) β decidable (p οΌ (Ξ» x β β)) β Ξ£ n κ π , (n οΌ β β p οΌ (Ξ» x β β)) lβ p (inl r) = (β , ((Ξ» _ β r) , Ξ» _ β refl)) lβ p (inr u) = (β , ((Ξ» s β π-elim (zero-is-not-one s)) , Ξ» r β π-elim (u r))) A : (X β π) β π A p = prβ (lβ p (c' p)) lβ : (p : X β π) β A p οΌ β β p οΌ (Ξ» x β β) lβ p = prβ (lβ p (c' p)) g : ((Ξ£ A κ ((X β π) β π), Ξβ£ A)) β Ξ -compact X g (A , Ο) = Ξ -compact'-gives-Ξ -compact c' where lβ : (p : X β π) β A p οΌ β β p οΌ (Ξ» x β β) lβ = prβ (Ξβ£-charac A) Ο lβ : (p : X β π) β decidable (A p οΌ β) β decidable (p οΌ (Ξ» x β β)) lβ p (inl r) = inl (prβ (lβ p) r) lβ p (inr u) = inr (contrapositive (prβ (lβ p)) u) c' : (p : X β π) β decidable (p οΌ (Ξ» x β β)) c' p = lβ p (π-is-discrete (A p) β) \end{code} Next we show that ΞΊ has a right adjoint iff it has a left adjoint, namely its De Morgan dual, which exists because π is a boolean algebra and hence so is the type (X β π) with the pointwise operations. \begin{code} π-DeMorgan-dual : {X : π€ Μ } β ((X β π) β π) β ((X β π) β π) π-DeMorgan-dual Ο = Ξ» p β complement (Ο (Ξ» x β complement (p x))) π-DeMorgan-dual-involutive : {X : π€ Μ } β (Ο : (X β π) β π) β π-DeMorgan-dual (π-DeMorgan-dual Ο) οΌ Ο π-DeMorgan-dual-involutive {π€} Ο = dfunext (fe π€ π€β) h where f : β p β complement (complement (Ο (Ξ» x β complement (complement (p x))))) οΌ Ο (Ξ» x β complement (complement (p x))) f p = complement-involutive (Ο (Ξ» x β complement (complement (p x)))) g : β p β Ο (Ξ» x β complement (complement (p x))) οΌ Ο p g p = ap Ο (dfunext (fe π€ π€β) (Ξ» x β complement-involutive (p x))) h : β p β π-DeMorgan-dual (π-DeMorgan-dual Ο) p οΌ Ο p h p = f p β g p Ξ -compact-is-π-overt : {X : π€ Μ } β (A : (X β π) β π) β Ξβ£ A β (π-DeMorgan-dual A) β£Ξ Ξ -compact-is-π-overt {π€} {X} A = f where E : (X β π) β π E = π-DeMorgan-dual A f : Ξβ£ A β E β£Ξ f Ο = Ξ³ where Ξ³ : (n : π) (p : X β π) β (E p β€ n) β (p β€Μ Ξ n) Ξ³ n p = (Ξ³β , Ξ³β ) where Ξ³β : E p β€ n β p β€Μ Ξ n Ξ³β l = mβ where mβ : complement n β€ A (Ξ» x β complement (p x)) mβ = complement-left l mβ : Ξ (complement n) β€Μ (Ξ» x β complement (p x)) mβ = prβ (Ο (complement n) (Ξ» x β complement (p x))) mβ mβ : (x : X) β complement n β€ complement (p x) mβ = mβ mβ : (x : X) β p x β€ n mβ x = complement-both-left (mβ x) Ξ³β : p β€Μ Ξ n β E p β€ n Ξ³β l = complement-left mβ where mβ : (x : X) β p x β€ n mβ = l mβ : (x : X) β complement n β€ complement (p x) mβ x = complement-both-right (mβ x) mβ : Ξ (complement n) β€Μ (Ξ» x β complement (p x)) mβ = mβ mβ : complement n β€ A (Ξ» x β complement (p x)) mβ = prβ (Ο (complement n) (Ξ» x β complement (p x))) mβ π-overt-is-Ξ -compact : {X : π€ Μ } β (E : (X β π) β π) β E β£Ξ β Ξβ£ (π-DeMorgan-dual E) π-overt-is-Ξ -compact {π€} {X} E = g where A : (X β π) β π A = π-DeMorgan-dual E g : E β£Ξ β Ξβ£ A g Ξ³ = Ο where Ο : (n : π) (p : X β π) β Ξ n β€Μ p β n β€ A p Ο n p = (Οβ , Οβ ) where Οβ : Ξ n β€Μ p β n β€ A p Οβ l = complement-right mβ where mβ : (x : X) β n β€ p x mβ = l mβ : (x : X) β complement (p x) β€ complement n mβ x = complement-both-right (mβ x) mβ : (Ξ» x β complement (p x)) β€Μ Ξ (complement n) mβ = mβ mβ : E (Ξ» x β complement (p x)) β€ complement n mβ = prβ (Ξ³ (complement n) (Ξ» x β complement (p x))) mβ Οβ : n β€ A p β Ξ n β€Μ p Οβ l = mβ where mβ : E (Ξ» x β complement (p x)) β€ complement n mβ = complement-right l mβ : (Ξ» x β complement (p x)) β€Μ Ξ (complement n) mβ = prβ (Ξ³ (complement n) (Ξ» x β complement (p x))) mβ mβ : (x : X) β complement (p x) β€ complement n mβ = mβ mβ : (x : X) β n β€ p x mβ x = complement-both-left (mβ x) \end{code} We have the following corollaries: \begin{code} Ξ -compact-iff-π-overt : {X : π€ Μ } β (Ξ£ A κ ((X β π) β π), Ξβ£ A) β (Ξ£ E κ ((X β π) β π), E β£Ξ) Ξ -compact-iff-π-overt {π€} {X} = (f , g) where f : (Ξ£ A κ ((X β π) β π), Ξβ£ A) β (Ξ£ E κ ((X β π) β π), E β£Ξ) f (A , Ο) = (π-DeMorgan-dual A , Ξ -compact-is-π-overt A Ο) g : (Ξ£ E κ ((X β π) β π), E β£Ξ) β (Ξ£ A κ ((X β π) β π), Ξβ£ A) g (E , Ξ³) = (π-DeMorgan-dual E , π-overt-is-Ξ -compact E Ξ³) \end{code} In this corollary we record explicitly that a type is Ξ -compact iff it is π-overt: \begin{code} Ξ -compact-iff-Ξ-has-left-adjoint : {X : π€ Μ } β Ξ -compact X β (Ξ£ E κ ((X β π) β π), E β£Ξ) Ξ -compact-iff-Ξ-has-left-adjoint {π€} {X} = (f , g) where f : Ξ -compact X β (Ξ£ E κ ((X β π) β π), E β£Ξ) f c = prβ Ξ -compact-iff-π-overt (prβ Ξ -compact-iff-Ξ-has-right-adjoint c) g : (Ξ£ E κ ((X β π) β π), E β£Ξ) β Ξ -compact X g o = prβ Ξ -compact-iff-Ξ-has-right-adjoint (prβ Ξ -compact-iff-π-overt o) \end{code} TODO. We get as a corollary that E β£Ξ β ((p : X β π) β E p οΌ β β p οΌ (Ξ» x β β)). TODO. Find the appropriate place in this file to remark that decidable propositions are closed under Ξ -compact/overt meets and joins. And then clopen sets (or π-open sets, or complemented subsets) are closed under Ξ -compact/over unions and intersections. 20 Feb 2018. In classical topology, a space X is compact iff the projection A Γ X β A is a closed map for every space A, meaning that the image of every closed set is closed. In our case, because of the use of decidable truth-values in the definition of Ξ -compactness, the appropriate notion is that of clopen map, that is, a map that sends clopen sets to clopen sets. As in our setup, clopen sets correspond to decidable subsets, or sets with π-valued characteristic functions. In our case, the clopeness of the projections characterize the notion of β-compactness, which is stronger than compactness. There is a certain asymmetry in the following definition, in that the input decidable predicate (or clopen subtype) is given as a π-valued function, whereas instead of saying that the image predicate factors through the embedding π of into the type of truth values, we say that it has decidable truth-values, which is equivalent. Such an asymmetry is already present in our formulation of the notion of compactness. We have defined image with lower case in the module UF. We now need Images with upper case: \begin{code} Image : {X : π€ Μ } {Y : π₯ Μ } β (X β Y) β (X β π¦ Μ ) β (Y β π€ β π₯ β π¦ Μ ) Image f A = Ξ» y β β x κ domain f , A x Γ (f x οΌ y) is-clopen-map : {X : π€ Μ } {Y : π₯ Μ } β (X β Y) β π€ β π₯ Μ is-clopen-map {π€} {π₯} {X} {Y} f = (p : X β π) (y : Y) β decidable (Image f (Ξ» x β p x οΌ β) y) being-clopen-map-is-prop : {X : π€ Μ } {Y : π₯ Μ } β (f : X β Y) β is-prop (is-clopen-map f) being-clopen-map-is-prop {π€} {π₯} f = Ξ -is-prop (fe π€ (π€ β π₯)) (Ξ» p β Ξ -is-prop (fe π₯ (π€ β π₯)) (Ξ» y β decidability-of-prop-is-prop (fe (π€ β π₯) π€β) β₯β₯-is-prop)) fst : (A : π€ Μ ) (X : π₯ Μ ) β A Γ X β A fst _ _ = prβ β-compact-clopen-projections : (X : π€ Μ ) β β-compact X β (β {π₯} (A : π₯ Μ ) β is-clopen-map (fst A X)) β-compact-clopen-projections X c A p a = g (c (Ξ» x β p (a , x))) where g : decidable (β x κ X , p (a , x) οΌ β) β decidable (β z κ A Γ X , (p z οΌ β) Γ (prβ z οΌ a)) g (inl e) = inl ((β₯β₯-functor h) e) where h : (Ξ£ x κ X , p (a , x) οΌ β) β Ξ£ z κ A Γ X , (p z οΌ β) Γ (prβ z οΌ a) h (x , r) = (a , x) , (r , refl) g (inr u) = inr (contrapositive (β₯β₯-functor h) u) where h : (Ξ£ z κ A Γ X , (p z οΌ β) Γ (prβ z οΌ a)) β Ξ£ x κ X , p (a , x) οΌ β h ((a' , x) , (r , s)) = x , transport (Ξ» - β p (- , x) οΌ β) s r clopen-projections-β-compact : β {π€ π¦} (X : π€ Μ ) β (β {π₯} (A : π₯ Μ ) β is-clopen-map (fst A X)) β β-compact X clopen-projections-β-compact {π€} {π¦} X ΞΊ p = g (ΞΊ π (Ξ» z β p (prβ z)) β) where g : decidable (β z κ π {π¦} Γ X , (p (prβ z) οΌ β) Γ (prβ z οΌ β)) β decidable (β x κ X , p x οΌ β) g (inl e) = inl (β₯β₯-functor h e) where h : (Ξ£ z κ π Γ X , (p (prβ z) οΌ β) Γ (prβ z οΌ β)) β Ξ£ x κ X , p x οΌ β h ((β , x) , r , _) = x , r g (inr u) = inr (contrapositive (β₯β₯-functor h) u) where h : (Ξ£ x κ X , p x οΌ β) β Ξ£ z κ π Γ X , (p (prβ z) οΌ β) Γ (prβ z οΌ β) h (x , r) = (β , x) , (r , refl) \end{code} TODO. β Consider π-perfect maps. β β-compactness: attainability of minima. Existence of potential maxima. β Relation of Ξ -compactness with finiteness and discreteness. β Non-classical cotaboos Every Ξ -compact subtype of β is finite. Every Ξ -compact subtype of a discrete type is finite. What are the cotaboos necessary (and sufficient) to prove that the type of decidable subsingletons of ββββ is Ξ -compact? Continuity principles are enough. β π-subspace: e:XβY such that every clopen Xβπ extends to some clopen Yβπ (formulated with Ξ£ and β). Or to a largest such clopen, or a smallest such clopen (right and left adjoints to the restriction map (Yβπ)β(Xβπ) that maps v to v β e and could be written e β»ΒΉ[ v ]. A π-subspace-embedding of totally separated types should be a (homotopy) embedding, but not conversely (find a counter-example). β π-injective types (injectives wrt to π-subspace-embeddigs). They should be the retracts of powers of π. Try to characterize them "intrinsically". β Relation of π-subspaces with Ξ -compact subtypes. β π-Hofmann-Mislove theorem: clopen filters of clopens should correspond to Ξ -compact (π-saturated) π-subspaces. Are cotaboos needed for this? β Which results here depend on the particular dominance π, and which ones generalize to any dominance, or to any "suitable" dominance? In particular, it is of interest to generalize this to "Sierpinki like" dominances. And what is "Sierpinski like" in precise (internal) terms? This should be formulated in terms of cotaboos.