Connected types
Content created by Egbert Rijke, Fredrik Bakke and Jonathan Prieto-Cubides
Created on 2022-02-18.
Last modified on 2023-06-10.
module foundation.connected-types where
Imports
open import foundation.contractible-types open import foundation.truncations open import foundation.universe-levels open import foundation-core.constant-maps open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.propositions open import foundation-core.truncated-types open import foundation-core.truncation-levels
Idea
A type is said to be k
-connected if its k
-truncation is contractible.
Definition
is-connected-Prop : {l : Level} (k : 𝕋) → UU l → Prop l is-connected-Prop k A = is-contr-Prop (type-trunc k A) is-connected : {l : Level} (k : 𝕋) → UU l → UU l is-connected k A = type-Prop (is-connected-Prop k A) is-prop-is-connected : {l : Level} (k : 𝕋) (A : UU l) → is-prop (is-connected k A) is-prop-is-connected k A = is-prop-type-Prop (is-connected-Prop k A)
Properties
A type A
is k
-connected if and only if the map B → (A → B)
is an equivalence for every k
-truncated type B
is-equiv-diagonal-is-connected : {l1 l2 : Level} {k : 𝕋} {A : UU l1} (B : Truncated-Type l2 k) → is-connected k A → is-equiv (λ (b : type-Truncated-Type B) → const A (type-Truncated-Type B) b) is-equiv-diagonal-is-connected B H = is-equiv-comp ( precomp unit-trunc (type-Truncated-Type B)) ( λ b → const _ _ b) ( is-equiv-diagonal-is-contr H (type-Truncated-Type B)) ( is-truncation-trunc B)
Recent changes
- 2023-06-10. Egbert Rijke. cleaning up transport and dependent identifications files (#650).
- 2023-06-08. Fredrik Bakke. Remove empty
foundation
modules and replace them by their core counterparts (#644). - 2023-03-14. Fredrik Bakke. Remove all unused imports (#502).
- 2023-03-07. Fredrik Bakke. Add blank lines between
<details>
tags and markdown syntax (#490). - 2023-03-07. Jonathan Prieto-Cubides. Show module declarations (#488).