diff --git a/src/foundation-core/equality-dependent-pair-types.lagda.md b/src/foundation-core/equality-dependent-pair-types.lagda.md index 2fbbc0134c..327d02e327 100644 --- a/src/foundation-core/equality-dependent-pair-types.lagda.md +++ b/src/foundation-core/equality-dependent-pair-types.lagda.md @@ -109,13 +109,13 @@ module _ η-pair : (t : Σ A B) → (pair (pr1 t) (pr2 t)) = t η-pair t = refl - eq-base-eq-pair : {s t : Σ A B} → (s = t) → (pr1 s = pr1 t) - eq-base-eq-pair p = pr1 (pair-eq-Σ p) + eq-base-eq-pair-Σ : {s t : Σ A B} → (s = t) → (pr1 s = pr1 t) + eq-base-eq-pair-Σ p = pr1 (pair-eq-Σ p) - dependent-eq-family-eq-pair : + dependent-eq-family-eq-pair-Σ : {s t : Σ A B} → (p : s = t) → - dependent-identification B (eq-base-eq-pair p) (pr2 s) (pr2 t) - dependent-eq-family-eq-pair p = pr2 (pair-eq-Σ p) + dependent-identification B (eq-base-eq-pair-Σ p) (pr2 s) (pr2 t) + dependent-eq-family-eq-pair-Σ p = pr2 (pair-eq-Σ p) ``` ### Lifting equality to the total space diff --git a/src/foundation-core/equivalences.lagda.md b/src/foundation-core/equivalences.lagda.md index f46566c15a..23da5bdb52 100644 --- a/src/foundation-core/equivalences.lagda.md +++ b/src/foundation-core/equivalences.lagda.md @@ -713,6 +713,10 @@ module _ (e : A ≃ B) (x y : A) → (x = y) ≃ (map-equiv e x = map-equiv e y) pr1 (equiv-ap e x y) = ap (map-equiv e) pr2 (equiv-ap e x y) = is-emb-is-equiv (is-equiv-map-equiv e) x y + + map-inv-equiv-ap : + (e : A ≃ B) (x y : A) → (map-equiv e x = map-equiv e y) → (x = y) + map-inv-equiv-ap e x y = map-inv-equiv (equiv-ap e x y) ``` ## Equivalence reasoning diff --git a/src/foundation/epimorphisms.lagda.md b/src/foundation/epimorphisms.lagda.md index 5505581352..d3a14b661e 100644 --- a/src/foundation/epimorphisms.lagda.md +++ b/src/foundation/epimorphisms.lagda.md @@ -182,6 +182,25 @@ If the map `f : A → B` is epi, then its codiagonal is an equivalence. is-epimorphism-is-pushout = is-epimorphism-is-equiv-codiagonal-map ``` +### The class of epimorphisms is closed under composition and has the right cancellation property + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + (g : B → C) (f : A → B) + where + + is-epimorphism-comp : + is-epimorphism g → is-epimorphism f → is-epimorphism (g ∘ f) + is-epimorphism-comp eg ef X = + is-emb-comp (precomp f X) (precomp g X) (ef X) (eg X) + + is-epimorphism-left-factor : + is-epimorphism (g ∘ f) → is-epimorphism f → is-epimorphism g + is-epimorphism-left-factor ec ef X = + is-emb-right-factor (precomp f X) (precomp g X) (ef X) (ec X) +``` + ## See also - [Acyclic maps](synthetic-homotopy-theory.acyclic-maps.md) diff --git a/src/foundation/equality-dependent-pair-types.lagda.md b/src/foundation/equality-dependent-pair-types.lagda.md index b6f577eadc..7048e772fa 100644 --- a/src/foundation/equality-dependent-pair-types.lagda.md +++ b/src/foundation/equality-dependent-pair-types.lagda.md @@ -11,9 +11,15 @@ open import foundation-core.equality-dependent-pair-types public ```agda open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions +open import foundation.contractible-types open import foundation.dependent-identifications open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies open import foundation.transport-along-identifications +open import foundation.type-arithmetic-dependent-pair-types open import foundation.universe-levels open import foundation-core.function-types @@ -113,7 +119,7 @@ module _ ```agda module _ - { l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {Y : UU l3} (f : Σ A B → Y) + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {Y : UU l3} (f : Σ A B → Y) where compute-ap-eq-pair-Σ : @@ -127,7 +133,7 @@ module _ ```agda module _ - { l1 l2 : Level} {A : UU l1} (B : A → UU l2) + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) where triangle-eq-pair-Σ : @@ -137,6 +143,96 @@ module _ triangle-eq-pair-Σ refl q = refl ``` +### Computing identifications in iterated dependent pair types + +#### Computing identifications in dependent pair types of the form Σ (Σ A B) C + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : Σ A B → UU l3} + where + + Eq-Σ²' : (s t : Σ (Σ A B) C) → UU (l1 ⊔ l2 ⊔ l3) + Eq-Σ²' s t = + Σ ( Eq-Σ (pr1 s) (pr1 t)) + ( λ q → dependent-identification C (eq-pair-Σ' q) (pr2 s) (pr2 t)) + + equiv-triple-eq-Σ' : + (s t : Σ (Σ A B) C) → + (s = t) ≃ Eq-Σ²' s t + equiv-triple-eq-Σ' s t = + ( equiv-Σ + ( λ q → + ( dependent-identification + ( C) + ( eq-pair-Σ' q) + ( pr2 s) + ( pr2 t))) + ( equiv-pair-eq-Σ (pr1 s) (pr1 t)) + ( λ p → + ( equiv-tr + ( λ q → dependent-identification C q (pr2 s) (pr2 t)) + ( inv (is-section-pair-eq-Σ (pr1 s) (pr1 t) p))))) ∘e + ( equiv-pair-eq-Σ s t) + + triple-eq-Σ' : + (s t : Σ (Σ A B) C) → + (s = t) → Eq-Σ²' s t + triple-eq-Σ' s t = map-equiv (equiv-triple-eq-Σ' s t) +``` + +#### Computing dependent identifications on the second component + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : (x : A) → B x → UU l3} + where + + coh-eq-base-Σ² : + {s t : Σ A (λ x → Σ (B x) λ y → C x y)} (p : s = t) → + eq-base-eq-pair-Σ p = + eq-base-eq-pair-Σ (eq-base-eq-pair-Σ (ap (map-inv-associative-Σ' A B C) p)) + coh-eq-base-Σ² refl = refl + + dependent-eq-second-component-eq-Σ² : + {s t : Σ A (λ x → Σ (B x) λ y → C x y)} (p : s = t) → + dependent-identification B (eq-base-eq-pair-Σ p) (pr1 (pr2 s)) (pr1 (pr2 t)) + dependent-eq-second-component-eq-Σ² {s = s} {t = t} p = + ( ap (λ q → tr B q (pr1 (pr2 s))) (coh-eq-base-Σ² p)) ∙ + ( pr2 + ( pr1 + ( triple-eq-Σ' + ( map-inv-associative-Σ' A B C s) + ( map-inv-associative-Σ' A B C t) + ( ap (map-inv-associative-Σ' A B C) p)))) +``` + +#### Computing dependent identifications on the third component + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3} + (D : (x : A) → B x → C x → UU l4) + where + + coh-eq-base-Σ³ : + { s t : Σ A (λ x → Σ (B x) (λ y → Σ (C x) (D x y)))} (p : s = t) → + eq-base-eq-pair-Σ p = + eq-base-eq-pair-Σ (ap (map-equiv (interchange-Σ-Σ-Σ D)) p) + coh-eq-base-Σ³ refl = refl + + dependent-eq-third-component-eq-Σ³ : + { s t : Σ A (λ x → Σ (B x) (λ y → Σ (C x) (D x y)))} (p : s = t) → + dependent-identification C + ( eq-base-eq-pair-Σ p) + ( pr1 (pr2 (pr2 s))) + ( pr1 (pr2 (pr2 t))) + dependent-eq-third-component-eq-Σ³ {s = s} {t = t} p = + ( ap (λ q → tr C q (pr1 (pr2 (pr2 s)))) (coh-eq-base-Σ³ p)) ∙ + ( dependent-eq-second-component-eq-Σ² + ( ap (map-equiv (interchange-Σ-Σ-Σ D)) p)) +``` + ## See also - Equality proofs in cartesian product types are characterized in diff --git a/src/foundation/type-arithmetic-dependent-pair-types.lagda.md b/src/foundation/type-arithmetic-dependent-pair-types.lagda.md index 4e4bbb8561..ebca7b4ac0 100644 --- a/src/foundation/type-arithmetic-dependent-pair-types.lagda.md +++ b/src/foundation/type-arithmetic-dependent-pair-types.lagda.md @@ -305,6 +305,14 @@ module _ pr1 interchange-Σ-Σ = map-interchange-Σ-Σ pr2 interchange-Σ-Σ = is-equiv-map-interchange-Σ-Σ + interchange-Σ-Σ-Σ : + Σ A (λ x → Σ (B x) (λ y → Σ (C x) (D x y))) ≃ + Σ A (λ x → Σ (C x) (λ z → Σ (B x) λ y → D x y z)) + interchange-Σ-Σ-Σ = + associative-Σ' A C (λ x z → Σ (B x) λ y → D x y z) ∘e + interchange-Σ-Σ ∘e + inv-associative-Σ' A B (λ x y → Σ (C x) (D x y)) + eq-interchange-Σ-Σ-is-contr : {a : A} {b : B a} → is-torsorial (D a b) → {x y : Σ (C a) (D a b)} → diff --git a/src/synthetic-homotopy-theory/acyclic-maps.lagda.md b/src/synthetic-homotopy-theory/acyclic-maps.lagda.md index a9f4aa9d5b..bceb52e31a 100644 --- a/src/synthetic-homotopy-theory/acyclic-maps.lagda.md +++ b/src/synthetic-homotopy-theory/acyclic-maps.lagda.md @@ -209,6 +209,34 @@ module _ ( is-acyclic-map-is-epimorphism f e) ``` +### The class of acyclic maps is closed under composition and has the right cancellation property + +Since the acyclic maps are precisely the epimorphisms this follows from the +corresponding facts about [epimorphisms](foundation.epimorphisms.md). + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + (g : B → C) (f : A → B) + where + + is-acyclic-map-comp : + is-acyclic-map g → is-acyclic-map f → is-acyclic-map (g ∘ f) + is-acyclic-map-comp ag af = + is-acyclic-map-is-epimorphism (g ∘ f) + ( is-epimorphism-comp g f + ( is-epimorphism-is-acyclic-map g ag) + ( is-epimorphism-is-acyclic-map f af)) + + is-acyclic-map-left-factor : + is-acyclic-map (g ∘ f) → is-acyclic-map f → is-acyclic-map g + is-acyclic-map-left-factor ac af = + is-acyclic-map-is-epimorphism g + ( is-epimorphism-left-factor g f + ( is-epimorphism-is-acyclic-map (g ∘ f) ac) + ( is-epimorphism-is-acyclic-map f af)) +``` + ## See also - [Dependent epimorphisms](foundation.dependent-epimorphisms.md)