module Homotopy.Connectedness where
Connectedness🔗
We say that a type is -connected if its -truncation is contractible.
While being -truncated expresses that all homotopy groups above are trivial, being -connected means that all homotopy groups below are trivial. A type that is both -truncated and -connected is contractible.
We give definitions in terms of the propositional truncation and set truncation for the lower levels, and then defer to the general “hubs and spokes” truncation. Note that our indices are offset by 2, just like h-levels.
: ∀ {ℓ} → Type ℓ → Nat → Type _
is-n-connected = Lift _ ⊤
is-n-connected A zero (suc zero) = ∥ A ∥
is-n-connected A (suc (suc zero)) = is-contr ∥ A ∥₀
is-n-connected A @(suc (suc (suc _))) = is-contr (n-Tr A n) is-n-connected A n
Being -connected is a proposition:
: (n : Nat) → is-prop (is-n-connected A n)
is-n-connected-is-prop 0 _ _ = refl
is-n-connected-is-prop 1 = is-prop-∥-∥
is-n-connected-is-prop 2 = is-contr-is-prop
is-n-connected-is-prop (suc (suc (suc n))) = is-contr-is-prop is-n-connected-is-prop
For short, we say that a type is connected if it is -connected, and simply connected if it is -connected:
: ∀ {ℓ} → Type ℓ → Type _
is-connected = is-n-connected A 2
is-connected A
: ∀ {ℓ} → Type ℓ → Type _
is-simply-connected = is-n-connected A 3 is-simply-connected A
Pointed connected types🔗
In the case of pointed types, there is an equivalent definition of being connected that is sometimes easier to work with: a pointed type is connected if every point is merely equal to the base point.
: ∀ {ℓ} → Type∙ ℓ → Type _
is-connected∙ (X , pt) = (x : X) → ∥ x ≡ pt ∥
is-connected∙
module _ {ℓ} {X@(_ , pt) : Type∙ ℓ} where
: is-connected∙ X → is-connected ⌞ X ⌟
is-connected∙→is-connected .centre = inc pt
is-connected∙→is-connected c .paths =
is-connected∙→is-connected c (λ _ → is-hlevel-suc 2 squash (inc pt) _) λ x →
∥-∥₀-elim {pprop = squash _ _} (ap inc ∘ sym) (c x)
∥-∥-rec!
: is-connected ⌞ X ⌟ → is-connected∙ X
is-connected→is-connected∙ =
is-connected→is-connected∙ c x .to (is-contr→is-prop c (inc x) (inc pt)) ∥-∥₀-path
This alternative definition lets us formulate a nice elimination principle for pointed connected types: any family of propositions that holds on the base point holds everywhere.
In particular, since being a proposition is a proposition
,
we only need to check that
is a proposition.
connected∙-elim-prop: ∀ {ℓ ℓ'} {X : Type∙ ℓ} {P : ⌞ X ⌟ → Type ℓ'}
→ is-connected∙ X
→ is-prop (P (X .snd))
→ P (X .snd)
→ ∀ x → P x
{X = X} {P} conn prop pb x =
connected∙-elim-prop (P-is-prop x) (λ e → subst P (sym e) pb) (conn x)
∥-∥-rec where abstract
: ∀ x → is-prop (P x)
P-is-prop = ∥-∥-rec is-prop-is-prop (λ e → subst (is-prop ∘ P) (sym e) prop) (conn x) P-is-prop x
Examples of pointed connected types include the circle and the delooping of a group.
Closure properties🔗
As a property of types, -connectedness enjoys closure properties very similar to those of -truncatedness; and we establish them in much the same way, by proving that -connectedness is preserved by retractions.
However, the definition of is-n-connected
, which uses the explicit
constructions of truncations for the base cases, is slightly annoying to
work when
is arbitrary: that’s the trade-off for it being easy to work with when
is a known, and usually small, number. Therefore, we prove the following
lemma by the recursion, establishing that the notion of connectedness
could very well have been defined using the general
-truncation
uniformly.
: ∀ {ℓ} {A : Type ℓ} n → is-n-connected A (suc n) → is-contr (n-Tr A (suc n))
is-n-connected-Tr = ∥-∥-proj! do
is-n-connected-Tr zero a-conn
pt ← a-conn(inc pt) (λ x → n-Tr-is-hlevel 0 _ _)
pure $ contr (suc zero) a-conn =
is-n-connected-Tr 0
retract→is-hlevel (∥-∥₀-rec (n-Tr-is-hlevel 1) inc)
(n-Tr-rec squash inc)
(n-Tr-elim _ (λ _ → is-prop→is-set (n-Tr-is-hlevel 1 _ _)) λ _ → refl)
a-conn(suc (suc n)) a-conn = a-conn is-n-connected-Tr
The first closure property we’ll prove is very applicable: it says that any retract of an -connected type is again -connected. Intuitively this is because any retraction can be extended to a retraction , and contractible types are closed under retractions.
retract→is-n-connected: (n : Nat) (f : A → B) (g : B → A)
→ is-left-inverse f g → is-n-connected A n → is-n-connected B n
0 = _
retract→is-n-connected 1 f g h a-conn = f <$> a-conn
retract→is-n-connected (suc (suc n)) f g h a-conn =
retract→is-n-connected .from (suc n) $ retract→is-contr
n-connected(n-Tr-rec! (inc ∘ f))
(n-Tr-rec! (inc ∘ g))
(n-Tr-elim! _ λ x → ap n-Tr.inc (h x))
(is-n-connected-Tr (suc n) a-conn)
Since the truncation operator also preserves products, a remarkably similar argument shows that if and are -connected, then so is is.
×-is-n-connected: ∀ n → is-n-connected A n → is-n-connected B n → is-n-connected (A × B) n
0 = _
×-is-n-connected (suc n) a-conn b-conn = n-connected.from n $ is-hlevel≃ 0
×-is-n-connected (×-is-hlevel 0 (n-connected.to n a-conn) (n-connected.to n b-conn)) n-Tr-product
Finally, we show the dual of two properties of truncations: if
is
-connected,
then each path space
is
-connected;
And if
is
-connected,
then it is also
)
connected. This latter property lets us count down in precisely
the same way that is-hlevel-suc
lets us count up.
: ∀ n → is-n-connected A (suc n) → is-n-connected (Path A x y) n
Path-is-connected 0 = _
Path-is-connected {x = x} (suc n) conn = n-connected.from n (contr (ps _ _) $
Path-is-connected _ (λ _ → hlevel!)
n-Tr-elim (J (λ y p → ps x y ≡ inc p) (Equiv.injective (Equiv.inverse n-Tr-path-equiv)
(is-contr→is-set (is-n-connected-Tr _ conn) _ _ _ _))))
where
: ∀ x y → n-Tr (x ≡ y) (suc n)
ps = Equiv.to n-Tr-path-equiv (is-contr→is-prop (is-n-connected-Tr _ conn) _ _)
ps x y
is-connected-suc: ∀ {ℓ} {A : Type ℓ} n → is-n-connected A (2 + n) → is-n-connected A (suc n)
{A = A} n w = n-connected.from n $ n-Tr-elim! _
is-connected-suc (λ x → contr (inc x) (n-Tr-elim _ (λ _ → hlevel!) (rem₁ n w x)))
(is-n-connected-Tr (suc n) w .centre)
where
: ∀ n → is-n-connected A (2 + n) → ∀ x y → Path (n-Tr A (suc n)) (inc x) (inc y)
rem₁ = n-Tr-is-hlevel 0 _ _
rem₁ zero a-conn x y (suc n) a-conn x y = Equiv.from n-Tr-path-equiv
rem₁ (n-Tr-rec (is-hlevel-suc _ (n-Tr-is-hlevel n)) inc
(is-n-connected-Tr _ (Path-is-connected (2 + n) a-conn) .centre))
In relation to truncatedness🔗
The following lemmas define -connectedness of a type by how it can map into the -truncated types. We then expand this idea to cover the -connected maps, too. At the level of types, the “relative” characterisation of -connectedness says that, if is -connected and is -truncated, then the function is an equivalence.
The intuitive idea behind this theorem is as follows: we have assumed that has no interesting information below dimension , and that has no interesting information in the dimensions and above. Therefore, the functions can’t be interesting, since they’re mapping boring data into boring data.
The proof is a direct application of the definition of -connectedness and the universal property of truncations: is equivalent to , but this is equivalent to since is contractible.
is-n-connected→n-type-const: ∀ n → is-hlevel B (suc n) → is-n-connected A (suc n)
→ is-equiv {B = A → B} (λ b a → b)
{B = B} {A = A} n B-hl A-conn =
is-n-connected→n-type-const (λ i x z → transp (λ i → B) i x) $ snd $
subst is-equiv (is-n-connected-Tr n A-conn) e⁻¹ ⟩
B ≃⟨ Π-contr-eqv (n-Tr A (suc n) → B) ≃⟨ n-Tr-univ n B-hl ⟩
(A → B) ≃∎
This direction of the theorem is actually half of an equivalence. In the other direction, since is -truncated by definition, we suppose that is an equivalence. From the constructor we obtain a point , and, for all , we have
But by induction on truncation, this says precisely that any is equal to ; so is a centre of contraction, and is -connected.
n-type-const→is-n-connected: ∀ {ℓ} {A : Type ℓ} n
→ (∀ {B : Type ℓ} → is-hlevel B (suc n) → is-equiv {B = A → B} (λ b a → b))
→ is-n-connected A (suc n)
{A = A} n eqv =
n-type-const→is-n-connected .from n $ contr (rem.from inc) $ n-Tr-elim _
n-connected(λ h → Path-is-hlevel (suc n) (n-Tr-is-hlevel n)) (rem.ε inc $ₚ_)
where module rem = Equiv (_ , eqv {n-Tr A (suc n)} hlevel!)
We can now generalise this to work with an -connected map and a family of -types over : in this setting, precomposition with is an equivalence
This is somewhat analogous to generalising from a recursion principle to an elimination principle. When we were limited to talking about types, we could extend points to functions , but no dependency was possible. With the extra generality, we think of as including a space of “constructors”, and the equivalence says that exhibiting at these constructors is equivalent to exhibiting it for the whole type.
relative-n-type-const: (P : B → Type ℓ'') (f : A → B)
→ ∀ n → is-n-connected-map f n
→ (∀ x → is-hlevel (P x) n)
→ is-equiv {A = (∀ b → P b)} {B = (∀ a → P (f a))} (_∘ f)
Despite the generality, the proof is just a calculation: observing that we have the following chain of equivalences suffices.
=
rem₁ ((b : B) → P b) ≃⟨ Π-cod≃ (λ x → Π-contr-eqv {B = λ _ → P x} (is-n-connected-Tr _ (n-conn x)) e⁻¹) ⟩
((b : B) → n-Tr (fibre f b) (suc n) → P b) ≃⟨ Π-cod≃ (λ x → n-Tr-univ n (phl _)) ⟩
((b : B) → fibre f b → P b) ≃⟨ shuffle ⟩
((a : A) → P (f a)) ≃∎
There’s also the shuffle
isomorphism that eliminates the fibre
argument using path induction, but
its construction is mechanical.
= Iso→Equiv λ where
shuffle .fst g a → g (f a) (a , refl)
.snd .inv g b (a , p) → subst P p (g a)
.snd .rinv g → funext λ a → transport-refl _
.snd .linv g → funext λ b → funext λ { (a , p) →
(λ b p → subst P p (g (f a) (a , refl)) ≡ g b (a , p))
J (transport-refl _) p }
We can specialise this to get a literal elimination principle: If is a family of -types over an -connected pointed type , then holds everywhere if holds. Moreover, this has the expected computation rule.
module _ n (P : A → Type ℓ) (tr : ∀ x → is-hlevel (P x) n)
{a₀ : A} (pa₀ : P a₀)
(a-conn : is-n-connected A (suc n))
where
: fibre (λ z x → z a₀) (λ _ → pa₀)
connected-elimination-principle = relative-n-type-const {A = ⊤} P (λ _ → a₀) n
connected-elimination-principle (λ x → retract→is-n-connected n (λ p → tt , p) snd (λ _ → refl)
(Path-is-connected n a-conn))
.is-eqv (λ _ → pa₀) .centre
tr
opaque: ∀ x → P x
connected-elim = connected-elimination-principle .fst
connected-elim
: connected-elim a₀ ≡ pa₀
connected-elim-β = connected-elimination-principle .snd $ₚ tt connected-elim-β
Using these elimination principles, we can prove that a pointed type is -connected if and only if the inclusion of is -connected when considered as a map .
is-n-connected-point: ∀ {ℓ} {A : Type ℓ} (a₀ : A) n
→ is-n-connected-map {A = ⊤} (λ _ → a₀) n
→ is-n-connected A (suc n)
{A = A} a₀ n pt-conn = done where
is-n-connected-point : ∀ {B : Type ℓ} → is-hlevel B (suc n) → ∀ (f : A → B) a → f a₀ ≡ f a
rem = equiv→inverse
rem b-hl f (relative-n-type-const (λ a → f a₀ ≡ f a) _ n pt-conn λ x → Path-is-hlevel' n b-hl _ _)
(λ _ → refl)
: is-n-connected A (suc n)
done = n-type-const→is-n-connected n λ hl → is-iso→is-equiv $
done (λ f → f a₀) (λ f → funext λ a → rem hl f a) λ _ → refl
iso
point-is-n-connected: ∀ {ℓ} {A : Type ℓ} (a₀ : A) n
→ is-n-connected A (2 + n)
→ is-n-connected-map {A = ⊤} (λ _ → a₀) (suc n)
=
point-is-n-connected a₀ n a-conn (suc n) (λ x → is-n-connected (⊤ × (a₀ ≡ x)) (suc n))
connected-elim (λ x → is-prop→is-hlevel-suc (is-n-connected-is-prop (suc n)))
(retract→is-n-connected (suc n) (tt ,_) snd (λ _ → refl)
(Path-is-connected {y = a₀} (suc n) a-conn))
a-conn