The universal property of the unit type
Content created by Fredrik Bakke, Egbert Rijke and Jonathan Prieto-Cubides
Created on 2022-02-08.
Last modified on 2023-09-11.
module foundation.universal-property-unit-type where
Imports
open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.unit-type 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.functoriality-function-types open import foundation-core.homotopies
Idea
The universal property of the unit type characterizes maps out of the unit type. Similarly, the dependent universal property of the unit type characterizes dependent functions out of the unit type.
In foundation.contractible-types
we have alread proven related universal
properties of contractible types.
Properties
ev-star : {l : Level} (P : unit → UU l) → ((x : unit) → P x) → P star ev-star P f = f star ev-star' : {l : Level} (Y : UU l) → (unit → Y) → Y ev-star' Y = ev-star (λ t → Y) abstract dependent-universal-property-unit : {l : Level} (P : unit → UU l) → is-equiv (ev-star P) dependent-universal-property-unit = dependent-universal-property-contr-is-contr star is-contr-unit equiv-dependent-universal-property-unit : {l : Level} (P : unit → UU l) → ((x : unit) → P x) ≃ P star pr1 (equiv-dependent-universal-property-unit P) = ev-star P pr2 (equiv-dependent-universal-property-unit P) = dependent-universal-property-unit P abstract universal-property-unit : {l : Level} (Y : UU l) → is-equiv (ev-star' Y) universal-property-unit Y = dependent-universal-property-unit (λ t → Y) equiv-universal-property-unit : {l : Level} (Y : UU l) → (unit → Y) ≃ Y pr1 (equiv-universal-property-unit Y) = ev-star' Y pr2 (equiv-universal-property-unit Y) = universal-property-unit Y abstract is-equiv-point-is-contr : {l1 : Level} {X : UU l1} (x : X) → is-contr X → is-equiv (point x) is-equiv-point-is-contr x is-contr-X = is-equiv-is-contr (point x) is-contr-unit is-contr-X abstract is-equiv-point-universal-property-unit : {l1 : Level} (X : UU l1) (x : X) → ((l2 : Level) (Y : UU l2) → is-equiv (λ (f : X → Y) → f x)) → is-equiv (point x) is-equiv-point-universal-property-unit X x H = is-equiv-is-equiv-precomp ( point x) ( λ l Y → is-equiv-right-factor ( ev-star' Y) ( precomp (point x) Y) ( universal-property-unit Y) ( H _ Y)) abstract universal-property-unit-is-equiv-point : {l1 : Level} {X : UU l1} (x : X) → is-equiv (point x) → ({l2 : Level} (Y : UU l2) → is-equiv (λ (f : X → Y) → f x)) universal-property-unit-is-equiv-point x is-equiv-point Y = is-equiv-comp ( ev-star' Y) ( precomp (point x) Y) ( is-equiv-precomp-is-equiv (point x) is-equiv-point Y) ( universal-property-unit Y) abstract universal-property-unit-is-contr : {l1 : Level} {X : UU l1} (x : X) → is-contr X → ({l2 : Level} (Y : UU l2) → is-equiv (λ (f : X → Y) → f x)) universal-property-unit-is-contr x is-contr-X = universal-property-unit-is-equiv-point x ( is-equiv-point-is-contr x is-contr-X) abstract is-equiv-diagonal-is-equiv-point : {l1 : Level} {X : UU l1} (x : X) → is-equiv (point x) → ({l2 : Level} (Y : UU l2) → is-equiv (λ y → const X Y y)) is-equiv-diagonal-is-equiv-point {X = X} x is-equiv-point Y = is-equiv-is-section ( universal-property-unit-is-equiv-point x is-equiv-point Y) ( refl-htpy)
Recent changes
- 2023-09-11. Fredrik Bakke and Egbert Rijke. Some computations for different notions of equivalence (#711).
- 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-05-03. Fredrik Bakke. Define the smash product of pointed types and refactor pointed types (#583).
- 2023-03-21. Fredrik Bakke. Formatting fixes (#530).