Ian Ray, 23rd July 2024 Modifications made by Ian Ray on 26th September 2024. Modifications made by Ian Ray on 17th December 2024. We define connected types and maps. We then explore relationships, closure properties and characterizations of interest pertaining to the concept of connectedness. \begin{code} {-# OPTIONS --safe --without-K #-} open import UF.FunExt module UF.ConnectedTypes (fe : Fun-Ext) where open import MLTT.Spartan hiding (_+_) open import Notation.Order open import UF.Base open import UF.Equiv open import UF.EquivalenceExamples open import UF.PropTrunc open import UF.Retracts open import UF.Subsingletons open import UF.Subsingletons-FunExt open import UF.TruncatedTypes fe open import UF.Truncations fe open import UF.TruncationLevels open import UF.Univalence open import UF.Yoneda \end{code} We now define the notion of connectedness for types and functions with respect to truncation levels. \begin{code} private fe' : FunExt fe' ๐ค ๐ฅ = fe {๐ค} {๐ฅ} module connectedness-results (te : general-truncations-exist) where private pt : propositional-truncations-exist pt = general-truncations-give-propositional-truncations te open general-truncations-exist te open propositional-truncations-exist pt open import UF.ImageAndSurjection pt _is_connected : ๐ค ฬ โ โโโ โ ๐ค ฬ X is k connected = is-contr (โฅ X โฅ[ k ]) _is_connected-map : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} โ (f : X โ Y) โ โโโ โ ๐ค โ ๐ฅ ฬ f is k connected-map = each-fiber-of f (ฮป - โ - is k connected) being-connected-is-prop : {k : โโโ} {X : ๐ค ฬ} โ is-prop (X is k connected) being-connected-is-prop = being-truncated-is-prop being-connected-map-is-prop : {k : โโโ} {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {f : X โ Y} โ is-prop (f is k connected-map) being-connected-map-is-prop = ฮ -is-prop fe (ฮป y โ being-connected-is-prop) connected-if-contr : {X : ๐ค ฬ} {k : โโโ} โ is-contr X โ X is k connected connected-if-contr {_} {X} {โ2} X-is-contr = โ2-trunc-is-contr connected-if-contr {_} {X} {succ k} (c , C) = (โฃ c โฃ[ k + 1 ] , C') where C'' : (x : X) โ โฃ c โฃ[ k + 1 ] ๏ผ โฃ x โฃ[ k + 1 ] C'' x = canonical-identity-trunc-map โฃ C x โฃ[ k ] C' : is-central โฅ X โฅ[ k + 1 ] โฃ c โฃ[ k + 1 ] C' = โฅโฅโ-ind (ฮป v โ ฮป p q โ truncation-levels-closed-under-Id (โฅโฅโ-is-truncated โฃ c โฃ[ succ k ] v) p q) C'' \end{code} TODO: show that connectedness as defined elsewhere in the library is a special case of k-connectedness. Connectedness typically means set connectedness, by our convention it will mean 0-connectedness. We characterize โ1-connected types as inhabited types and โ1-connected maps as surjections. \begin{code} โ1-connected-types-are-inhabited : {X : ๐ค ฬ} โ X is โ1 connected โ โฅ X โฅ โ1-connected-types-are-inhabited X-1-conn = โ1-trunc-to-prop-trunc pt (center X-1-conn) inhabited-types-are-โ1-connected : {X : ๐ค ฬ} โ โฅ X โฅ โ X is โ1 connected inhabited-types-are-โ1-connected x-anon = pointed-props-are-singletons (prop-trunc-to-โ1-trunc pt x-anon) โ1-trunc-is-prop โ1-connected-iff-inhabited : {X : ๐ค ฬ} โ X is โ1 connected โ โฅ X โฅ โ1-connected-iff-inhabited = (โ1-connected-types-are-inhabited , inhabited-types-are-โ1-connected) โ1-connected-maps-are-surjections : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {f : X โ Y} โ f is โ1 connected-map โ is-surjection f โ1-connected-maps-are-surjections m y = โ1-connected-types-are-inhabited (m y) surjections-are-โ1-connected : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {f : X โ Y} โ is-surjection f โ f is โ1 connected-map surjections-are-โ1-connected f-is-surj y = inhabited-types-are-โ1-connected (f-is-surj y) โ1-connected-map-iff-surjection : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {f : X โ Y} โ f is โ1 connected-map โ is-surjection f โ1-connected-map-iff-surjection = (โ1-connected-maps-are-surjections , surjections-are-โ1-connected) \end{code} We develop some closure conditions pertaining to connectedness. Connectedness is closed under retracts and equivalence as expected, but more importantly connectedness extends below: more precisely if a type is k connected then it is l connected for all l โค k. We provide a few incarnations of this fact below which may prove useful. \begin{code} retract-is-connected : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {k : โโโ} โ retract Y of X โ X is k connected โ Y is k connected retract-is-connected R X-conn = retract-of-singleton (truncation-closed-under-retract R) X-conn connectedness-closed-under-equiv : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {k : โโโ} โ X โ Y โ Y is k connected โ X is k connected connectedness-closed-under-equiv e Y-con = equiv-to-singleton (truncation-closed-under-equiv e) Y-con contractible-types-are-connected : {X : ๐ค ฬ} {k : โโโ} โ is-contr X โ X is k connected contractible-types-are-connected {๐ค} {X} {k} (c , C) = ((โฃ c โฃ[ k ]) , C') where C' : (s : โฅ X โฅ[ k ]) โ (โฃ c โฃ[ k ]) ๏ผ s C' = โฅโฅโ-ind (truncation-levels-closed-under-Id โฅโฅโ-is-truncated (โฃ c โฃ[ k ])) (ฮป x โ ap โฃ_โฃ[ k ] (C x)) connectedness-is-lower-closed : {X : ๐ค ฬ} {k : โโโ} โ X is (k + 1) connected โ X is k connected connectedness-is-lower-closed {๐ค} {X} {k} X-succ-con = equiv-to-singleton successive-truncations-equiv (contractible-types-are-connected X-succ-con) connectedness-is-lower-closed-+ : {X : ๐ค ฬ} {l : โโโ} {k : โ} โ X is (l + k) connected โ X is l connected connectedness-is-lower-closed-+ {๐ค} {X} {l} {0} X-con = X-con connectedness-is-lower-closed-+ {๐ค} {X} {l} {succ k} X-con = connectedness-is-lower-closed-+ (connectedness-is-lower-closed X-con) connectedness-is-lower-closed' : {X : ๐ค ฬ} {k l : โโโ} โ (l โค k) โ X is k connected โ X is l connected connectedness-is-lower-closed' {๐ค} {X} {k} {l} o X-con = connectedness-is-lower-closed-+ (transport (ฮป z โ X is z connected) p X-con) where m : โ m = subtraction-โโโ-term l k o p = k ๏ผโจ subtraction-โโโ-identification l k o โปยน โฉ l + m โ map-connectedness-is-lower-closed : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {f : X โ Y} {k l : โโโ} โ (l โค k) โ f is k connected-map โ f is l connected-map map-connectedness-is-lower-closed o f-k-con y = connectedness-is-lower-closed' o (f-k-con y) \end{code} We characterize connected types in terms of inhabitedness and connectedness of the identity type at one level below. We will assume univalence only when necessary. \begin{code} connected-types-are-inhabited : {X : ๐ค ฬ} {k : โโโ} โ X is (k + 1) connected โ โฅ X โฅ connected-types-are-inhabited {_} {_} {k} X-conn = โ1-connected-types-are-inhabited (connectedness-is-lower-closed' โ X-conn) _is-locally_connected : (X : ๐ค ฬ) (k : โโโ) โ ๐ค ฬ X is-locally k connected = (x y : X) โ (x ๏ผ y) is k connected connected-types-are-locally-connected : {X : ๐ค ฬ} {k : โโโ} โ is-univalent ๐ค โ X is (k + 1) connected โ X is-locally k connected connected-types-are-locally-connected {_} {_} {k} ua X-conn x y = equiv-to-singleton (eliminated-trunc-identity-char ua) (is-prop-implies-is-prop' (singletons-are-props X-conn) โฃ x โฃ[ succ k ] โฃ y โฃ[ succ k ]) connected-types-are-inhabited-and-locally-connected : {X : ๐ค ฬ} {k : โโโ} โ is-univalent ๐ค โ X is (k + 1) connected โ โฅ X โฅ ร X is-locally k connected connected-types-are-inhabited-and-locally-connected ua X-conn = (connected-types-are-inhabited X-conn , connected-types-are-locally-connected ua X-conn) inhabited-and-locally-connected-types-are-connected : {X : ๐ค ฬ} {k : โโโ} โ is-univalent ๐ค โ โฅ X โฅ โ X is-locally k connected โ X is (k + 1) connected inhabited-and-locally-connected-types-are-connected {_} {_} {โ2} ua anon-x id-conn = pointed-props-are-singletons (prop-trunc-to-โ1-trunc pt anon-x) โ1-trunc-is-prop inhabited-and-locally-connected-types-are-connected {_} {_} {succ k} ua anon-x id-conn = โฅโฅ-rec (being-singleton-is-prop fe) (ฮป x โ (โฃ x โฃ[ (k + 1) + 1 ] , โฅโฅโ-ind (ฮป v โ truncation-levels-are-upper-closed (ฮป p q โ โฅโฅโ-is-truncated โฃ x โฃ[ (k + 1) + 1 ] v p q)) (ฮป y โ forth-trunc-id-char ua (center (id-conn x y))))) anon-x connected-characterization : {X : ๐ค ฬ} {k : โโโ} โ is-univalent ๐ค โ X is (k + 1) connected โ โฅ X โฅ ร X is-locally k connected connected-characterization ua = (connected-types-are-inhabited-and-locally-connected ua , uncurry (inhabited-and-locally-connected-types-are-connected ua)) ap-is-less-connected : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {k : โโโ} โ (ua : is-univalent (๐ค โ ๐ฅ)) โ (f : X โ Y) โ f is (k + 1) connected-map โ {x x' : X} โ (ap f {x} {x'}) is k connected-map ap-is-less-connected ua f f-conn {x} {x'} p = equiv-to-singleton (truncation-closed-under-equiv (fiber-of-ap-โ f p)) (connected-types-are-locally-connected ua (f-conn (f x')) (x , p) (x' , refl)) \end{code} We prove a characterization of connectedness from the HoTT book (see Corollary 7.5.9.). Notice we choose to directly prove this result rather than instantiate it as a special case of a more general result (Lemma 7.5.7.) which is stated below. NOTE: We will NOT state the corollary as an iff statement due to a large quantification issue. \begin{code} private consts : (X : ๐ค ฬ) (Y : ๐ฅ ฬ) โ Y โ (X โ Y) consts X Y y x = y maps-with-connected-dom-truncated-cod-are-constant : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {n : โโโ} โ X is n connected โ Y is n truncated โ Y โ (X โ Y) maps-with-connected-dom-truncated-cod-are-constant {๐ค} {_} {X} {Y} {n} X-conn Y-trunc = e where e : Y โ (X โ Y) e = Y โโจ ๐โ fe โฉ (๐ {๐ค} โ Y) โโจ โcong'' fe fe I โฉ (โฅ X โฅ[ n ] โ Y) โโจ โฅโฅโ-universal-property Y-trunc โฉ (X โ Y) โ where I : ๐ {๐ค} โ โฅ X โฅ[ n ] I = singleton-โ-๐' X-conn observation : consts X Y ๏ผ โ e โ observation = refl constants-map-from-truncated-type-is-equiv : {X : ๐ค ฬ} {n : โโโ} โ (Y : ๐ฅ ฬ) โ X is n connected โ Y is n truncated โ is-equiv (consts X Y) constants-map-from-truncated-type-is-equiv Y X-conn Y-trunc = โโ-is-equiv (maps-with-connected-dom-truncated-cod-are-constant X-conn Y-trunc) connected-when-consts-is-equiv : {X : ๐ค ฬ} {n : โโโ} โ ({๐ฅ : Universe} (Y : ๐ฅ ฬ) โ Y is n truncated โ is-equiv (consts X Y)) โ X is n connected connected-when-consts-is-equiv {_} {X} {n} is-equiv-from-trunc = (c , G) where s : (X โ โฅ X โฅ[ n ]) โ โฅ X โฅ[ n ] s = section-of (consts X โฅ X โฅ[ n ]) (equivs-have-sections (consts X โฅ X โฅ[ n ]) (is-equiv-from-trunc โฅ X โฅ[ n ] โฅโฅโ-is-truncated)) H : (consts X โฅ X โฅ[ n ]) โ s โผ id H = section-equation (consts X โฅ X โฅ[ n ]) (equivs-have-sections (consts X โฅ X โฅ[ n ]) (is-equiv-from-trunc โฅ X โฅ[ n ] โฅโฅโ-is-truncated)) c : โฅ X โฅ[ n ] c = s โฃ_โฃ[ n ] H' : (consts X โฅ X โฅ[ n ]) c ๏ผ โฃ_โฃ[ n ] H' = H โฃ_โฃ[ n ] G : (v : โฅ X โฅ[ n ]) โ c ๏ผ v G = โฅโฅโ-ind (ฮป - โ truncation-levels-are-upper-closed โฅโฅโ-is-truncated c -) (ฮป x โ happly H' x) \end{code} We will now prove a general result from the HoTT book that characterizes when a map is connected (see Lemma 7.5.7.) \begin{code} dependent-equiv-from-truncated-fam-connected-map : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {n : โโโ} โ (f : X โ Y) โ (P : Y โ ๐ฆ ฬ) โ ((y : Y) โ (P y) is n truncated) โ f is n connected-map โ ((y : Y) โ P y) โ ((x : X) โ P (f x)) dependent-equiv-from-truncated-fam-connected-map {_} {_} {_} {X} {Y} {n} f P P-trunc f-conn = e where e : ((y : Y) โ P y) โ ((x : X) โ P (f x)) e = ((y : Y) โ P y) โโจ I โฉ ((y : Y) โ (fiber f y โ P y)) โโจ II โฉ ((y : Y) โ (x : X) โ (p : f x ๏ผ y) โ P y) โโจ ฮ -flip' โฉ ((x : X) โ (y : Y) โ (p : f x ๏ผ y) โ P y) โโจ III โฉ ((x : X) โ P (f x)) โ where I = ฮ -cong fe fe (ฮป - โ maps-with-connected-dom-truncated-cod-are-constant (f-conn -) (P-trunc -)) II = ฮ -cong fe fe (ฮป - โ curry-uncurry' fe fe) III = ฮ -cong fe fe (ฮป - โ โ-sym (Yoneda-equivalence fe' (f -) P)) observation : โ e โ ๏ผ dprecomp P f observation = refl dep-precomp-from-truncated-family-is-equiv : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {n : โโโ} โ (f : X โ Y) โ (P : Y โ ๐ฆ ฬ) โ ((y : Y) โ (P y) is n truncated) โ f is n connected-map โ is-equiv (dprecomp P f) dep-precomp-from-truncated-family-is-equiv f P P-trunc f-conn = โโ-is-equiv (dependent-equiv-from-truncated-fam-connected-map f P P-trunc f-conn) dep-precomp-has-section-if-is-equiv : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} โ (f : X โ Y) โ (P : Y โ ๐ฆ ฬ) โ is-equiv (dprecomp P f) โ has-section (dprecomp P f) dep-precomp-has-section-if-is-equiv f P = equivs-have-sections (dprecomp P f) map-is-connected-if-dep-precomp-from-truncated-family-has-section : {X : ๐ค ฬ} {Y : ๐ฅ ฬ} {n : โโโ} โ (f : X โ Y) โ ({๐ฆ : Universe} (P : Y โ ๐ฆ ฬ) โ ((y : Y) โ (P y) is n truncated) โ has-section (dprecomp P f)) โ f is n connected-map map-is-connected-if-dep-precomp-from-truncated-family-has-section {๐ค} {๐ฅ} {X} {Y} {n} f sec-from-trunc y = (c y , C) where Q : Y โ ๐ค โ ๐ฅ ฬ Q y = โฅ fiber f y โฅ[ n ] c' : ((x : X) โ โฅ fiber f (f x) โฅ[ n ]) โ ((y : Y) โ โฅ fiber f y โฅ[ n ]) c' = section-of (dprecomp Q f) (sec-from-trunc Q (ฮป - โ โฅโฅโ-is-truncated)) c : (y : Y) โ โฅ fiber f y โฅ[ n ] c = c' (ฮป - โ โฃ (- , refl) โฃ[ n ]) H' : (dprecomp Q f) โ c' โผ id H' = section-equation (dprecomp Q f) (sec-from-trunc Q (ฮป - โ โฅโฅโ-is-truncated)) H : (x : X) โ c (f x) ๏ผ โฃ (x , refl) โฃ[ n ] H = happly' ((dprecomp Q f โ c') (ฮป - โ โฃ (- , refl) โฃ[ n ])) (ฮป - โ โฃ (- , refl) โฃ[ n ]) (H' (ฮป - โ โฃ (- , refl) โฃ[ n ])) C : (w : โฅ fiber f y โฅ[ n ]) โ c y ๏ผ w C = โฅโฅโ-ind (ฮป - โ truncation-levels-are-upper-closed โฅโฅโ-is-truncated (c y) -) C' where C' : ((x , p) : fiber f y) โ c y ๏ผ โฃ (x , p) โฃ[ n ] C' (x , refl) = H x \end{code} We show that the n-truncation map is n-connected. \begin{code} trunc-map-is-connected : {X : ๐ค ฬ} {n : โโโ} โ โฃ_โฃ[ n ] is n connected-map trunc-map-is-connected {_} {X} {n} = map-is-connected-if-dep-precomp-from-truncated-family-has-section โฃ_โฃ[ n ] has-sec where has-sec : {๐ฆ : Universe} โ (P : โฅ X โฅ[ n ] โ ๐ฆ ฬ) โ ((v : โฅ X โฅ[ n ]) โ P v is n truncated) โ has-section (dprecomp P โฃ_โฃ[ n ]) has-sec {_} P P-trunc = (โฅโฅโ-ind P-trunc , comp-rule) where comp-rule : dprecomp P โฃ_โฃ[ n ] โ โฅโฅโ-ind P-trunc โผ id comp-rule h = (dprecomp P โฃ_โฃ[ n ]) (โฅโฅโ-ind P-trunc h) ๏ผโจ refl โฉ (โฅโฅโ-ind P-trunc h) โ โฃ_โฃ[ n ] ๏ผโจ I โฉ h โ where I = dfunext fe (โฅโฅโ-ind-comp P-trunc h) \end{code} We provide a useful result about maps from the constant type to a connected type (see Lemma 7.5.11 of the HoTT book for the full statement). Observe we do assume univalence locally here (the HoTT book assumes it implicitly). \begin{code} basepoint-map-is-less-connected : {X : ๐ค ฬ} {n : โโโ} โ is-univalent ๐ค โ (xโ : ๐ {๐ค} โ X) โ X is (n + 1) connected โ xโ is n connected-map basepoint-map-is-less-connected {_} {X} {n} ua xโ X-con x = connectedness-closed-under-equiv ๐-lneutral (connected-types-are-locally-connected ua X-con (xโ โ) x) \end{code}