diff --git a/src/category-theory/split-essentially-surjective-functors-precategories.lagda.md b/src/category-theory/split-essentially-surjective-functors-precategories.lagda.md index 1a8bb8a958..37871d352f 100644 --- a/src/category-theory/split-essentially-surjective-functors-precategories.lagda.md +++ b/src/category-theory/split-essentially-surjective-functors-precategories.lagda.md @@ -151,9 +151,8 @@ module _ where is-essentially-surjective-is-split-essentially-surjective-functor-Precategory : - (F : functor-Precategory C D) - (is-split-ess-surj-F : - is-split-essentially-surjective-functor-Precategory C D F) → + (F : functor-Precategory C D) → + is-split-essentially-surjective-functor-Precategory C D F → is-essentially-surjective-functor-Precategory C D F is-essentially-surjective-is-split-essentially-surjective-functor-Precategory F is-split-ess-surj-F = @@ -179,7 +178,7 @@ Rezk completion_. ([arXiv:1303.0584](https://arxiv.org/abs/1303.0584), [DOI:10.1017/S0960129514000486](https://doi.org/10.1017/S0960129514000486)) -## External link +## External links - [Essential Fibres](https://1lab.dev/Cat.Functor.Properties.html#essential-fibres) at 1lab diff --git a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md index dba98f35f5..1788b0f6b9 100644 --- a/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md +++ b/src/elementary-number-theory/fundamental-theorem-of-arithmetic.lagda.md @@ -748,8 +748,7 @@ is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H nil D y d p = ( 1) ( inv (is-decomposition-list-is-prime-decomposition-list-ℕ x nil D)))) is-in-prime-decomposition-is-nontrivial-prime-divisor-ℕ x H (cons z l) D y d p = - ind-coprod - ( λ _ → y ∈-list (cons z l)) + rec-coprod ( λ e → tr (λ w → w ∈-list (cons z l)) (inv e) (is-head z l)) ( λ e → is-in-tail @@ -947,7 +946,7 @@ pr2 (fundamental-theorem-arithmetic-list-ℕ x H) d = ```agda is-prime-list-concat-list-ℕ : (p q : list ℕ) → is-prime-list-ℕ p → is-prime-list-ℕ q → - is-prime-list-ℕ (concat-list p q) + is-prime-list-ℕ (concat-list p q) is-prime-list-concat-list-ℕ nil q Pp Pq = Pq is-prime-list-concat-list-ℕ (cons x p) q Pp Pq = pr1 Pp , is-prime-list-concat-list-ℕ p q (pr2 Pp) Pq diff --git a/src/foundation-core/coproduct-types.lagda.md b/src/foundation-core/coproduct-types.lagda.md index 2849c4751f..850d17ff10 100644 --- a/src/foundation-core/coproduct-types.lagda.md +++ b/src/foundation-core/coproduct-types.lagda.md @@ -19,6 +19,8 @@ of `A` and `B`. ## Definition +### Coproduct types + ```agda infixr 10 _+_ @@ -26,7 +28,11 @@ data _+_ {l1 l2 : Level} (A : UU l1) (B : UU l2) : UU (l1 ⊔ l2) where inl : A → A + B inr : B → A + B +``` +### The induction principle for coproduct types + +```agda ind-coprod : {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : A + B → UU l3) → ((x : A) → C (inl x)) → ((y : B) → C (inr y)) → @@ -34,3 +40,12 @@ ind-coprod : ind-coprod C f g (inl x) = f x ind-coprod C f g (inr x) = g x ``` + +### The recursion principle for coproduct types + +```agda +rec-coprod : + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} → + (A → C) → (B → C) → (A + B) → C +rec-coprod {C = C} = ind-coprod (λ _ → C) +``` diff --git a/src/foundation-core/fibers-of-maps.lagda.md b/src/foundation-core/fibers-of-maps.lagda.md index 5d25a4ad41..3533e4d0b5 100644 --- a/src/foundation-core/fibers-of-maps.lagda.md +++ b/src/foundation-core/fibers-of-maps.lagda.md @@ -15,6 +15,8 @@ open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.retractions +open import foundation-core.sections open import foundation-core.transport-along-identifications ``` @@ -79,11 +81,13 @@ module _ eq-Eq-fiber α β = eq-Eq-fiber-uncurry (α , β) is-section-eq-Eq-fiber : - {s t : fiber f b} → Eq-eq-fiber {s} {t} ∘ eq-Eq-fiber-uncurry {s} {t} ~ id + {s t : fiber f b} → + is-section (Eq-eq-fiber {s} {t}) (eq-Eq-fiber-uncurry {s} {t}) is-section-eq-Eq-fiber (refl , refl) = refl is-retraction-eq-Eq-fiber : - {s t : fiber f b} → eq-Eq-fiber-uncurry {s} {t} ∘ Eq-eq-fiber {s} {t} ~ id + {s t : fiber f b} → + is-retraction (Eq-eq-fiber {s} {t}) (eq-Eq-fiber-uncurry {s} {t}) is-retraction-eq-Eq-fiber refl = refl abstract @@ -144,12 +148,12 @@ module _ is-section-eq-Eq-fiber' : {s t : fiber' f b} → - Eq-eq-fiber' {s} {t} ∘ eq-Eq-fiber-uncurry' {s} {t} ~ id + is-section (Eq-eq-fiber' {s} {t}) (eq-Eq-fiber-uncurry' {s} {t}) is-section-eq-Eq-fiber' {x , refl} (refl , refl) = refl is-retraction-eq-Eq-fiber' : {s t : fiber' f b} → - eq-Eq-fiber-uncurry' {s} {t} ∘ Eq-eq-fiber' {s} {t} ~ id + is-retraction (Eq-eq-fiber' {s} {t}) (eq-Eq-fiber-uncurry' {s} {t}) is-retraction-eq-Eq-fiber' {x , refl} refl = refl abstract @@ -187,17 +191,19 @@ module _ where map-equiv-fiber : fiber f y → fiber' f y - pr1 (map-equiv-fiber (x , refl)) = x - pr2 (map-equiv-fiber (x , refl)) = refl + pr1 (map-equiv-fiber (x , _)) = x + pr2 (map-equiv-fiber (x , p)) = inv p map-inv-equiv-fiber : fiber' f y → fiber f y - pr1 (map-inv-equiv-fiber (x , refl)) = x - pr2 (map-inv-equiv-fiber (x , refl)) = refl + pr1 (map-inv-equiv-fiber (x , _)) = x + pr2 (map-inv-equiv-fiber (x , p)) = inv p - is-section-map-inv-equiv-fiber : map-equiv-fiber ∘ map-inv-equiv-fiber ~ id + is-section-map-inv-equiv-fiber : + is-section map-equiv-fiber map-inv-equiv-fiber is-section-map-inv-equiv-fiber (x , refl) = refl - is-retraction-map-inv-equiv-fiber : map-inv-equiv-fiber ∘ map-equiv-fiber ~ id + is-retraction-map-inv-equiv-fiber : + is-retraction map-equiv-fiber map-inv-equiv-fiber is-retraction-map-inv-equiv-fiber (x , refl) = refl is-equiv-map-equiv-fiber : is-equiv map-equiv-fiber @@ -227,19 +233,21 @@ module _ pr2 (pr1 (map-inv-fiber-pr1 b)) = b pr2 (map-inv-fiber-pr1 b) = refl - is-section-map-inv-fiber-pr1 : (map-inv-fiber-pr1 ∘ map-fiber-pr1) ~ id - is-section-map-inv-fiber-pr1 ((.a , y) , refl) = refl + is-section-map-inv-fiber-pr1 : + is-section map-fiber-pr1 map-inv-fiber-pr1 + is-section-map-inv-fiber-pr1 b = refl - is-retraction-map-inv-fiber-pr1 : (map-fiber-pr1 ∘ map-inv-fiber-pr1) ~ id - is-retraction-map-inv-fiber-pr1 b = refl + is-retraction-map-inv-fiber-pr1 : + is-retraction map-fiber-pr1 map-inv-fiber-pr1 + is-retraction-map-inv-fiber-pr1 ((.a , y) , refl) = refl abstract is-equiv-map-fiber-pr1 : is-equiv map-fiber-pr1 is-equiv-map-fiber-pr1 = is-equiv-is-invertible map-inv-fiber-pr1 - is-retraction-map-inv-fiber-pr1 is-section-map-inv-fiber-pr1 + is-retraction-map-inv-fiber-pr1 equiv-fiber-pr1 : fiber (pr1 {B = B}) a ≃ B a pr1 equiv-fiber-pr1 = map-fiber-pr1 @@ -250,8 +258,8 @@ module _ is-equiv-map-inv-fiber-pr1 = is-equiv-is-invertible map-fiber-pr1 - is-section-map-inv-fiber-pr1 is-retraction-map-inv-fiber-pr1 + is-section-map-inv-fiber-pr1 inv-equiv-fiber-pr1 : B a ≃ fiber (pr1 {B = B}) a pr1 inv-equiv-fiber-pr1 = map-inv-fiber-pr1 @@ -265,10 +273,10 @@ module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) where - map-equiv-total-fiber : (Σ B (fiber f)) → A + map-equiv-total-fiber : Σ B (fiber f) → A map-equiv-total-fiber t = pr1 (pr2 t) - triangle-map-equiv-total-fiber : pr1 ~ (f ∘ map-equiv-total-fiber) + triangle-map-equiv-total-fiber : pr1 ~ f ∘ map-equiv-total-fiber triangle-map-equiv-total-fiber t = inv (pr2 (pr2 t)) map-inv-equiv-total-fiber : A → Σ B (fiber f) @@ -277,11 +285,11 @@ module _ pr2 (pr2 (map-inv-equiv-total-fiber x)) = refl is-retraction-map-inv-equiv-total-fiber : - (map-inv-equiv-total-fiber ∘ map-equiv-total-fiber) ~ id + is-retraction map-equiv-total-fiber map-inv-equiv-total-fiber is-retraction-map-inv-equiv-total-fiber (.(f x) , x , refl) = refl is-section-map-inv-equiv-total-fiber : - (map-equiv-total-fiber ∘ map-inv-equiv-total-fiber) ~ id + is-section map-equiv-total-fiber map-inv-equiv-total-fiber is-section-map-inv-equiv-total-fiber x = refl abstract @@ -329,11 +337,11 @@ module _ pr2 (inv-map-compute-fiber-comp t) = ap g (pr2 (pr2 t)) ∙ pr2 (pr1 t) is-section-inv-map-compute-fiber-comp : - (map-compute-fiber-comp ∘ inv-map-compute-fiber-comp) ~ id + is-section map-compute-fiber-comp inv-map-compute-fiber-comp is-section-inv-map-compute-fiber-comp ((.(h a) , refl) , (a , refl)) = refl is-retraction-inv-map-compute-fiber-comp : - (inv-map-compute-fiber-comp ∘ map-compute-fiber-comp) ~ id + is-retraction map-compute-fiber-comp inv-map-compute-fiber-comp is-retraction-inv-map-compute-fiber-comp (a , refl) = refl abstract diff --git a/src/foundation-core/functoriality-dependent-function-types.lagda.md b/src/foundation-core/functoriality-dependent-function-types.lagda.md index d2ccfdb959..cc5b313869 100644 --- a/src/foundation-core/functoriality-dependent-function-types.lagda.md +++ b/src/foundation-core/functoriality-dependent-function-types.lagda.md @@ -52,8 +52,7 @@ compute-fiber-map-Π : {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} (f : (i : I) → A i → B i) (h : (i : I) → B i) → ((i : I) → fiber (f i) (h i)) ≃ fiber (map-Π f) h -compute-fiber-map-Π f h = - equiv-tot (λ _ → equiv-eq-htpy) ∘e distributive-Π-Σ +compute-fiber-map-Π f h = equiv-tot (λ _ → equiv-eq-htpy) ∘e distributive-Π-Σ compute-fiber-map-Π' : {l1 l2 l3 l4 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} @@ -109,7 +108,7 @@ module _ compute-inv-equiv-Π-equiv-family : (e : (x : A) → B x ≃ C x) → ( map-inv-equiv (equiv-Π-equiv-family e)) ~ - ( map-equiv (equiv-Π-equiv-family λ x → (inv-equiv (e x)))) + ( map-equiv (equiv-Π-equiv-family (λ x → inv-equiv (e x)))) compute-inv-equiv-Π-equiv-family e f = is-injective-map-equiv ( equiv-Π-equiv-family e) diff --git a/src/foundation-core/pullbacks.lagda.md b/src/foundation-core/pullbacks.lagda.md index 3463c68207..a776424bfe 100644 --- a/src/foundation-core/pullbacks.lagda.md +++ b/src/foundation-core/pullbacks.lagda.md @@ -19,6 +19,7 @@ open import foundation.type-arithmetic-dependent-pair-types open import foundation.universe-levels open import foundation-core.cartesian-product-types +open import foundation-core.commuting-squares-of-maps open import foundation-core.diagonal-maps-of-types open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalences @@ -27,12 +28,48 @@ open import foundation-core.fibers-of-maps open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies +open import foundation-core.retractions +open import foundation-core.sections open import foundation-core.type-theoretic-principle-of-choice open import foundation-core.universal-property-pullbacks ``` +## Idea + +Given a [cospan of types](foundation.cospans.md) + +```text + f : A → X ← B : g, +``` + +we can form the +{{#concept "standard pullback" Disambiguation="types" Agda=standard-pullback}} +`A ×_X B` satisfying +[the universal property of the pullback](foundation-core.universal-property-pullbacks.md) +of the cospan, completing the diagram + +```text + A ×_X B ------> B + | ⌟ | + | | g + | | + v v + A ---------> X. + f +``` + +The standard pullback consists of [pairs](foundation.dependent-pair-types.md) +`a : A` and `b : B` such that `f a` and `g b` agree: + +```text + A ×_X B := Σ (a : A) (b : B), (f a = g b). +``` + +Thus the standard [cone](foundation.cones-over-cospans.md) consists of the +canonical projections. + ## Definitions ### The standard pullback of a cospan @@ -56,8 +93,11 @@ module _ horizontal-map-standard-pullback t = pr1 (pr2 t) coherence-square-standard-pullback : - ( f ∘ vertical-map-standard-pullback) ~ - ( g ∘ horizontal-map-standard-pullback) + coherence-square-maps + horizontal-map-standard-pullback + vertical-map-standard-pullback + g + f coherence-square-standard-pullback t = pr2 (pr2 t) ``` @@ -76,8 +116,9 @@ module _ ### The gap map into the standard pullback -The **gap map** of a square is the map fron the vertex of the -[cone](foundation.cones-over-cospans.md) into the standard pullback. +The {{#concept "standard gap map" Disambiguation="cone over a cospan" Agda=gap}} +of a [commuting square](foundation-core.commuting-squares-of-maps.md) is the map +from the domain of the cone into the standard pullback. ```agda module _ @@ -91,7 +132,7 @@ module _ pr2 (pr2 (gap c z)) = coherence-square-cone f g c z ``` -### The property of being a pullback +### The small property of being a pullback The [proposition](foundation-core.propositions.md) `is-pullback` is the assertion that the gap map is an [equivalence](foundation-core.equivalences.md). @@ -136,8 +177,7 @@ module _ (t t' : standard-pullback f g) → (t = t') ≃ Eq-standard-pullback t t' extensionality-standard-pullback (a , b , p) = extensionality-Σ - ( λ {a'} bp' α → - Σ (b = pr1 bp') (λ β → ap f α ∙ pr2 bp' = p ∙ ap g β)) + ( λ bp' α → Σ (b = pr1 bp') (λ β → ap f α ∙ pr2 bp' = p ∙ ap g β)) ( refl) ( refl , inv right-unit) ( λ x → id-equiv) @@ -157,9 +197,7 @@ module _ ( coherence-square-standard-pullback s ∙ ap g β)) → s = t map-extensionality-standard-pullback {s} {t} α β γ = - map-inv-equiv - ( extensionality-standard-pullback s t) - ( α , β , γ) + map-inv-equiv (extensionality-standard-pullback s t) (α , β , γ) ``` ### The standard pullback satisfies the universal property of pullbacks @@ -180,7 +218,7 @@ module _ ( is-equiv-tot-is-fiberwise-equiv (λ _ → is-equiv-map-distributive-Π-Σ)) ``` -### A cone is equal to the value of cone-map at its own gap map +### A cone is equal to the value of `cone-map` at its own gap map ```agda module _ @@ -231,6 +269,21 @@ module _ ### The pullback of a Σ-type +Squares of the form + +```text + Σ (x : A) (Q (f x)) ----> Σ (y : B) Q + | | + | | + pr1 | | pr1 + | | + V V + A -----------------------> B + f +``` + +are pullbacks. + ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (Q : B → UU l3) @@ -246,21 +299,22 @@ module _ inv-gap-cone-standard-pullback-Σ : standard-pullback f (pr1 {B = Q}) → standard-pullback-Σ - pr1 (inv-gap-cone-standard-pullback-Σ (x , (.(f x) , q) , refl)) = x + pr1 (inv-gap-cone-standard-pullback-Σ (x , _)) = x pr2 (inv-gap-cone-standard-pullback-Σ (x , (.(f x) , q) , refl)) = q abstract is-section-inv-gap-cone-standard-pullback-Σ : - ( gap f pr1 cone-standard-pullback-Σ ∘ inv-gap-cone-standard-pullback-Σ) ~ - ( id) - is-section-inv-gap-cone-standard-pullback-Σ (x , (.(f x) , q) , refl) = - refl + is-section + ( gap f pr1 cone-standard-pullback-Σ) + ( inv-gap-cone-standard-pullback-Σ) + is-section-inv-gap-cone-standard-pullback-Σ (x , (.(f x) , q) , refl) = refl abstract is-retraction-inv-gap-cone-standard-pullback-Σ : - ( inv-gap-cone-standard-pullback-Σ ∘ - gap f pr1 cone-standard-pullback-Σ) ~ id - is-retraction-inv-gap-cone-standard-pullback-Σ (x , q) = refl + is-retraction + ( gap f pr1 cone-standard-pullback-Σ) + ( inv-gap-cone-standard-pullback-Σ) + is-retraction-inv-gap-cone-standard-pullback-Σ = refl-htpy abstract is-pullback-cone-standard-pullback-Σ : @@ -275,10 +329,19 @@ module _ standard-pullback-Σ ≃ standard-pullback f pr1 pr1 compute-standard-pullback-Σ = gap f pr1 cone-standard-pullback-Σ pr2 compute-standard-pullback-Σ = is-pullback-cone-standard-pullback-Σ + + universal-property-pullback-standard-pullback-Σ : + universal-property-pullback f pr1 cone-standard-pullback-Σ + universal-property-pullback-standard-pullback-Σ = + universal-property-pullback-is-pullback f pr1 + ( cone-standard-pullback-Σ) + ( is-pullback-cone-standard-pullback-Σ) ``` ### Pullbacks are symmetric +The pullback of `f : A → X ← B : g` is also the pullback of `g : B → X ← A : f`. + ```agda map-commutative-standard-pullback : {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} @@ -322,9 +385,9 @@ pr2 (commutative-standard-pullback f g) = triangle-map-commutative-standard-pullback : {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} (f : A → X) (g : B → X) (c : cone f g C) → - ( gap g f (swap-cone f g c)) ~ - ( map-commutative-standard-pullback f g ∘ gap f g c) -triangle-map-commutative-standard-pullback f g (p , q , H) x = refl + gap g f (swap-cone f g c) ~ + map-commutative-standard-pullback f g ∘ gap f g c +triangle-map-commutative-standard-pullback f g c = refl-htpy abstract is-pullback-swap-cone : @@ -345,291 +408,285 @@ abstract {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} (f : A → X) (g : B → X) (c : cone f g C) → is-pullback g f (swap-cone f g c) → is-pullback f g c - is-pullback-swap-cone' f g c is-pb-c' = + is-pullback-swap-cone' f g c = is-equiv-top-map-triangle ( gap g f (swap-cone f g c)) ( map-commutative-standard-pullback f g) ( gap f g c) ( triangle-map-commutative-standard-pullback f g c) ( is-equiv-map-commutative-standard-pullback f g) - ( is-pb-c') ``` ### Pullbacks can be "folded" +Given a pullback square + +```text + f' + C -------> B + | ⌟ | + g'| | g + v v + A -------> X + f +``` + +we can "fold" the vertical edge onto the horizontal one and get a new pullback +square + +```text + C ---------> X + | ⌟ | + (f' , g') | | + v v + A × B -----> X × X, + f × g +``` + +moreover, this folded square is a pullback if and only if the original one is. + ```agda -fold-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) → - cone f g C → cone (map-prod f g) (diagonal X) C -pr1 (pr1 (fold-cone f g c) z) = vertical-map-cone f g c z -pr2 (pr1 (fold-cone f g c) z) = horizontal-map-cone f g c z -pr1 (pr2 (fold-cone f g c)) = g ∘ horizontal-map-cone f g c -pr2 (pr2 (fold-cone f g c)) z = eq-pair (coherence-square-cone f g c z) refl +module _ + {l1 l2 l3 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} + (f : A → X) (g : B → X) + where -map-fold-cone : - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) → (g : B → X) → - standard-pullback f g → standard-pullback (map-prod f g) (diagonal X) -pr1 (pr1 (map-fold-cone f g x)) = vertical-map-standard-pullback x -pr2 (pr1 (map-fold-cone f g x)) = horizontal-map-standard-pullback x -pr1 (pr2 (map-fold-cone f g x)) = g (horizontal-map-standard-pullback x) -pr2 (pr2 (map-fold-cone f g x)) = - eq-pair (coherence-square-standard-pullback x) refl - -inv-map-fold-cone : - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) → (g : B → X) → - standard-pullback (map-prod f g) (diagonal X) → standard-pullback f g -pr1 (inv-map-fold-cone f g ((a , b) , x , α)) = a -pr1 (pr2 (inv-map-fold-cone f g ((a , b) , x , α))) = b -pr2 (pr2 (inv-map-fold-cone f g ((a , b) , x , α))) = ap pr1 α ∙ inv (ap pr2 α) + fold-cone : + {l4 : Level} {C : UU l4} → + cone f g C → cone (map-prod f g) (diagonal X) C + pr1 (pr1 (fold-cone c) z) = vertical-map-cone f g c z + pr2 (pr1 (fold-cone c) z) = horizontal-map-cone f g c z + pr1 (pr2 (fold-cone c)) = g ∘ horizontal-map-cone f g c + pr2 (pr2 (fold-cone c)) z = eq-pair (coherence-square-cone f g c z) refl + + map-fold-cone : + standard-pullback f g → standard-pullback (map-prod f g) (diagonal X) + pr1 (pr1 (map-fold-cone x)) = vertical-map-standard-pullback x + pr2 (pr1 (map-fold-cone x)) = horizontal-map-standard-pullback x + pr1 (pr2 (map-fold-cone x)) = g (horizontal-map-standard-pullback x) + pr2 (pr2 (map-fold-cone x)) = + eq-pair (coherence-square-standard-pullback x) refl + + inv-map-fold-cone : + standard-pullback (map-prod f g) (diagonal X) → standard-pullback f g + pr1 (inv-map-fold-cone ((a , b) , x , α)) = a + pr1 (pr2 (inv-map-fold-cone ((a , b) , x , α))) = b + pr2 (pr2 (inv-map-fold-cone ((a , b) , x , α))) = ap pr1 α ∙ inv (ap pr2 α) -abstract - is-section-inv-map-fold-cone : - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) (g : B → X) → - ((map-fold-cone f g) ∘ (inv-map-fold-cone f g)) ~ id - is-section-inv-map-fold-cone {A = A} {B} {X} f g ((a , b) , (x , α)) = - map-extensionality-standard-pullback - ( map-prod f g) - ( diagonal X) - refl - ( ap pr2 α) - ( ( inv (is-section-pair-eq α)) ∙ - ( ap - ( λ t → eq-pair t (ap pr2 α)) - ( ( inv right-unit) ∙ - ( inv (ap (concat (ap pr1 α) x) (left-inv (ap pr2 α)))) ∙ - ( inv (assoc (ap pr1 α) (inv (ap pr2 α)) (ap pr2 α))))) ∙ - ( eq-pair-concat - ( ap pr1 α ∙ inv (ap pr2 α)) - ( ap pr2 α) - ( refl) - ( ap pr2 α)) ∙ - ( ap - ( concat - ( eq-pair (ap pr1 α ∙ inv (ap pr2 α)) refl) - ( x , x)) - ( inv (ap-diagonal (ap pr2 α))))) + abstract + is-section-inv-map-fold-cone : + is-section (map-fold-cone) (inv-map-fold-cone) + is-section-inv-map-fold-cone ((a , b) , (x , α)) = + map-extensionality-standard-pullback + ( map-prod f g) + ( diagonal X) + ( refl) + ( ap pr2 α) + ( ( inv (is-section-pair-eq α)) ∙ + ( ap + ( λ t → eq-pair t (ap pr2 α)) + ( ( inv right-unit) ∙ + ( inv (ap (concat (ap pr1 α) x) (left-inv (ap pr2 α)))) ∙ + ( inv (assoc (ap pr1 α) (inv (ap pr2 α)) (ap pr2 α))))) ∙ + ( eq-pair-concat + ( ap pr1 α ∙ inv (ap pr2 α)) + ( ap pr2 α) + ( refl) + ( ap pr2 α)) ∙ + ( ap + ( concat (eq-pair (ap pr1 α ∙ inv (ap pr2 α)) refl) (x , x)) + ( inv (ap-diagonal (ap pr2 α))))) -abstract - is-retraction-inv-map-fold-cone : - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) (g : B → X) → - inv-map-fold-cone f g ∘ map-fold-cone f g ~ id - is-retraction-inv-map-fold-cone f g (a , b , p) = - map-extensionality-standard-pullback f g - ( refl) - ( refl) - ( inv - ( ( ap - ( concat' (f a) refl) - ( ( ap - ( λ t → t ∙ inv (ap pr2 (eq-pair p refl))) - ( ap-pr1-eq-pair p refl)) ∙ - ( ap (λ t → p ∙ inv t) (ap-pr2-eq-pair p refl)) ∙ - ( right-unit))) ∙ - ( right-unit))) + abstract + is-retraction-inv-map-fold-cone : + is-retraction (map-fold-cone) (inv-map-fold-cone) + is-retraction-inv-map-fold-cone (a , b , p) = + map-extensionality-standard-pullback f g + ( refl) + ( refl) + ( inv + ( ( ap + ( concat' (f a) refl) + ( ( ap + ( _∙ inv (ap pr2 (eq-pair p refl))) + ( ap-pr1-eq-pair p refl)) ∙ + ( ap (λ t → p ∙ inv t) (ap-pr2-eq-pair p refl)) ∙ + ( right-unit))) ∙ + ( right-unit))) -abstract - is-equiv-map-fold-cone : - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) (g : B → X) → is-equiv (map-fold-cone f g) - is-equiv-map-fold-cone f g = - is-equiv-is-invertible - ( inv-map-fold-cone f g) - ( is-section-inv-map-fold-cone f g) - ( is-retraction-inv-map-fold-cone f g) + abstract + is-equiv-map-fold-cone : is-equiv map-fold-cone + is-equiv-map-fold-cone = + is-equiv-is-invertible + ( inv-map-fold-cone) + ( is-section-inv-map-fold-cone) + ( is-retraction-inv-map-fold-cone) -triangle-map-fold-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → - ( gap (map-prod f g) (diagonal X) (fold-cone f g c)) ~ - ( map-fold-cone f g ∘ gap f g c) -triangle-map-fold-cone f g c z = refl + triangle-map-fold-cone : + {l4 : Level} {C : UU l4} (c : cone f g C) → + gap (map-prod f g) (diagonal X) (fold-cone c) ~ + map-fold-cone ∘ gap f g c + triangle-map-fold-cone c = refl-htpy -abstract - is-pullback-fold-cone-is-pullback : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → - is-pullback f g c → - is-pullback (map-prod f g) (diagonal X) (fold-cone f g c) - is-pullback-fold-cone-is-pullback {X = X} f g c is-pb-c = - is-equiv-left-map-triangle - ( gap (map-prod f g) (diagonal X) (fold-cone f g c)) - ( map-fold-cone f g) - ( gap f g c) - ( triangle-map-fold-cone f g c) - ( is-pb-c) - ( is-equiv-map-fold-cone f g) + abstract + is-pullback-fold-cone-is-pullback : + {l4 : Level} {C : UU l4} (c : cone f g C) → + is-pullback f g c → + is-pullback (map-prod f g) (diagonal X) (fold-cone c) + is-pullback-fold-cone-is-pullback c is-pb-c = + is-equiv-left-map-triangle + ( gap (map-prod f g) (diagonal X) (fold-cone c)) + ( map-fold-cone) + ( gap f g c) + ( triangle-map-fold-cone c) + ( is-pb-c) + ( is-equiv-map-fold-cone) -abstract - is-pullback-is-pullback-fold-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → - is-pullback (map-prod f g) (diagonal X) (fold-cone f g c) → - is-pullback f g c - is-pullback-is-pullback-fold-cone {X = X} f g c is-pb-fold = - is-equiv-top-map-triangle - ( gap (map-prod f g) (diagonal X) (fold-cone f g c)) - ( map-fold-cone f g) - ( gap f g c) - ( triangle-map-fold-cone f g c) - ( is-equiv-map-fold-cone f g) - ( is-pb-fold) + abstract + is-pullback-is-pullback-fold-cone : + {l4 : Level} {C : UU l4} (c : cone f g C) → + is-pullback (map-prod f g) (diagonal X) (fold-cone c) → + is-pullback f g c + is-pullback-is-pullback-fold-cone c = + is-equiv-top-map-triangle + ( gap (map-prod f g) (diagonal X) (fold-cone c)) + ( map-fold-cone) + ( gap f g c) + ( triangle-map-fold-cone c) + ( is-equiv-map-fold-cone) ``` ### Products of pullbacks are pullbacks ```agda -prod-cone : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} - (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') → - cone f g C → cone f' g' C' → - cone (map-prod f f') (map-prod g g') (C × C') -pr1 (prod-cone f g f' g' (p , q , H) (p' , q' , H')) = map-prod p p' -pr1 (pr2 (prod-cone f g f' g' (p , q , H) (p' , q' , H'))) = map-prod q q' -pr2 (pr2 (prod-cone f g f' g' (p , q , H) (p' , q' , H'))) = - ( inv-htpy (map-prod-comp p p' f f')) ∙h - ( htpy-map-prod H H') ∙h - ( map-prod-comp q q' g g') - -map-prod-cone : +module _ {l1 l2 l3 l1' l2' l3' : Level} {A : UU l1} {B : UU l2} {X : UU l3} {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} - (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') → - (standard-pullback f g) × (standard-pullback f' g') → - standard-pullback (map-prod f f') (map-prod g g') -map-prod-cone {A' = A'} {B'} f g f' g' = - ( tot - ( λ t → - ( tot (λ s → eq-pair')) ∘ - ( map-interchange-Σ-Σ (λ y p y' → f' (pr2 t) = g' y')))) ∘ - ( map-interchange-Σ-Σ (λ x t x' → Σ B' (λ y' → f' x' = g' y'))) - -triangle-map-prod-cone : + (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') + where + + map-prod-cone : + (standard-pullback f g) × (standard-pullback f' g') → + standard-pullback (map-prod f f') (map-prod g g') + map-prod-cone = + ( tot + ( λ t → + ( tot (λ s → eq-pair')) ∘ + ( map-interchange-Σ-Σ (λ y p y' → f' (pr2 t) = g' y')))) ∘ + ( map-interchange-Σ-Σ (λ x t x' → Σ B' (λ y' → f' x' = g' y'))) + + abstract + is-equiv-map-prod-cone : is-equiv map-prod-cone + is-equiv-map-prod-cone = + is-equiv-comp + ( tot (λ t → tot (λ s → eq-pair') ∘ map-interchange-Σ-Σ _)) + ( map-interchange-Σ-Σ _) + ( is-equiv-map-interchange-Σ-Σ _) + ( is-equiv-tot-is-fiberwise-equiv + ( λ t → + is-equiv-comp + ( tot (λ s → eq-pair')) + ( map-interchange-Σ-Σ (λ y p y' → f' (pr2 t) = g' y')) + ( is-equiv-map-interchange-Σ-Σ _) + ( is-equiv-tot-is-fiberwise-equiv + ( λ s → is-equiv-eq-pair (map-prod f f' t) (map-prod g g' s))))) + +module _ {l1 l2 l3 l4 l1' l2' l3' l4' : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} - (f : A → X) (g : B → X) (c : cone f g C) - (f' : A' → X') (g' : B' → X') (c' : cone f' g' C') → - ( gap (map-prod f f') (map-prod g g') (prod-cone f g f' g' c c')) ~ - ( map-prod-cone f g f' g' ∘ map-prod (gap f g c) (gap f' g' c')) -triangle-map-prod-cone f g c f' g' c' z = - eq-pair-eq-pr2 (eq-pair-eq-pr2 right-unit) - -abstract - is-equiv-map-prod-cone : - {l1 l2 l3 l1' l2' l3' : Level} - {A : UU l1} {B : UU l2} {X : UU l3} - {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} - (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') → - is-equiv (map-prod-cone f g f' g') - is-equiv-map-prod-cone f g f' g' = - is-equiv-comp - ( tot - ( λ t → - ( tot (λ s → eq-pair')) ∘ - ( map-interchange-Σ-Σ _))) - ( map-interchange-Σ-Σ _) - ( is-equiv-map-interchange-Σ-Σ _) - ( is-equiv-tot-is-fiberwise-equiv - ( λ t → - is-equiv-comp - ( tot (λ s → eq-pair')) - ( map-interchange-Σ-Σ - ( λ y p y' → f' (pr2 t) = g' y')) - ( is-equiv-map-interchange-Σ-Σ _) - ( is-equiv-tot-is-fiberwise-equiv - ( λ s → - is-equiv-comp - ( eq-pair') - ( id) - ( is-equiv-id) - ( is-equiv-eq-pair - ( map-prod f f' t) - ( map-prod g g' s)))))) + (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') + where -abstract - is-pullback-prod-is-pullback-pair : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} - (f : A → X) (g : B → X) (c : cone f g C) - (f' : A' → X') (g' : B' → X') (c' : cone f' g' C') → - is-pullback f g c → is-pullback f' g' c' → - is-pullback - ( map-prod f f') (map-prod g g') (prod-cone f g f' g' c c') - is-pullback-prod-is-pullback-pair f g c f' g' c' is-pb-c is-pb-c' = - is-equiv-left-map-triangle - ( gap (map-prod f f') (map-prod g g') (prod-cone f g f' g' c c')) - ( map-prod-cone f g f' g') - ( map-prod (gap f g c) (gap f' g' c')) - ( triangle-map-prod-cone f g c f' g' c') - ( is-equiv-map-prod (gap f g c) (gap f' g' c') is-pb-c is-pb-c') - ( is-equiv-map-prod-cone f g f' g') + prod-cone : + cone f g C → cone f' g' C' → + cone (map-prod f f') (map-prod g g') (C × C') + pr1 (prod-cone (p , q , H) (p' , q' , H')) = map-prod p p' + pr1 (pr2 (prod-cone (p , q , H) (p' , q' , H'))) = map-prod q q' + pr2 (pr2 (prod-cone (p , q , H) (p' , q' , H'))) = + ( inv-htpy (preserves-comp-map-prod p p' f f')) ∙h + ( htpy-map-prod H H') ∙h + ( preserves-comp-map-prod q q' g g') + + triangle-map-prod-cone : + (c : cone f g C) (c' : cone f' g' C') → + gap (map-prod f f') (map-prod g g') (prod-cone c c') ~ + map-prod-cone f g f' g' ∘ map-prod (gap f g c) (gap f' g' c') + triangle-map-prod-cone c c' z = + eq-pair-eq-pr2 (eq-pair-eq-pr2 right-unit) -abstract - is-pullback-left-factor-is-pullback-prod : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} - (f : A → X) (g : B → X) (c : cone f g C) - (f' : A' → X') (g' : B' → X') (c' : cone f' g' C') → - is-pullback - ( map-prod f f') - ( map-prod g g') - ( prod-cone f g f' g' c c') → - standard-pullback f' g' → is-pullback f g c - is-pullback-left-factor-is-pullback-prod f g c f' g' c' is-pb-cc' t = - is-equiv-left-factor-is-equiv-map-prod (gap f g c) (gap f' g' c') t - ( is-equiv-top-map-triangle - ( gap - ( map-prod f f') - ( map-prod g g') - ( prod-cone f g f' g' c c')) - ( map-prod-cone f g f' g') + abstract + is-pullback-prod-is-pullback-pair : + (c : cone f g C) (c' : cone f' g' C') → + is-pullback f g c → is-pullback f' g' c' → + is-pullback (map-prod f f') (map-prod g g') (prod-cone c c') + is-pullback-prod-is-pullback-pair c c' is-pb-c is-pb-c' = + is-equiv-left-map-triangle + ( gap (map-prod f f') (map-prod g g') (prod-cone c c')) + ( map-prod-cone f g f' g') ( map-prod (gap f g c) (gap f' g' c')) - ( triangle-map-prod-cone f g c f' g' c') + ( triangle-map-prod-cone c c') + ( is-equiv-map-prod (gap f g c) (gap f' g' c') is-pb-c is-pb-c') ( is-equiv-map-prod-cone f g f' g') - ( is-pb-cc')) -abstract - is-pullback-right-factor-is-pullback-prod : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} - (f : A → X) (g : B → X) (c : cone f g C) - (f' : A' → X') (g' : B' → X') (c' : cone f' g' C') → - is-pullback - ( map-prod f f') - ( map-prod g g') - ( prod-cone f g f' g' c c') → - standard-pullback f g → is-pullback f' g' c' - is-pullback-right-factor-is-pullback-prod f g c f' g' c' is-pb-cc' t = - is-equiv-right-factor-is-equiv-map-prod (gap f g c) (gap f' g' c') t - ( is-equiv-top-map-triangle - ( gap - ( map-prod f f') - ( map-prod g g') - ( prod-cone f g f' g' c c')) + abstract + is-pullback-left-factor-is-pullback-prod : + (c : cone f g C) (c' : cone f' g' C') → + is-pullback (map-prod f f') (map-prod g g') (prod-cone c c') → + standard-pullback f' g' → is-pullback f g c + is-pullback-left-factor-is-pullback-prod c c' is-pb-cc' t = + is-equiv-left-factor-is-equiv-map-prod (gap f g c) (gap f' g' c') t + ( is-equiv-top-map-triangle + ( gap (map-prod f f') (map-prod g g') (prod-cone c c')) ( map-prod-cone f g f' g') - ( map-prod (gap f g c) (gap f' g' c')) - ( triangle-map-prod-cone f g c f' g' c') - ( is-equiv-map-prod-cone f g f' g') - ( is-pb-cc')) + ( map-prod (gap f g c) (gap f' g' c')) + ( triangle-map-prod-cone c c') + ( is-equiv-map-prod-cone f g f' g') + ( is-pb-cc')) + + abstract + is-pullback-right-factor-is-pullback-prod : + (c : cone f g C) (c' : cone f' g' C') → + is-pullback (map-prod f f') (map-prod g g') (prod-cone c c') → + standard-pullback f g → is-pullback f' g' c' + is-pullback-right-factor-is-pullback-prod c c' is-pb-cc' t = + is-equiv-right-factor-is-equiv-map-prod (gap f g c) (gap f' g' c') t + ( is-equiv-top-map-triangle + ( gap + ( map-prod f f') + ( map-prod g g') + ( prod-cone c c')) + ( map-prod-cone f g f' g') + ( map-prod (gap f g c) (gap f' g' c')) + ( triangle-map-prod-cone c c') + ( is-equiv-map-prod-cone f g f' g') + ( is-pb-cc')) ``` ### A family of maps over a base map induces a pullback square if and only if it is a family of equivalences +Given a map `f : A → B` with a family of maps over it +`g : (x : A) → P x → Q (f x)`, then the square + +```text + map-Σ f g + Σ A P ----------> Σ B Q + | | + | | + v v + A -------------> B + f +``` + +is a pullback if and only if `g` is a +[fiberwise equivalence](foundation-core.families-of-equivalences.md). + ```agda module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {P : A → UU l3} - (Q : B → UU l4) (f : A → B) (g : (x : A) → (P x) → (Q (f x))) + (Q : B → UU l4) (f : A → B) (g : (x : A) → P x → Q (f x)) where cone-map-Σ : cone f pr1 (Σ A P) @@ -683,52 +740,63 @@ module _ (f : A → X) (g : B → X) (c : cone f g C) where - square-tot-map-fiber-cone : - ( gap f g c ∘ map-equiv-total-fiber (pr1 c)) ~ - ( tot (λ a → tot (λ b → inv)) ∘ tot (map-fiber-cone f g c)) - square-tot-map-fiber-cone (.(vertical-map-cone f g c x) , x , refl) = + square-tot-map-fiber-vertical-map-cone : + gap f g c ∘ map-equiv-total-fiber (pr1 c) ~ + tot (λ _ → tot (λ _ → inv)) ∘ tot (map-fiber-vertical-map-cone f g c) + square-tot-map-fiber-vertical-map-cone + (.(vertical-map-cone f g c x) , x , refl) = eq-pair-eq-pr2 ( eq-pair-eq-pr2 ( inv (ap inv right-unit ∙ inv-inv (coherence-square-cone f g c x)))) abstract - is-fiberwise-equiv-map-fiber-cone-is-pullback : - is-pullback f g c → is-fiberwise-equiv (map-fiber-cone f g c) - is-fiberwise-equiv-map-fiber-cone-is-pullback pb = + is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback : + is-pullback f g c → is-fiberwise-equiv (map-fiber-vertical-map-cone f g c) + is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback pb = is-fiberwise-equiv-is-equiv-tot ( is-equiv-top-is-equiv-bottom-square ( map-equiv-total-fiber (vertical-map-cone f g c)) - ( tot (λ x → tot (λ y → inv))) - ( tot (map-fiber-cone f g c)) + ( tot (λ _ → tot (λ _ → inv))) + ( tot (map-fiber-vertical-map-cone f g c)) ( gap f g c) - ( square-tot-map-fiber-cone) + ( square-tot-map-fiber-vertical-map-cone) ( is-equiv-map-equiv-total-fiber (vertical-map-cone f g c)) ( is-equiv-tot-is-fiberwise-equiv ( λ x → - is-equiv-tot-is-fiberwise-equiv - ( λ y → is-equiv-inv (g y) (f x)))) + is-equiv-tot-is-fiberwise-equiv (λ y → is-equiv-inv (g y) (f x)))) ( pb)) abstract - is-pullback-is-fiberwise-equiv-map-fiber-cone : - is-fiberwise-equiv (map-fiber-cone f g c) → is-pullback f g c - is-pullback-is-fiberwise-equiv-map-fiber-cone is-equiv-fsq = + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone : + is-fiberwise-equiv (map-fiber-vertical-map-cone f g c) → is-pullback f g c + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone is-equiv-fsq = is-equiv-bottom-is-equiv-top-square ( map-equiv-total-fiber (vertical-map-cone f g c)) - ( tot (λ x → tot (λ y → inv))) - ( tot (map-fiber-cone f g c)) + ( tot (λ _ → tot (λ _ → inv))) + ( tot (map-fiber-vertical-map-cone f g c)) ( gap f g c) - ( square-tot-map-fiber-cone) + ( square-tot-map-fiber-vertical-map-cone) ( is-equiv-map-equiv-total-fiber (vertical-map-cone f g c)) ( is-equiv-tot-is-fiberwise-equiv ( λ x → - is-equiv-tot-is-fiberwise-equiv - ( λ y → is-equiv-inv (g y) (f x)))) + is-equiv-tot-is-fiberwise-equiv (λ y → is-equiv-inv (g y) (f x)))) ( is-equiv-tot-is-fiberwise-equiv is-equiv-fsq) ``` ### The horizontal pullback pasting property +Given a diagram as follows where the right-hand square is a pullback + +```text + ∙ -------> ∙ -------> ∙ + | | ⌟ | + | | | + v v v + ∙ -------> ∙ -------> ∙, +``` + +then the left-hand square is a pullback if and only if the composite square is. + ```agda module _ {l1 l2 l3 l4 l5 l6 : Level} @@ -742,20 +810,32 @@ module _ is-pullback j h c → is-pullback i (vertical-map-cone j h c) d → is-pullback (j ∘ i) h (pasting-horizontal-cone i j h c d) is-pullback-rectangle-is-pullback-left-square c d is-pb-c is-pb-d = - is-pullback-is-fiberwise-equiv-map-fiber-cone (j ∘ i) h + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone (j ∘ i) h ( pasting-horizontal-cone i j h c d) ( λ x → is-equiv-left-map-triangle - ( map-fiber-cone (j ∘ i) h (pasting-horizontal-cone i j h c d) x) - ( map-fiber-cone j h c (i x)) - ( map-fiber-cone i (vertical-map-cone j h c) d x) - ( preserves-pasting-horizontal-map-fiber-cone i j h c d x) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback i + ( map-fiber-vertical-map-cone + ( j ∘ i) h (pasting-horizontal-cone i j h c d) x) + ( map-fiber-vertical-map-cone j h c (i x)) + ( map-fiber-vertical-map-cone i (vertical-map-cone j h c) d x) + ( preserves-pasting-horizontal-map-fiber-vertical-map-cone + ( i) + ( j) + ( h) + ( c) + ( d) + ( x)) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( i) ( vertical-map-cone j h c) ( d) ( is-pb-d) ( x)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback j h c is-pb-c + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( j) + ( h) + ( c) + ( is-pb-c) ( i x))) abstract @@ -765,17 +845,29 @@ module _ is-pullback (j ∘ i) h (pasting-horizontal-cone i j h c d) → is-pullback i (vertical-map-cone j h c) d is-pullback-left-square-is-pullback-rectangle c d is-pb-c is-pb-rect = - is-pullback-is-fiberwise-equiv-map-fiber-cone i + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone i ( vertical-map-cone j h c) ( d) ( λ x → is-equiv-top-map-triangle - ( map-fiber-cone (j ∘ i) h (pasting-horizontal-cone i j h c d) x) - ( map-fiber-cone j h c (i x)) - ( map-fiber-cone i (vertical-map-cone j h c) d x) - ( preserves-pasting-horizontal-map-fiber-cone i j h c d x) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback j h c is-pb-c (i x)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback + ( map-fiber-vertical-map-cone + ( j ∘ i) h (pasting-horizontal-cone i j h c d) x) + ( map-fiber-vertical-map-cone j h c (i x)) + ( map-fiber-vertical-map-cone i (vertical-map-cone j h c) d x) + ( preserves-pasting-horizontal-map-fiber-vertical-map-cone + ( i) + ( j) + ( h) + ( c) + ( d) + ( x)) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( j) + ( h) + ( c) + ( is-pb-c) + ( i x)) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback ( j ∘ i) ( h) ( pasting-horizontal-cone i j h c d) @@ -785,6 +877,22 @@ module _ ### The vertical pullback pasting property +Given a diagram as follows where the lower square is a pullback + +```text + ∙ -------> ∙ + | | + | | + v v + ∙ -------> ∙ + | ⌟ | + | | + v v + ∙ -------> ∙, +``` + +then the upper square is a pullback if and only if the composite square is. + ```agda module _ {l1 l2 l3 l4 l5 l6 : Level} @@ -799,16 +907,25 @@ module _ is-pullback f (g ∘ h) (pasting-vertical-cone f g h c d) → is-pullback (horizontal-map-cone f g c) h d is-pullback-top-is-pullback-rectangle c d is-pb-c is-pb-dc = - is-pullback-is-fiberwise-equiv-map-fiber-cone + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone ( horizontal-map-cone f g c) ( h) ( d) ( λ x → is-fiberwise-equiv-is-equiv-map-Σ ( λ t → fiber h (pr1 t)) - ( map-fiber-cone f g c (vertical-map-cone f g c x)) - ( λ t → map-fiber-cone (horizontal-map-cone f g c) h d (pr1 t)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback f g c is-pb-c + ( map-fiber-vertical-map-cone f g c (vertical-map-cone f g c x)) + ( λ t → + map-fiber-vertical-map-cone + ( horizontal-map-cone f g c) + ( h) + ( d) + ( pr1 t)) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( f) + ( g) + ( c) + ( is-pb-c) ( vertical-map-cone f g c x)) ( is-equiv-top-is-equiv-bottom-square ( inv-map-compute-fiber-comp @@ -818,13 +935,15 @@ module _ ( inv-map-compute-fiber-comp g h (f (vertical-map-cone f g c x))) ( map-Σ ( λ t → fiber h (pr1 t)) - ( map-fiber-cone f g c (vertical-map-cone f g c x)) - ( λ t → map-fiber-cone (horizontal-map-cone f g c) h d (pr1 t))) - ( map-fiber-cone f + ( map-fiber-vertical-map-cone f g c (vertical-map-cone f g c x)) + ( λ t → + map-fiber-vertical-map-cone + ( horizontal-map-cone f g c) h d (pr1 t))) + ( map-fiber-vertical-map-cone f ( g ∘ h) ( pasting-vertical-cone f g h c d) ( vertical-map-cone f g c x)) - ( preserves-pasting-vertical-map-fiber-cone f g h c d + ( preserves-pasting-vertical-map-fiber-vertical-map-cone f g h c d ( vertical-map-cone f g c x)) ( is-equiv-inv-map-compute-fiber-comp ( vertical-map-cone f g c) @@ -832,7 +951,7 @@ module _ ( vertical-map-cone f g c x)) ( is-equiv-inv-map-compute-fiber-comp g h ( f (vertical-map-cone f g c x))) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback f + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback f ( g ∘ h) ( pasting-vertical-cone f g h c d) ( is-pb-dc) @@ -846,7 +965,9 @@ module _ is-pullback (horizontal-map-cone f g c) h d → is-pullback f (g ∘ h) (pasting-vertical-cone f g h c d) is-pullback-rectangle-is-pullback-top c d is-pb-c is-pb-d = - is-pullback-is-fiberwise-equiv-map-fiber-cone f (g ∘ h) + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone + ( f) + ( g ∘ h) ( pasting-vertical-cone f g h c d) ( λ x → is-equiv-bottom-is-equiv-top-square @@ -857,10 +978,25 @@ module _ ( inv-map-compute-fiber-comp g h (f x)) ( map-Σ ( λ t → fiber h (pr1 t)) - ( map-fiber-cone f g c x) - ( λ t → map-fiber-cone (horizontal-map-cone f g c) h d (pr1 t))) - ( map-fiber-cone f (g ∘ h) (pasting-vertical-cone f g h c d) x) - ( preserves-pasting-vertical-map-fiber-cone f g h c d x) + ( map-fiber-vertical-map-cone f g c x) + ( λ t → + map-fiber-vertical-map-cone + ( horizontal-map-cone f g c) + ( h) + ( d) + ( pr1 t))) + ( map-fiber-vertical-map-cone + ( f) + ( g ∘ h) + ( pasting-vertical-cone f g h c d) + ( x)) + ( preserves-pasting-vertical-map-fiber-vertical-map-cone + ( f) + ( g) + ( h) + ( c) + ( d) + ( x)) ( is-equiv-inv-map-compute-fiber-comp ( vertical-map-cone f g c) ( vertical-map-cone (horizontal-map-cone f g c) h d) @@ -868,9 +1004,14 @@ module _ ( is-equiv-inv-map-compute-fiber-comp g h (f x)) ( is-equiv-map-Σ ( λ t → fiber h (pr1 t)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback f g c is-pb-c x) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( f) + ( g) + ( c) + ( is-pb-c) + ( x)) ( λ t → - is-fiberwise-equiv-map-fiber-cone-is-pullback + is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback ( horizontal-map-cone f g c) ( h) ( d) diff --git a/src/foundation-core/universal-property-pullbacks.lagda.md b/src/foundation-core/universal-property-pullbacks.lagda.md index adab6c57fc..b1b3c95921 100644 --- a/src/foundation-core/universal-property-pullbacks.lagda.md +++ b/src/foundation-core/universal-property-pullbacks.lagda.md @@ -46,21 +46,43 @@ module _ map-universal-property-pullback : universal-property-pullback f g c → {C' : UU l5} (c' : cone f g C') → C' → C - map-universal-property-pullback up-c {C'} c' = - map-inv-is-equiv (up-c C') c' + map-universal-property-pullback up-c {C'} = + map-inv-is-equiv (up-c C') compute-map-universal-property-pullback : (up-c : universal-property-pullback f g c) → {C' : UU l5} (c' : cone f g C') → cone-map f g c (map-universal-property-pullback up-c c') = c' - compute-map-universal-property-pullback up-c {C'} c' = - is-section-map-inv-is-equiv (up-c C') c' + compute-map-universal-property-pullback up-c {C'} = + is-section-map-inv-is-equiv (up-c C') ``` ## Properties ### 3-for-2 property of pullbacks +Given two cones `c` and `c'` over a common cospan `f : A → X ← B : g`, and a map +between them `h` such that the diagram + +```text + h + C ----> C' + / \ / \ + / / \ \ + / / \ \ + v v v v + A --------> X <-------- B + f g +``` + +is coherent, then if two out of the three conditions + +- `c` is a pullback cone +- `c'` is a pullback cone +- `h` is an equivalence + +hold, then so does the third. + ```agda module _ {l1 l2 l3 l4 l5 : Level} @@ -78,8 +100,7 @@ module _ ( eq-htpy-cone f g (cone-map f g c h) c' KLM) triangle-cone-cone : - {l6 : Level} (D : UU l6) → - cone-map f g c' ~ cone-map f g c ∘ postcomp D h + {l6 : Level} (D : UU l6) → cone-map f g c' ~ cone-map f g c ∘ postcomp D h triangle-cone-cone D k = inv (inv-triangle-cone-cone D k) abstract diff --git a/src/foundation-core/whiskering-homotopies.lagda.md b/src/foundation-core/whiskering-homotopies.lagda.md index 0abb45d11a..1cf39a10e4 100644 --- a/src/foundation-core/whiskering-homotopies.lagda.md +++ b/src/foundation-core/whiskering-homotopies.lagda.md @@ -266,7 +266,7 @@ We have the coherence (h ·l H) ·r h' ~ h ·l (H ·r h') ``` -and, in fact, this equation holds definitionally. +definitionally. ```agda module _ diff --git a/src/foundation.lagda.md b/src/foundation.lagda.md index 859535f3cc..e9498652a5 100644 --- a/src/foundation.lagda.md +++ b/src/foundation.lagda.md @@ -36,6 +36,7 @@ open import foundation.binary-transport public open import foundation.booleans public open import foundation.cantor-schroder-bernstein-escardo public open import foundation.cantors-diagonal-argument public +open import foundation.cartesian-morphisms-arrows public open import foundation.cartesian-product-types public open import foundation.cartesian-products-set-quotients public open import foundation.category-of-families-of-sets public @@ -131,6 +132,7 @@ open import foundation.equivalence-extensionality public open import foundation.equivalence-induction public open import foundation.equivalence-relations public open import foundation.equivalences public +open import foundation.equivalences-arrows public open import foundation.equivalences-inverse-sequential-diagrams public open import foundation.equivalences-maybe public open import foundation.exclusive-disjunction public diff --git a/src/foundation/axiom-of-choice.lagda.md b/src/foundation/axiom-of-choice.lagda.md index b78a7d0fb8..5c7ddfae1a 100644 --- a/src/foundation/axiom-of-choice.lagda.md +++ b/src/foundation/axiom-of-choice.lagda.md @@ -9,6 +9,7 @@ module foundation.axiom-of-choice where ```agda open import foundation.dependent-pair-types open import foundation.functoriality-propositional-truncation +open import foundation.postcomposition-functions open import foundation.projective-types open import foundation.propositional-truncations open import foundation.sections @@ -69,7 +70,7 @@ is-set-projective-AC-0 ac X A B f h = ( ( map-Σ ( λ g → ((map-surjection f) ∘ g) = h) ( precomp h A) - ( λ s H → eq-htpy (H ·r h))) ∘ + ( λ s H → htpy-postcomp X H h)) ∘ ( section-is-split-surjective (map-surjection f))) ( ac B (fiber (map-surjection f)) (is-surjective-map-surjection f)) diff --git a/src/foundation/cartesian-morphisms-arrows.lagda.md b/src/foundation/cartesian-morphisms-arrows.lagda.md new file mode 100644 index 0000000000..e20aa4ce38 --- /dev/null +++ b/src/foundation/cartesian-morphisms-arrows.lagda.md @@ -0,0 +1,144 @@ +# Cartesian morphisms of arrows + +```agda +module foundation.cartesian-morphisms-arrows where +``` + +
Imports + +```agda +open import foundation.cones-over-cospans +open import foundation.dependent-pair-types +open import foundation.fibers-of-maps +open import foundation.morphisms-arrows +open import foundation.pullbacks +open import foundation.unit-type +open import foundation.universe-levels + +open import foundation-core.commuting-squares-of-maps +open import foundation-core.propositions +open import foundation-core.universal-property-pullbacks +``` + +
+ +## Idea + +A [morphism of arrows](foundation.morphisms-arrows.md) `h : f → g` is said to be +{{#concept "cartesian" Disambiguation="morphism of arrows" Agda=is-cartesian-hom-arrow}} +if the [commuting square](foundation-core.commuting-squares-of-maps.md) + +```text + i + A -----> X + | | + f | h | g + V V + B -----> Y + j +``` + +is a [pullback](foundation.pullbacks.md) square. + +## Definitions + +### The predicate of being a cartesian morphism of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (h : hom-arrow f g) + where + + is-cartesian-hom-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-cartesian-hom-arrow = + is-pullback (map-codomain-hom-arrow f g h) g (cone-hom-arrow f g h) + + is-prop-is-cartesian-hom-arrow : is-prop is-cartesian-hom-arrow + is-prop-is-cartesian-hom-arrow = + is-prop-is-pullback (map-codomain-hom-arrow f g h) g (cone-hom-arrow f g h) + + is-cartesian-hom-arrow-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + pr1 is-cartesian-hom-arrow-Prop = is-cartesian-hom-arrow + pr2 is-cartesian-hom-arrow-Prop = is-prop-is-cartesian-hom-arrow +``` + +### The type of cartesian morphisms of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + cartesian-hom-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + cartesian-hom-arrow = Σ (hom-arrow f g) (is-cartesian-hom-arrow f g) + +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (h : cartesian-hom-arrow f g) + where + + hom-arrow-cartesian-hom-arrow : hom-arrow f g + hom-arrow-cartesian-hom-arrow = pr1 h + + is-cartesian-cartesian-hom-arrow : + is-cartesian-hom-arrow f g hom-arrow-cartesian-hom-arrow + is-cartesian-cartesian-hom-arrow = pr2 h + + map-domain-cartesian-hom-arrow : A → X + map-domain-cartesian-hom-arrow = + map-domain-hom-arrow f g hom-arrow-cartesian-hom-arrow + + map-codomain-cartesian-hom-arrow : B → Y + map-codomain-cartesian-hom-arrow = + map-codomain-hom-arrow f g hom-arrow-cartesian-hom-arrow + + coh-cartesian-hom-arrow : + coherence-square-maps + ( map-domain-cartesian-hom-arrow) + ( f) + ( g) + ( map-codomain-cartesian-hom-arrow) + coh-cartesian-hom-arrow = + coh-hom-arrow f g hom-arrow-cartesian-hom-arrow + + cone-cartesian-hom-arrow : + cone map-codomain-cartesian-hom-arrow g A + cone-cartesian-hom-arrow = + cone-hom-arrow f g hom-arrow-cartesian-hom-arrow + + universal-property-cartesian-hom-arrow : + universal-property-pullback + ( map-codomain-cartesian-hom-arrow) + ( g) + ( cone-cartesian-hom-arrow) + universal-property-cartesian-hom-arrow = + universal-property-pullback-is-pullback + ( map-codomain-cartesian-hom-arrow) + ( g) + ( cone-cartesian-hom-arrow) + ( is-cartesian-cartesian-hom-arrow) +``` + +### Cartesian morphisms of arrows arising from fibers + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) (y : B) + where + + fiber-cartesian-hom-arrow : + cartesian-hom-arrow (terminal-map (fiber f y)) f + pr1 fiber-cartesian-hom-arrow = + hom-arrow-cone (point y) f (swap-cone f (point y) (cone-fiber f y)) + pr2 fiber-cartesian-hom-arrow = + is-pullback-swap-cone f (point y) + ( cone-fiber f y) + ( is-pullback-cone-fiber f y) +``` + +## See also + +- [Cocartesian morphisms of arrows](synthetic-homotopy-theory.cocartesian-morphisms-arrows.md) + for the dual. diff --git a/src/foundation/commuting-prisms-of-maps.lagda.md b/src/foundation/commuting-prisms-of-maps.lagda.md index ad629b1178..5ad0bd595c 100644 --- a/src/foundation/commuting-prisms-of-maps.lagda.md +++ b/src/foundation/commuting-prisms-of-maps.lagda.md @@ -9,13 +9,10 @@ open import foundation-core.commuting-prisms-of-maps public
Imports ```agda -open import foundation.action-on-identifications-binary-functions open import foundation.action-on-identifications-functions open import foundation.commuting-squares-of-maps open import foundation.commuting-triangles-of-maps open import foundation.composition-algebra -open import foundation.dependent-pair-types -open import foundation.function-extensionality open import foundation.identity-types open import foundation.path-algebra open import foundation.postcomposition-functions @@ -24,6 +21,7 @@ open import foundation.universe-levels open import foundation.whiskering-homotopies open import foundation-core.equivalences +open import foundation-core.function-extensionality open import foundation-core.function-types open import foundation-core.functoriality-dependent-function-types open import foundation-core.homotopies diff --git a/src/foundation/commuting-squares-of-maps.lagda.md b/src/foundation/commuting-squares-of-maps.lagda.md index 15ebedb291..3df33693bc 100644 --- a/src/foundation/commuting-squares-of-maps.lagda.md +++ b/src/foundation/commuting-squares-of-maps.lagda.md @@ -825,7 +825,7 @@ module _ ( (h ∘ bottom-right) ·l H) ∙ ap ( precomp top-left W) - ( eq-htpy (h ·l K)) + ( htpy-precomp K W h) by ap-binary ( λ L q → eq-htpy L ∙ q) @@ -911,7 +911,7 @@ module _ ( h ·l (right-bottom ·l H)) = ap ( precomp left-top W) - ( eq-htpy (h ·l K)) ∙ + ( htpy-precomp K W h) ∙ eq-htpy ( (h ∘ right-bottom) ·l H) by diff --git a/src/foundation/composition-algebra.lagda.md b/src/foundation/composition-algebra.lagda.md index afb8e079dc..da04122fb6 100644 --- a/src/foundation/composition-algebra.lagda.md +++ b/src/foundation/composition-algebra.lagda.md @@ -9,16 +9,14 @@ module foundation.composition-algebra where ```agda open import foundation.action-on-identifications-functions open import foundation.function-extensionality -open import foundation.function-types open import foundation.homotopy-induction -open import foundation.identity-types open import foundation.postcomposition-functions open import foundation.precomposition-functions open import foundation.universe-levels -open import foundation-core.commuting-squares-of-maps -open import foundation-core.commuting-triangles-of-maps +open import foundation-core.function-types open import foundation-core.homotopies +open import foundation-core.identity-types open import foundation-core.whiskering-homotopies ``` diff --git a/src/foundation/cones-over-cospans.lagda.md b/src/foundation/cones-over-cospans.lagda.md index 2b7ca15a4e..79c663903f 100644 --- a/src/foundation/cones-over-cospans.lagda.md +++ b/src/foundation/cones-over-cospans.lagda.md @@ -219,6 +219,20 @@ module _ ( coherence-square-cone f g c) ``` +### Cones obtained from morphisms of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (h : hom-arrow f g) + where + + cone-hom-arrow : cone (map-codomain-hom-arrow f g h) g A + pr1 cone-hom-arrow = f + pr1 (pr2 cone-hom-arrow) = map-domain-hom-arrow f g h + pr2 (pr2 cone-hom-arrow) = coh-hom-arrow f g h +``` + ### The swapping function on cones over cospans ```agda @@ -265,6 +279,15 @@ module _ ( fam-htpy-parallel-cone c c') ``` +### The identity cone over the identity cospan + +```agda +id-cone : {l : Level} (A : UU l) → cone (id {A = A}) (id {A = A}) A +pr1 (id-cone A) = id +pr1 (pr2 (id-cone A)) = id +pr2 (pr2 (id-cone A)) = refl-htpy +``` + ## Table of files about pullbacks The following table lists files that are about pullbacks as a general concept. diff --git a/src/foundation/coproduct-types.lagda.md b/src/foundation/coproduct-types.lagda.md index 90b990a87f..5d00804d5a 100644 --- a/src/foundation/coproduct-types.lagda.md +++ b/src/foundation/coproduct-types.lagda.md @@ -111,11 +111,12 @@ module _ where map-equiv-left-summand : Σ (X + Y) is-left → X - map-equiv-left-summand (pair (inl x) star) = x - map-equiv-left-summand (pair (inr x) ()) + map-equiv-left-summand (inl x , star) = x + map-equiv-left-summand (inr x , ()) map-inv-equiv-left-summand : X → Σ (X + Y) is-left - map-inv-equiv-left-summand x = pair (inl x) star + pr1 (map-inv-equiv-left-summand x) = inl x + pr2 (map-inv-equiv-left-summand x) = star is-section-map-inv-equiv-left-summand : (map-equiv-left-summand ∘ map-inv-equiv-left-summand) ~ id @@ -123,8 +124,8 @@ module _ is-retraction-map-inv-equiv-left-summand : (map-inv-equiv-left-summand ∘ map-equiv-left-summand) ~ id - is-retraction-map-inv-equiv-left-summand (pair (inl x) star) = refl - is-retraction-map-inv-equiv-left-summand (pair (inr x) ()) + is-retraction-map-inv-equiv-left-summand (inl x , star) = refl + is-retraction-map-inv-equiv-left-summand (inr x , ()) equiv-left-summand : (Σ (X + Y) is-left) ≃ X pr1 equiv-left-summand = map-equiv-left-summand @@ -143,22 +144,23 @@ module _ where map-equiv-right-summand : Σ (X + Y) is-right → Y - map-equiv-right-summand (pair (inl x) ()) - map-equiv-right-summand (pair (inr x) star) = x + map-equiv-right-summand (inl x , ()) + map-equiv-right-summand (inr x , star) = x map-inv-equiv-right-summand : Y → Σ (X + Y) is-right - map-inv-equiv-right-summand y = pair (inr y) star + pr1 (map-inv-equiv-right-summand y) = inr y + pr2 (map-inv-equiv-right-summand y) = star is-section-map-inv-equiv-right-summand : - (map-equiv-right-summand ∘ map-inv-equiv-right-summand) ~ id + map-equiv-right-summand ∘ map-inv-equiv-right-summand ~ id is-section-map-inv-equiv-right-summand y = refl is-retraction-map-inv-equiv-right-summand : - (map-inv-equiv-right-summand ∘ map-equiv-right-summand) ~ id - is-retraction-map-inv-equiv-right-summand (pair (inl x) ()) - is-retraction-map-inv-equiv-right-summand (pair (inr x) star) = refl + map-inv-equiv-right-summand ∘ map-equiv-right-summand ~ id + is-retraction-map-inv-equiv-right-summand (inl x , ()) + is-retraction-map-inv-equiv-right-summand (inr x , star) = refl - equiv-right-summand : (Σ (X + Y) is-right) ≃ Y + equiv-right-summand : Σ (X + Y) is-right ≃ Y pr1 equiv-right-summand = map-equiv-right-summand pr2 equiv-right-summand = is-equiv-is-invertible diff --git a/src/foundation/cospans.lagda.md b/src/foundation/cospans.lagda.md index 2438a1676d..1328e7dacd 100644 --- a/src/foundation/cospans.lagda.md +++ b/src/foundation/cospans.lagda.md @@ -27,7 +27,12 @@ open import foundation-core.torsorial-type-families ## Idea -A **cospan** is a pair of functions with a common codomain. +A {{#concept "cospan" Disambiguation="types" Agda=cospan}} is a +[pair](foundation.dependent-pair-types.md) of functions with a common codomain + +```text + f : A → X ← B : g. +``` ## Definition diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md index 2c9c8bb91a..d3ef6b55b1 100644 --- a/src/foundation/decidable-embeddings.lagda.md +++ b/src/foundation/decidable-embeddings.lagda.md @@ -238,8 +238,7 @@ abstract pr1 (is-decidable-emb-comp {g = g} {f} H K) = is-emb-comp _ _ (pr1 K) (pr1 H) pr2 (is-decidable-emb-comp {g = g} {f} H K) x = - ind-coprod - ( λ t → is-decidable (fiber (g ∘ f) x)) + rec-coprod ( λ u → is-decidable-equiv ( compute-fiber-comp g f x) @@ -249,7 +248,7 @@ abstract ( is-prop-map-is-emb (is-emb-is-decidable-emb K) x) ( u)) ( u)) ( is-decidable-map-is-decidable-emb H (pr1 u)))) - ( λ α → inr (λ t → α (pair (f (pr1 t)) (pr2 t)))) + ( λ α → inr (λ t → α (f (pr1 t) , pr2 t))) ( pr2 K x) ``` diff --git a/src/foundation/decidable-types.lagda.md b/src/foundation/decidable-types.lagda.md index f63c806b34..516877063a 100644 --- a/src/foundation/decidable-types.lagda.md +++ b/src/foundation/decidable-types.lagda.md @@ -95,7 +95,7 @@ is-decidable-coprod : is-decidable A → is-decidable B → is-decidable (A + B) is-decidable-coprod (inl a) y = inl (inl a) is-decidable-coprod (inr na) (inl b) = inl (inr b) -is-decidable-coprod (inr na) (inr nb) = inr (ind-coprod (λ x → empty) na nb) +is-decidable-coprod (inr na) (inr nb) = inr (rec-coprod na nb) ``` ### Cartesian products of decidable types are decidable diff --git a/src/foundation/descent-coproduct-types.lagda.md b/src/foundation/descent-coproduct-types.lagda.md index 243578e098..49319f7477 100644 --- a/src/foundation/descent-coproduct-types.lagda.md +++ b/src/foundation/descent-coproduct-types.lagda.md @@ -35,38 +35,30 @@ module _ {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} (f : A' → A) (g : B' → B) (h : X' → X) (αA : A → X) (αB : B → X) (αA' : A' → X') (αB' : B' → X') - (HA : (αA ∘ f) ~ (h ∘ αA')) (HB : (αB ∘ g) ~ (h ∘ αB')) + (HA : αA ∘ f ~ h ∘ αA') (HB : αB ∘ g ~ h ∘ αB') where triangle-descent-square-fiber-map-coprod-inl-fiber : (x : A) → - (map-fiber-cone αA h (triple f αA' HA) x) ~ - ( ( map-fiber-cone (ind-coprod _ αA αB) h - ( triple - ( map-coprod f g) - ( ind-coprod _ αA' αB') - ( ind-coprod _ HA HB)) - ( inl x)) ∘ - ( fiber-map-coprod-inl-fiber f g x)) - triangle-descent-square-fiber-map-coprod-inl-fiber x (pair a' p) = - eq-pair-Σ refl - ( ap (concat (inv (HA a')) (αA x)) - ( ap-comp (ind-coprod _ αA αB) inl p)) + ( map-fiber-vertical-map-cone αA h (f , αA' , HA) x) ~ + ( map-fiber-vertical-map-cone (ind-coprod _ αA αB) h + ( map-coprod f g , ind-coprod _ αA' αB' , ind-coprod _ HA HB) + ( inl x)) ∘ + ( fiber-map-coprod-inl-fiber f g x) + triangle-descent-square-fiber-map-coprod-inl-fiber x (a' , p) = + eq-pair-eq-pr2 + ( ap (concat (inv (HA a')) (αA x)) (ap-comp (ind-coprod _ αA αB) inl p)) triangle-descent-square-fiber-map-coprod-inr-fiber : (y : B) → - (map-fiber-cone αB h (triple g αB' HB) y) ~ - ( ( map-fiber-cone (ind-coprod _ αA αB) h - ( triple - ( map-coprod f g) - ( ind-coprod _ αA' αB') - ( ind-coprod _ HA HB)) - ( inr y)) ∘ - ( fiber-map-coprod-inr-fiber f g y)) - triangle-descent-square-fiber-map-coprod-inr-fiber y ( pair b' p) = - eq-pair-Σ refl - ( ap (concat (inv (HB b')) (αB y)) - ( ap-comp (ind-coprod _ αA αB) inr p)) + ( map-fiber-vertical-map-cone αB h (g , αB' , HB) y) ~ + ( map-fiber-vertical-map-cone (ind-coprod _ αA αB) h + ( map-coprod f g , ind-coprod _ αA' αB' , ind-coprod _ HA HB) + ( inr y)) ∘ + ( fiber-map-coprod-inr-fiber f g y) + triangle-descent-square-fiber-map-coprod-inr-fiber y (b' , p) = + eq-pair-eq-pr2 + ( ap (concat (inv (HB b')) (αB y)) (ap-comp (ind-coprod _ αA αB) inr p)) module _ {l1 l2 l3 l1' l2' l3' : Level} @@ -77,20 +69,11 @@ module _ cone-descent-coprod : (cone-A' : cone f i A') (cone-B' : cone g i B') → cone (ind-coprod _ f g) i (A' + B') - pr1 (cone-descent-coprod (pair h (pair f' H)) (pair k (pair g' K))) = - map-coprod h k - pr1 - ( pr2 (cone-descent-coprod (pair h (pair f' H)) (pair k (pair g' K)))) - ( inl a') = f' a' - pr1 - ( pr2 (cone-descent-coprod (pair h (pair f' H)) (pair k (pair g' K)))) - ( inr b') = g' b' - pr2 - ( pr2 (cone-descent-coprod (pair h (pair f' H)) (pair k (pair g' K)))) - ( inl a') = H a' - pr2 - ( pr2 (cone-descent-coprod (pair h (pair f' H)) (pair k (pair g' K)))) - ( inr b') = K b' + pr1 (cone-descent-coprod (h , f' , H) (k , g' , K)) = map-coprod h k + pr1 (pr2 (cone-descent-coprod (h , f' , H) (k , g' , K))) (inl a') = f' a' + pr1 (pr2 (cone-descent-coprod (h , f' , H) (k , g' , K))) (inr b') = g' b' + pr2 (pr2 (cone-descent-coprod (h , f' , H) (k , g' , K))) (inl a') = H a' + pr2 (pr2 (cone-descent-coprod (h , f' , H) (k , g' , K))) (inr b') = K b' abstract descent-coprod : @@ -98,44 +81,43 @@ module _ is-pullback f i cone-A' → is-pullback g i cone-B' → is-pullback (ind-coprod _ f g) i (cone-descent-coprod cone-A' cone-B') - descent-coprod (pair h (pair f' H)) (pair k (pair g' K)) - is-pb-cone-A' is-pb-cone-B' = - is-pullback-is-fiberwise-equiv-map-fiber-cone + descent-coprod (h , f' , H) (k , g' , K) is-pb-cone-A' is-pb-cone-B' = + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone ( ind-coprod _ f g) ( i) - ( cone-descent-coprod (triple h f' H) (triple k g' K)) + ( cone-descent-coprod (h , f' , H) (k , g' , K)) ( α) where α : is-fiberwise-equiv - ( map-fiber-cone + ( map-fiber-vertical-map-cone ( ind-coprod (λ _ → X) f g) ( i) - ( cone-descent-coprod (triple h f' H) (triple k g' K))) + ( cone-descent-coprod (h , f' , H) (k , g' , K))) α (inl x) = is-equiv-right-map-triangle - ( map-fiber-cone f i (triple h f' H) x) - ( map-fiber-cone (ind-coprod _ f g) i - ( cone-descent-coprod (triple h f' H) (triple k g' K)) + ( map-fiber-vertical-map-cone f i (h , f' , H) x) + ( map-fiber-vertical-map-cone (ind-coprod _ f g) i + ( cone-descent-coprod (h , f' , H) (k , g' , K)) ( inl x)) ( fiber-map-coprod-inl-fiber h k x) ( triangle-descent-square-fiber-map-coprod-inl-fiber h k i f g f' g' H K x) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback f i - ( triple h f' H) is-pb-cone-A' x) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback f i + ( h , f' , H) is-pb-cone-A' x) ( is-equiv-fiber-map-coprod-inl-fiber h k x) α (inr y) = is-equiv-right-map-triangle - ( map-fiber-cone g i (triple k g' K) y) - ( map-fiber-cone + ( map-fiber-vertical-map-cone g i (k , g' , K) y) + ( map-fiber-vertical-map-cone ( ind-coprod _ f g) i - ( cone-descent-coprod (triple h f' H) (triple k g' K)) + ( cone-descent-coprod (h , f' , H) (k , g' , K)) ( inr y)) ( fiber-map-coprod-inr-fiber h k y) ( triangle-descent-square-fiber-map-coprod-inr-fiber h k i f g f' g' H K y) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback g i - ( triple k g' K) is-pb-cone-B' y) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback g i + ( k , g' , K) is-pb-cone-B' y) ( is-equiv-fiber-map-coprod-inr-fiber h k y) abstract @@ -143,42 +125,48 @@ module _ (cone-A' : cone f i A') (cone-B' : cone g i B') → is-pullback (ind-coprod _ f g) i (cone-descent-coprod cone-A' cone-B') → is-pullback f i cone-A' - descent-coprod-inl (pair h (pair f' H)) (pair k (pair g' K)) is-pb-dsq = - is-pullback-is-fiberwise-equiv-map-fiber-cone f i (triple h f' H) - ( λ a → - is-equiv-left-map-triangle - ( map-fiber-cone f i (triple h f' H) a) - ( map-fiber-cone (ind-coprod _ f g) i - ( cone-descent-coprod (triple h f' H) (triple k g' K)) - ( inl a)) - ( fiber-map-coprod-inl-fiber h k a) - ( triangle-descent-square-fiber-map-coprod-inl-fiber - h k i f g f' g' H K a) - ( is-equiv-fiber-map-coprod-inl-fiber h k a) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback (ind-coprod _ f g) i - ( cone-descent-coprod ( triple h f' H) (triple k g' K)) - ( is-pb-dsq) - ( inl a))) + descent-coprod-inl (h , f' , H) (k , g' , K) is-pb-dsq = + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone f i + ( h , f' , H) + ( λ a → + is-equiv-left-map-triangle + ( map-fiber-vertical-map-cone f i (h , f' , H) a) + ( map-fiber-vertical-map-cone (ind-coprod _ f g) i + ( cone-descent-coprod (h , f' , H) (k , g' , K)) + ( inl a)) + ( fiber-map-coprod-inl-fiber h k a) + ( triangle-descent-square-fiber-map-coprod-inl-fiber + h k i f g f' g' H K a) + ( is-equiv-fiber-map-coprod-inl-fiber h k a) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( ind-coprod _ f g) + ( i) + ( cone-descent-coprod ( h , f' , H) (k , g' , K)) + ( is-pb-dsq) + ( inl a))) abstract descent-coprod-inr : (cone-A' : cone f i A') (cone-B' : cone g i B') → is-pullback (ind-coprod _ f g) i (cone-descent-coprod cone-A' cone-B') → is-pullback g i cone-B' - descent-coprod-inr (pair h (pair f' H)) (pair k (pair g' K)) is-pb-dsq = - is-pullback-is-fiberwise-equiv-map-fiber-cone g i (triple k g' K) - ( λ b → - is-equiv-left-map-triangle - ( map-fiber-cone g i (triple k g' K) b) - ( map-fiber-cone (ind-coprod _ f g) i - ( cone-descent-coprod (triple h f' H) (triple k g' K)) - ( inr b)) - ( fiber-map-coprod-inr-fiber h k b) - ( triangle-descent-square-fiber-map-coprod-inr-fiber - h k i f g f' g' H K b) - ( is-equiv-fiber-map-coprod-inr-fiber h k b) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback (ind-coprod _ f g) i - ( cone-descent-coprod (triple h f' H) (triple k g' K)) - ( is-pb-dsq) - ( inr b))) + descent-coprod-inr (h , f' , H) (k , g' , K) is-pb-dsq = + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone g i + ( k , g' , K) + ( λ b → + is-equiv-left-map-triangle + ( map-fiber-vertical-map-cone g i (k , g' , K) b) + ( map-fiber-vertical-map-cone (ind-coprod _ f g) i + ( cone-descent-coprod (h , f' , H) (k , g' , K)) + ( inr b)) + ( fiber-map-coprod-inr-fiber h k b) + ( triangle-descent-square-fiber-map-coprod-inr-fiber + h k i f g f' g' H K b) + ( is-equiv-fiber-map-coprod-inr-fiber h k b) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback + ( ind-coprod _ f g) + ( i) + ( cone-descent-coprod (h , f' , H) (k , g' , K)) + ( is-pb-dsq) + ( inr b))) ``` diff --git a/src/foundation/descent-dependent-pair-types.lagda.md b/src/foundation/descent-dependent-pair-types.lagda.md index ca2556629b..b111c50337 100644 --- a/src/foundation/descent-dependent-pair-types.lagda.md +++ b/src/foundation/descent-dependent-pair-types.lagda.md @@ -41,8 +41,8 @@ module _ triangle-descent-Σ : (i : I) (a : A i) → - ( map-fiber-cone (f i) h (c i) a) ~ - ( ( map-fiber-cone (ind-Σ f) h cone-descent-Σ (pair i a)) ∘ + ( map-fiber-vertical-map-cone (f i) h (c i) a) ~ + ( ( map-fiber-vertical-map-cone (ind-Σ f) h cone-descent-Σ (pair i a)) ∘ ( map-inv-compute-fiber-tot (λ i → (pr1 (c i))) (pair i a))) triangle-descent-Σ i .(pr1 (c i) a') (pair a' refl) = refl @@ -51,18 +51,18 @@ module _ ((i : I) → is-pullback (f i) h (c i)) → is-pullback (ind-Σ f) h cone-descent-Σ descent-Σ is-pb-c = - is-pullback-is-fiberwise-equiv-map-fiber-cone + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone ( ind-Σ f) ( h) ( cone-descent-Σ) ( ind-Σ ( λ i a → is-equiv-right-map-triangle - ( map-fiber-cone (f i) h (c i) a) - ( map-fiber-cone (ind-Σ f) h cone-descent-Σ (pair i a)) + ( map-fiber-vertical-map-cone (f i) h (c i) a) + ( map-fiber-vertical-map-cone (ind-Σ f) h cone-descent-Σ (pair i a)) ( map-inv-compute-fiber-tot (λ i → pr1 (c i)) (pair i a)) ( triangle-descent-Σ i a) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback (f i) h (c i) (is-pb-c i) a) ( is-equiv-map-inv-compute-fiber-tot (λ i → pr1 (c i)) (pair i a)))) @@ -71,15 +71,15 @@ module _ is-pullback (ind-Σ f) h cone-descent-Σ → ((i : I) → is-pullback (f i) h (c i)) descent-Σ' is-pb-dsq i = - is-pullback-is-fiberwise-equiv-map-fiber-cone (f i) h (c i) + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone (f i) h (c i) ( λ a → is-equiv-left-map-triangle - ( map-fiber-cone (f i) h (c i) a) - ( map-fiber-cone (ind-Σ f) h cone-descent-Σ (pair i a)) + ( map-fiber-vertical-map-cone (f i) h (c i) a) + ( map-fiber-vertical-map-cone (ind-Σ f) h cone-descent-Σ (pair i a)) ( map-inv-compute-fiber-tot (λ i → pr1 (c i)) (pair i a)) ( triangle-descent-Σ i a) ( is-equiv-map-inv-compute-fiber-tot (λ i → pr1 (c i)) (pair i a)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback ( ind-Σ f) ( h) ( cone-descent-Σ) diff --git a/src/foundation/descent-empty-types.lagda.md b/src/foundation/descent-empty-types.lagda.md index 54258e0723..4f80cf53a5 100644 --- a/src/foundation/descent-empty-types.lagda.md +++ b/src/foundation/descent-empty-types.lagda.md @@ -25,12 +25,14 @@ module _ where cone-empty : is-empty C → (C → B) → cone ex-falso g C - cone-empty p q = triple p q (λ c → ex-falso (p c)) + pr1 (cone-empty p q) = p + pr1 (pr2 (cone-empty p q)) = q + pr2 (pr2 (cone-empty p q)) c = ex-falso (p c) abstract descent-empty : (c : cone ex-falso g C) → is-pullback ex-falso g c descent-empty c = - is-pullback-is-fiberwise-equiv-map-fiber-cone _ g c ind-empty + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone _ g c ind-empty abstract descent-empty' : diff --git a/src/foundation/descent-equivalences.lagda.md b/src/foundation/descent-equivalences.lagda.md index c72997443c..c835b57cc9 100644 --- a/src/foundation/descent-equivalences.lagda.md +++ b/src/foundation/descent-equivalences.lagda.md @@ -53,25 +53,31 @@ module _ is-pullback j h c descent-is-equiv i j h c d is-equiv-i is-equiv-k is-pb-rectangle = - is-pullback-is-fiberwise-equiv-map-fiber-cone j h c + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone j h c ( map-inv-is-equiv-precomp-Π-is-equiv ( is-equiv-i) - ( λ y → is-equiv (map-fiber-cone j h c y)) + ( λ y → is-equiv (map-fiber-vertical-map-cone j h c y)) ( λ x → is-equiv-right-map-triangle - ( map-fiber-cone (j ∘ i) h + ( map-fiber-vertical-map-cone (j ∘ i) h ( pasting-horizontal-cone i j h c d) x) - ( map-fiber-cone j h c (i x)) - ( map-fiber-cone i (vertical-map-cone j h c) d x) - ( preserves-pasting-horizontal-map-fiber-cone i j h c d x) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback (j ∘ i) h + ( map-fiber-vertical-map-cone j h c (i x)) + ( map-fiber-vertical-map-cone i (vertical-map-cone j h c) d x) + ( preserves-pasting-horizontal-map-fiber-vertical-map-cone + ( i) + ( j) + ( h) + ( c) + ( d) + ( x)) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback (j ∘ i) h ( pasting-horizontal-cone i j h c d) ( is-pb-rectangle) ( x)) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback i + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback i ( vertical-map-cone j h c) ( d) - ( is-pullback-is-equiv' i + ( is-pullback-is-equiv-horizontal-maps i ( vertical-map-cone j h c) ( d) ( is-equiv-i) diff --git a/src/foundation/disjunction.lagda.md b/src/foundation/disjunction.lagda.md index 3aa838aa02..e845578552 100644 --- a/src/foundation/disjunction.lagda.md +++ b/src/foundation/disjunction.lagda.md @@ -101,7 +101,7 @@ elim-disjunction-Prop : ( conjunction-Prop (hom-Prop P R) (hom-Prop Q R)) ( hom-Prop (disjunction-Prop P Q) R) elim-disjunction-Prop P Q R (pair f g) = - map-universal-property-trunc-Prop R (ind-coprod (λ t → type-Prop R) f g) + map-universal-property-trunc-Prop R (rec-coprod f g) abstract is-equiv-ev-disjunction-Prop : diff --git a/src/foundation/embeddings.lagda.md b/src/foundation/embeddings.lagda.md index ecce3bbcc6..24bac8e1e7 100644 --- a/src/foundation/embeddings.lagda.md +++ b/src/foundation/embeddings.lagda.md @@ -361,14 +361,15 @@ module _ is-pullback f g c → is-emb g → is-emb (vertical-map-cone f g c) is-emb-vertical-map-cone-is-pullback pb is-emb-g = is-emb-is-prop-map - ( is-trunc-is-pullback neg-one-𝕋 f g c pb (is-prop-map-is-emb is-emb-g)) + ( is-trunc-vertical-map-is-pullback neg-one-𝕋 f g c pb + ( is-prop-map-is-emb is-emb-g)) abstract is-emb-horizontal-map-cone-is-pullback : is-pullback f g c → is-emb f → is-emb (horizontal-map-cone f g c) is-emb-horizontal-map-cone-is-pullback pb is-emb-f = is-emb-is-prop-map - ( is-trunc-is-pullback' neg-one-𝕋 f g c pb + ( is-trunc-horizontal-map-is-pullback neg-one-𝕋 f g c pb ( is-prop-map-is-emb is-emb-f)) ``` diff --git a/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md b/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md index 4760d7ea8c..193d6fb9ed 100644 --- a/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md +++ b/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md @@ -206,8 +206,7 @@ module _ ( diagonal-into-fibers-precomp f (type-Truncated-Type X)) ( is-equiv-diagonal-into-fibers-precomp-is-epimorphism-Truncated-Type e X) ( is-equiv-map-equiv - ( compute-total-fiber-precomp f - ( type-Truncated-Type X))) + ( compute-total-fiber-precomp f (type-Truncated-Type X))) is-equiv-horizontal-map-cocone-is-epimorphism-Truncated-Type : is-epimorphism-Truncated-Type k f → diff --git a/src/foundation/equality-coproduct-types.lagda.md b/src/foundation/equality-coproduct-types.lagda.md index bdb3e69575..6252418e66 100644 --- a/src/foundation/equality-coproduct-types.lagda.md +++ b/src/foundation/equality-coproduct-types.lagda.md @@ -8,13 +8,13 @@ module foundation.equality-coproduct-types where ```agda open import foundation.action-on-identifications-functions +open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.fundamental-theorem-of-identity-types open import foundation.negated-equality open import foundation.universe-levels open import foundation-core.contractible-types -open import foundation-core.coproduct-types open import foundation-core.embeddings open import foundation-core.empty-types open import foundation-core.equivalences @@ -22,6 +22,8 @@ open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.retractions +open import foundation-core.sections open import foundation-core.sets open import foundation-core.torsorial-type-families open import foundation-core.truncated-types @@ -77,20 +79,14 @@ module _ (x : A + B) → is-torsorial (Eq-coprod x) pr1 (pr1 (is-torsorial-Eq-coprod (inl x))) = inl x pr2 (pr1 (is-torsorial-Eq-coprod (inl x))) = Eq-eq-coprod-inl refl - pr2 - ( is-torsorial-Eq-coprod (inl x)) - ( pair (inl .x) (Eq-eq-coprod-inl refl)) = refl + pr2 (is-torsorial-Eq-coprod (inl x)) (.(inl x) , Eq-eq-coprod-inl refl) = refl pr1 (pr1 (is-torsorial-Eq-coprod (inr x))) = inr x pr2 (pr1 (is-torsorial-Eq-coprod (inr x))) = Eq-eq-coprod-inr refl - pr2 - ( is-torsorial-Eq-coprod (inr x)) - ( pair .(inr x) (Eq-eq-coprod-inr refl)) = refl + pr2 (is-torsorial-Eq-coprod (inr x)) (.(inr x) , Eq-eq-coprod-inr refl) = refl is-equiv-Eq-eq-coprod : (x y : A + B) → is-equiv (Eq-eq-coprod x y) is-equiv-Eq-eq-coprod x = - fundamental-theorem-id - ( is-torsorial-Eq-coprod x) - ( Eq-eq-coprod x) + fundamental-theorem-id (is-torsorial-Eq-coprod x) (Eq-eq-coprod x) extensionality-coprod : (x y : A + B) → (x = y) ≃ Eq-coprod x y pr1 (extensionality-coprod x y) = Eq-eq-coprod x y @@ -256,6 +252,34 @@ module _ pr2 emb-inr = is-emb-inr ``` +Moreover, `is-injective-inl` and `is-injective-inr` are the inverses. + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-retraction-is-injective-inl : + {x y : A} → + is-retraction (ap (inl {A = A} {B = B}) {x} {y}) (is-injective-inl) + is-retraction-is-injective-inl refl = refl + + is-section-is-injective-inl : + {x y : A} → + is-section (ap (inl {A = A} {B = B}) {x} {y}) (is-injective-inl) + is-section-is-injective-inl refl = refl + + is-retraction-is-injective-inr : + {x y : B} → + is-retraction (ap (inr {A = A} {B = B}) {x} {y}) (is-injective-inr) + is-retraction-is-injective-inr refl = refl + + is-section-is-injective-inr : + {x y : B} → + is-section (ap (inr {A = A} {B = B}) {x} {y}) (is-injective-inr) + is-section-is-injective-inr refl = refl +``` + ### A map `A + B → C` defined by maps `f : A → C` and `B → C` is an embedding if both `f` and `g` are embeddings and they don't overlap ```agda @@ -265,30 +289,30 @@ module _ is-emb-coprod : is-emb f → is-emb g → ((a : A) (b : B) → f a ≠ g b) → - is-emb (ind-coprod (λ x → C) f g) + is-emb (rec-coprod f g) is-emb-coprod H K L (inl a) (inl a') = is-equiv-right-map-triangle ( ap f) - ( ap (ind-coprod (λ x → C) f g)) + ( ap (rec-coprod f g)) ( ap inl) - ( λ p → ap-comp (ind-coprod (λ x → C) f g) inl p) + ( ap-comp (rec-coprod f g) inl) ( H a a') ( is-emb-inl A B a a') is-emb-coprod H K L (inl a) (inr b') = - is-equiv-is-empty (ap (ind-coprod (λ x → C) f g)) (L a b') + is-equiv-is-empty (ap (rec-coprod f g)) (L a b') is-emb-coprod H K L (inr b) (inl a') = - is-equiv-is-empty (ap (ind-coprod (λ x → C) f g)) (L a' b ∘ inv) + is-equiv-is-empty (ap (rec-coprod f g)) (L a' b ∘ inv) is-emb-coprod H K L (inr b) (inr b') = is-equiv-right-map-triangle ( ap g) - ( ap (ind-coprod (λ x → C) f g)) + ( ap (rec-coprod f g)) ( ap inr) - ( λ p → ap-comp (ind-coprod (λ x → C) f g) inr p) + ( ap-comp (rec-coprod f g) inr) ( K b b') ( is-emb-inr A B b b') ``` -### Coproducts of (k+2)-truncated types are (k+2)-truncated +### Coproducts of `k+2`-truncated types are `k+2`-truncated ```agda module _ @@ -326,8 +350,8 @@ abstract coprod-Set : {l1 l2 : Level} (A : Set l1) (B : Set l2) → Set (l1 ⊔ l2) -pr1 (coprod-Set (pair A is-set-A) (pair B is-set-B)) = A + B -pr2 (coprod-Set (pair A is-set-A) (pair B is-set-B)) = +pr1 (coprod-Set (A , is-set-A) (B , is-set-B)) = A + B +pr2 (coprod-Set (A , is-set-A) (B , is-set-B)) = is-set-coprod is-set-A is-set-B ``` diff --git a/src/foundation/equivalences-arrows.lagda.md b/src/foundation/equivalences-arrows.lagda.md new file mode 100644 index 0000000000..000c4c22ef --- /dev/null +++ b/src/foundation/equivalences-arrows.lagda.md @@ -0,0 +1,256 @@ +# Equivalences of arrows + +```agda +module foundation.equivalences-arrows where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-morphisms-arrows +open import foundation.commuting-squares-of-homotopies +open import foundation.commuting-squares-of-identifications +open import foundation.commuting-squares-of-maps +open import foundation.commuting-triangles-of-identifications +open import foundation.cones-over-cospans +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopy-induction +open import foundation.morphisms-arrows +open import foundation.path-algebra +open import foundation.structure-identity-principle +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.function-types +open import foundation-core.homotopies +open import foundation-core.identity-types +open import foundation-core.torsorial-type-families +open import foundation-core.whiskering-homotopies +``` + +
+ +## Idea + +An {{#concept "equivalence of arrows"}} from a function `f : A → B` to a +function `g : X → Y` is a [morphism of arrows](foundation.morphisms-arrows.md) + +```text + i + A ----> X + | | + f | H | g + v v + B ----> Y + j +``` + +such that `i` and `j` are [equivalences](foundation-core.equivalences.md). + +## Definitions + +### The predicate of being an equivalence of arrows on morphisms of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (α : hom-arrow f g) + where + + is-equiv-hom-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-equiv-hom-arrow = + is-equiv (map-domain-hom-arrow f g α) × + is-equiv (map-codomain-hom-arrow f g α) + + is-equiv-map-domain-is-equiv-hom-arrow : + is-equiv-hom-arrow → is-equiv (map-domain-hom-arrow f g α) + is-equiv-map-domain-is-equiv-hom-arrow = pr1 + + is-equiv-map-codomain-is-equiv-hom-arrow : + is-equiv-hom-arrow → is-equiv (map-codomain-hom-arrow f g α) + is-equiv-map-codomain-is-equiv-hom-arrow = pr2 +``` + +### The type of equivalences of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + equiv-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + equiv-arrow = Σ (hom-arrow f g) (is-equiv-hom-arrow f g) + +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (α : equiv-arrow f g) + where + + hom-arrow-equiv-arrow : hom-arrow f g + hom-arrow-equiv-arrow = pr1 α + + is-equiv-hom-arrow-equiv-arrow : is-equiv-hom-arrow f g hom-arrow-equiv-arrow + is-equiv-hom-arrow-equiv-arrow = pr2 α + + map-domain-equiv-arrow : A → X + map-domain-equiv-arrow = map-domain-hom-arrow f g hom-arrow-equiv-arrow + + is-equiv-map-domain-equiv-arrow : is-equiv map-domain-equiv-arrow + is-equiv-map-domain-equiv-arrow = + is-equiv-map-domain-is-equiv-hom-arrow f g + ( hom-arrow-equiv-arrow) + ( is-equiv-hom-arrow-equiv-arrow) + + equiv-domain-equiv-arrow : A ≃ X + pr1 equiv-domain-equiv-arrow = map-domain-equiv-arrow + pr2 equiv-domain-equiv-arrow = is-equiv-map-domain-equiv-arrow + + map-codomain-equiv-arrow : B → Y + map-codomain-equiv-arrow = map-codomain-hom-arrow f g hom-arrow-equiv-arrow + + is-equiv-map-codomain-equiv-arrow : is-equiv map-codomain-equiv-arrow + is-equiv-map-codomain-equiv-arrow = + is-equiv-map-codomain-is-equiv-hom-arrow f g + ( hom-arrow-equiv-arrow) + ( is-equiv-hom-arrow-equiv-arrow) + + equiv-codomain-equiv-arrow : B ≃ Y + pr1 equiv-codomain-equiv-arrow = map-codomain-equiv-arrow + pr2 equiv-codomain-equiv-arrow = is-equiv-map-codomain-equiv-arrow + + coh-equiv-arrow : + coherence-square-maps map-domain-equiv-arrow f g map-codomain-equiv-arrow + coh-equiv-arrow = coh-hom-arrow f g hom-arrow-equiv-arrow +``` + +### The identity equivalence of arrows + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + id-equiv-arrow : equiv-arrow f f + pr1 id-equiv-arrow = id-hom-arrow + pr1 (pr2 id-equiv-arrow) = is-equiv-id + pr2 (pr2 id-equiv-arrow) = is-equiv-id +``` + +### Composition of equivalences of arrows + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {U : UU l5} {V : UU l6} + (f : A → B) (g : X → Y) (h : U → V) + (b : equiv-arrow g h) (a : equiv-arrow f g) + where + + comp-equiv-arrow : equiv-arrow f h + pr1 comp-equiv-arrow = + comp-hom-arrow f g h + ( hom-arrow-equiv-arrow g h b) + ( hom-arrow-equiv-arrow f g a) + pr1 (pr2 comp-equiv-arrow) = + is-equiv-comp + ( map-domain-equiv-arrow g h b) + ( map-domain-equiv-arrow f g a) + ( is-equiv-map-domain-equiv-arrow f g a) + ( is-equiv-map-domain-equiv-arrow g h b) + pr2 (pr2 comp-equiv-arrow) = + is-equiv-comp + ( map-codomain-equiv-arrow g h b) + ( map-codomain-equiv-arrow f g a) + ( is-equiv-map-codomain-equiv-arrow f g a) + ( is-equiv-map-codomain-equiv-arrow g h b) +``` + +### Inverses of equivalences of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (α : equiv-arrow f g) + where + + hom-inv-equiv-arrow : hom-arrow g f + pr1 hom-inv-equiv-arrow = + map-inv-is-equiv (is-equiv-map-domain-equiv-arrow f g α) + pr1 (pr2 hom-inv-equiv-arrow) = + map-inv-is-equiv (is-equiv-map-codomain-equiv-arrow f g α) + pr2 (pr2 hom-inv-equiv-arrow) = + coherence-square-inv-horizontal + ( equiv-domain-equiv-arrow f g α) + ( f) + ( g) + ( equiv-codomain-equiv-arrow f g α) + ( coh-equiv-arrow f g α) + + is-equiv-inv-equiv-arrow : is-equiv-hom-arrow g f hom-inv-equiv-arrow + pr1 is-equiv-inv-equiv-arrow = + is-equiv-map-inv-is-equiv (is-equiv-map-domain-equiv-arrow f g α) + pr2 is-equiv-inv-equiv-arrow = + is-equiv-map-inv-is-equiv (is-equiv-map-codomain-equiv-arrow f g α) + + inv-equiv-arrow : equiv-arrow g f + pr1 inv-equiv-arrow = hom-inv-equiv-arrow + pr2 inv-equiv-arrow = is-equiv-inv-equiv-arrow +``` + +### If a map is equivalent to an equivalence, then it is an equivalence + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (α : equiv-arrow f g) + where + + is-equiv-source-is-equiv-target-equiv-arrow : is-equiv g → is-equiv f + is-equiv-source-is-equiv-target-equiv-arrow = + is-equiv-left-is-equiv-right-square + ( f) + ( g) + ( map-domain-equiv-arrow f g α) + ( map-codomain-equiv-arrow f g α) + ( coh-equiv-arrow f g α) + ( is-equiv-map-domain-equiv-arrow f g α) + ( is-equiv-map-codomain-equiv-arrow f g α) + + is-equiv-target-is-equiv-source-equiv-arrow : is-equiv f → is-equiv g + is-equiv-target-is-equiv-source-equiv-arrow = + is-equiv-right-is-equiv-left-square + ( f) + ( g) + ( map-domain-equiv-arrow f g α) + ( map-codomain-equiv-arrow f g α) + ( coh-equiv-arrow f g α) + ( is-equiv-map-domain-equiv-arrow f g α) + ( is-equiv-map-codomain-equiv-arrow f g α) +``` + +### Equivalences of arrows are cartesian + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (α : equiv-arrow f g) + where + + is-cartesian-equiv-arrow : + is-cartesian-hom-arrow f g (hom-arrow-equiv-arrow f g α) + is-cartesian-equiv-arrow = + is-pullback-is-equiv-horizontal-maps + ( map-codomain-equiv-arrow f g α) + ( g) + ( cone-hom-arrow f g (hom-arrow-equiv-arrow f g α)) + ( is-equiv-map-codomain-equiv-arrow f g α) + ( is-equiv-map-domain-equiv-arrow f g α) + + cartesian-hom-arrow-equiv-arrow : cartesian-hom-arrow f g + pr1 cartesian-hom-arrow-equiv-arrow = hom-arrow-equiv-arrow f g α + pr2 cartesian-hom-arrow-equiv-arrow = is-cartesian-equiv-arrow +``` diff --git a/src/foundation/equivalences.lagda.md b/src/foundation/equivalences.lagda.md index 60b77a2f67..0923096b18 100644 --- a/src/foundation/equivalences.lagda.md +++ b/src/foundation/equivalences.lagda.md @@ -753,20 +753,23 @@ module _ where abstract - is-equiv-is-pullback : is-equiv g → is-pullback f g c → is-equiv (pr1 c) - is-equiv-is-pullback is-equiv-g pb = + is-equiv-vertical-map-is-pullback : + is-equiv g → is-pullback f g c → is-equiv (vertical-map-cone f g c) + is-equiv-vertical-map-is-pullback is-equiv-g pb = is-equiv-is-contr-map - ( is-trunc-is-pullback neg-two-𝕋 f g c pb + ( is-trunc-vertical-map-is-pullback neg-two-𝕋 f g c pb ( is-contr-map-is-equiv is-equiv-g)) abstract - is-pullback-is-equiv : is-equiv g → is-equiv (pr1 c) → is-pullback f g c - is-pullback-is-equiv is-equiv-g is-equiv-p = - is-pullback-is-fiberwise-equiv-map-fiber-cone f g c - ( λ a → is-equiv-is-contr - ( map-fiber-cone f g c a) - ( is-contr-map-is-equiv is-equiv-p a) - ( is-contr-map-is-equiv is-equiv-g (f a))) + is-pullback-is-equiv-vertical-maps : + is-equiv g → is-equiv (vertical-map-cone f g c) → is-pullback f g c + is-pullback-is-equiv-vertical-maps is-equiv-g is-equiv-p = + is-pullback-is-fiberwise-equiv-map-fiber-vertical-map-cone f g c + ( λ a → + is-equiv-is-contr + ( map-fiber-vertical-map-cone f g c a) + ( is-contr-map-is-equiv is-equiv-p a) + ( is-contr-map-is-equiv is-equiv-g (f a))) module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} @@ -774,19 +777,19 @@ module _ where abstract - is-equiv-is-pullback' : - is-equiv f → is-pullback f g c → is-equiv (pr1 (pr2 c)) - is-equiv-is-pullback' is-equiv-f pb = + is-equiv-horizontal-map-is-pullback : + is-equiv f → is-pullback f g c → is-equiv (horizontal-map-cone f g c) + is-equiv-horizontal-map-is-pullback is-equiv-f pb = is-equiv-is-contr-map - ( is-trunc-is-pullback' neg-two-𝕋 f g c pb + ( is-trunc-horizontal-map-is-pullback neg-two-𝕋 f g c pb ( is-contr-map-is-equiv is-equiv-f)) abstract - is-pullback-is-equiv' : - is-equiv f → is-equiv (pr1 (pr2 c)) → is-pullback f g c - is-pullback-is-equiv' is-equiv-f is-equiv-q = + is-pullback-is-equiv-horizontal-maps : + is-equiv f → is-equiv (horizontal-map-cone f g c) → is-pullback f g c + is-pullback-is-equiv-horizontal-maps is-equiv-f is-equiv-q = is-pullback-swap-cone' f g c - ( is-pullback-is-equiv g f + ( is-pullback-is-equiv-vertical-maps g f ( swap-cone f g c) is-equiv-f is-equiv-q) diff --git a/src/foundation/fibered-equivalences.lagda.md b/src/foundation/fibered-equivalences.lagda.md index 5b09867e3b..c7dc8a27f4 100644 --- a/src/foundation/fibered-equivalences.lagda.md +++ b/src/foundation/fibered-equivalences.lagda.md @@ -316,13 +316,16 @@ module _ is-fibered-equiv-fibered-map f g ihH pr1 (is-fibered-equiv-is-pullback is-equiv-i pb) = is-equiv-i pr2 (is-fibered-equiv-is-pullback is-equiv-i pb) = - is-equiv-is-pullback' (pr1 ihH) g (cone-fibered-map f g ihH) is-equiv-i pb + is-equiv-horizontal-map-is-pullback (pr1 ihH) g + ( cone-fibered-map f g ihH) + ( is-equiv-i) + ( pb) is-pullback-is-fibered-equiv : is-fibered-equiv-fibered-map f g ihH → is-pullback (pr1 ihH) g (cone-fibered-map f g ihH) is-pullback-is-fibered-equiv (is-equiv-i , is-equiv-h) = - is-pullback-is-equiv' + is-pullback-is-equiv-horizontal-maps (pr1 ihH) g (cone-fibered-map f g ihH) is-equiv-i is-equiv-h equiv-is-fibered-equiv-is-pullback : @@ -331,7 +334,7 @@ module _ is-fibered-equiv-fibered-map f g ihH equiv-is-fibered-equiv-is-pullback is-equiv-i = equiv-prop - ( is-property-is-pullback (pr1 ihH) g (cone-fibered-map f g ihH)) + ( is-prop-is-pullback (pr1 ihH) g (cone-fibered-map f g ihH)) ( is-prop-is-fibered-equiv-fibered-map f g ihH) ( is-fibered-equiv-is-pullback is-equiv-i) ( is-pullback-is-fibered-equiv) diff --git a/src/foundation/fibered-maps.lagda.md b/src/foundation/fibered-maps.lagda.md index 35509c3867..73e852a3ea 100644 --- a/src/foundation/fibered-maps.lagda.md +++ b/src/foundation/fibered-maps.lagda.md @@ -550,5 +550,7 @@ module _ ## See also +- [Morphisms of arrows](foundation.morphisms-arrows.md) for the same concept + under a different name. - For the pullback property of the type of fibered maps, see [the pullback-hom](orthogonal-factorization-systems.pullback-hom.md) diff --git a/src/foundation/fibers-of-maps.lagda.md b/src/foundation/fibers-of-maps.lagda.md index 0f11beebae..e77588d386 100644 --- a/src/foundation/fibers-of-maps.lagda.md +++ b/src/foundation/fibers-of-maps.lagda.md @@ -17,7 +17,6 @@ open import foundation.type-arithmetic-unit-type open import foundation.unit-type open import foundation.universe-levels -open import foundation-core.constant-maps open import foundation-core.contractible-types open import foundation-core.equivalences open import foundation-core.function-types diff --git a/src/foundation/function-extensionality.lagda.md b/src/foundation/function-extensionality.lagda.md index 3352a3cf7b..910709f6f4 100644 --- a/src/foundation/function-extensionality.lagda.md +++ b/src/foundation/function-extensionality.lagda.md @@ -11,17 +11,13 @@ open import foundation-core.function-extensionality public ```agda open import foundation.action-on-identifications-functions open import foundation.commuting-squares-of-maps -open import foundation.dependent-pair-types -open import foundation.implicit-function-types -open import foundation.injective-maps open import foundation.postcomposition-dependent-functions -open import foundation.postcomposition-functions open import foundation.universe-levels -open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.postcomposition-functions open import foundation-core.precomposition-dependent-functions open import foundation-core.precomposition-functions open import foundation-core.whiskering-homotopies diff --git a/src/foundation/functoriality-cartesian-product-types.lagda.md b/src/foundation/functoriality-cartesian-product-types.lagda.md index 2aef1d8066..415ac2b111 100644 --- a/src/foundation/functoriality-cartesian-product-types.lagda.md +++ b/src/foundation/functoriality-cartesian-product-types.lagda.md @@ -99,11 +99,11 @@ map-prod-id (a , b) = refl ### Functoriality of products preserves composition ```agda -map-prod-comp : +preserves-comp-map-prod : {l1 l2 l3 l4 l5 l6 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} {E : UU l5} {F : UU l6} (f : A → C) (g : B → D) (h : C → E) (k : D → F) → map-prod (h ∘ f) (k ∘ g) ~ map-prod h k ∘ map-prod f g -map-prod-comp f g h k t = refl +preserves-comp-map-prod f g h k t = refl ``` ### Functoriality of products preserves homotopies diff --git a/src/foundation/functoriality-coproduct-types.lagda.md b/src/foundation/functoriality-coproduct-types.lagda.md index a0f1373328..342c95baa2 100644 --- a/src/foundation/functoriality-coproduct-types.lagda.md +++ b/src/foundation/functoriality-coproduct-types.lagda.md @@ -86,7 +86,7 @@ module _ where preserves-comp-map-coprod : - (map-coprod (f' ∘ f) (g' ∘ g)) ~ ((map-coprod f' g') ∘ (map-coprod f g)) + map-coprod (f' ∘ f) (g' ∘ g) ~ map-coprod f' g' ∘ map-coprod f g preserves-comp-map-coprod (inl x) = refl preserves-comp-map-coprod (inr y) = refl ``` @@ -99,7 +99,7 @@ module _ {f f' : A → A'} (H : f ~ f') {g g' : B → B'} (K : g ~ g') where - htpy-map-coprod : (map-coprod f g) ~ (map-coprod f' g') + htpy-map-coprod : map-coprod f g ~ map-coprod f' g' htpy-map-coprod (inl x) = ap inl (H x) htpy-map-coprod (inr y) = ap inr (K y) ``` @@ -114,29 +114,29 @@ module _ fiber-map-coprod-inl-fiber : (x : A) → fiber f x → fiber (map-coprod f g) (inl x) - pr1 (fiber-map-coprod-inl-fiber x (pair a' p)) = inl a' - pr2 (fiber-map-coprod-inl-fiber x (pair a' p)) = ap inl p + pr1 (fiber-map-coprod-inl-fiber x (a' , p)) = inl a' + pr2 (fiber-map-coprod-inl-fiber x (a' , p)) = ap inl p fiber-fiber-map-coprod-inl : (x : A) → fiber (map-coprod f g) (inl x) → fiber f x - fiber-fiber-map-coprod-inl x (pair (inl a') p) = - pair a' (map-compute-eq-coprod-inl-inl (f a') x p) - fiber-fiber-map-coprod-inl x (pair (inr b') p) = + pr1 (fiber-fiber-map-coprod-inl x (inl a' , p)) = a' + pr2 (fiber-fiber-map-coprod-inl x (inl a' , p)) = + map-compute-eq-coprod-inl-inl (f a') x p + fiber-fiber-map-coprod-inl x (inr b' , p) = ex-falso (is-empty-eq-coprod-inr-inl (g b') x p) abstract is-section-fiber-fiber-map-coprod-inl : - (x : A) → - (fiber-map-coprod-inl-fiber x ∘ fiber-fiber-map-coprod-inl x) ~ id - is-section-fiber-fiber-map-coprod-inl .(f a') (pair (inl a') refl) = refl - is-section-fiber-fiber-map-coprod-inl x (pair (inr b') p) = + (x : A) → fiber-map-coprod-inl-fiber x ∘ fiber-fiber-map-coprod-inl x ~ id + is-section-fiber-fiber-map-coprod-inl .(f a') (inl a' , refl) = refl + is-section-fiber-fiber-map-coprod-inl x (inr b' , p) = ex-falso (is-empty-eq-coprod-inr-inl (g b') x p) abstract is-retraction-fiber-fiber-map-coprod-inl : (x : A) → (fiber-fiber-map-coprod-inl x ∘ fiber-map-coprod-inl-fiber x) ~ id - is-retraction-fiber-fiber-map-coprod-inl .(f a') (pair a' refl) = refl + is-retraction-fiber-fiber-map-coprod-inl .(f a') (a' , refl) = refl abstract is-equiv-fiber-map-coprod-inl-fiber : @@ -149,30 +149,30 @@ module _ fiber-map-coprod-inr-fiber : (y : B) → fiber g y → fiber (map-coprod f g) (inr y) - pr1 (fiber-map-coprod-inr-fiber y (pair b' p)) = inr b' - pr2 (fiber-map-coprod-inr-fiber y (pair b' p)) = ap inr p + pr1 (fiber-map-coprod-inr-fiber y (b' , p)) = inr b' + pr2 (fiber-map-coprod-inr-fiber y (b' , p)) = ap inr p fiber-fiber-map-coprod-inr : (y : B) → fiber (map-coprod f g) (inr y) → fiber g y - fiber-fiber-map-coprod-inr y (pair (inl a') p) = + fiber-fiber-map-coprod-inr y (inl a' , p) = ex-falso (is-empty-eq-coprod-inl-inr (f a') y p) - pr1 (fiber-fiber-map-coprod-inr y (pair (inr b') p)) = b' - pr2 (fiber-fiber-map-coprod-inr y (pair (inr b') p)) = + pr1 (fiber-fiber-map-coprod-inr y (inr b' , p)) = b' + pr2 (fiber-fiber-map-coprod-inr y (inr b' , p)) = map-compute-eq-coprod-inr-inr (g b') y p abstract is-section-fiber-fiber-map-coprod-inr : (y : B) → (fiber-map-coprod-inr-fiber y ∘ fiber-fiber-map-coprod-inr y) ~ id - is-section-fiber-fiber-map-coprod-inr .(g b') (pair (inr b') refl) = refl - is-section-fiber-fiber-map-coprod-inr y (pair (inl a') p) = + is-section-fiber-fiber-map-coprod-inr .(g b') (inr b' , refl) = refl + is-section-fiber-fiber-map-coprod-inr y (inl a' , p) = ex-falso (is-empty-eq-coprod-inl-inr (f a') y p) abstract is-retraction-fiber-fiber-map-coprod-inr : (y : B) → (fiber-fiber-map-coprod-inr y ∘ fiber-map-coprod-inr-fiber y) ~ id - is-retraction-fiber-fiber-map-coprod-inr .(g b') (pair b' refl) = refl + is-retraction-fiber-fiber-map-coprod-inr .(g b') (b' , refl) = refl abstract is-equiv-fiber-map-coprod-inr-fiber : @@ -198,36 +198,34 @@ module _ pr1 ( pr1 ( is-equiv-map-coprod - ( pair (pair sf Sf) (pair rf Rf)) - ( pair (pair sg Sg) (pair rg Rg)))) = map-coprod sf sg + ( (sf , Sf) , (rf , Rf)) + ( (sg , Sg) , (rg , Rg)))) = map-coprod sf sg pr2 ( pr1 ( is-equiv-map-coprod {f} {g} - ( pair (pair sf Sf) (pair rf Rf)) - ( pair (pair sg Sg) (pair rg Rg)))) = + ( (sf , Sf) , (rf , Rf)) + ( (sg , Sg) , (rg , Rg)))) = ( ( inv-htpy (preserves-comp-map-coprod sf f sg g)) ∙h ( htpy-map-coprod Sf Sg)) ∙h ( id-map-coprod A' B') pr1 ( pr2 ( is-equiv-map-coprod - ( pair (pair sf Sf) (pair rf Rf)) - ( pair (pair sg Sg) (pair rg Rg)))) = map-coprod rf rg + ( (sf , Sf) , (rf , Rf)) + ( (sg , Sg) , (rg , Rg)))) = map-coprod rf rg pr2 ( pr2 ( is-equiv-map-coprod {f} {g} - ( pair (pair sf Sf) (pair rf Rf)) - ( pair (pair sg Sg) (pair rg Rg)))) = + ( (sf , Sf) , (rf , Rf)) + ( (sg , Sg) , (rg , Rg)))) = ( ( inv-htpy (preserves-comp-map-coprod f rf g rg)) ∙h ( htpy-map-coprod Rf Rg)) ∙h ( id-map-coprod A B) - map-equiv-coprod : - (A ≃ A') → (B ≃ B') → A + B → A' + B' + map-equiv-coprod : A ≃ A' → B ≃ B' → A + B → A' + B' map-equiv-coprod e e' = map-coprod (map-equiv e) (map-equiv e') - equiv-coprod : - (A ≃ A') → (B ≃ B') → (A + B) ≃ (A' + B') + equiv-coprod : A ≃ A' → B ≃ B' → (A + B) ≃ (A' + B') pr1 (equiv-coprod e e') = map-equiv-coprod e e' pr2 (equiv-coprod e e') = is-equiv-map-coprod (is-equiv-map-equiv e) (is-equiv-map-equiv e') @@ -262,8 +260,9 @@ is-contr-fiber-map-coprod : {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} (f : A → B) (g : C → D) → is-contr - ( fiber ( λ (fg' : (A → B) × (C → D)) → map-coprod (pr1 fg') (pr2 fg')) - ( map-coprod f g)) + ( fiber + ( λ (fg' : (A → B) × (C → D)) → map-coprod (pr1 fg') (pr2 fg')) + ( map-coprod f g)) is-contr-fiber-map-coprod {A = A} {B} {C} {D} f g = is-contr-equiv ( Σ ( (A → B) × (C → D)) @@ -271,30 +270,18 @@ is-contr-fiber-map-coprod {A = A} {B} {C} {D} f g = ((a : A) → pr1 fg' a = f a) × ((c : C) → pr2 fg' c = g c))) ( equiv-tot ( λ fg' → - ( ( equiv-prod - ( equiv-Π-equiv-family - ( λ a → compute-eq-coprod-inl-inl (pr1 fg' a) (f a))) - ( equiv-Π-equiv-family - ( λ c → compute-eq-coprod-inr-inr (pr2 fg' c) (g c)))) ∘e - ( equiv-dependent-universal-property-coprod - ( λ x → - map-coprod (pr1 fg') (pr2 fg') x = map-coprod f g x))) ∘e + ( equiv-prod + ( equiv-Π-equiv-family + ( λ a → compute-eq-coprod-inl-inl (pr1 fg' a) (f a))) + ( equiv-Π-equiv-family + ( λ c → compute-eq-coprod-inr-inr (pr2 fg' c) (g c)))) ∘e + ( equiv-dependent-universal-property-coprod + ( λ x → map-coprod (pr1 fg') (pr2 fg') x = map-coprod f g x)) ∘e ( equiv-funext))) ( is-torsorial-Eq-structure ( is-torsorial-htpy' f) - ( pair f refl-htpy) + ( f , refl-htpy) ( is-torsorial-htpy' g)) - -{- -is-emb-map-coprod : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} → - is-emb (λ (fg : (A → B) × (C → D)) → map-coprod (pr1 fg) (pr2 fg)) -is-emb-map-coprod (pair f g) = - fundamental-theorem-id (pair f g) - ( refl) - {! is-contr-fiber-map-coprod f g!} - {!!} --} ``` ### For any equivalence `f : A + B ≃ A + B` and `g : B ≃ B` such that `f` and `g` coincide on `B`, we construct an `h : A ≃ A` such that `htpy-equiv (equiv-coprod h d) f` @@ -311,9 +298,9 @@ module _ equiv-coproduct-induce-equiv-disjoint f g p x y q = neq-inr-inl ( is-injective-map-equiv f - ( ( p (map-equiv (inv-equiv g) y) ∙ - ( ( ap (λ z → inr (map-equiv z y)) (right-inverse-law-equiv g)) ∙ - ( inv q))))) + ( ( p ( map-equiv (inv-equiv g) y)) ∙ + ( ap (λ z → inr (map-equiv z y)) (right-inverse-law-equiv g)) ∙ + ( inv q))) inv-commutative-square-inr : (f : (A + B) ≃ (A + B)) (g : B ≃ B) @@ -322,8 +309,8 @@ module _ inv-commutative-square-inr f g p b = is-injective-map-equiv f ( ( ap (λ z → map-equiv z (inr b)) (right-inverse-law-equiv f)) ∙ - ( ( inv (ap (λ z → inr (map-equiv z b)) (right-inverse-law-equiv g))) ∙ - ( inv (p (map-inv-equiv g b))))) + ( inv (ap (λ z → inr (map-equiv z b)) (right-inverse-law-equiv g))) ∙ + ( inv (p (map-inv-equiv g b)))) cases-retraction-equiv-coprod : (f : (A + B) ≃ (A + B)) (g : B ≃ B) @@ -381,8 +368,8 @@ module _ section-cases-retraction-equiv-coprod f g p x (inl y) (inl z) q r = is-injective-inl ( ( inv r) ∙ - ( ( ap (map-equiv f) (inv q)) ∙ - ( ap (λ w → map-equiv w (inl x)) (right-inverse-law-equiv f)))) + ( ap (map-equiv f) (inv q)) ∙ + ( ap (λ w → map-equiv w (inl x)) (right-inverse-law-equiv f))) section-cases-retraction-equiv-coprod f g p x (inl y) (inr z) q r = ex-falso (equiv-coproduct-induce-equiv-disjoint f g p y z r) section-cases-retraction-equiv-coprod f g p x (inr y) z q r = @@ -508,11 +495,11 @@ module _ inv-equiv (equiv-right-summand)) map-inv-mutually-exclusive-coprod : (P ≃ P') × (Q ≃ Q') → (P + Q) ≃ (P' + Q') - map-inv-mutually-exclusive-coprod (pair e₁ e₂) = equiv-coprod e₁ e₂ + map-inv-mutually-exclusive-coprod (e₁ , e₂) = equiv-coprod e₁ e₂ is-retraction-map-inv-mutually-exclusive-coprod : (map-mutually-exclusive-coprod ∘ map-inv-mutually-exclusive-coprod) ~ id - is-retraction-map-inv-mutually-exclusive-coprod (pair e₁ e₂) = + is-retraction-map-inv-mutually-exclusive-coprod (e₁ , e₂) = eq-pair (eq-equiv-eq-map-equiv refl) (eq-equiv-eq-map-equiv refl) diff --git a/src/foundation/functoriality-dependent-pair-types.lagda.md b/src/foundation/functoriality-dependent-pair-types.lagda.md index da486d448c..929bf7f24c 100644 --- a/src/foundation/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation/functoriality-dependent-pair-types.lagda.md @@ -59,9 +59,9 @@ module _ is-trunc-map k (tot f) → ((x : A) → is-trunc-map k (f x)) is-trunc-map-is-trunc-map-tot is-trunc-tot-f x z = is-trunc-equiv k - ( fiber (tot f) (pair x z)) - ( inv-compute-fiber-tot f (pair x z)) - ( is-trunc-tot-f (pair x z)) + ( fiber (tot f) (x , z)) + ( inv-compute-fiber-tot f (x , z)) + ( is-trunc-tot-f (x , z)) module _ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3} @@ -137,7 +137,10 @@ module _ is-prop-map-map-Σ = is-trunc-map-map-Σ neg-one-𝕋 D ``` -### A family of squares over a pullback squares is a family of pullback squares if and only if the induced square of total spaces is a pullback square +### Pullbacks are preserved by dependent sums + +A family of squares over a pullback square is a family of pullback squares if +and only if the induced square of total spaces is a pullback square. ```agda module _ @@ -180,29 +183,26 @@ module _ f' (vertical-map-standard-pullback t)) ( g' (horizontal-map-standard-pullback t))) → Σ ( Σ A PA) - ( λ aa' → Σ (Σ B (λ b → Id (f (pr1 aa')) (g b))) + ( λ aa' → Σ (Σ B (λ b → f (pr1 aa') = g b)) ( λ bα → Σ (PB (pr1 bα)) - ( λ b' → Id - ( tr PX (pr2 bα) (f' (pr1 aa') (pr2 aa'))) - ( g' (pr1 bα) b')))) + ( λ b' → tr PX (pr2 bα) (f' (pr1 aa') (pr2 aa')) = g' (pr1 bα) b'))) map-standard-pullback-tot-cone-cone-fam-right-factor = map-interchange-Σ-Σ ( λ a bα a' → Σ (PB (pr1 bα)) - ( λ b' → Id (tr PX (pr2 bα) (f' a a')) (g' (pr1 bα) b'))) + ( λ b' → tr PX (pr2 bα) (f' a a') = g' (pr1 bα) b')) map-standard-pullback-tot-cone-cone-fam-left-factor : (aa' : Σ A PA) → - Σ (Σ B (λ b → Id (f (pr1 aa')) (g b))) - ( λ bα → Σ (PB (pr1 bα)) - ( λ b' → Id - ( tr PX (pr2 bα) (f' (pr1 aa') (pr2 aa'))) - ( g' (pr1 bα) b'))) → + Σ (Σ B (λ b → f (pr1 aa') = g b)) + ( λ bα → + Σ ( PB (pr1 bα)) + ( λ b' → tr PX (pr2 bα) (f' (pr1 aa') (pr2 aa')) = g' (pr1 bα) b')) → Σ ( Σ B PB) - ( λ bb' → Σ (Id (f (pr1 aa')) (g (pr1 bb'))) - ( λ α → Id (tr PX α (f' (pr1 aa') (pr2 aa'))) (g' (pr1 bb') (pr2 bb')))) + ( λ bb' → Σ (f (pr1 aa') = g (pr1 bb')) + ( λ α → tr PX α (f' (pr1 aa') (pr2 aa')) = g' (pr1 bb') (pr2 bb'))) map-standard-pullback-tot-cone-cone-fam-left-factor aa' = ( map-interchange-Σ-Σ - ( λ b α b' → Id (tr PX α (f' (pr1 aa') (pr2 aa'))) (g' b b'))) + ( λ b α b' → tr PX α (f' (pr1 aa') (pr2 aa')) = g' b b')) map-standard-pullback-tot-cone-cone-family : Σ ( standard-pullback f g) @@ -229,14 +229,18 @@ module _ ( map-standard-pullback-tot-cone-cone-fam-right-factor) ( is-equiv-map-interchange-Σ-Σ ( λ a bα a' → Σ (PB (pr1 bα)) - ( λ b' → Id (tr PX (pr2 bα) (f' a a')) (g' (pr1 bα) b')))) - ( is-equiv-tot-is-fiberwise-equiv (λ aa' → is-equiv-comp - ( tot (λ bb' → eq-pair-Σ')) - ( map-standard-pullback-tot-cone-cone-fam-left-factor aa') - ( is-equiv-map-interchange-Σ-Σ _) - ( is-equiv-tot-is-fiberwise-equiv (λ bb' → is-equiv-eq-pair-Σ - ( pair (f (pr1 aa')) (f' (pr1 aa') (pr2 aa'))) - ( pair (g (pr1 bb')) (g' (pr1 bb') (pr2 bb'))))))) + ( λ b' → tr PX (pr2 bα) (f' a a') = g' (pr1 bα) b'))) + ( is-equiv-tot-is-fiberwise-equiv + ( λ aa' → + is-equiv-comp + ( tot (λ bb' → eq-pair-Σ')) + ( map-standard-pullback-tot-cone-cone-fam-left-factor aa') + ( is-equiv-map-interchange-Σ-Σ _) + ( is-equiv-tot-is-fiberwise-equiv + ( λ bb' → + is-equiv-eq-pair-Σ + ( f (pr1 aa') , f' (pr1 aa') (pr2 aa')) + ( g (pr1 bb') , g' (pr1 bb') (pr2 bb')))))) triangle-standard-pullback-tot-cone-cone-family : ( gap (map-Σ PX f f') (map-Σ PX g g') tot-cone-cone-family) ~ @@ -244,11 +248,11 @@ module _ ( map-Σ _ ( gap f g c) ( λ x → gap - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) ( c' x)))) - triangle-standard-pullback-tot-cone-cone-family x = - refl + triangle-standard-pullback-tot-cone-cone-family = refl-htpy is-pullback-family-is-pullback-tot : is-pullback f g c → @@ -256,26 +260,31 @@ module _ (map-Σ PX f f') (map-Σ PX g g') tot-cone-cone-family → (x : C) → is-pullback - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) ( c' x) is-pullback-family-is-pullback-tot is-pb-c is-pb-tot = is-fiberwise-equiv-is-equiv-map-Σ _ ( gap f g c) - ( λ x → gap - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) - ( c' x)) + ( λ x → + gap + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) + ( c' x)) ( is-pb-c) ( is-equiv-top-map-triangle ( gap (map-Σ PX f f') (map-Σ PX g g') tot-cone-cone-family) ( map-standard-pullback-tot-cone-cone-family) ( map-Σ _ ( gap f g c) - ( λ x → gap - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) - ( c' x))) + ( λ x → + gap + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) + ( c' x))) ( triangle-standard-pullback-tot-cone-cone-family) ( is-equiv-map-standard-pullback-tot-cone-cone-family) ( is-pb-tot)) @@ -284,8 +293,9 @@ module _ is-pullback f g c → ( (x : C) → is-pullback - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) ( c' x)) → is-pullback (map-Σ PX f f') (map-Σ PX g g') tot-cone-cone-family @@ -296,13 +306,12 @@ module _ ( map-Σ _ ( gap f g c) ( λ x → gap - ( (tr PX (pr2 (pr2 c) x)) ∘ (f' (pr1 c x))) - ( g' (pr1 (pr2 c) x)) + ( ( tr PX (coherence-square-cone f g c x)) ∘ + ( f' (vertical-map-cone f g c x))) + ( g' (horizontal-map-cone f g c x)) ( c' x))) ( triangle-standard-pullback-tot-cone-cone-family) - ( is-equiv-map-Σ _ - ( is-pb-c) - ( is-pb-c')) + ( is-equiv-map-Σ _ is-pb-c is-pb-c') ( is-equiv-map-standard-pullback-tot-cone-cone-family) ``` @@ -418,8 +427,8 @@ module _ compute-inv-equiv-tot : (e : (x : A) → B x ≃ C x) → - ( map-inv-equiv (equiv-tot e)) ~ - ( map-equiv (equiv-tot (λ x → inv-equiv (e x)))) + map-inv-equiv (equiv-tot e) ~ + map-equiv (equiv-tot (λ x → inv-equiv (e x))) compute-inv-equiv-tot e (a , c) = is-injective-map-equiv ( equiv-tot e) diff --git a/src/foundation/functoriality-fibers-of-maps.lagda.md b/src/foundation/functoriality-fibers-of-maps.lagda.md index 176f65f124..1430ba9386 100644 --- a/src/foundation/functoriality-fibers-of-maps.lagda.md +++ b/src/foundation/functoriality-fibers-of-maps.lagda.md @@ -111,8 +111,9 @@ module _ (f : A → X) (g : B → X) (c : cone f g C) (a : A) where - map-fiber-cone : fiber (vertical-map-cone f g c) a → fiber g (f a) - map-fiber-cone = + map-fiber-vertical-map-cone : + fiber (vertical-map-cone f g c) a → fiber g (f a) + map-fiber-vertical-map-cone = map-domain-hom-arrow-fiber ( vertical-map-cone f g c) ( g) @@ -417,7 +418,7 @@ module _ pr2 (pr2 htpy-hom-arrow-fiber) = coh-htpy-hom-arrow-fiber ``` -### Computing `map-fiber-cone` of a horizontal pasting of cones +### Computing `map-fiber-vertical-map-cone` of a horizontal pasting of cones ```agda module _ @@ -426,12 +427,14 @@ module _ (i : X → Y) (j : Y → Z) (h : C → Z) where - preserves-pasting-horizontal-map-fiber-cone : + preserves-pasting-horizontal-map-fiber-vertical-map-cone : (c : cone j h B) (d : cone i (vertical-map-cone j h c) A) (x : X) → - ( map-fiber-cone (j ∘ i) h (pasting-horizontal-cone i j h c d) x) ~ - ( map-fiber-cone j h c (i x) ∘ - map-fiber-cone i (vertical-map-cone j h c) d x) - preserves-pasting-horizontal-map-fiber-cone c d = + ( map-fiber-vertical-map-cone (j ∘ i) h + ( pasting-horizontal-cone i j h c d) + ( x)) ~ + ( map-fiber-vertical-map-cone j h c (i x) ∘ + map-fiber-vertical-map-cone i (vertical-map-cone j h c) d x) + preserves-pasting-horizontal-map-fiber-vertical-map-cone c d = preserves-comp-map-fiber ( vertical-map-cone i (vertical-map-cone j h c) d) ( vertical-map-cone j h c) @@ -440,7 +443,7 @@ module _ ( hom-arrow-cone i (vertical-map-cone j h c) d) ``` -### Computing `map-fiber-cone` of a horizontal pasting of cones +### Computing `map-fiber-vertical-map-cone` of a horizontal pasting of cones Note: There should be a definition of vertical composition of morphisms of arrows, and a proof that `hom-arrow-fiber` preserves vertical composition. @@ -452,16 +455,19 @@ module _ (f : C → Z) (g : Y → Z) (h : X → Y) where - preserves-pasting-vertical-map-fiber-cone : - (c : cone f g B) (d : cone (pr1 (pr2 c)) h A) (x : C) → - ( ( map-fiber-cone f (g ∘ h) (pasting-vertical-cone f g h c d) x) ∘ + preserves-pasting-vertical-map-fiber-vertical-map-cone : + (c : cone f g B) (d : cone (horizontal-map-cone f g c) h A) (x : C) → + ( ( map-fiber-vertical-map-cone f (g ∘ h) + ( pasting-vertical-cone f g h c d) + ( x)) ∘ ( inv-map-compute-fiber-comp (pr1 c) (pr1 d) x)) ~ ( ( inv-map-compute-fiber-comp g h (f x)) ∘ ( map-Σ ( λ t → fiber h (pr1 t)) - ( map-fiber-cone f g c x) - ( λ t → map-fiber-cone (pr1 (pr2 c)) h d (pr1 t)))) - preserves-pasting-vertical-map-fiber-cone + ( map-fiber-vertical-map-cone f g c x) + ( λ t → + map-fiber-vertical-map-cone (horizontal-map-cone f g c) h d (pr1 t)))) + preserves-pasting-vertical-map-fiber-vertical-map-cone (p , q , H) (p' , q' , H') .(p (p' a)) ((.(p' a) , refl) , (a , refl)) = eq-pair-eq-pr2 @@ -473,7 +479,7 @@ module _ ( ap ( concat' (g (h (q' a))) ( pr2 - ( map-fiber-cone f g + ( map-fiber-vertical-map-cone f g ( p , q , H) ( p (p' a)) ( p' a , refl)))) diff --git a/src/foundation/functoriality-pullbacks.lagda.md b/src/foundation/functoriality-pullbacks.lagda.md index 737434ef51..c8c6533b7c 100644 --- a/src/foundation/functoriality-pullbacks.lagda.md +++ b/src/foundation/functoriality-pullbacks.lagda.md @@ -48,7 +48,7 @@ module _ (c : cone f g C) (c' : cone f' g' C') → is-pullback f g c → is-pullback f' g' c' → hom-cospan f' g' f g → C' → C - map-is-pullback c c' is-pb-c is-pb-c' h x = + map-is-pullback c c' is-pb-c _ h x = map-inv-is-equiv is-pb-c (map-standard-pullback h (gap f' g' c' x)) ``` diff --git a/src/foundation/impredicative-encodings.lagda.md b/src/foundation/impredicative-encodings.lagda.md index d599081c26..c7e62481a1 100644 --- a/src/foundation/impredicative-encodings.lagda.md +++ b/src/foundation/impredicative-encodings.lagda.md @@ -135,8 +135,7 @@ map-impredicative-disjunction-Prop : map-impredicative-disjunction-Prop {l1} {l2} P1 P2 = map-universal-property-trunc-Prop ( impredicative-disjunction-Prop P1 P2) - ( ind-coprod - ( λ x → type-impredicative-disjunction-Prop P1 P2) + ( rec-coprod ( λ x Q f1 f2 → f1 x) ( λ y Q f1 f2 → f2 y)) diff --git a/src/foundation/morphisms-arrows.lagda.md b/src/foundation/morphisms-arrows.lagda.md index 714c29475b..dc339db174 100644 --- a/src/foundation/morphisms-arrows.lagda.md +++ b/src/foundation/morphisms-arrows.lagda.md @@ -588,3 +588,5 @@ module _ ## See also - [Morphisms of twisted arrows](foundation.morphisms-twisted-arrows.md). +- [Fibered maps](foundation.fibered-maps.md) for the same concept under a + different name. diff --git a/src/foundation/perfect-images.lagda.md b/src/foundation/perfect-images.lagda.md index 6426bc1492..21b4e37a63 100644 --- a/src/foundation/perfect-images.lagda.md +++ b/src/foundation/perfect-images.lagda.md @@ -13,6 +13,7 @@ open import foundation.action-on-identifications-functions open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.double-negation +open import foundation.iterated-dependent-product-types open import foundation.iterating-functions open import foundation.law-of-excluded-middle open import foundation.negated-equality @@ -36,13 +37,17 @@ open import foundation-core.transport-along-identifications ## Idea -Consider two maps `f : A → B` and `g : B → A`. For `(g ◦ f ) ^ n (a₀) = a`, -consider also the following chain +Consider two maps `f : A → B` and `g : B → A`. For `(g ◦ f)ⁿ(a₀) = a`, consider +also the following chain -`a₀ --> f (a₀) --> g (f (a₀)) --> f (g (f (a₀))) --> ... --> (g ◦ f ) ^ n (a₀) = a` +```text + f g f g g + a₀ --> f (a₀) --> g(f(a₀)) --> f(g(f(a₀))) --> ... --> (g ◦ f)ⁿ(a₀) = a +``` -We say `a₀` is an origin for `a`, and `a` is `perfect image` for `g` if any -origin of `a` is in the image of `g`. +We say `a₀` is an {{#concept "origin"}} for `a`, and `a` is a +{{#concept "perfect image" Agda=is-perfect-image}} for `g` if any origin of `a` +is in the [image](foundation.images.md) of `g`. ## Definition @@ -58,9 +63,12 @@ module _ ## Properties -If `g` is an embedding, then `is-perfect-image a` is a proposition. In this -case, if we assume law of exluded middle, we can show `is-perfect-image a` is a -decidable type for any `a : A`. +If `g` is an [embedding](foundation-core.embeddings.md), then +`is-perfect-image a` is a [proposition](foundation-core.propositions.md). In +this case, if we assume the +[law of exluded middle](foundation.law-of-excluded-middle.md), we can show +`is-perfect-image a` is a [decidable type](foundation.decidable-types.md) for +any `a : A`. ```agda module _ @@ -71,8 +79,7 @@ module _ is-prop-is-perfect-image-is-emb : (a : A) → is-prop (is-perfect-image f g a) is-prop-is-perfect-image-is-emb a = - is-prop-Π (λ a₀ → (is-prop-Π λ n → - is-prop-Π (λ p → (is-prop-map-is-emb is-emb-g a₀)))) + is-prop-iterated-Π 3 (λ a₀ n p → is-prop-map-is-emb is-emb-g a₀) is-perfect-image-Prop : A → Prop (l1 ⊔ l2) pr1 (is-perfect-image-Prop a) = is-perfect-image f g a @@ -85,7 +92,7 @@ module _ ``` If `a` is a perfect image for `g`, then `a` has a preimage under `g`. Just take -n=zero in the definition. +`n = zero` in the definition. ```agda module _ @@ -99,8 +106,9 @@ module _ ``` One can define a map from `A` to `B` restricting the domain to the perfect -images of `g`. This gives a kind of section of g. When g is also an embedding, -the map gives a kind of retraction of g. +images of `g`. This gives a kind of [section](foundation-core.sections.md) of g. +When g is also an embedding, the map gives a kind of +[retraction](foundation-core.retractions.md) of g. ```agda module _ @@ -134,7 +142,7 @@ module _ (is-section-inverse-of-perfect-image (g b) ρ) ``` -If `g (f (a))` is a perfect image for `g`, so is `a`. +If `g(f(a))` is a perfect image for `g`, so is `a`. ```agda module _ @@ -185,8 +193,8 @@ module _ Σ A (λ a₀ → (Σ ℕ (λ n → ((iterate n (g ∘ f)) a₀ = a) × ¬ (fiber g a₀)))) ``` -If we assume law of excluded middle and `g` is embedding, we can prove that if -`is-not-perfect-image a` does not hold, we have `is-perfect-image a`. +If we assume the law of excluded middle and `g` is an embedding, we can prove +that if `is-not-perfect-image a` does not hold, we have `is-perfect-image a`. ```agda module _ @@ -195,15 +203,13 @@ module _ (lem : LEM (l1 ⊔ l2)) where - not-not-perfect-is-perfect : - (a : A) → - ¬ (is-not-perfect-image a) → - (is-perfect-image f g a) - not-not-perfect-is-perfect a nρ a₀ n p = - ind-coprod _ - (id) - (λ a₁ → ex-falso (nρ (pair a₀ (pair n (pair p a₁))))) - (lem (pair (fiber g a₀) (is-prop-map-is-emb is-emb-g a₀))) + is-perfect-not-not-is-perfect-image : + (a : A) → ¬ (is-not-perfect-image a) → is-perfect-image f g a + is-perfect-not-not-is-perfect-image a nρ a₀ n p = + rec-coprod + ( id) + ( λ a₁ → ex-falso (nρ (a₀ , n , p , a₁))) + ( lem (fiber g a₀ , is-prop-map-is-emb is-emb-g a₀)) ``` The following property states that if `g (b)` is not a perfect image, then `b` @@ -225,21 +231,21 @@ module _ not-perfect-image-has-not-perfect-fiber b nρ = v where i : ¬¬ (is-not-perfect-image {f = f} (g b)) - i = λ nμ → nρ (not-not-perfect-is-perfect is-emb-g lem (g b) nμ) + i = λ nμ → nρ (is-perfect-not-not-is-perfect-image is-emb-g lem (g b) nμ) ii : is-not-perfect-image (g b) → Σ (fiber f b) (λ s → ¬ (is-perfect-image f g (pr1 s))) - ii (pair x₀ (pair zero-ℕ u)) = - ex-falso (pr2 u (pair b (inv (pr1 u)))) - ii (pair x₀ (pair (succ-ℕ n) u)) = - pair a w + ii (x₀ , 0 , u) = + ex-falso (pr2 u (b , inv (pr1 u))) + ii (x₀ , succ-ℕ n , u) = + a , w where - q : f ((iterate n (g ∘ f)) x₀) = b + q : f (iterate n (g ∘ f) x₀) = b q = is-injective-is-emb is-emb-g (pr1 u) a : fiber f b - a = pair ((iterate n (g ∘ f)) x₀) q + a = iterate n (g ∘ f) x₀ , q w : ¬ (is-perfect-image f g ((iterate n (g ∘ f)) x₀)) w = λ s → pr2 u (s x₀ n refl) @@ -254,5 +260,5 @@ module _ (λ s → is-prop-neg {A = is-perfect-image f g (pr1 s)}) v : Σ (fiber f b) (λ s → ¬ (is-perfect-image f g (pr1 s))) - v = double-negation-elim-is-decidable (lem (pair _ iv)) iii + v = double-negation-elim-is-decidable (lem (_ , iv)) iii ``` diff --git a/src/foundation/postcomposition-functions.lagda.md b/src/foundation/postcomposition-functions.lagda.md index 6df3fc32b6..9ca67261ed 100644 --- a/src/foundation/postcomposition-functions.lagda.md +++ b/src/foundation/postcomposition-functions.lagda.md @@ -13,6 +13,8 @@ open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types open import foundation.universe-levels +open import foundation-core.commuting-squares-of-maps +open import foundation-core.commuting-triangles-of-maps open import foundation-core.contractible-maps open import foundation-core.contractible-types open import foundation-core.equivalences @@ -48,7 +50,7 @@ is defined by `λ h x → f (h x)`. htpy-postcomp : {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} (A : UU l3) → {f g : X → Y} → (f ~ g) → postcomp A f ~ postcomp A g -htpy-postcomp A H h = eq-htpy (H ∘ h) +htpy-postcomp A H h = eq-htpy (H ·r h) compute-htpy-postcomp-refl-htpy : {l1 l2 l3 : Level} (A : UU l1) {B : UU l2} {C : UU l3} (f : B → C) → @@ -56,15 +58,54 @@ compute-htpy-postcomp-refl-htpy : compute-htpy-postcomp-refl-htpy A f h = eq-htpy-refl-htpy (f ∘ h) ``` -### The fibers of `postcomp` +### Computations of the fibers of `postcomp` + +We give three computations of the fibers of a postcomposition function: + +1. `fiber (postcomp A f) h ≃ ((x : A) → fiber f (h x))` +2. `fiber (postcomp A f) h ≃ Σ (A → X) (coherence-triangle-maps h f)`, and +3. `fiber (postcomp A f) (f ∘ h) ≃ Σ (A → X) (λ g → coherence-square-maps g h f f)` ```agda -compute-fiber-postcomp : - {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} (A : UU l3) → - (f : X → Y) (h : A → Y) → - ((x : A) → fiber f (h x)) ≃ fiber (postcomp A f) h -compute-fiber-postcomp A f h = - equiv-tot (λ _ → equiv-eq-htpy) ∘e distributive-Π-Σ +module _ + {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} (A : UU l3) + (f : X → Y) + where + + inv-compute-Π-fiber-postcomp : + (h : A → Y) → fiber (postcomp A f) h ≃ ((x : A) → fiber f (h x)) + inv-compute-Π-fiber-postcomp h = + inv-distributive-Π-Σ ∘e equiv-tot (λ _ → equiv-funext) + + compute-Π-fiber-postcomp : + (h : A → Y) → ((x : A) → fiber f (h x)) ≃ fiber (postcomp A f) h + compute-Π-fiber-postcomp h = + equiv-tot (λ _ → equiv-eq-htpy) ∘e distributive-Π-Σ + + inv-compute-coherence-triangle-fiber-postcomp : + (h : A → Y) → + fiber (postcomp A f) h ≃ Σ (A → X) (coherence-triangle-maps h f) + inv-compute-coherence-triangle-fiber-postcomp h = + equiv-tot (λ _ → equiv-funext) ∘e equiv-fiber (postcomp A f) h + + compute-coherence-triangle-fiber-postcomp : + (h : A → Y) → + Σ (A → X) (coherence-triangle-maps h f) ≃ fiber (postcomp A f) h + compute-coherence-triangle-fiber-postcomp h = + inv-equiv (inv-compute-coherence-triangle-fiber-postcomp h) + + inv-compute-fiber-postcomp : + (h : A → X) → + fiber (postcomp A f) (f ∘ h) ≃ + Σ (A → X) (λ g → coherence-square-maps g h f f) + inv-compute-fiber-postcomp h = + inv-compute-coherence-triangle-fiber-postcomp (f ∘ h) + + compute-fiber-postcomp : + (h : A → X) → + Σ (A → X) (λ g → coherence-square-maps g h f f) ≃ + fiber (postcomp A f) (f ∘ h) + compute-fiber-postcomp h = compute-coherence-triangle-fiber-postcomp (f ∘ h) ``` ### Postcomposition and equivalences @@ -94,10 +135,9 @@ module _ ( eq-is-contr ( is-contr-map-is-equiv (H X) f) { x = - pair - ( map-inv-is-equiv-is-equiv-postcomp ∘ f) - ( ap (λ u → u ∘ f) (is-section-map-inv-is-equiv (H Y) id))} - { y = pair id refl})) + ( map-inv-is-equiv-is-equiv-postcomp ∘ f) , + ( ap (_∘ f) (is-section-map-inv-is-equiv (H Y) id))} + { y = id , refl})) abstract is-equiv-is-equiv-postcomp : is-equiv f @@ -116,8 +156,7 @@ simplified to that universe. is-equiv-is-equiv-postcomp' : {l : Level} {X : UU l} {Y : UU l} (f : X → Y) → ((A : UU l) → is-equiv (postcomp A f)) → is-equiv f -is-equiv-is-equiv-postcomp' - {l} {X} {Y} f is-equiv-postcomp-f = +is-equiv-is-equiv-postcomp' {l} {X} {Y} f is-equiv-postcomp-f = let section-f = center (is-contr-map-is-equiv (is-equiv-postcomp-f Y) id) in is-equiv-is-invertible @@ -128,8 +167,8 @@ is-equiv-is-equiv-postcomp' ( pr1) ( eq-is-contr' ( is-contr-map-is-equiv (is-equiv-postcomp-f X) f) - ( pair ((pr1 section-f) ∘ f) (ap (λ t → t ∘ f) (pr2 section-f))) - ( pair id refl)))) + ( pr1 section-f ∘ f , ap (_∘ f) (pr2 section-f)) + ( id , refl)))) abstract is-equiv-postcomp-is-equiv : @@ -165,6 +204,5 @@ module _ equiv-htpy-postcomp-htpy : (e : B ≃ C) (f g : A → B) → (f ~ g) ≃ (map-equiv e ∘ f ~ map-equiv e ∘ g) equiv-htpy-postcomp-htpy e f g = - equiv-Π-equiv-family - ( λ a → equiv-ap e (f a) (g a)) + equiv-Π-equiv-family (λ a → equiv-ap e (f a) (g a)) ``` diff --git a/src/foundation/precomposition-functions.lagda.md b/src/foundation/precomposition-functions.lagda.md index 92d4af7361..b3e28081d2 100644 --- a/src/foundation/precomposition-functions.lagda.md +++ b/src/foundation/precomposition-functions.lagda.md @@ -16,6 +16,7 @@ open import foundation.sections open import foundation.universe-levels open import foundation-core.commuting-squares-of-maps +open import foundation-core.commuting-triangles-of-maps open import foundation-core.equivalences open import foundation-core.fibers-of-maps open import foundation-core.function-extensionality @@ -104,25 +105,33 @@ module _ equiv-htpy-precomp-htpy-Π f g e ``` -### The fibers of `precomp` +### Computations of the fibers of `precomp` ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (X : UU l3) where + compute-coherence-triangle-fiber-precomp' : + (g : A → X) → + fiber (precomp f X) g ≃ Σ (B → X) (λ h → coherence-triangle-maps' g h f) + compute-coherence-triangle-fiber-precomp' g = equiv-tot (λ _ → equiv-funext) + + compute-coherence-triangle-fiber-precomp : + (g : A → X) → + fiber (precomp f X) g ≃ Σ (B → X) (λ h → coherence-triangle-maps g h f) + compute-coherence-triangle-fiber-precomp g = + equiv-tot (λ _ → equiv-funext) ∘e equiv-fiber (precomp f X) g + compute-fiber-precomp : (g : B → X) → fiber (precomp f X) (g ∘ f) ≃ Σ (B → X) (λ h → coherence-square-maps f f h g) - compute-fiber-precomp g = - equiv-tot ( λ h → equiv-funext) ∘e - equiv-fiber (precomp f X) (g ∘ f) + compute-fiber-precomp g = compute-coherence-triangle-fiber-precomp (g ∘ f) compute-total-fiber-precomp : Σ (B → X) (λ g → fiber (precomp f X) (g ∘ f)) ≃ cocone f f X - compute-total-fiber-precomp = - equiv-tot compute-fiber-precomp + compute-total-fiber-precomp = equiv-tot compute-fiber-precomp diagonal-into-fibers-precomp : (B → X) → Σ (B → X) (λ g → fiber (precomp f X) (g ∘ f)) diff --git a/src/foundation/propositional-truncations.lagda.md b/src/foundation/propositional-truncations.lagda.md index 8e0c239e69..05fef93d64 100644 --- a/src/foundation/propositional-truncations.lagda.md +++ b/src/foundation/propositional-truncations.lagda.md @@ -374,8 +374,7 @@ module _ map-trunc-Prop-diagonal-coprod = map-universal-property-trunc-Prop ( trunc-Prop A) - ( unit-trunc ∘ - ind-coprod (λ _ → A) id id) + ( unit-trunc ∘ rec-coprod id id) map-inv-trunc-Prop-diagonal-coprod : type-trunc-Prop A → type-trunc-Prop (A + A) diff --git a/src/foundation/pullback-squares.lagda.md b/src/foundation/pullback-squares.lagda.md index 6ded80adcd..47aa0b3886 100644 --- a/src/foundation/pullback-squares.lagda.md +++ b/src/foundation/pullback-squares.lagda.md @@ -9,10 +9,10 @@ module foundation.pullback-squares where ```agda open import foundation.cones-over-cospans open import foundation.dependent-pair-types -open import foundation.pullbacks open import foundation.universe-levels open import foundation-core.commuting-squares-of-maps +open import foundation-core.pullbacks open import foundation-core.universal-property-pullbacks ``` @@ -48,13 +48,7 @@ module _ where pullback-square : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) - pullback-square = - Σ ( A → C) - ( λ f → - Σ ( B → C) - ( λ g → - Σ ( cone f g X) - ( is-pullback f g))) + pullback-square = Σ (A → C) (λ f → Σ (B → C) (λ g → pullback-cone f g X)) ``` ### Components of a pullback cone diff --git a/src/foundation/pullbacks.lagda.md b/src/foundation/pullbacks.lagda.md index ffa6e02590..a4444e0ac5 100644 --- a/src/foundation/pullbacks.lagda.md +++ b/src/foundation/pullbacks.lagda.md @@ -12,16 +12,19 @@ open import foundation-core.pullbacks public open import foundation.action-on-identifications-functions open import foundation.commuting-cubes-of-maps open import foundation.cones-over-cospans +open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences open import foundation.descent-equivalences +open import foundation.equality-coproduct-types open import foundation.equivalences +open import foundation.functoriality-coproduct-types open import foundation.functoriality-dependent-pair-types open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.homotopy-induction open import foundation.identity-types -open import foundation.postcomposition-functions +open import foundation.multivariable-homotopies open import foundation.unit-type open import foundation.universe-levels @@ -33,7 +36,10 @@ open import foundation-core.equality-dependent-pair-types open import foundation-core.function-extensionality open import foundation-core.function-types open import foundation-core.functoriality-dependent-function-types +open import foundation-core.postcomposition-functions open import foundation-core.propositions +open import foundation-core.retractions +open import foundation-core.sections open import foundation-core.torsorial-type-families open import foundation-core.transport-along-identifications open import foundation-core.whiskering-homotopies @@ -41,6 +47,40 @@ open import foundation-core.whiskering-homotopies
+## Idea + +Given a [cospan of types](foundation.cospans.md) + +```text + f : A → X ← B : g, +``` + +we can form the +{{#concept "standard pullback" Disambiguation="types" Agda=standard-pullback}} +`A ×_X B` satisfying +[the universal property of the pullback](foundation-core.universal-property-pullbacks.md) +of the cospan, completing the diagram + +```text + A ×_X B ------> B + | ⌟ | + | | g + | | + v v + A ---------> X. + f +``` + +The standard pullback consists of [pairs](foundation.dependent-pair-types.md) +`a : A` and `b : B` such that `f a = g b` agree + +```text + A ×_X B := Σ (a : A) (b : B), (f a = g b), +``` + +thus the standard [cone](foundation.cones-over-cospans.md) consists of the +canonical projections. + ## Properties ### Being a pullback is a property @@ -51,173 +91,191 @@ module _ (f : A → X) (g : B → X) where - is-property-is-pullback : (c : cone f g C) → is-prop (is-pullback f g c) - is-property-is-pullback c = is-property-is-equiv (gap f g c) + is-prop-is-pullback : (c : cone f g C) → is-prop (is-pullback f g c) + is-prop-is-pullback c = is-property-is-equiv (gap f g c) is-pullback-Prop : (c : cone f g C) → Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) pr1 (is-pullback-Prop c) = is-pullback f g c - pr2 (is-pullback-Prop c) = is-property-is-pullback c + pr2 (is-pullback-Prop c) = is-prop-is-pullback c +``` + +### Pullbacks are closed under exponentiation + +Given a pullback square + +```text + f' + C ---------> B + | ⌟ | + g'| | g + | | + v v + A ---------> X + f ``` -### Exponents of pullbacks are pullbacks +then the exponentiated square given by postcomposition + +```text + f' ∘ - + (S → C) ---------> (S → B) + | | + g' ∘ - | | g ∘ - + | | + v v + (S → A) ---------> (S → X) + f ∘ - +``` + +is a pullback square for any type `S`. ```agda -exponent-cone : +postcomp-cone : {l1 l2 l3 l4 l5 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {X : UU l4} (T : UU l5) (f : A → X) (g : B → X) (c : cone f g C) → cone (postcomp T f) (postcomp T g) (T → C) -pr1 (exponent-cone T f g c) h = vertical-map-cone f g c ∘ h -pr1 (pr2 (exponent-cone T f g c)) h = horizontal-map-cone f g c ∘ h -pr2 (pr2 (exponent-cone T f g c)) h = eq-htpy (coherence-square-cone f g c ·r h) +pr1 (postcomp-cone T f g c) h = vertical-map-cone f g c ∘ h +pr1 (pr2 (postcomp-cone T f g c)) h = horizontal-map-cone f g c ∘ h +pr2 (pr2 (postcomp-cone T f g c)) h = eq-htpy (coherence-square-cone f g c ·r h) -map-standard-pullback-exponent : +map-standard-pullback-postcomp : {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} (f : A → X) (g : B → X) - (T : UU l4) → - standard-pullback (postcomp T f) (postcomp T g) → - cone f g T -map-standard-pullback-exponent f g T = - tot (λ p → tot (λ q → htpy-eq)) + (T : UU l4) → standard-pullback (postcomp T f) (postcomp T g) → cone f g T +map-standard-pullback-postcomp f g T = tot (λ _ → tot (λ _ → htpy-eq)) abstract - is-equiv-map-standard-pullback-exponent : + is-equiv-map-standard-pullback-postcomp : {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} (f : A → X) (g : B → X) - (T : UU l4) → is-equiv (map-standard-pullback-exponent f g T) - is-equiv-map-standard-pullback-exponent f g T = + (T : UU l4) → is-equiv (map-standard-pullback-postcomp f g T) + is-equiv-map-standard-pullback-postcomp f g T = is-equiv-tot-is-fiberwise-equiv - ( λ p → is-equiv-tot-is-fiberwise-equiv - ( λ q → funext (f ∘ p) (g ∘ q))) + ( λ p → is-equiv-tot-is-fiberwise-equiv (λ q → funext (f ∘ p) (g ∘ q))) -triangle-map-standard-pullback-exponent : +triangle-map-standard-pullback-postcomp : {l1 l2 l3 l4 l5 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} (T : UU l5) (f : A → X) (g : B → X) (c : cone f g C) → - ( cone-map f g c {T}) ~ - ( ( map-standard-pullback-exponent f g T) ∘ - ( gap - ( postcomp T f) - ( postcomp T g) - ( exponent-cone T f g c))) -triangle-map-standard-pullback-exponent - {A = A} {B} T f g c h = + cone-map f g c {T} ~ + map-standard-pullback-postcomp f g T ∘ + gap (postcomp T f) (postcomp T g) (postcomp-cone T f g c) +triangle-map-standard-pullback-postcomp T f g c h = eq-pair-eq-pr2 ( eq-pair-eq-pr2 ( inv (is-section-eq-htpy (coherence-square-cone f g c ·r h)))) abstract - is-pullback-exponent-is-pullback : + is-pullback-postcomp-is-pullback : {l1 l2 l3 l4 l5 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} (f : A → X) (g : B → X) (c : cone f g C) → is-pullback f g c → (T : UU l5) → - is-pullback - ( postcomp T f) - ( postcomp T g) - ( exponent-cone T f g c) - is-pullback-exponent-is-pullback f g c is-pb-c T = + is-pullback (postcomp T f) (postcomp T g) (postcomp-cone T f g c) + is-pullback-postcomp-is-pullback f g c is-pb-c T = is-equiv-top-map-triangle ( cone-map f g c) - ( map-standard-pullback-exponent f g T) - ( gap (f ∘_) (g ∘_) (exponent-cone T f g c)) - ( triangle-map-standard-pullback-exponent T f g c) - ( is-equiv-map-standard-pullback-exponent f g T) + ( map-standard-pullback-postcomp f g T) + ( gap (f ∘_) (g ∘_) (postcomp-cone T f g c)) + ( triangle-map-standard-pullback-postcomp T f g c) + ( is-equiv-map-standard-pullback-postcomp f g T) ( universal-property-pullback-is-pullback f g c is-pb-c T) abstract - is-pullback-is-pullback-exponent : + is-pullback-is-pullback-postcomp : {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} (f : A → X) (g : B → X) (c : cone f g C) → - ((l5 : Level) (T : UU l5) → is-pullback - ( postcomp T f) - ( postcomp T g) - ( exponent-cone T f g c)) → + ( {l5 : Level} (T : UU l5) → + is-pullback (postcomp T f) (postcomp T g) (postcomp-cone T f g c)) → is-pullback f g c - is-pullback-is-pullback-exponent f g c is-pb-exp = + is-pullback-is-pullback-postcomp f g c is-pb-postcomp = is-pullback-universal-property-pullback f g c - ( λ T → is-equiv-left-map-triangle - ( cone-map f g c) - ( map-standard-pullback-exponent f g T) - ( gap (f ∘_) (g ∘_) (exponent-cone T f g c)) - ( triangle-map-standard-pullback-exponent T f g c) - ( is-pb-exp _ T) - ( is-equiv-map-standard-pullback-exponent f g T)) + ( λ T → + is-equiv-left-map-triangle + ( cone-map f g c) + ( map-standard-pullback-postcomp f g T) + ( gap (f ∘_) (g ∘_) (postcomp-cone T f g c)) + ( triangle-map-standard-pullback-postcomp T f g c) + ( is-pb-postcomp T) + ( is-equiv-map-standard-pullback-postcomp f g T)) ``` ### Identity types can be presented as pullbacks +Identity types fit into pullback squares + +```text + (x = y) ----> 1 + | ⌟ | + | | y + v v + 1 --------> A. + x +``` + ```agda -cone-Id : - {l : Level} {A : UU l} (x y : A) → - cone (point x) (point y) (x = y) -pr1 (cone-Id x y) = terminal-map (x = y) -pr1 (pr2 (cone-Id x y)) = terminal-map (x = y) -pr2 (pr2 (cone-Id x y)) = id - -inv-gap-cone-Id : - {l : Level} {A : UU l} (x y : A) → - standard-pullback (point x) (point y) → x = y -inv-gap-cone-Id x y (star , star , p) = p +module _ + {l : Level} {A : UU l} (x y : A) + where -abstract - is-section-inv-gap-cone-Id : - {l : Level} {A : UU l} (x y : A) → - ( gap (point x) (point y) (cone-Id x y) ∘ - inv-gap-cone-Id x y) ~ - id - is-section-inv-gap-cone-Id x y (star , star , p) = refl + cone-Id : cone (point x) (point y) (x = y) + pr1 cone-Id = terminal-map (x = y) + pr1 (pr2 cone-Id) = terminal-map (x = y) + pr2 (pr2 cone-Id) = id -abstract - is-retraction-inv-gap-cone-Id : - {l : Level} {A : UU l} (x y : A) → - ( ( inv-gap-cone-Id x y) ∘ - ( gap (point x) (point y) (cone-Id x y))) ~ id - is-retraction-inv-gap-cone-Id x y p = refl + inv-gap-cone-Id : standard-pullback (point x) (point y) → x = y + inv-gap-cone-Id (star , star , p) = p -abstract - is-pullback-cone-Id : - {l : Level} {A : UU l} (x y : A) → - is-pullback (point x) (point y) (cone-Id x y) - is-pullback-cone-Id x y = - is-equiv-is-invertible - ( inv-gap-cone-Id x y) - ( is-section-inv-gap-cone-Id x y) - ( is-retraction-inv-gap-cone-Id x y) - -cone-Id' : - {l : Level} {A : UU l} (t : A × A) → - cone (point t) (diagonal A) (pr1 t = pr2 t) -pr1 (cone-Id' {A = A} (x , y)) = terminal-map (x = y) -pr1 (pr2 (cone-Id' {A = A} (x , y))) = const (x = y) A x -pr2 (pr2 (cone-Id' {A = A} (x , y))) p = eq-pair-eq-pr2 (inv p) - -inv-gap-cone-Id' : - {l : Level} {A : UU l} (t : A × A) → - standard-pullback (point t) (diagonal A) → (pr1 t = pr2 t) -inv-gap-cone-Id' t (star , z , p) = ap pr1 p ∙ inv (ap pr2 p) + abstract + is-section-inv-gap-cone-Id : + is-section (gap (point x) (point y) cone-Id) (inv-gap-cone-Id) + is-section-inv-gap-cone-Id (star , star , p) = refl -abstract - is-section-inv-gap-cone-Id' : - {l : Level} {A : UU l} (t : A × A) → - ( ( gap (point t) (diagonal A) (cone-Id' t)) ∘ - ( inv-gap-cone-Id' t)) ~ id - is-section-inv-gap-cone-Id' .(z , z) (star , z , refl) = refl + abstract + is-retraction-inv-gap-cone-Id : + is-retraction (gap (point x) (point y) cone-Id) (inv-gap-cone-Id) + is-retraction-inv-gap-cone-Id p = refl -abstract - is-retraction-inv-gap-cone-Id' : - {l : Level} {A : UU l} (t : A × A) → - ( ( inv-gap-cone-Id' t) ∘ - ( gap (point t) (diagonal A) (cone-Id' t))) ~ id - is-retraction-inv-gap-cone-Id' (x , .x) refl = refl + abstract + is-pullback-cone-Id : is-pullback (point x) (point y) cone-Id + is-pullback-cone-Id = + is-equiv-is-invertible + ( inv-gap-cone-Id) + ( is-section-inv-gap-cone-Id) + ( is-retraction-inv-gap-cone-Id) -abstract - is-pullback-cone-Id' : - {l : Level} {A : UU l} (t : A × A) → - is-pullback (point t) (diagonal A) (cone-Id' t) - is-pullback-cone-Id' t = - is-equiv-is-invertible - ( inv-gap-cone-Id' t) - ( is-section-inv-gap-cone-Id' t) - ( is-retraction-inv-gap-cone-Id' t) +module _ + {l : Level} {A : UU l} ((x , y) : A × A) + where + + cone-Id' : cone (point (x , y)) (diagonal A) (x = y) + pr1 cone-Id' = terminal-map (x = y) + pr1 (pr2 cone-Id') = const (x = y) A x + pr2 (pr2 cone-Id') p = eq-pair-eq-pr2 (inv p) + + inv-gap-cone-Id' : + standard-pullback (point (x , y)) (diagonal A) → x = y + inv-gap-cone-Id' (star , z , p) = ap pr1 p ∙ inv (ap pr2 p) + + abstract + is-section-inv-gap-cone-Id' : + gap (point (x , y)) (diagonal A) cone-Id' ∘ + inv-gap-cone-Id' ~ id + is-section-inv-gap-cone-Id' (star , z , refl) = refl + + abstract + is-retraction-inv-gap-cone-Id' : + inv-gap-cone-Id' ∘ + gap (point (x , y)) (diagonal A) cone-Id' ~ id + is-retraction-inv-gap-cone-Id' refl = refl + + abstract + is-pullback-cone-Id' : + is-pullback (point (x , y)) (diagonal A) cone-Id' + is-pullback-cone-Id' = + is-equiv-is-invertible + ( inv-gap-cone-Id') + ( is-section-inv-gap-cone-Id') + ( is-retraction-inv-gap-cone-Id') ``` ### The equivalence on canonical pullbacks induced by parallel cospans @@ -231,20 +289,21 @@ module _ map-equiv-standard-pullback-htpy : standard-pullback f' g' → standard-pullback f g map-equiv-standard-pullback-htpy = - tot (λ a → tot (λ b → - ( concat' (f a) (inv (Hg b))) ∘ (concat (Hf a) (g' b)))) + tot (λ a → tot (λ b → concat' (f a) (inv (Hg b)) ∘ concat (Hf a) (g' b))) abstract is-equiv-map-equiv-standard-pullback-htpy : is-equiv map-equiv-standard-pullback-htpy is-equiv-map-equiv-standard-pullback-htpy = - is-equiv-tot-is-fiberwise-equiv (λ a → - is-equiv-tot-is-fiberwise-equiv (λ b → - is-equiv-comp - ( concat' (f a) (inv (Hg b))) - ( concat (Hf a) (g' b)) - ( is-equiv-concat (Hf a) (g' b)) - ( is-equiv-concat' (f a) (inv (Hg b))))) + is-equiv-tot-is-fiberwise-equiv + ( λ a → + is-equiv-tot-is-fiberwise-equiv + ( λ b → + is-equiv-comp + ( concat' (f a) (inv (Hg b))) + ( concat (Hf a) (g' b)) + ( is-equiv-concat (Hf a) (g' b)) + ( is-equiv-concat' (f a) (inv (Hg b))))) equiv-standard-pullback-htpy : standard-pullback f' g' ≃ standard-pullback f g @@ -256,497 +315,521 @@ module _ ```agda module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} {f f' : A → X} (Hf : f ~ f') {g g' : B → X} (Hg : g ~ g') where triangle-is-pullback-htpy : - {l4 : Level} {C : UU l4} {c : cone f g C} {c' : cone f' g' C} (Hc : htpy-parallel-cone Hf Hg c c') → - (gap f g c) ~ (map-equiv-standard-pullback-htpy Hf Hg ∘ (gap f' g' c')) - triangle-is-pullback-htpy {c = p , q , H} {p' , q' , H'} (Hp , Hq , HH) z = + gap f g c ~ map-equiv-standard-pullback-htpy Hf Hg ∘ gap f' g' c' + triangle-is-pullback-htpy {p , q , H} {p' , q' , H'} (Hp , Hq , HH) z = map-extensionality-standard-pullback f g ( Hp z) ( Hq z) - ( ( inv - ( assoc - ( ap f (Hp z)) - ( (Hf (p' z)) ∙ (H' z)) - ( inv (Hg (q' z))))) ∙ + ( ( inv (assoc (ap f (Hp z)) (Hf (p' z) ∙ H' z) (inv (Hg (q' z))))) ∙ ( inv ( right-transpose-eq-concat - ( (H z) ∙ (ap g (Hq z))) + ( H z ∙ ap g (Hq z)) ( Hg (q' z)) - ( ( ap f (Hp z)) ∙ ((Hf (p' z)) ∙ (H' z))) + ( ap f (Hp z) ∙ (Hf (p' z) ∙ H' z)) ( ( assoc (H z) (ap g (Hq z)) (Hg (q' z))) ∙ - ( ( HH z) ∙ - ( assoc (ap f (Hp z)) (Hf (p' z)) (H' z))))))) + ( HH z) ∙ + ( assoc (ap f (Hp z)) (Hf (p' z)) (H' z)))))) abstract is-pullback-htpy : - {l4 : Level} {C : UU l4} {c : cone f g C} (c' : cone f' g' C) + {c : cone f g C} (c' : cone f' g' C) (Hc : htpy-parallel-cone Hf Hg c c') → is-pullback f' g' c' → is-pullback f g c - is-pullback-htpy {c = p , q , H} (p' , q' , H') (Hp , Hq , HH) is-pb-c' = + is-pullback-htpy {c} c' H is-pb-c' = is-equiv-left-map-triangle - ( gap f g (p , q , H)) + ( gap f g c) ( map-equiv-standard-pullback-htpy Hf Hg) - ( gap f' g' (p' , q' , H')) - ( triangle-is-pullback-htpy - {c = p , q , H} {p' , q' , H'} (Hp , Hq , HH)) + ( gap f' g' c') + ( triangle-is-pullback-htpy H) ( is-pb-c') ( is-equiv-map-equiv-standard-pullback-htpy Hf Hg) abstract is-pullback-htpy' : - {l4 : Level} {C : UU l4} (c : cone f g C) {c' : cone f' g' C} + (c : cone f g C) {c' : cone f' g' C} (Hc : htpy-parallel-cone Hf Hg c c') → is-pullback f g c → is-pullback f' g' c' - is-pullback-htpy' (p , q , H) {p' , q' , H'} (Hp , Hq , HH) is-pb-c = + is-pullback-htpy' c {c'} H = is-equiv-top-map-triangle - ( gap f g (p , q , H)) + ( gap f g c) ( map-equiv-standard-pullback-htpy Hf Hg) - ( gap f' g' (p' , q' , H')) - ( triangle-is-pullback-htpy - {c = p , q , H} {p' , q' , H'} (Hp , Hq , HH)) + ( gap f' g' c') + ( triangle-is-pullback-htpy H) ( is-equiv-map-equiv-standard-pullback-htpy Hf Hg) - ( is-pb-c) ``` -In the following part we will relate the type htpy-parallel-cone to the identity -type of cones. Here we will rely on function extensionality. +In the following part we will relate the type `htpy-parallel-cone` to the +identity type of cones. Here we will rely on +[function extensionality](foundation-core.function-extensionality.md). ```agda -refl-htpy-parallel-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → - htpy-parallel-cone (refl-htpy {f = f}) (refl-htpy {f = g}) c c -pr1 (refl-htpy-parallel-cone f g c) = refl-htpy -pr1 (pr2 (refl-htpy-parallel-cone f g c)) = refl-htpy -pr2 (pr2 (refl-htpy-parallel-cone f g c)) = right-unit-htpy - -htpy-eq-square : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c c' : cone f g C) → - c = c' → htpy-parallel-cone (refl-htpy {f = f}) (refl-htpy {f = g}) c c' -htpy-eq-square f g c .c refl = refl-htpy-parallel-cone f g c +module _ + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} + {f : A → X} {g : B → X} + where -htpy-parallel-cone-refl-htpy-htpy-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) → - (c c' : cone f g C) → - htpy-cone f g c c' → - htpy-parallel-cone (refl-htpy {f = f}) (refl-htpy {f = g}) c c' -htpy-parallel-cone-refl-htpy-htpy-cone f g (p , q , H) (p' , q' , H') = - tot - ( λ K → tot - ( λ L M → - ( ap-concat-htpy H right-unit-htpy) ∙h - ( M ∙h ap-concat-htpy' H' inv-htpy-right-unit-htpy))) + refl-htpy-parallel-cone : + (c : cone f g C) → + htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c + pr1 (refl-htpy-parallel-cone c) = refl-htpy + pr1 (pr2 (refl-htpy-parallel-cone c)) = refl-htpy + pr2 (pr2 (refl-htpy-parallel-cone c)) = right-unit-htpy -abstract - is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) → + htpy-eq-square : (c c' : cone f g C) → - is-equiv (htpy-parallel-cone-refl-htpy-htpy-cone f g c c') - is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone f g - (p , q , H) (p' , q' , H') = - is-equiv-tot-is-fiberwise-equiv - ( λ K → is-equiv-tot-is-fiberwise-equiv - ( λ L → is-equiv-comp - ( concat-htpy - ( ap-concat-htpy H right-unit-htpy) - ( (f ·l K) ∙h refl-htpy ∙h H')) - ( concat-htpy' - ( H ∙h (g ·l L)) - ( ap-concat-htpy' H' inv-htpy-right-unit-htpy)) - ( is-equiv-concat-htpy' - ( H ∙h (g ·l L)) - ( λ x → ap (λ z → z ∙ H' x) (inv right-unit))) - ( is-equiv-concat-htpy - ( λ x → ap (H x ∙_) right-unit) - ( (f ·l K) ∙h refl-htpy ∙h H')))) + c = c' → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c' + htpy-eq-square c .c refl = refl-htpy-parallel-cone c + + htpy-parallel-cone-refl-htpy-htpy-cone : + (c c' : cone f g C) → htpy-cone f g c c' → + htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c' + htpy-parallel-cone-refl-htpy-htpy-cone (p , q , H) (p' , q' , H') = + tot + ( λ K → + tot + ( λ L M → + ( ap-concat-htpy H right-unit-htpy) ∙h + ( M ∙h ap-concat-htpy' H' inv-htpy-right-unit-htpy))) -abstract - is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) → - (c : cone f g C) → - is-torsorial (htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c) - is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy {C = C} f g c = - is-contr-is-equiv' - ( Σ (cone f g C) (htpy-cone f g c)) - ( tot (htpy-parallel-cone-refl-htpy-htpy-cone f g c)) - ( is-equiv-tot-is-fiberwise-equiv - ( is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone f g c)) - ( is-torsorial-htpy-cone f g c) + abstract + is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone : + (c c' : cone f g C) → + is-equiv (htpy-parallel-cone-refl-htpy-htpy-cone c c') + is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone (p , q , H) (p' , q' , H') = + is-equiv-tot-is-fiberwise-equiv + ( λ K → + is-equiv-tot-is-fiberwise-equiv + ( λ L → + is-equiv-comp + ( concat-htpy + ( ap-concat-htpy H right-unit-htpy) + ( (f ·l K) ∙h refl-htpy ∙h H')) + ( concat-htpy' + ( H ∙h (g ·l L)) + ( ap-concat-htpy' H' inv-htpy-right-unit-htpy)) + ( is-equiv-concat-htpy' + ( H ∙h (g ·l L)) + ( λ x → ap (_∙ H' x) (inv right-unit))) + ( is-equiv-concat-htpy + ( λ x → ap (H x ∙_) right-unit) + ( (f ·l K) ∙h refl-htpy ∙h H')))) -abstract - is-torsorial-htpy-parallel-cone-refl-htpy : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) {g g' : B → X} (Hg : g ~ g') → - (c : cone f g C) → - is-torsorial (htpy-parallel-cone (refl-htpy' f) Hg c) - is-torsorial-htpy-parallel-cone-refl-htpy {C = C} f {g} = - ind-htpy g - ( λ g'' Hg' → - (c : cone f g C) → - is-torsorial (htpy-parallel-cone (refl-htpy' f) Hg' c)) - ( is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy f g) + abstract + is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy : + (c : cone f g C) → + is-torsorial (htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c) + is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy c = + is-contr-is-equiv' + ( Σ (cone f g C) (htpy-cone f g c)) + ( tot (htpy-parallel-cone-refl-htpy-htpy-cone c)) + ( is-equiv-tot-is-fiberwise-equiv + ( is-equiv-htpy-parallel-cone-refl-htpy-htpy-cone c)) + ( is-torsorial-htpy-cone f g c) -abstract - is-torsorial-htpy-parallel-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {f f' : A → X} (Hf : f ~ f') {g g' : B → X} (Hg : g ~ g') → - (c : cone f g C) → - is-torsorial (htpy-parallel-cone Hf Hg c) - is-torsorial-htpy-parallel-cone - {A = A} {B} {X} {C} {f} {f'} Hf {g} {g'} Hg = - ind-htpy - { A = A} - { B = λ t → X} - ( f) - ( λ f'' Hf' → (g g' : B → X) (Hg : g ~ g') (c : cone f g C) → - is-contr (Σ (cone f'' g' C) (htpy-parallel-cone Hf' Hg c))) - ( λ g g' Hg → is-torsorial-htpy-parallel-cone-refl-htpy f Hg) - Hf g g' Hg - -tr-tr-refl-htpy-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → - let - tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy {f = f})) c - tr-tr-c = tr (λ y → cone f y C) (eq-htpy (refl-htpy {f = g})) tr-c - in - tr-tr-c = c -tr-tr-refl-htpy-cone {C = C} f g c = - let - tr-c = tr (λ f''' → cone f''' g C) (eq-htpy refl-htpy) c - tr-tr-c = tr (λ g'' → cone f g'' C) (eq-htpy refl-htpy) tr-c - α : tr-tr-c = tr-c - α = ap (λ t → tr (λ g'' → cone f g'' C) t tr-c) (eq-htpy-refl-htpy g) - β : tr-c = c - β = ap (λ t → tr (λ f''' → cone f''' g C) t c) (eq-htpy-refl-htpy f) - in - α ∙ β - -htpy-eq-square-refl-htpy : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c c' : cone f g C) → - let tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy {f = f})) c - tr-tr-c = tr (λ y → cone f y C) (eq-htpy (refl-htpy {f = g})) tr-c - in - tr-tr-c = c' → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c' -htpy-eq-square-refl-htpy f g c c' = - map-inv-is-equiv-precomp-Π-is-equiv - ( is-equiv-concat (tr-tr-refl-htpy-cone f g c) c') - ( λ p → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c') - ( htpy-eq-square f g c c') - -left-map-triangle-eq-square-refl-htpy : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c c' : cone f g C) → - ( htpy-eq-square-refl-htpy f g c c' ∘ - concat (tr-tr-refl-htpy-cone f g c) c') ~ - ( htpy-eq-square f g c c') -left-map-triangle-eq-square-refl-htpy f g c c' = - is-section-map-inv-is-equiv-precomp-Π-is-equiv - ( is-equiv-concat (tr-tr-refl-htpy-cone f g c) c') - ( λ p → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c') - ( htpy-eq-square f g c c') + abstract + is-torsorial-htpy-parallel-cone-refl-htpy : + {g' : B → X} (Hg : g ~ g') (c : cone f g C) → + is-torsorial (htpy-parallel-cone (refl-htpy' f) Hg c) + is-torsorial-htpy-parallel-cone-refl-htpy = + ind-htpy g + ( λ g'' Hg' → + (c : cone f g C) → + is-torsorial (htpy-parallel-cone (refl-htpy' f) Hg' c)) + ( is-torsorial-htpy-parallel-cone-refl-htpy-refl-htpy) -abstract - htpy-parallel-cone-eq' : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) {g g' : B → X} (Hg : g ~ g') → - (c : cone f g C) (c' : cone f g' C) → + abstract + is-torsorial-htpy-parallel-cone : + {f' : A → X} (Hf : f ~ f') {g' : B → X} (Hg : g ~ g') (c : cone f g C) → + is-torsorial (htpy-parallel-cone Hf Hg c) + is-torsorial-htpy-parallel-cone + {f'} Hf {g'} Hg = + ind-htpy + ( f) + ( λ f'' Hf' → (g' : B → X) (Hg : g ~ g') (c : cone f g C) → + is-contr (Σ (cone f'' g' C) (htpy-parallel-cone Hf' Hg c))) + ( λ g' Hg → is-torsorial-htpy-parallel-cone-refl-htpy Hg) + Hf g' Hg + + tr-tr-refl-htpy-cone : + (c : cone f g C) → let - tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy {f = f})) c - tr-tr-c = tr (λ y → cone f y C) (eq-htpy Hg) tr-c + tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy' f)) c + tr-tr-c = tr (λ y → cone f y C) (eq-htpy (refl-htpy' g)) tr-c in - tr-tr-c = c' → htpy-parallel-cone (refl-htpy' f) Hg c c' - htpy-parallel-cone-eq' {C = C} f {g} = - ind-htpy g - ( λ g'' Hg' → - ( c : cone f g C) (c' : cone f g'' C) → - Id - ( tr - ( λ g'' → cone f g'' C) - ( eq-htpy Hg') - ( tr (λ f''' → cone f''' g C) (eq-htpy (refl-htpy' f)) c)) - ( c') → - htpy-parallel-cone refl-htpy Hg' c c') - ( htpy-eq-square-refl-htpy f g) - - left-map-triangle-parallel-cone-eq' : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c c' : cone f g C) → - ( ( htpy-parallel-cone-eq' f refl-htpy c c') ∘ - ( concat (tr-tr-refl-htpy-cone f g c) c')) ~ - ( htpy-eq-square f g c c') - left-map-triangle-parallel-cone-eq' {A = A} {B} {X} {C} f g c c' = - htpy-right-whisk - ( htpy-eq (htpy-eq (htpy-eq (compute-ind-htpy g + tr-tr-c = c + tr-tr-refl-htpy-cone c = + let + tr-c = tr (λ f''' → cone f''' g C) (eq-htpy refl-htpy) c + tr-tr-c = tr (λ g'' → cone f g'' C) (eq-htpy refl-htpy) tr-c + α : tr-tr-c = tr-c + α = ap (λ t → tr (λ g'' → cone f g'' C) t tr-c) (eq-htpy-refl-htpy g) + β : tr-c = c + β = ap (λ t → tr (λ f''' → cone f''' g C) t c) (eq-htpy-refl-htpy f) + in + α ∙ β + + htpy-eq-square-refl-htpy : + (c c' : cone f g C) → + let tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy' f)) c + tr-tr-c = tr (λ y → cone f y C) (eq-htpy (refl-htpy' g)) tr-c + in + tr-tr-c = c' → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c' + htpy-eq-square-refl-htpy c c' = + map-inv-is-equiv-precomp-Π-is-equiv + ( is-equiv-concat (tr-tr-refl-htpy-cone c) c') + ( λ p → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c') + ( htpy-eq-square c c') + + left-map-triangle-eq-square-refl-htpy : + (c c' : cone f g C) → + htpy-eq-square-refl-htpy c c' ∘ concat (tr-tr-refl-htpy-cone c) c' ~ + htpy-eq-square c c' + left-map-triangle-eq-square-refl-htpy c c' = + is-section-map-inv-is-equiv-precomp-Π-is-equiv + ( is-equiv-concat (tr-tr-refl-htpy-cone c) c') + ( λ p → htpy-parallel-cone (refl-htpy' f) (refl-htpy' g) c c') + ( htpy-eq-square c c') + + abstract + htpy-parallel-cone-eq' : + {g' : B → X} (Hg : g ~ g') → + (c : cone f g C) (c' : cone f g' C) → + let + tr-c = tr (λ x → cone x g C) (eq-htpy (refl-htpy' f)) c + tr-tr-c = tr (λ y → cone f y C) (eq-htpy Hg) tr-c + in + tr-tr-c = c' → htpy-parallel-cone (refl-htpy' f) Hg c c' + htpy-parallel-cone-eq' = + ind-htpy g ( λ g'' Hg' → ( c : cone f g C) (c' : cone f g'' C) → - Id - ( tr (λ g'' → cone f g'' C) (eq-htpy Hg') - ( tr (λ f''' → cone f''' g C) (eq-htpy (refl-htpy' f)) c)) c' → + Id + ( tr + ( λ g'' → cone f g'' C) + ( eq-htpy Hg') + ( tr (λ f''' → cone f''' g C) (eq-htpy (refl-htpy' f)) c)) + ( c') → htpy-parallel-cone refl-htpy Hg' c c') - ( htpy-eq-square-refl-htpy f g)) c) c')) - ( concat (tr-tr-refl-htpy-cone f g c) c') ∙h - ( left-map-triangle-eq-square-refl-htpy f g c c') + ( htpy-eq-square-refl-htpy) + + left-map-triangle-parallel-cone-eq' : + (c c' : cone f g C) → + ( ( htpy-parallel-cone-eq' refl-htpy c c') ∘ + ( concat (tr-tr-refl-htpy-cone c) c')) ~ + ( htpy-eq-square c c') + left-map-triangle-parallel-cone-eq' c c' = + ( htpy-right-whisk + ( multivariable-htpy-eq 3 + ( compute-ind-htpy g + ( λ g'' Hg' → + ( c : cone f g C) (c' : cone f g'' C) → + ( tr + ( λ g'' → cone f g'' C) + ( eq-htpy Hg') + ( tr (λ f''' → cone f''' g C) (eq-htpy (refl-htpy' f)) c)) = + ( c') → + htpy-parallel-cone refl-htpy Hg' c c') + ( htpy-eq-square-refl-htpy)) + ( c) + ( c')) + ( concat (tr-tr-refl-htpy-cone c) c')) ∙h + ( left-map-triangle-eq-square-refl-htpy c c') -abstract - htpy-parallel-cone-eq : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {f f' : A → X} (Hf : f ~ f') {g g' : B → X} (Hg : g ~ g') → - (c : cone f g C) (c' : cone f' g' C) → - let tr-c = tr (λ x → cone x g C) (eq-htpy Hf) c - tr-tr-c = tr (λ y → cone f' y C) (eq-htpy Hg) tr-c - in - tr-tr-c = c' → htpy-parallel-cone Hf Hg c c' - htpy-parallel-cone-eq {A = A} {B} {X} {C} {f} {f'} Hf {g} {g'} Hg c c' p = - ind-htpy f - ( λ f'' Hf' → - ( g g' : B → X) (Hg : g ~ g') (c : cone f g C) (c' : cone f'' g' C) → - ( Id - ( tr (λ g'' → cone f'' g'' C) (eq-htpy Hg) - ( tr (λ f''' → cone f''' g C) (eq-htpy Hf') c)) - ( c')) → - htpy-parallel-cone Hf' Hg c c') - ( λ g g' → htpy-parallel-cone-eq' f {g = g} {g' = g'}) - Hf g g' Hg c c' p - - left-map-triangle-parallel-cone-eq : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c c' : cone f g C) → - ( ( htpy-parallel-cone-eq refl-htpy refl-htpy c c') ∘ - ( concat (tr-tr-refl-htpy-cone f g c) c')) ~ - ( htpy-eq-square f g c c') - left-map-triangle-parallel-cone-eq {A = A} {B} {X} {C} f g c c' = - htpy-right-whisk - ( htpy-eq (htpy-eq (htpy-eq (htpy-eq (htpy-eq (htpy-eq (compute-ind-htpy f - ( λ f'' Hf' → - ( g g' : B → X) (Hg : g ~ g') (c : cone f g C) (c' : cone f'' g' C) → - ( Id + abstract + htpy-parallel-cone-eq : + {f' : A → X} (Hf : f ~ f') {g' : B → X} (Hg : g ~ g') → + (c : cone f g C) (c' : cone f' g' C) → + let tr-c = tr (λ x → cone x g C) (eq-htpy Hf) c + tr-tr-c = tr (λ y → cone f' y C) (eq-htpy Hg) tr-c + in + tr-tr-c = c' → htpy-parallel-cone Hf Hg c c' + htpy-parallel-cone-eq {f'} Hf {g'} Hg c c' p = + ind-htpy f + ( λ f'' Hf' → + ( g' : B → X) (Hg : g ~ g') (c : cone f g C) (c' : cone f'' g' C) → + ( Id + ( tr (λ g'' → cone f'' g'' C) (eq-htpy Hg) + ( tr (λ f''' → cone f''' g C) (eq-htpy Hf') c)) + ( c')) → + htpy-parallel-cone Hf' Hg c c') + ( λ g' → htpy-parallel-cone-eq' {g' = g'}) + Hf g' Hg c c' p + + left-map-triangle-parallel-cone-eq : + (c c' : cone f g C) → + ( htpy-parallel-cone-eq refl-htpy refl-htpy c c') ∘ + ( concat (tr-tr-refl-htpy-cone c) c') ~ + ( htpy-eq-square c c') + left-map-triangle-parallel-cone-eq c c' = + ( htpy-right-whisk + ( multivariable-htpy-eq 5 + ( compute-ind-htpy f + ( λ f'' Hf' → + ( g' : B → X) (Hg : g ~ g') + (c : cone f g C) (c' : cone f'' g' C) → ( tr - ( λ g'' → cone f'' g'' C) - ( eq-htpy Hg) - ( tr (λ f''' → cone f''' g C) (eq-htpy Hf') c)) - ( c')) → - htpy-parallel-cone Hf' Hg c c') - ( λ g g' → htpy-parallel-cone-eq' f {g = g} {g' = g'})) g) g) - refl-htpy) c) c')) - ( concat (tr-tr-refl-htpy-cone f g c) c') ∙h - ( left-map-triangle-parallel-cone-eq' f g c c') + ( λ g'' → cone f'' g'' C) + ( eq-htpy Hg) + ( tr (λ f''' → cone f''' g C) (eq-htpy Hf') c) = + ( c')) → + htpy-parallel-cone Hf' Hg c c') + ( λ g' → htpy-parallel-cone-eq' {g' = g'})) + ( g) + ( refl-htpy) + ( c) + ( c')) + ( concat (tr-tr-refl-htpy-cone c) c')) ∙h + ( left-map-triangle-parallel-cone-eq' c c') -abstract - is-fiberwise-equiv-htpy-parallel-cone-eq : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {f f' : A → X} (Hf : f ~ f') {g g' : B → X} (Hg : g ~ g') → + abstract + is-fiberwise-equiv-htpy-parallel-cone-eq : + {f' : A → X} (Hf : f ~ f') {g' : B → X} (Hg : g ~ g') → + (c : cone f g C) (c' : cone f' g' C) → + is-equiv (htpy-parallel-cone-eq Hf Hg c c') + is-fiberwise-equiv-htpy-parallel-cone-eq {f'} Hf {g'} Hg c c' = + ind-htpy f + ( λ f' Hf → + ( g' : B → X) (Hg : g ~ g') (c : cone f g C) (c' : cone f' g' C) → + is-equiv (htpy-parallel-cone-eq Hf Hg c c')) + ( λ g' Hg c c' → + ind-htpy g + ( λ g' Hg → + ( c : cone f g C) (c' : cone f g' C) → + is-equiv (htpy-parallel-cone-eq refl-htpy Hg c c')) + ( λ c c' → + is-equiv-right-map-triangle + ( htpy-eq-square c c') + ( htpy-parallel-cone-eq refl-htpy refl-htpy c c') + ( concat (tr-tr-refl-htpy-cone c) c') + ( inv-htpy (left-map-triangle-parallel-cone-eq c c')) + ( fundamental-theorem-id + ( is-torsorial-htpy-parallel-cone + ( refl-htpy' f) + ( refl-htpy' g) + ( c)) + ( htpy-eq-square c) c') + ( is-equiv-concat (tr-tr-refl-htpy-cone c) c')) + Hg c c') + Hf g' Hg c c' + + eq-htpy-parallel-cone : + {f' : A → X} (Hf : f ~ f') {g' : B → X} (Hg : g ~ g') → (c : cone f g C) (c' : cone f' g' C) → - is-equiv (htpy-parallel-cone-eq Hf Hg c c') - is-fiberwise-equiv-htpy-parallel-cone-eq - {A = A} {B} {X} {C} {f} {f'} Hf {g} {g'} Hg c c' = - ind-htpy f - ( λ f' Hf → - ( g g' : B → X) (Hg : g ~ g') (c : cone f g C) (c' : cone f' g' C) → - is-equiv (htpy-parallel-cone-eq Hf Hg c c')) - ( λ g g' Hg c c' → - ind-htpy g - ( λ g' Hg → - ( c : cone f g C) (c' : cone f g' C) → - is-equiv (htpy-parallel-cone-eq refl-htpy Hg c c')) - ( λ c c' → - is-equiv-right-map-triangle - ( htpy-eq-square f g c c') - ( htpy-parallel-cone-eq refl-htpy refl-htpy c c') - ( concat (tr-tr-refl-htpy-cone f g c) c') - ( inv-htpy (left-map-triangle-parallel-cone-eq f g c c')) - ( fundamental-theorem-id - ( is-torsorial-htpy-parallel-cone - ( refl-htpy' f) - ( refl-htpy' g) - ( c)) - ( htpy-eq-square f g c) c') - ( is-equiv-concat (tr-tr-refl-htpy-cone f g c) c')) - Hg c c') - Hf g g' Hg c c' - -eq-htpy-parallel-cone : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - {f f' : A → X} (Hf : f ~ f') {g g' : B → X} (Hg : g ~ g') → - (c : cone f g C) (c' : cone f' g' C) → - let - tr-c = tr (λ x → cone x g C) (eq-htpy Hf) c - tr-tr-c = tr (λ y → cone f' y C) (eq-htpy Hg) tr-c - in - htpy-parallel-cone Hf Hg c c' → tr-tr-c = c' -eq-htpy-parallel-cone Hf Hg c c' = - map-inv-is-equiv (is-fiberwise-equiv-htpy-parallel-cone-eq Hf Hg c c') + let + tr-c = tr (λ x → cone x g C) (eq-htpy Hf) c + tr-tr-c = tr (λ y → cone f' y C) (eq-htpy Hg) tr-c + in + htpy-parallel-cone Hf Hg c c' → tr-tr-c = c' + eq-htpy-parallel-cone Hf Hg c c' = + map-inv-is-equiv (is-fiberwise-equiv-htpy-parallel-cone-eq Hf Hg c c') ``` ### Dependent products of pullbacks are pullbacks +Given a family of pullback squares, their dependent product is again a pullback +square. + ```agda -cone-Π : +module _ + {l1 l2 l3 l4 : Level} {I : UU l1} + {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} + (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) + where + + map-standard-pullback-Π : + standard-pullback (map-Π f) (map-Π g) → + (i : I) → standard-pullback (f i) (g i) + pr1 (map-standard-pullback-Π (α , β , γ) i) = α i + pr1 (pr2 (map-standard-pullback-Π (α , β , γ) i)) = β i + pr2 (pr2 (map-standard-pullback-Π (α , β , γ) i)) = htpy-eq γ i + + inv-map-standard-pullback-Π : + ((i : I) → standard-pullback (f i) (g i)) → + standard-pullback (map-Π f) (map-Π g) + pr1 (inv-map-standard-pullback-Π h) i = pr1 (h i) + pr1 (pr2 (inv-map-standard-pullback-Π h)) i = pr1 (pr2 (h i)) + pr2 (pr2 (inv-map-standard-pullback-Π h)) = eq-htpy (λ i → pr2 (pr2 (h i))) + + abstract + is-section-inv-map-standard-pullback-Π : + is-section (map-standard-pullback-Π) (inv-map-standard-pullback-Π) + is-section-inv-map-standard-pullback-Π h = + eq-htpy + ( λ i → + map-extensionality-standard-pullback (f i) (g i) refl refl + ( inv + ( ( right-unit) ∙ + ( htpy-eq (is-section-eq-htpy (λ i → pr2 (pr2 (h i)))) i)))) + + abstract + is-retraction-inv-map-standard-pullback-Π : + is-retraction (map-standard-pullback-Π) (inv-map-standard-pullback-Π) + is-retraction-inv-map-standard-pullback-Π (α , β , γ) = + map-extensionality-standard-pullback + ( map-Π f) + ( map-Π g) + ( refl) + ( refl) + ( inv (right-unit ∙ is-retraction-eq-htpy γ)) + + abstract + is-equiv-map-standard-pullback-Π : + is-equiv (map-standard-pullback-Π) + is-equiv-map-standard-pullback-Π = + is-equiv-is-invertible + ( inv-map-standard-pullback-Π) + ( is-section-inv-map-standard-pullback-Π) + ( is-retraction-inv-map-standard-pullback-Π) + +module _ {l1 l2 l3 l4 l5 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} {C : I → UU l5} (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) - (c : (i : I) → cone (f i) (g i) (C i)) → - cone (map-Π f) (map-Π g) ((i : I) → C i) -pr1 (cone-Π f g c) = map-Π (λ i → pr1 (c i)) -pr1 (pr2 (cone-Π f g c)) = map-Π (λ i → pr1 (pr2 (c i))) -pr2 (pr2 (cone-Π f g c)) = htpy-map-Π (λ i → pr2 (pr2 (c i))) + (c : (i : I) → cone (f i) (g i) (C i)) + where -map-standard-pullback-Π : - {l1 l2 l3 l4 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) → - standard-pullback (map-Π f) (map-Π g) → - (i : I) → standard-pullback (f i) (g i) -pr1 (map-standard-pullback-Π f g (α , β , γ) i) = α i -pr1 (pr2 (map-standard-pullback-Π f g (α , β , γ) i)) = β i -pr2 (pr2 (map-standard-pullback-Π f g (α , β , γ) i)) = htpy-eq γ i - -inv-map-standard-pullback-Π : - {l1 l2 l3 l4 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) → - ((i : I) → standard-pullback (f i) (g i)) → - standard-pullback (map-Π f) (map-Π g) -pr1 (inv-map-standard-pullback-Π f g h) i = pr1 (h i) -pr1 (pr2 (inv-map-standard-pullback-Π f g h)) i = pr1 (pr2 (h i)) -pr2 (pr2 (inv-map-standard-pullback-Π f g h)) = - eq-htpy (λ i → (pr2 (pr2 (h i)))) + cone-Π : cone (map-Π f) (map-Π g) ((i : I) → C i) + pr1 cone-Π = map-Π (λ i → pr1 (c i)) + pr1 (pr2 cone-Π) = map-Π (λ i → pr1 (pr2 (c i))) + pr2 (pr2 cone-Π) = htpy-map-Π (λ i → pr2 (pr2 (c i))) -abstract - is-section-inv-map-standard-pullback-Π : - {l1 l2 l3 l4 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) → - map-standard-pullback-Π f g ∘ inv-map-standard-pullback-Π f g ~ id - is-section-inv-map-standard-pullback-Π f g h = + triangle-map-standard-pullback-Π : + map-Π (λ i → gap (f i) (g i) (c i)) ~ + map-standard-pullback-Π f g ∘ gap (map-Π f) (map-Π g) cone-Π + triangle-map-standard-pullback-Π h = eq-htpy ( λ i → - map-extensionality-standard-pullback (f i) (g i) refl refl - ( inv - ( ( right-unit) ∙ - ( htpy-eq (is-section-eq-htpy (λ i → pr2 (pr2 (h i)))) i)))) + map-extensionality-standard-pullback + ( f i) + ( g i) + ( refl) + ( refl) + ( htpy-eq (is-section-eq-htpy _) i ∙ inv right-unit)) -abstract - is-retraction-inv-map-standard-pullback-Π : - {l1 l2 l3 l4 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) → - inv-map-standard-pullback-Π f g ∘ map-standard-pullback-Π f g ~ id - is-retraction-inv-map-standard-pullback-Π f g (α , β , γ) = - map-extensionality-standard-pullback - ( map-Π f) - ( map-Π g) - refl - refl - ( inv (right-unit ∙ (is-retraction-eq-htpy γ))) + abstract + is-pullback-Π : + ((i : I) → is-pullback (f i) (g i) (c i)) → + is-pullback (map-Π f) (map-Π g) cone-Π + is-pullback-Π is-pb-c = + is-equiv-top-map-triangle + ( map-Π (λ i → gap (f i) (g i) (c i))) + ( map-standard-pullback-Π f g) + ( gap (map-Π f) (map-Π g) cone-Π) + ( triangle-map-standard-pullback-Π) + ( is-equiv-map-standard-pullback-Π f g) + ( is-equiv-map-Π-is-fiberwise-equiv is-pb-c) +``` -abstract - is-equiv-map-standard-pullback-Π : - {l1 l2 l3 l4 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) → - is-equiv (map-standard-pullback-Π f g) - is-equiv-map-standard-pullback-Π f g = - is-equiv-is-invertible - ( inv-map-standard-pullback-Π f g) - ( is-section-inv-map-standard-pullback-Π f g) - ( is-retraction-inv-map-standard-pullback-Π f g) - -triangle-map-standard-pullback-Π : - {l1 l2 l3 l4 l5 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} {C : I → UU l5} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) - (c : (i : I) → cone (f i) (g i) (C i)) → - ( map-Π (λ i → gap (f i) (g i) (c i))) ~ - ( map-standard-pullback-Π f g ∘ gap (map-Π f) (map-Π g) (cone-Π f g c)) -triangle-map-standard-pullback-Π f g c h = - eq-htpy (λ i → - map-extensionality-standard-pullback - (f i) - (g i) - refl refl - ( htpy-eq (is-section-eq-htpy _) i ∙ inv right-unit)) +### Coproducts of pullbacks are pullbacks -abstract - is-pullback-cone-Π : - {l1 l2 l3 l4 l5 : Level} {I : UU l1} - {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} {C : I → UU l5} - (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) - (c : (i : I) → cone (f i) (g i) (C i)) → - ((i : I) → is-pullback (f i) (g i) (c i)) → - is-pullback (map-Π f) (map-Π g) (cone-Π f g c) - is-pullback-cone-Π f g c is-pb-c = - is-equiv-top-map-triangle - ( map-Π (λ i → gap (f i) (g i) (c i))) - ( map-standard-pullback-Π f g) - ( gap (map-Π f) (map-Π g) (cone-Π f g c)) - ( triangle-map-standard-pullback-Π f g c) - ( is-equiv-map-standard-pullback-Π f g) - ( is-equiv-map-Π-is-fiberwise-equiv is-pb-c) +```agda +module _ + {l1 l2 l3 l1' l2' l3' : Level} + {A : UU l1} {B : UU l2} {X : UU l3} + {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} + (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') + where + + map-coprod-cone-inl : + standard-pullback f g → + standard-pullback (map-coprod f f') (map-coprod g g') + pr1 (map-coprod-cone-inl (x , y , p)) = inl x + pr1 (pr2 (map-coprod-cone-inl (x , y , p))) = inl y + pr2 (pr2 (map-coprod-cone-inl (x , y , p))) = ap inl p + + map-coprod-cone-inr : + standard-pullback f' g' → + standard-pullback (map-coprod f f') (map-coprod g g') + pr1 (map-coprod-cone-inr (x , y , p)) = inr x + pr1 (pr2 (map-coprod-cone-inr (x , y , p))) = inr y + pr2 (pr2 (map-coprod-cone-inr (x , y , p))) = ap inr p + + map-coprod-cone : + standard-pullback f g + standard-pullback f' g' → + standard-pullback (map-coprod f f') (map-coprod g g') + map-coprod-cone (inl v) = map-coprod-cone-inl v + map-coprod-cone (inr u) = map-coprod-cone-inr u + + map-inv-coprod-cone : + standard-pullback (map-coprod f f') (map-coprod g g') → + standard-pullback f g + standard-pullback f' g' + map-inv-coprod-cone (inl x , inl y , p) = inl (x , y , is-injective-inl p) + map-inv-coprod-cone (inr x , inr y , p) = inr (x , y , is-injective-inr p) + + is-section-map-inv-coprod-cone : + is-section map-coprod-cone map-inv-coprod-cone + is-section-map-inv-coprod-cone (inl x , inl y , p) = + eq-pair-eq-pr2 (eq-pair-eq-pr2 (is-section-is-injective-inl p)) + is-section-map-inv-coprod-cone (inr x , inr y , p) = + eq-pair-eq-pr2 (eq-pair-eq-pr2 (is-section-is-injective-inr p)) + + is-retraction-map-inv-coprod-cone : + is-retraction map-coprod-cone map-inv-coprod-cone + is-retraction-map-inv-coprod-cone (inl (x , y , p)) = + ap inl (eq-pair-eq-pr2 (eq-pair-eq-pr2 (is-retraction-is-injective-inl p))) + is-retraction-map-inv-coprod-cone (inr (x , y , p)) = + ap inr (eq-pair-eq-pr2 (eq-pair-eq-pr2 (is-retraction-is-injective-inr p))) + + abstract + is-equiv-map-coprod-cone : is-equiv map-coprod-cone + is-equiv-map-coprod-cone = + is-equiv-is-invertible + map-inv-coprod-cone + is-section-map-inv-coprod-cone + is-retraction-map-inv-coprod-cone ``` ```agda -is-pullback-back-left-is-pullback-back-right-cube : +module _ {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} - {f : A → B} {g : A → C} {h : B → D} {k : C → D} - {A' : UU l1'} {B' : UU l2'} {C' : UU l3'} {D' : UU l4'} - {f' : A' → B'} {g' : A' → C'} {h' : B' → D'} {k' : C' → D'} - {hA : A' → A} {hB : B' → B} {hC : C' → C} {hD : D' → D} - (top : h' ∘ f' ~ k' ∘ g') - (back-left : f ∘ hA ~ hB ∘ f') - (back-right : g ∘ hA ~ hC ∘ g') - (front-left : h ∘ hB ~ hD ∘ h') - (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → - (c : - coherence-cube-maps - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-pullback h hD (hB , h' , front-left) → - is-pullback k hD (hC , k' , front-right) → - is-pullback g hC (hA , g' , back-right) → - is-pullback f hB (hA , f' , back-left) -is-pullback-back-left-is-pullback-back-right-cube - {f = f} {g} {h} {k} {f' = f'} {g'} {h'} {k'} {hA = hA} {hB} {hC} {hD} - top back-left back-right front-left front-right bottom c - is-pb-front-left is-pb-front-right is-pb-back-right = - is-pullback-left-square-is-pullback-rectangle f h hD - ( hB , h' , front-left) - ( hA , f' , back-left) - ( is-pb-front-left) - ( is-pullback-htpy - ( bottom) - ( refl-htpy) - ( triple - ( hA) - ( k' ∘ g') - ( rectangle-right-cube - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom)) - ( triple - ( refl-htpy) - ( top) - ( coherence-htpy-parallel-cone-rectangle-left-rectangle-right-cube - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c)) - ( is-pullback-rectangle-is-pullback-left-square g k hD - ( hC , k' , front-right) - ( hA , g' , back-right) - ( is-pb-front-right) - ( is-pb-back-right))) - -is-pullback-back-right-is-pullback-back-left-cube : + {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} + {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} + (f : A → X) (g : B → X) (f' : A' → X') (g' : B' → X') + where + + coprod-cone : + cone f g C → cone f' g' C' → + cone (map-coprod f f') (map-coprod g g') (C + C') + pr1 (coprod-cone (p , q , H) (p' , q' , H')) = map-coprod p p' + pr1 (pr2 (coprod-cone (p , q , H) (p' , q' , H'))) = map-coprod q q' + pr2 (pr2 (coprod-cone (p , q , H) (p' , q' , H'))) = + ( inv-htpy (preserves-comp-map-coprod p f p' f')) ∙h + ( htpy-map-coprod H H') ∙h + ( preserves-comp-map-coprod q g q' g') + + triangle-map-coprod-cone : + (c : cone f g C) (c' : cone f' g' C') → + gap (map-coprod f f') (map-coprod g g') (coprod-cone c c') ~ + map-coprod-cone f g f' g' ∘ map-coprod (gap f g c) (gap f' g' c') + triangle-map-coprod-cone c c' (inl _) = + eq-pair-eq-pr2 (eq-pair-eq-pr2 right-unit) + triangle-map-coprod-cone c c' (inr _) = + eq-pair-eq-pr2 (eq-pair-eq-pr2 right-unit) + + abstract + is-pullback-coprod-is-pullback-pair : + (c : cone f g C) (c' : cone f' g' C') → + is-pullback f g c → is-pullback f' g' c' → + is-pullback (map-coprod f f') (map-coprod g g') (coprod-cone c c') + is-pullback-coprod-is-pullback-pair c c' is-pb-c is-pb-c' = + is-equiv-left-map-triangle + ( gap (map-coprod f f') (map-coprod g g') (coprod-cone c c')) + ( map-coprod-cone f g f' g') + ( map-coprod (gap f g c) (gap f' g' c')) + ( triangle-map-coprod-cone c c') + ( is-equiv-map-coprod is-pb-c is-pb-c') + ( is-equiv-map-coprod-cone f g f' g') +``` + +```agda +module _ {l1 l2 l3 l4 l1' l2' l3' l4' : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} {f : A → B} {g : A → C} {h : B → D} {k : C → D} @@ -758,49 +841,82 @@ is-pullback-back-right-is-pullback-back-left-cube : (back-right : g ∘ hA ~ hC ∘ g') (front-left : h ∘ hB ~ hD ∘ h') (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → + (bottom : h ∘ f ~ k ∘ g) (c : coherence-cube-maps f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-pullback h hD (hB , h' , front-left) → - is-pullback k hD (hC , k' , front-right) → - is-pullback f hB (hA , f' , back-left) → - is-pullback g hC (hA , g' , back-right) -is-pullback-back-right-is-pullback-back-left-cube - {f = f} {g} {h} {k} {f' = f'} {g'} {h'} {k'} {hA = hA} {hB} {hC} {hD} - top back-left back-right front-left front-right bottom c - is-pb-front-left is-pb-front-right is-pb-back-left = - is-pullback-left-square-is-pullback-rectangle g k hD - ( hC , k' , front-right) - ( hA , g' , back-right) - ( is-pb-front-right) - ( is-pullback-htpy' - ( bottom) - ( refl-htpy) - ( triple - ( hA) - ( h' ∘ f') - ( rectangle-left-cube - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom)) - ( triple + top back-left back-right front-left front-right bottom) + where + + is-pullback-back-left-is-pullback-back-right-cube : + is-pullback h hD (hB , h' , front-left) → + is-pullback k hD (hC , k' , front-right) → + is-pullback g hC (hA , g' , back-right) → + is-pullback f hB (hA , f' , back-left) + is-pullback-back-left-is-pullback-back-right-cube + is-pb-front-left is-pb-front-right is-pb-back-right = + is-pullback-left-square-is-pullback-rectangle f h hD + ( hB , h' , front-left) + ( hA , f' , back-left) + ( is-pb-front-left) + ( is-pullback-htpy + ( bottom) ( refl-htpy) - ( top) - ( coherence-htpy-parallel-cone-rectangle-left-rectangle-right-cube - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c)) - ( is-pullback-rectangle-is-pullback-left-square f h hD - ( hB , h' , front-left) - ( hA , f' , back-left) - ( is-pb-front-left) - ( is-pb-back-left))) + ( triple + ( hA) + ( k' ∘ g') + ( rectangle-right-cube + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom)) + ( triple + ( refl-htpy) + ( top) + ( coherence-htpy-parallel-cone-rectangle-left-rectangle-right-cube + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom c)) + ( is-pullback-rectangle-is-pullback-left-square g k hD + ( hC , k' , front-right) + ( hA , g' , back-right) + ( is-pb-front-right) + ( is-pb-back-right))) + + is-pullback-back-right-is-pullback-back-left-cube : + is-pullback h hD (hB , h' , front-left) → + is-pullback k hD (hC , k' , front-right) → + is-pullback f hB (hA , f' , back-left) → + is-pullback g hC (hA , g' , back-right) + is-pullback-back-right-is-pullback-back-left-cube + is-pb-front-left is-pb-front-right is-pb-back-left = + is-pullback-left-square-is-pullback-rectangle g k hD + ( hC , k' , front-right) + ( hA , g' , back-right) + ( is-pb-front-right) + ( is-pullback-htpy' + ( bottom) + ( refl-htpy) + ( triple + ( hA) + ( h' ∘ f') + ( rectangle-left-cube + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom)) + ( triple + ( refl-htpy) + ( top) + ( coherence-htpy-parallel-cone-rectangle-left-rectangle-right-cube + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom c)) + ( is-pullback-rectangle-is-pullback-left-square f h hD + ( hB , h' , front-left) + ( hA , f' , back-left) + ( is-pb-front-left) + ( is-pb-back-left))) ``` ### In a commuting cube where the vertical maps are equivalences, the bottom square is a pullback if and only if the top square is a pullback ```agda -is-pullback-bottom-is-pullback-top-cube-is-equiv : +module _ {l1 l2 l3 l4 l1' l2' l3' l4' : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} (f : A → B) (g : A → C) (h : B → D) (k : C → D) @@ -812,106 +928,92 @@ is-pullback-bottom-is-pullback-top-cube-is-equiv : (back-right : g ∘ hA ~ hC ∘ g') (front-left : h ∘ hB ~ hD ∘ h') (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → + (bottom : h ∘ f ~ k ∘ g) (c : coherence-cube-maps f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-equiv hA → is-equiv hB → is-equiv hC → is-equiv hD → - is-pullback h' k' (f' , g' , top) → - is-pullback h k (f , g , bottom) -is-pullback-bottom-is-pullback-top-cube-is-equiv - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c - is-equiv-hA is-equiv-hB is-equiv-hC is-equiv-hD is-pb-top = - descent-is-equiv hB h k - ( f , g , bottom) - ( f' , hA , inv-htpy (back-left)) - ( is-equiv-hB) - ( is-equiv-hA) - ( is-pullback-htpy - ( front-left) - ( refl-htpy' k) - ( triple - ( f') - ( hC ∘ g') - ( rectangle-top-front-right-cube - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom)) - ( triple - ( refl-htpy' f') - ( back-right) - ( coherence-htpy-parallel-cone-coherence-cube-maps - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c)) - ( is-pullback-rectangle-is-pullback-left-square - ( h') - ( hD) - ( k) - ( k' , hC , inv-htpy front-right) - ( f' , g' , top) - ( is-pullback-is-equiv' hD k - ( k' , hC , inv-htpy front-right) - ( is-equiv-hD) - ( is-equiv-hC)) - ( is-pb-top))) + top back-left back-right front-left front-right bottom) + where -is-pullback-top-is-pullback-bottom-cube-is-equiv : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} - (f : A → B) (g : A → C) (h : B → D) (k : C → D) - {A' : UU l1'} {B' : UU l2'} {C' : UU l3'} {D' : UU l4'} - (f' : A' → B') (g' : A' → C') (h' : B' → D') (k' : C' → D') - (hA : A' → A) (hB : B' → B) (hC : C' → C) (hD : D' → D) - (top : h' ∘ f' ~ k' ∘ g') - (back-left : f ∘ hA ~ hB ∘ f') - (back-right : g ∘ hA ~ hC ∘ g') - (front-left : h ∘ hB ~ hD ∘ h') - (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → - (c : - coherence-cube-maps - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-equiv hA → is-equiv hB → is-equiv hC → is-equiv hD → - is-pullback h k (f , g , bottom) → - is-pullback h' k' (f' , g' , top) -is-pullback-top-is-pullback-bottom-cube-is-equiv - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c - is-equiv-hA is-equiv-hB is-equiv-hC is-equiv-hD is-pb-bottom = - is-pullback-top-is-pullback-rectangle h hD k' - ( hB , h' , front-left) - ( f' , g' , top) - ( is-pullback-is-equiv h hD + is-pullback-bottom-is-pullback-top-cube-is-equiv : + is-equiv hA → is-equiv hB → is-equiv hC → is-equiv hD → + is-pullback h' k' (f' , g' , top) → + is-pullback h k (f , g , bottom) + is-pullback-bottom-is-pullback-top-cube-is-equiv + is-equiv-hA is-equiv-hB is-equiv-hC is-equiv-hD is-pb-top = + descent-is-equiv hB h k + ( f , g , bottom) + ( f' , hA , inv-htpy (back-left)) + ( is-equiv-hB) + ( is-equiv-hA) + ( is-pullback-htpy + ( front-left) + ( refl-htpy' k) + ( triple + ( f') + ( hC ∘ g') + ( rectangle-top-front-right-cube + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom)) + ( triple + ( refl-htpy' f') + ( back-right) + ( coherence-htpy-parallel-cone-coherence-cube-maps + f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom c)) + ( is-pullback-rectangle-is-pullback-left-square + ( h') + ( hD) + ( k) + ( k' , hC , inv-htpy front-right) + ( f' , g' , top) + ( is-pullback-is-equiv-horizontal-maps hD k + ( k' , hC , inv-htpy front-right) + ( is-equiv-hD) + ( is-equiv-hC)) + ( is-pb-top))) + + is-pullback-top-is-pullback-bottom-cube-is-equiv : + is-equiv hA → is-equiv hB → is-equiv hC → is-equiv hD → + is-pullback h k (f , g , bottom) → + is-pullback h' k' (f' , g' , top) + is-pullback-top-is-pullback-bottom-cube-is-equiv + is-equiv-hA is-equiv-hB is-equiv-hC is-equiv-hD is-pb-bottom = + is-pullback-top-is-pullback-rectangle h hD k' ( hB , h' , front-left) - is-equiv-hD is-equiv-hB) - ( is-pullback-htpy' refl-htpy front-right - ( pasting-vertical-cone h k hC - ( f , g , bottom) - ( hA , g' , back-right)) - ( triple - ( back-left) - ( refl-htpy) - ( ( assoc-htpy (bottom ·r hA) (k ·l back-right) (front-right ·r g')) ∙h - ( inv-htpy c) ∙h - ( assoc-htpy (h ·l back-left) (front-left ·r f') (hD ·l top)) ∙h - ( ap-concat-htpy' - ( (front-left ·r f') ∙h (hD ·l top)) - ( inv-htpy-right-unit-htpy {H = h ·l back-left})))) - ( is-pullback-rectangle-is-pullback-top h k hC - ( f , g , bottom) - ( hA , g' , back-right) - ( is-pb-bottom) - ( is-pullback-is-equiv g hC + ( f' , g' , top) + ( is-pullback-is-equiv-vertical-maps h hD + ( hB , h' , front-left) + is-equiv-hD is-equiv-hB) + ( is-pullback-htpy' refl-htpy front-right + ( pasting-vertical-cone h k hC + ( f , g , bottom) + ( hA , g' , back-right)) + ( triple + ( back-left) + ( refl-htpy) + ( ( assoc-htpy + ( bottom ·r hA) + ( k ·l back-right) + ( front-right ·r g')) ∙h + ( inv-htpy c) ∙h + ( assoc-htpy (h ·l back-left) (front-left ·r f') (hD ·l top)) ∙h + ( ap-concat-htpy' + ( (front-left ·r f') ∙h (hD ·l top)) + ( inv-htpy-right-unit-htpy {H = h ·l back-left})))) + ( is-pullback-rectangle-is-pullback-top h k hC + ( f , g , bottom) ( hA , g' , back-right) - is-equiv-hC is-equiv-hA))) + ( is-pb-bottom) + ( is-pullback-is-equiv-vertical-maps g hC + ( hA , g' , back-right) + is-equiv-hC is-equiv-hA))) ``` ### In a commuting cube where the maps from back-right to front-left are equivalences, the back-right square is a pullback if and only if the front-left square is a pullback ```agda -is-pullback-front-left-is-pullback-back-right-cube-is-equiv : +module _ {l1 l2 l3 l4 l1' l2' l3' l4' : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} (f : A → B) (g : A → C) (h : B → D) (k : C → D) @@ -923,141 +1025,132 @@ is-pullback-front-left-is-pullback-back-right-cube-is-equiv : (back-right : g ∘ hA ~ hC ∘ g') (front-left : h ∘ hB ~ hD ∘ h') (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → + (bottom : h ∘ f ~ k ∘ g) (c : coherence-cube-maps f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-equiv f' → is-equiv f → is-equiv k' → is-equiv k → - is-pullback g hC (hA , g' , back-right) → - is-pullback h hD (hB , h' , front-left) -is-pullback-front-left-is-pullback-back-right-cube-is-equiv - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c - is-equiv-f' is-equiv-f is-equiv-k' is-equiv-k is-pb-back-right = - is-pullback-bottom-is-pullback-top-cube-is-equiv - hB h' h hD hA g' g hC f' f k' k - back-right (inv-htpy back-left) top bottom (inv-htpy front-right) front-left - ( coherence-cube-maps-mirror-B f g h k f' g' h' k' hA hB hC hD top - back-left back-right front-left front-right bottom c) - is-equiv-f' is-equiv-f is-equiv-k' is-equiv-k is-pb-back-right + top back-left back-right front-left front-right bottom) + where -is-pullback-front-right-is-pullback-back-left-cube-is-equiv : - {l1 l2 l3 l4 l1' l2' l3' l4' : Level} - {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} - (f : A → B) (g : A → C) (h : B → D) (k : C → D) - {A' : UU l1'} {B' : UU l2'} {C' : UU l3'} {D' : UU l4'} - (f' : A' → B') (g' : A' → C') (h' : B' → D') (k' : C' → D') - (hA : A' → A) (hB : B' → B) (hC : C' → C) (hD : D' → D) - (top : h' ∘ f' ~ k' ∘ g') - (back-left : f ∘ hA ~ hB ∘ f') - (back-right : g ∘ hA ~ hC ∘ g') - (front-left : h ∘ hB ~ hD ∘ h') - (front-right : k ∘ hC ~ hD ∘ k') - (bottom : h ∘ f ~ k ∘ g) → - (c : - coherence-cube-maps - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom) → - is-equiv g' → is-equiv h' → is-equiv g → is-equiv h → - is-pullback f hB (hA , f' , back-left) → - is-pullback k hD (hC , k' , front-right) -is-pullback-front-right-is-pullback-back-left-cube-is-equiv - f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c - is-equiv-g' is-equiv-h' is-equiv-g is-equiv-h is-pb-back-left = - is-pullback-bottom-is-pullback-top-cube-is-equiv - hC k' k hD hA f' f hB g' g h' h - back-left (inv-htpy back-right) (inv-htpy top) - ( inv-htpy bottom) (inv-htpy front-left) front-right - ( coherence-cube-maps-rotate-120 f g h k f' g' h' k' hA hB hC hD - top back-left back-right front-left front-right bottom c) - is-equiv-g' is-equiv-g is-equiv-h' is-equiv-h is-pb-back-left + is-pullback-front-left-is-pullback-back-right-cube-is-equiv : + is-equiv f' → is-equiv f → is-equiv k' → is-equiv k → + is-pullback g hC (hA , g' , back-right) → + is-pullback h hD (hB , h' , front-left) + is-pullback-front-left-is-pullback-back-right-cube-is-equiv + is-equiv-f' is-equiv-f is-equiv-k' is-equiv-k is-pb-back-right = + is-pullback-bottom-is-pullback-top-cube-is-equiv + hB h' h hD hA g' g hC f' f k' k + back-right (inv-htpy back-left) top bottom (inv-htpy front-right) + ( front-left) + ( coherence-cube-maps-mirror-B f g h k f' g' h' k' hA hB hC hD top + back-left back-right front-left front-right bottom c) + is-equiv-f' is-equiv-f is-equiv-k' is-equiv-k is-pb-back-right + + is-pullback-front-right-is-pullback-back-left-cube-is-equiv : + is-equiv g' → is-equiv h' → is-equiv g → is-equiv h → + is-pullback f hB (hA , f' , back-left) → + is-pullback k hD (hC , k' , front-right) + is-pullback-front-right-is-pullback-back-left-cube-is-equiv + is-equiv-g' is-equiv-h' is-equiv-g is-equiv-h is-pb-back-left = + is-pullback-bottom-is-pullback-top-cube-is-equiv + hC k' k hD hA f' f hB g' g h' h + back-left (inv-htpy back-right) (inv-htpy top) + ( inv-htpy bottom) (inv-htpy front-left) front-right + ( coherence-cube-maps-rotate-120 f g h k f' g' h' k' hA hB hC hD + top back-left back-right front-left front-right bottom c) + is-equiv-g' is-equiv-g is-equiv-h' is-equiv-h is-pb-back-left ``` ### Identity types commute with pullbacks ```agda -cone-ap : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) (c1 c2 : C) → - let p = pr1 c - q = pr1 (pr2 c) - H = pr2 (pr2 c) - in - cone - ( λ (α : p c1 = p c2) → ap f α ∙ H c2) - ( λ (β : q c1 = q c2) → H c1 ∙ ap g β) - ( c1 = c2) -pr1 (cone-ap f g (p , q , H) c1 c2) = ap p -pr1 (pr2 (cone-ap f g (p , q , H) c1 c2)) = ap q -pr2 (pr2 (cone-ap f g (p , q , H) c1 c2)) γ = - ( ap (λ t → t ∙ (H c2)) (inv (ap-comp f p γ))) ∙ - ( ( inv-nat-htpy H γ) ∙ - ( ap (λ t → (H c1) ∙ t) (ap-comp g q γ))) - -cone-ap' : - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) (c1 c2 : C) → - let p = pr1 c - q = pr1 (pr2 c) - H = pr2 (pr2 c) - in - cone - ( λ (α : p c1 = p c2) → tr (f (p c1) =_) (H c2) (ap f α)) - ( λ (β : q c1 = q c2) → H c1 ∙ ap g β) - ( c1 = c2) -pr1 (cone-ap' f g (p , q , H) c1 c2) = ap p -pr1 (pr2 (cone-ap' f g (p , q , H) c1 c2)) = ap q -pr2 (pr2 (cone-ap' f g (p , q , H) c1 c2)) γ = - ( tr-Id-right (H c2) (ap f (ap p γ))) ∙ - ( ( ap (_∙ H c2) (inv (ap-comp f p γ))) ∙ - ( ( inv-nat-htpy H γ) ∙ - ( ap (H c1 ∙_) (ap-comp g q γ)))) - -is-pullback-cone-ap : +module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} - (f : A → X) (g : B → X) (c : cone f g C) → is-pullback f g c → - (c1 c2 : C) → - let p = pr1 c - q = pr1 (pr2 c) - H = pr2 (pr2 c) - in - is-pullback - ( λ (α : vertical-map-cone f g c c1 = vertical-map-cone f g c c2) → - ap f α ∙ coherence-square-cone f g c c2) - ( λ (β : horizontal-map-cone f g c c1 = horizontal-map-cone f g c c2) → - H c1 ∙ ap g β) - ( cone-ap f g c c1 c2) -is-pullback-cone-ap f g (p , q , H) is-pb-c c1 c2 = - is-pullback-htpy' - ( λ α → tr-Id-right (H c2) (ap f α)) - ( refl-htpy) - ( cone-ap' f g (p , q , H) c1 c2) - ( refl-htpy , refl-htpy , right-unit-htpy) - ( is-pullback-family-is-pullback-tot - ( f (p c1) =_) - ( λ a → ap f {x = p c1} {y = a}) - ( λ b β → H c1 ∙ ap g β) - ( p , q , H) - ( cone-ap' f g (p , q , H) c1) - ( is-pb-c) - ( is-pullback-is-equiv - ( map-Σ _ f (λ a α → ap f α)) - ( map-Σ _ g (λ b β → H c1 ∙ ap g β)) - ( tot-cone-cone-family - ( f (p c1) =_) - ( λ a → ap f) - ( λ b β → H c1 ∙ ap g β) - ( p , q , H) - ( cone-ap' f g (p , q , H) c1)) - ( is-equiv-is-contr _ - ( is-torsorial-path (q c1)) - ( is-torsorial-path (f (p c1)))) - ( is-equiv-is-contr _ - ( is-torsorial-path c1) - ( is-torsorial-path (p c1)))) - ( c2)) + (f : A → X) (g : B → X) (c : cone f g C) + where + + cone-ap : + (c1 c2 : C) → + cone + ( λ (α : vertical-map-cone f g c c1 = vertical-map-cone f g c c2) → + ap f α ∙ coherence-square-cone f g c c2) + ( λ (β : horizontal-map-cone f g c c1 = horizontal-map-cone f g c c2) → + coherence-square-cone f g c c1 ∙ ap g β) + ( c1 = c2) + pr1 (cone-ap c1 c2) = ap (vertical-map-cone f g c) + pr1 (pr2 (cone-ap c1 c2)) = ap (horizontal-map-cone f g c) + pr2 (pr2 (cone-ap c1 c2)) γ = + ( ap + ( _∙ coherence-square-cone f g c c2) + ( inv (ap-comp f (vertical-map-cone f g c) γ))) ∙ + ( ( inv-nat-htpy (coherence-square-cone f g c) γ) ∙ + ( ap + ( coherence-square-cone f g c c1 ∙_) + ( ap-comp g (horizontal-map-cone f g c) γ))) + + cone-ap' : + (c1 c2 : C) → + cone + ( λ (α : vertical-map-cone f g c c1 = vertical-map-cone f g c c2) → + tr + ( f (vertical-map-cone f g c c1) =_) + ( coherence-square-cone f g c c2) + ( ap f α)) + ( λ (β : horizontal-map-cone f g c c1 = horizontal-map-cone f g c c2) → + coherence-square-cone f g c c1 ∙ ap g β) + ( c1 = c2) + pr1 (cone-ap' c1 c2) = ap (vertical-map-cone f g c) + pr1 (pr2 (cone-ap' c1 c2)) = ap (horizontal-map-cone f g c) + pr2 (pr2 (cone-ap' c1 c2)) γ = + ( tr-Id-right + ( coherence-square-cone f g c c2) + ( ap f (ap (vertical-map-cone f g c) γ))) ∙ + ( ( ap + ( _∙ coherence-square-cone f g c c2) + ( inv (ap-comp f (vertical-map-cone f g c) γ))) ∙ + ( ( inv-nat-htpy (coherence-square-cone f g c) γ) ∙ + ( ap + ( coherence-square-cone f g c c1 ∙_) + ( ap-comp g (horizontal-map-cone f g c) γ)))) + + is-pullback-cone-ap : + is-pullback f g c → + (c1 c2 : C) → + is-pullback + ( λ (α : vertical-map-cone f g c c1 = vertical-map-cone f g c c2) → + ap f α ∙ coherence-square-cone f g c c2) + ( λ (β : horizontal-map-cone f g c c1 = horizontal-map-cone f g c c2) → + coherence-square-cone f g c c1 ∙ ap g β) + ( cone-ap c1 c2) + is-pullback-cone-ap is-pb-c c1 c2 = + is-pullback-htpy' + ( λ α → tr-Id-right (coherence-square-cone f g c c2) (ap f α)) + ( refl-htpy) + ( cone-ap' c1 c2) + ( refl-htpy , refl-htpy , right-unit-htpy) + ( is-pullback-family-is-pullback-tot + ( f (vertical-map-cone f g c c1) =_) + ( λ a → ap f {x = vertical-map-cone f g c c1} {y = a}) + ( λ b β → coherence-square-cone f g c c1 ∙ ap g β) + ( c) + ( cone-ap' c1) + ( is-pb-c) + ( is-pullback-is-equiv-vertical-maps + ( map-Σ _ f (λ a α → ap f α)) + ( map-Σ _ g (λ b β → coherence-square-cone f g c c1 ∙ ap g β)) + ( tot-cone-cone-family + ( f (vertical-map-cone f g c c1) =_) + ( λ a → ap f) + ( λ b β → coherence-square-cone f g c c1 ∙ ap g β) + ( c) + ( cone-ap' c1)) + ( is-equiv-is-contr _ + ( is-torsorial-path (horizontal-map-cone f g c c1)) + ( is-torsorial-path (f (vertical-map-cone f g c c1)))) + ( is-equiv-is-contr _ + ( is-torsorial-path c1) + ( is-torsorial-path (vertical-map-cone f g c c1)))) + ( c2)) ``` ## Table of files about pullbacks diff --git a/src/foundation/retracts-of-maps.lagda.md b/src/foundation/retracts-of-maps.lagda.md index da92cc77d8..b0da218350 100644 --- a/src/foundation/retracts-of-maps.lagda.md +++ b/src/foundation/retracts-of-maps.lagda.md @@ -9,24 +9,24 @@ module foundation.retracts-of-maps where ```agda open import foundation.action-on-identifications-functions open import foundation.commuting-prisms-of-maps -open import foundation.commuting-squares-of-maps open import foundation.dependent-pair-types -open import foundation.fibers-of-maps -open import foundation.function-extensionality -open import foundation.function-types open import foundation.functoriality-fibers-of-maps -open import foundation.homotopies -open import foundation.identity-types open import foundation.morphisms-arrows open import foundation.postcomposition-functions open import foundation.precomposition-functions open import foundation.retracts-of-types open import foundation.universe-levels -open import foundation.whiskering-homotopies +open import foundation-core.commuting-squares-of-maps open import foundation-core.equivalences +open import foundation-core.fibers-of-maps +open import foundation-core.function-extensionality +open import foundation-core.function-types +open import foundation-core.homotopies +open import foundation-core.identity-types open import foundation-core.retractions open import foundation-core.sections +open import foundation-core.whiskering-homotopies ``` diff --git a/src/foundation/truncated-maps.lagda.md b/src/foundation/truncated-maps.lagda.md index 93b99bffdd..105179f41a 100644 --- a/src/foundation/truncated-maps.lagda.md +++ b/src/foundation/truncated-maps.lagda.md @@ -49,23 +49,25 @@ module _ where abstract - is-trunc-is-pullback : - is-pullback f g c → is-trunc-map k g → is-trunc-map k (pr1 c) - is-trunc-is-pullback pb is-trunc-g a = + is-trunc-vertical-map-is-pullback : + is-pullback f g c → + is-trunc-map k g → is-trunc-map k (vertical-map-cone f g c) + is-trunc-vertical-map-is-pullback pb is-trunc-g a = is-trunc-is-equiv k ( fiber g (f a)) - ( map-fiber-cone f g c a) - ( is-fiberwise-equiv-map-fiber-cone-is-pullback f g c pb a) + ( map-fiber-vertical-map-cone f g c a) + ( is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback f g c pb a) ( is-trunc-g (f a)) abstract - is-trunc-is-pullback' : + is-trunc-horizontal-map-is-pullback : {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {C : UU l3} {X : UU l4} (f : A → X) (g : B → X) (c : cone f g C) → - is-pullback f g c → is-trunc-map k f → is-trunc-map k (pr1 (pr2 c)) - is-trunc-is-pullback' k f g (pair p (pair q H)) pb is-trunc-f = - is-trunc-is-pullback k g f + is-pullback f g c → + is-trunc-map k f → is-trunc-map k (horizontal-map-cone f g c) + is-trunc-horizontal-map-is-pullback k f g (pair p (pair q H)) pb is-trunc-f = + is-trunc-vertical-map-is-pullback k g f ( swap-cone f g (triple p q H)) ( is-pullback-swap-cone f g (triple p q H) pb) is-trunc-f diff --git a/src/foundation/type-arithmetic-coproduct-types.lagda.md b/src/foundation/type-arithmetic-coproduct-types.lagda.md index a334b91e56..1b2eff0dcb 100644 --- a/src/foundation/type-arithmetic-coproduct-types.lagda.md +++ b/src/foundation/type-arithmetic-coproduct-types.lagda.md @@ -128,27 +128,27 @@ module _ map-right-distributive-Σ-coprod : Σ (A + B) C → (Σ A (λ x → C (inl x))) + (Σ B (λ y → C (inr y))) - map-right-distributive-Σ-coprod (pair (inl x) z) = inl (pair x z) - map-right-distributive-Σ-coprod (pair (inr y) z) = inr (pair y z) + map-right-distributive-Σ-coprod (inl x , z) = inl (x , z) + map-right-distributive-Σ-coprod (inr y , z) = inr (y , z) map-inv-right-distributive-Σ-coprod : (Σ A (λ x → C (inl x))) + (Σ B (λ y → C (inr y))) → Σ (A + B) C - pr1 (map-inv-right-distributive-Σ-coprod (inl (pair x z))) = inl x - pr2 (map-inv-right-distributive-Σ-coprod (inl (pair x z))) = z - pr1 (map-inv-right-distributive-Σ-coprod (inr (pair y z))) = inr y - pr2 (map-inv-right-distributive-Σ-coprod (inr (pair y z))) = z + pr1 (map-inv-right-distributive-Σ-coprod (inl (x , z))) = inl x + pr2 (map-inv-right-distributive-Σ-coprod (inl (x , z))) = z + pr1 (map-inv-right-distributive-Σ-coprod (inr (y , z))) = inr y + pr2 (map-inv-right-distributive-Σ-coprod (inr (y , z))) = z is-section-map-inv-right-distributive-Σ-coprod : ( map-right-distributive-Σ-coprod ∘ map-inv-right-distributive-Σ-coprod) ~ ( id) - is-section-map-inv-right-distributive-Σ-coprod (inl (pair x z)) = refl - is-section-map-inv-right-distributive-Σ-coprod (inr (pair y z)) = refl + is-section-map-inv-right-distributive-Σ-coprod (inl (x , z)) = refl + is-section-map-inv-right-distributive-Σ-coprod (inr (y , z)) = refl is-retraction-map-inv-right-distributive-Σ-coprod : ( map-inv-right-distributive-Σ-coprod ∘ map-right-distributive-Σ-coprod) ~ ( id) - is-retraction-map-inv-right-distributive-Σ-coprod (pair (inl x) z) = refl - is-retraction-map-inv-right-distributive-Σ-coprod (pair (inr y) z) = refl + is-retraction-map-inv-right-distributive-Σ-coprod (inl x , z) = refl + is-retraction-map-inv-right-distributive-Σ-coprod (inr y , z) = refl abstract is-equiv-map-right-distributive-Σ-coprod : @@ -174,25 +174,25 @@ module _ map-left-distributive-Σ-coprod : Σ A (λ x → B x + C x) → (Σ A B) + (Σ A C) - map-left-distributive-Σ-coprod (pair x (inl y)) = inl (pair x y) - map-left-distributive-Σ-coprod (pair x (inr z)) = inr (pair x z) + map-left-distributive-Σ-coprod (x , inl y) = inl (x , y) + map-left-distributive-Σ-coprod (x , inr z) = inr (x , z) map-inv-left-distributive-Σ-coprod : (Σ A B) + (Σ A C) → Σ A (λ x → B x + C x) - pr1 (map-inv-left-distributive-Σ-coprod (inl (pair x y))) = x - pr2 (map-inv-left-distributive-Σ-coprod (inl (pair x y))) = inl y - pr1 (map-inv-left-distributive-Σ-coprod (inr (pair x z))) = x - pr2 (map-inv-left-distributive-Σ-coprod (inr (pair x z))) = inr z + pr1 (map-inv-left-distributive-Σ-coprod (inl (x , y))) = x + pr2 (map-inv-left-distributive-Σ-coprod (inl (x , y))) = inl y + pr1 (map-inv-left-distributive-Σ-coprod (inr (x , z))) = x + pr2 (map-inv-left-distributive-Σ-coprod (inr (x , z))) = inr z is-section-map-inv-left-distributive-Σ-coprod : - ( map-left-distributive-Σ-coprod ∘ map-inv-left-distributive-Σ-coprod) ~ id - is-section-map-inv-left-distributive-Σ-coprod (inl (pair x y)) = refl - is-section-map-inv-left-distributive-Σ-coprod (inr (pair x z)) = refl + map-left-distributive-Σ-coprod ∘ map-inv-left-distributive-Σ-coprod ~ id + is-section-map-inv-left-distributive-Σ-coprod (inl (x , y)) = refl + is-section-map-inv-left-distributive-Σ-coprod (inr (x , z)) = refl is-retraction-map-inv-left-distributive-Σ-coprod : - ( map-inv-left-distributive-Σ-coprod ∘ map-left-distributive-Σ-coprod) ~ id - is-retraction-map-inv-left-distributive-Σ-coprod (pair x (inl y)) = refl - is-retraction-map-inv-left-distributive-Σ-coprod (pair x (inr z)) = refl + map-inv-left-distributive-Σ-coprod ∘ map-left-distributive-Σ-coprod ~ id + is-retraction-map-inv-left-distributive-Σ-coprod (x , inl y) = refl + is-retraction-map-inv-left-distributive-Σ-coprod (x , inr z) = refl is-equiv-map-left-distributive-Σ-coprod : is-equiv map-left-distributive-Σ-coprod @@ -217,33 +217,33 @@ module _ map-right-distributive-prod-coprod : (A + B) × C → (A × C) + (B × C) map-right-distributive-prod-coprod = - map-right-distributive-Σ-coprod A B (λ x → C) + map-right-distributive-Σ-coprod A B (λ _ → C) map-inv-right-distributive-prod-coprod : (A × C) + (B × C) → (A + B) × C map-inv-right-distributive-prod-coprod = - map-inv-right-distributive-Σ-coprod A B (λ x → C) + map-inv-right-distributive-Σ-coprod A B (λ _ → C) is-section-map-inv-right-distributive-prod-coprod : - ( map-right-distributive-prod-coprod ∘ - map-inv-right-distributive-prod-coprod) ~ id + map-right-distributive-prod-coprod ∘ + map-inv-right-distributive-prod-coprod ~ id is-section-map-inv-right-distributive-prod-coprod = - is-section-map-inv-right-distributive-Σ-coprod A B (λ x → C) + is-section-map-inv-right-distributive-Σ-coprod A B (λ _ → C) is-retraction-map-inv-right-distributive-prod-coprod : - ( map-inv-right-distributive-prod-coprod ∘ - map-right-distributive-prod-coprod) ~ id + map-inv-right-distributive-prod-coprod ∘ + map-right-distributive-prod-coprod ~ id is-retraction-map-inv-right-distributive-prod-coprod = - is-retraction-map-inv-right-distributive-Σ-coprod A B (λ x → C) + is-retraction-map-inv-right-distributive-Σ-coprod A B (λ _ → C) abstract is-equiv-map-right-distributive-prod-coprod : is-equiv map-right-distributive-prod-coprod is-equiv-map-right-distributive-prod-coprod = - is-equiv-map-right-distributive-Σ-coprod A B (λ x → C) + is-equiv-map-right-distributive-Σ-coprod A B (λ _ → C) right-distributive-prod-coprod : ((A + B) × C) ≃ ((A × C) + (B × C)) - right-distributive-prod-coprod = right-distributive-Σ-coprod A B (λ x → C) + right-distributive-prod-coprod = right-distributive-Σ-coprod A B (λ _ → C) ``` ### Left distributivity of products over coproducts @@ -255,33 +255,33 @@ module _ map-left-distributive-prod-coprod : A × (B + C) → (A × B) + (A × C) map-left-distributive-prod-coprod = - map-left-distributive-Σ-coprod A (λ x → B) (λ x → C) + map-left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) map-inv-left-distributive-prod-coprod : (A × B) + (A × C) → A × (B + C) map-inv-left-distributive-prod-coprod = - map-inv-left-distributive-Σ-coprod A (λ x → B) (λ x → C) + map-inv-left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) is-section-map-inv-left-distributive-prod-coprod : - ( map-left-distributive-prod-coprod ∘ - map-inv-left-distributive-prod-coprod) ~ id + map-left-distributive-prod-coprod ∘ + map-inv-left-distributive-prod-coprod ~ id is-section-map-inv-left-distributive-prod-coprod = - is-section-map-inv-left-distributive-Σ-coprod A (λ x → B) (λ x → C) + is-section-map-inv-left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) is-retraction-map-inv-left-distributive-prod-coprod : - ( map-inv-left-distributive-prod-coprod ∘ - map-left-distributive-prod-coprod) ~ id + map-inv-left-distributive-prod-coprod ∘ + map-left-distributive-prod-coprod ~ id is-retraction-map-inv-left-distributive-prod-coprod = - is-retraction-map-inv-left-distributive-Σ-coprod A (λ x → B) (λ x → C) + is-retraction-map-inv-left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) is-equiv-map-left-distributive-prod-coprod : is-equiv map-left-distributive-prod-coprod is-equiv-map-left-distributive-prod-coprod = - is-equiv-map-left-distributive-Σ-coprod A (λ x → B) (λ x → C) + is-equiv-map-left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) left-distributive-prod-coprod : (A × (B + C)) ≃ ((A × B) + (A × C)) left-distributive-prod-coprod = - left-distributive-Σ-coprod A (λ x → B) (λ x → C) + left-distributive-Σ-coprod A (λ _ → B) (λ _ → C) ``` ### If a coproduct is contractible then one summand is contractible and the other is empty diff --git a/src/foundation/uniqueness-set-truncations.lagda.md b/src/foundation/uniqueness-set-truncations.lagda.md index c86ccf4779..604c442e9f 100644 --- a/src/foundation/uniqueness-set-truncations.lagda.md +++ b/src/foundation/uniqueness-set-truncations.lagda.md @@ -83,7 +83,7 @@ module _ ( is-set-quotient-is-set-truncation B f Sf)) ``` -### The uniquely uniqueness of set truncations +### The unique uniqueness of set truncations ```agda module _ diff --git a/src/foundation/universal-property-cartesian-product-types.lagda.md b/src/foundation/universal-property-cartesian-product-types.lagda.md index 801bbb4dc3..7c68c09ab3 100644 --- a/src/foundation/universal-property-cartesian-product-types.lagda.md +++ b/src/foundation/universal-property-cartesian-product-types.lagda.md @@ -13,7 +13,6 @@ open import foundation.unit-type open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.constant-maps open import foundation-core.contractible-types open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalences @@ -36,26 +35,38 @@ product ### The universal property of cartesian products as pullbacks ```agda -map-up-product : - {l1 l2 l3 : Level} {X : UU l1} {A : X → UU l2} {B : X → UU l3} → - ((x : X) → A x × B x) → (((x : X) → A x) × ((x : X) → B x)) -pr1 (map-up-product f) x = pr1 (f x) -pr2 (map-up-product f) x = pr2 (f x) - -up-product : - {l1 l2 l3 : Level} {X : UU l1} {A : X → UU l2} {B : X → UU l3} → - is-equiv (map-up-product {A = A} {B}) -up-product = - is-equiv-is-invertible - ( λ (f , g) → (λ x → (f x , g x))) - ( refl-htpy) - ( refl-htpy) - -equiv-up-product : - {l1 l2 l3 : Level} {X : UU l1} {A : X → UU l2} {B : X → UU l3} → - ((x : X) → A x × B x) ≃ (((x : X) → A x) × ((x : X) → B x)) -pr1 equiv-up-product = map-up-product -pr2 equiv-up-product = up-product +module _ + {l1 l2 l3 : Level} {X : UU l1} {A : X → UU l2} {B : X → UU l3} + where + + map-up-product : + ((x : X) → A x × B x) → (((x : X) → A x) × ((x : X) → B x)) + pr1 (map-up-product f) x = pr1 (f x) + pr2 (map-up-product f) x = pr2 (f x) + + map-inv-up-product : + (((x : X) → A x) × ((x : X) → B x)) → (x : X) → A x × B x + pr1 (map-inv-up-product (f , g) x) = f x + pr2 (map-inv-up-product (f , g) x) = g x + + up-product : is-equiv map-up-product + up-product = + is-equiv-is-invertible + ( map-inv-up-product) + ( refl-htpy) + ( refl-htpy) + + is-equiv-map-inv-up-product : is-equiv map-inv-up-product + is-equiv-map-inv-up-product = is-equiv-map-inv-is-equiv up-product + + equiv-up-product : + ((x : X) → A x × B x) ≃ (((x : X) → A x) × ((x : X) → B x)) + pr1 equiv-up-product = map-up-product + pr2 equiv-up-product = up-product + + inv-equiv-up-product : + (((x : X) → A x) × ((x : X) → B x)) ≃ ((x : X) → A x × B x) + inv-equiv-up-product = inv-equiv equiv-up-product ``` We construct the cone for two maps into the unit type. diff --git a/src/foundation/universal-property-coproduct-types.lagda.md b/src/foundation/universal-property-coproduct-types.lagda.md index 1dd7ce000a..a6680f102a 100644 --- a/src/foundation/universal-property-coproduct-types.lagda.md +++ b/src/foundation/universal-property-coproduct-types.lagda.md @@ -17,6 +17,7 @@ open import foundation-core.coproduct-types open import foundation-core.equivalences open import foundation-core.function-extensionality open import foundation-core.function-types +open import foundation-core.homotopies open import foundation-core.identity-types open import foundation-core.precomposition-functions ``` @@ -36,7 +37,8 @@ module _ ev-inl-inr : {l3 : Level} (P : A + B → UU l3) → ((t : A + B) → P t) → ((x : A) → P (inl x)) × ((y : B) → P (inr y)) - ev-inl-inr P s = pair (λ x → s (inl x)) (λ y → s (inr y)) + pr1 (ev-inl-inr P s) x = s (inl x) + pr2 (ev-inl-inr P s) y = s (inr y) dependent-universal-property-coprod : {l3 : Level} (P : A + B → UU l3) → is-equiv (ev-inl-inr P) @@ -44,7 +46,7 @@ module _ is-equiv-is-invertible ( λ p → ind-coprod P (pr1 p) (pr2 p)) ( ind-Σ (λ f g → eq-pair refl refl)) - ( λ s → eq-htpy (ind-coprod _ (λ x → refl) λ y → refl)) + ( λ s → eq-htpy (ind-coprod _ refl-htpy refl-htpy)) equiv-dependent-universal-property-coprod : {l3 : Level} (P : A + B → UU l3) → @@ -55,44 +57,42 @@ module _ abstract universal-property-coprod : - {l3 : Level} (X : UU l3) → - is-equiv (ev-inl-inr (λ (t : A + B) → X)) - universal-property-coprod X = dependent-universal-property-coprod (λ t → X) + {l3 : Level} (X : UU l3) → is-equiv (ev-inl-inr (λ _ → X)) + universal-property-coprod X = dependent-universal-property-coprod (λ _ → X) equiv-universal-property-coprod : - {l3 : Level} (X : UU l3) → - (A + B → X) ≃ ((A → X) × (B → X)) + {l3 : Level} (X : UU l3) → (A + B → X) ≃ ((A → X) × (B → X)) equiv-universal-property-coprod X = - equiv-dependent-universal-property-coprod (λ t → X) + equiv-dependent-universal-property-coprod (λ _ → X) abstract uniqueness-coprod : {l3 : Level} {Y : UU l3} (i : A → Y) (j : B → Y) → ( {l : Level} (X : UU l) → is-equiv (λ (s : Y → X) → pair' (s ∘ i) (s ∘ j))) → - is-equiv (ind-coprod (λ t → Y) i j) - uniqueness-coprod {Y = Y} i j H = + is-equiv (rec-coprod i j) + uniqueness-coprod i j H = is-equiv-is-equiv-precomp - ( ind-coprod _ i j) + ( rec-coprod i j) ( λ X → is-equiv-right-factor - ( ev-inl-inr (λ t → X)) - ( precomp (ind-coprod (λ t → Y) i j) X) + ( ev-inl-inr (λ _ → X)) + ( precomp (rec-coprod i j) X) ( universal-property-coprod X) ( H X)) abstract - universal-property-coprod-is-equiv-ind-coprod : + universal-property-coprod-is-equiv-rec-coprod : {l3 : Level} (X : UU l3) (i : A → X) (j : B → X) → - is-equiv (ind-coprod (λ t → X) i j) → + is-equiv (rec-coprod i j) → (l4 : Level) (Y : UU l4) → - is-equiv (λ (s : X → Y) → pair' (s ∘ i) (s ∘ j)) - universal-property-coprod-is-equiv-ind-coprod X i j H l Y = + is-equiv (λ (s : X → Y) → pair' (s ∘ i) (s ∘ j)) + universal-property-coprod-is-equiv-rec-coprod X i j H l Y = is-equiv-comp - ( ev-inl-inr (λ t → Y)) - ( precomp (ind-coprod (λ t → X) i j) Y) + ( ev-inl-inr (λ _ → Y)) + ( precomp (rec-coprod i j) Y) ( is-equiv-precomp-is-equiv - ( ind-coprod (λ t → X) i j) + ( rec-coprod i j) ( H) ( Y)) ( universal-property-coprod Y) diff --git a/src/foundation/universal-property-dependent-pair-types.lagda.md b/src/foundation/universal-property-dependent-pair-types.lagda.md index 977be5bba8..ee775bfefc 100644 --- a/src/foundation/universal-property-dependent-pair-types.lagda.md +++ b/src/foundation/universal-property-dependent-pair-types.lagda.md @@ -42,7 +42,7 @@ module _ is-equiv-ind-Σ : is-equiv (ind-Σ {C = C}) is-equiv-ind-Σ = is-equiv-is-section is-equiv-ev-pair refl-htpy - equiv-ev-pair : ((x : Σ A B) → C x) ≃ ((a : A) (b : B a) → C (pair a b)) + equiv-ev-pair : ((x : Σ A B) → C x) ≃ ((a : A) (b : B a) → C (a , b)) pr1 equiv-ev-pair = ev-pair pr2 equiv-ev-pair = is-equiv-ev-pair @@ -61,10 +61,8 @@ equiv-ev-pair² : { A : UU l1} {B : A → UU l2} {C : UU l3} { X : UU l4} {Y : X → UU l5} { Z : ( Σ A B → C) → Σ X Y → UU l6} → - Σ ( Σ A B → C) - ( λ k → ( xy : Σ X Y) → Z k xy) ≃ - Σ ( (a : A) → B a → C) - ( λ k → (x : X) → (y : Y x) → Z (ind-Σ k) (x , y)) + Σ ( Σ A B → C) (λ k → ( xy : Σ X Y) → Z k xy) ≃ + Σ ( (a : A) → B a → C) (λ k → (x : X) → (y : Y x) → Z (ind-Σ k) (x , y)) equiv-ev-pair² {X = X} {Y = Y} {Z = Z} = equiv-Σ ( λ k → (x : X) (y : Y x) → Z (ind-Σ k) (x , y)) diff --git a/src/foundation/universal-property-fiber-products.lagda.md b/src/foundation/universal-property-fiber-products.lagda.md index 50a5befa1e..51957b589b 100644 --- a/src/foundation/universal-property-fiber-products.lagda.md +++ b/src/foundation/universal-property-fiber-products.lagda.md @@ -57,9 +57,9 @@ map. inv-gap-fiberwise-prod : standard-pullback (pr1 {B = P}) (pr1 {B = Q}) → Σ X (λ x → (P x) × (Q x)) - pr1 (inv-gap-fiberwise-prod ((x , p) , ((.x , q) , refl))) = x - pr1 (pr2 (inv-gap-fiberwise-prod ((x , p) , ((.x , q) , refl)))) = p - pr2 (pr2 (inv-gap-fiberwise-prod ((x , p) , ((.x , q) , refl)))) = q + pr1 (inv-gap-fiberwise-prod ((x , p) , (.x , q) , refl)) = x + pr1 (pr2 (inv-gap-fiberwise-prod ((x , p) , (.x , q) , refl))) = p + pr2 (pr2 (inv-gap-fiberwise-prod ((x , p) , (.x , q) , refl))) = q abstract is-section-inv-gap-fiberwise-prod : diff --git a/src/foundation/universal-property-pullbacks.lagda.md b/src/foundation/universal-property-pullbacks.lagda.md index 01e0994c99..64bcaf8830 100644 --- a/src/foundation/universal-property-pullbacks.lagda.md +++ b/src/foundation/universal-property-pullbacks.lagda.md @@ -12,19 +12,22 @@ open import foundation-core.universal-property-pullbacks public open import foundation.cones-over-cospans open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.functoriality-cartesian-product-types +open import foundation.pullbacks open import foundation.subtype-identity-principle open import foundation.universe-levels open import foundation-core.contractible-types -open import foundation-core.propositions +open import foundation-core.function-types +open import foundation-core.postcomposition-functions ``` ## Idea -The universal property of pullbacks describes the optimal way to complete a -diagram of the form +The {{#concept "universal property of pullbacks" Disambiguation="types"}} +describes the optimal way to complete a diagram of the form ```text B @@ -67,7 +70,7 @@ module _ ( compute-map-universal-property-pullback f g c up c') ``` -### Uniquely uniqueness of pullbacks +### Unique uniqueness of pullbacks ```agda module _ @@ -94,6 +97,226 @@ module _ up-c up-c') ``` +### The horizontal pullback pasting property + +Given a diagram + +```text + ∙ -------> ∙ -------> ∙ + | | ⌟ | + | | | + v v v + ∙ -------> ∙ -------> ∙ +``` + +where the right-hand square is a pullback, then the left-hand square is a +pullback if and only if the composite square is. + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {C : UU l3} {X : UU l4} {Y : UU l5} {Z : UU l6} + (i : X → Y) (j : Y → Z) (h : C → Z) + where + + abstract + up-pullback-rectangle-up-pullback-left-square : + (c : cone j h B) (d : cone i (vertical-map-cone j h c) A) → + universal-property-pullback j h c → + universal-property-pullback i (vertical-map-cone j h c) d → + universal-property-pullback (j ∘ i) h (pasting-horizontal-cone i j h c d) + up-pullback-rectangle-up-pullback-left-square c d up-pb-c up-pb-d = + universal-property-pullback-is-pullback (j ∘ i) h + ( pasting-horizontal-cone i j h c d) + ( is-pullback-rectangle-is-pullback-left-square i j h c d + ( is-pullback-universal-property-pullback j h c up-pb-c) + ( is-pullback-universal-property-pullback i + ( vertical-map-cone j h c) d up-pb-d)) + + abstract + up-pullback-left-square-up-pullback-rectangle : + (c : cone j h B) (d : cone i (vertical-map-cone j h c) A) → + universal-property-pullback j h c → + universal-property-pullback (j ∘ i) h + ( pasting-horizontal-cone i j h c d) → + universal-property-pullback i (vertical-map-cone j h c) d + up-pullback-left-square-up-pullback-rectangle c d up-pb-c up-pb-rect = + universal-property-pullback-is-pullback + ( i) + ( vertical-map-cone j h c) + ( d) + ( is-pullback-left-square-is-pullback-rectangle i j h c d + ( is-pullback-universal-property-pullback j h c up-pb-c) + ( is-pullback-universal-property-pullback (j ∘ i) h + ( pasting-horizontal-cone i j h c d) up-pb-rect)) +``` + +### The vertical pullback pasting property + +Given a diagram + +```text + ∙ -------> ∙ + | | + | | + v v + ∙ -------> ∙ + | ⌟ | + | | + v v + ∙ -------> ∙ +``` + +where the bottom square is a pullback, then the top square is a pullback if and +only if the composite square is. + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {C : UU l3} {X : UU l4} {Y : UU l5} {Z : UU l6} + (f : C → Z) (g : Y → Z) (h : X → Y) + where + + abstract + up-pullback-top-up-pullback-rectangle : + (c : cone f g B) (d : cone (horizontal-map-cone f g c) h A) → + universal-property-pullback f g c → + universal-property-pullback f (g ∘ h) (pasting-vertical-cone f g h c d) → + universal-property-pullback (horizontal-map-cone f g c) h d + up-pullback-top-up-pullback-rectangle c d up-pb-c up-pb-dc = + universal-property-pullback-is-pullback + ( horizontal-map-cone f g c) + ( h) + ( d) + ( is-pullback-top-is-pullback-rectangle f g h c d + ( is-pullback-universal-property-pullback f g c up-pb-c) + ( is-pullback-universal-property-pullback f (g ∘ h) + ( pasting-vertical-cone f g h c d) + ( up-pb-dc))) + + abstract + up-pullback-rectangle-up-pullback-top : + (c : cone f g B) (d : cone (horizontal-map-cone f g c) h A) → + universal-property-pullback f g c → + universal-property-pullback (horizontal-map-cone f g c) h d → + universal-property-pullback f (g ∘ h) (pasting-vertical-cone f g h c d) + up-pullback-rectangle-up-pullback-top c d up-pb-c up-pb-d = + universal-property-pullback-is-pullback + ( f) + ( g ∘ h) + ( pasting-vertical-cone f g h c d) + ( is-pullback-rectangle-is-pullback-top f g h c d + ( is-pullback-universal-property-pullback f g c up-pb-c) + ( is-pullback-universal-property-pullback + ( horizontal-map-cone f g c) + ( h) + ( d) + ( up-pb-d))) +``` + +### Pullbacks are closed under dependent products + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} {I : UU l1} + {A : I → UU l2} {B : I → UU l3} {X : I → UU l4} + (f : (i : I) → A i → X i) (g : (i : I) → B i → X i) + {C : I → UU l5} (c : (i : I) → cone (f i) (g i) (C i)) + where + + up-pullback-Π : + ((i : I) → universal-property-pullback (f i) (g i) (c i)) → + universal-property-pullback (map-Π f) (map-Π g) (cone-Π f g c) + up-pullback-Π H = + universal-property-pullback-is-pullback + ( map-Π f) + ( map-Π g) + ( cone-Π f g c) + ( is-pullback-Π f g c + ( λ i → + is-pullback-universal-property-pullback (f i) (g i) (c i) (H i))) +``` + +### Pullbacks are closed under cartesian products + +```agda +module _ + {l1 l2 l3 l4 l1' l2' l3' l4' : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} + {A' : UU l1'} {B' : UU l2'} {X' : UU l3'} {C' : UU l4'} + (f : A → X) (g : B → X) (c : cone f g C) + (f' : A' → X') (g' : B' → X') (c' : cone f' g' C') + where + + up-pullback-prod : + universal-property-pullback f g c → + universal-property-pullback f' g' c' → + universal-property-pullback + ( map-prod f f') + ( map-prod g g') + ( prod-cone f g f' g' c c') + up-pullback-prod H H' = + universal-property-pullback-is-pullback + ( map-prod f f') + ( map-prod g g') + ( prod-cone f g f' g' c c') + ( is-pullback-prod-is-pullback-pair f g f' g' c c' + ( is-pullback-universal-property-pullback f g c H) + ( is-pullback-universal-property-pullback f' g' c' H')) +``` + +### Pullbacks are closed under exponentiation + +Given a pullback square + +```text + f' + C ---------> B + | ⌟ | + g'| | g + | | + v v + A ---------> X + f +``` + +then the exponentiated square given by postcomposition + +```text + f' ∘ - + (S → C) ---------> (S → B) + | | + g' ∘ - | | g ∘ - + | | + v v + (S → A) ---------> (S → X) + f ∘ - +``` + +is a pullback square for any type `S`. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {C : UU l4} + (T : UU l5) (f : A → X) (g : B → X) (c : cone f g C) + where + + up-pullback-postcomp : + universal-property-pullback f g c → + universal-property-pullback + ( postcomp T f) + ( postcomp T g) + ( postcomp-cone T f g c) + up-pullback-postcomp H = + universal-property-pullback-is-pullback + ( postcomp T f) + ( postcomp T g) + ( postcomp-cone T f g c) + ( is-pullback-postcomp-is-pullback f g c + ( is-pullback-universal-property-pullback f g c H) + ( T)) +``` + ## Table of files about pullbacks The following table lists files that are about pullbacks as a general concept. diff --git a/src/foundation/whiskering-homotopies.lagda.md b/src/foundation/whiskering-homotopies.lagda.md index 6d3f49a16c..bfe1cc284d 100644 --- a/src/foundation/whiskering-homotopies.lagda.md +++ b/src/foundation/whiskering-homotopies.lagda.md @@ -14,7 +14,6 @@ open import foundation.commuting-squares-of-homotopies open import foundation.commuting-squares-of-identifications open import foundation.homotopy-induction open import foundation.path-algebra -open import foundation.postcomposition-functions open import foundation.universe-levels open import foundation-core.equivalences @@ -23,6 +22,7 @@ open import foundation-core.function-types open import foundation-core.functoriality-dependent-function-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.postcomposition-functions open import foundation-core.precomposition-functions ``` diff --git a/src/linear-algebra/functoriality-vectors.lagda.md b/src/linear-algebra/functoriality-vectors.lagda.md index fb997c2a83..86333bca37 100644 --- a/src/linear-algebra/functoriality-vectors.lagda.md +++ b/src/linear-algebra/functoriality-vectors.lagda.md @@ -15,10 +15,13 @@ open import foundation.function-extensionality open import foundation.function-types open import foundation.homotopies open import foundation.identity-types +open import foundation.postcomposition-functions open import foundation.universe-levels open import foundation.whiskering-homotopies open import linear-algebra.vectors + +open import univalent-combinatorics.standard-finite-types ``` @@ -73,7 +76,7 @@ module _ htpy-functional-vec : (n : ℕ) {f g : A → B} → (f ~ g) → map-functional-vec n f ~ map-functional-vec n g - htpy-functional-vec n H v = eq-htpy (H ·r v) + htpy-functional-vec n = htpy-postcomp (Fin n) ``` ### Binary functoriality of the type of functional vectors diff --git a/src/lists/arrays.lagda.md b/src/lists/arrays.lagda.md index cc929f553b..cd79b99a75 100644 --- a/src/lists/arrays.lagda.md +++ b/src/lists/arrays.lagda.md @@ -77,8 +77,7 @@ module _ cons-array : A → array A → array A cons-array a t = - ( succ-ℕ (length-array t) , - ind-coprod (λ _ → A) (functional-vec-array t) λ _ → a) + ( succ-ℕ (length-array t) , rec-coprod (functional-vec-array t) (λ _ → a)) revert-array : array A → array A revert-array (n , t) = (n , λ k → t (opposite-Fin n k)) diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 178dac2fd7..528701370e 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -32,7 +32,7 @@ open import orthogonal-factorization-systems.lifting-operations public open import orthogonal-factorization-systems.lifting-squares public open import orthogonal-factorization-systems.lifts-families-of-elements public open import orthogonal-factorization-systems.lifts-of-maps public -open import orthogonal-factorization-systems.local-families public +open import orthogonal-factorization-systems.local-families-of-types public open import orthogonal-factorization-systems.local-maps public open import orthogonal-factorization-systems.local-types public open import orthogonal-factorization-systems.localizations-maps public diff --git a/src/orthogonal-factorization-systems/extensions-of-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-of-maps.lagda.md index bda396694f..8719be817a 100644 --- a/src/orthogonal-factorization-systems/extensions-of-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-of-maps.lagda.md @@ -400,13 +400,13 @@ module _ is-contr-extension-dependent-type-is-local-dependent-type : is-local-dependent-type i P → - ((f : (x : A) → P (i x)) → is-contr (extension-dependent-type i P f)) + (f : (x : A) → P (i x)) → is-contr (extension-dependent-type i P f) is-contr-extension-dependent-type-is-local-dependent-type = map-equiv equiv-is-contr-extension-dependent-type-is-local-dependent-type is-local-dependent-type-is-contr-extension-dependent-type : - ((f : (x : A) → P (i x)) → - is-contr (extension-dependent-type i P f)) → is-local-dependent-type i P + ((f : (x : A) → P (i x)) → is-contr (extension-dependent-type i P f)) → + is-local-dependent-type i P is-local-dependent-type-is-contr-extension-dependent-type = map-inv-equiv equiv-is-contr-extension-dependent-type-is-local-dependent-type diff --git a/src/orthogonal-factorization-systems/factorizations-of-maps.lagda.md b/src/orthogonal-factorization-systems/factorizations-of-maps.lagda.md index d0469b27dd..9eca2a16c8 100644 --- a/src/orthogonal-factorization-systems/factorizations-of-maps.lagda.md +++ b/src/orthogonal-factorization-systems/factorizations-of-maps.lagda.md @@ -7,8 +7,6 @@ module orthogonal-factorization-systems.factorizations-of-maps where
Imports ```agda -open import foundation.action-on-identifications-functions -open import foundation.conjunction open import foundation.dependent-pair-types open import foundation.equivalences open import foundation.function-types @@ -16,10 +14,7 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.homotopy-induction open import foundation.identity-types -open import foundation.propositions -open import foundation.retractions open import foundation.retracts-of-types -open import foundation.sections open import foundation.structure-identity-principle open import foundation.torsorial-type-families open import foundation.univalence diff --git a/src/orthogonal-factorization-systems/local-families.lagda.md b/src/orthogonal-factorization-systems/local-families-of-types.lagda.md similarity index 51% rename from src/orthogonal-factorization-systems/local-families.lagda.md rename to src/orthogonal-factorization-systems/local-families-of-types.lagda.md index 8cf6afab72..18f7cb3737 100644 --- a/src/orthogonal-factorization-systems/local-families.lagda.md +++ b/src/orthogonal-factorization-systems/local-families-of-types.lagda.md @@ -1,7 +1,7 @@ -# Local families +# Local families of types ```agda -module orthogonal-factorization-systems.local-families where +module orthogonal-factorization-systems.local-families-of-types where ```
Imports @@ -14,46 +14,40 @@ open import foundation.propositions open import foundation.universe-levels open import orthogonal-factorization-systems.local-types +open import orthogonal-factorization-systems.orthogonal-maps ```
## Idea -A family `B : A → UU l` is said to be **local at** `f : Y → X`, or -**`f`-local**, if every [fiber](foundation-core.fibers-of-maps.md) is. +A family of types `B : A → UU l` is said to be +{{#concept "local" Disambiguation="family of types" Agda=is-local-family}} at +`f : Y → X`, or **`f`-local**, if every +[fiber](foundation-core.fibers-of-maps.md) is. ## Definition -```agda -is-local-family : - {l1 l2 l3 l4 : Level} {Y : UU l1} {X : UU l2} - (f : Y → X) {A : UU l3} → (A → UU l4) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) -is-local-family f {A} B = (x : A) → is-local f (B x) -``` - -## Properties - -### Being local is a property - ```agda module _ - {l1 l2 : Level} {Y : UU l1} {X : UU l2} (f : Y → X) + {l1 l2 l3 l4 : Level} {Y : UU l1} {X : UU l2} + (f : Y → X) {A : UU l3} (B : A → UU l4) where - is-property-is-local-family : - {l3 l4 : Level} {A : UU l3} - (B : A → UU l4) → is-prop (is-local-family f B) - is-property-is-local-family B = + is-local-family : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-local-family = (x : A) → is-local f (B x) + + is-property-is-local-family : is-prop is-local-family + is-property-is-local-family = is-prop-Π (λ x → is-property-is-equiv (precomp f (B x))) - is-local-family-Prop : - {l3 l4 : Level} {A : UU l3} - (B : A → UU l4) → Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) - pr1 (is-local-family-Prop B) = is-local-family f B - pr2 (is-local-family-Prop B) = is-property-is-local-family B + is-local-family-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + pr1 is-local-family-Prop = is-local-family + pr2 is-local-family-Prop = is-property-is-local-family ``` +## Properties + ### A family is `f`-local if and only if it is `f`-orthogonal This remains to be formalized. diff --git a/src/orthogonal-factorization-systems/local-maps.lagda.md b/src/orthogonal-factorization-systems/local-maps.lagda.md index 414018dfb5..5703f6ba16 100644 --- a/src/orthogonal-factorization-systems/local-maps.lagda.md +++ b/src/orthogonal-factorization-systems/local-maps.lagda.md @@ -7,60 +7,93 @@ module orthogonal-factorization-systems.local-maps where
Imports ```agda +open import foundation.cartesian-morphisms-arrows open import foundation.fibers-of-maps open import foundation.propositions +open import foundation.unit-type open import foundation.universe-levels -open import orthogonal-factorization-systems.local-families +open import orthogonal-factorization-systems.local-families-of-types +open import orthogonal-factorization-systems.local-types +open import orthogonal-factorization-systems.orthogonal-maps ```
## Idea -A map `g : A → B` is said to be **local at** `f : Y → X`, or **`f`-local**, if -all its [fibers](foundation-core.fibers-of-maps.md) are. Likewise, a family -`B : A → UU l` is local at `f` if each `B x` is. +A map `g : A → B` is said to be +{{#concept "local" Disambiguation="maps of types" Agda=is-local-map}} at +`f : Y → X`, or +{{#concept "`f`-local" Disambiguation="maps of types" Agda=is-local-map}}, if +all its [fibers](foundation-core.fibers-of-maps.md) are +[`f`-local types](orthogonal-factorization-systems.local-types.md). ## Definition ```agda module _ + {l1 l2 l3 l4 : Level} {Y : UU l1} {X : UU l2} {A : UU l3} {B : UU l4} + (f : Y → X) (g : A → B) where - is-local-map : - {l1 l2 l3 l4 : Level} {Y : UU l1} {X : UU l2} (f : Y → X) - {A : UU l3} {B : UU l4} → (A → B) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) - is-local-map f g = is-local-family f (fiber g) -``` + is-local-map : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-local-map = is-local-family f (fiber g) + + is-property-is-local-map : is-prop is-local-map + is-property-is-local-map = is-property-is-local-family f (fiber g) -## Properties + is-local-map-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-local-map-Prop = is-local-family-Prop f (fiber g) +``` -### Being local is a property +### A type `B` is `f`-local if and only if the terminal map `B → unit` is `f`-local ```agda module _ - {l1 l2 : Level} {Y : UU l1} {X : UU l2} (f : Y → X) + {l1 l2 l3 : Level} {Y : UU l1} {X : UU l2} {B : UU l3} + (f : Y → X) where - is-property-is-local-map : - {l3 l4 : Level} {A : UU l3} {B : UU l4} - (g : A → B) → is-prop (is-local-map f g) - is-property-is-local-map g = is-property-is-local-family f (fiber g) + is-local-is-local-terminal-map : + is-local-map f (terminal-map B) → is-local f B + is-local-is-local-terminal-map H = + is-local-equiv f (inv-equiv-fiber-terminal-map star) (H star) - is-local-map-Prop : - {l3 l4 : Level} {A : UU l3} {B : UU l4} - (g : A → B) → Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) - is-local-map-Prop g = is-local-family-Prop f (fiber g) + is-local-terminal-map-is-local : + is-local f B → is-local-map f (terminal-map B) + is-local-terminal-map-is-local H u = + is-local-equiv f (equiv-fiber-terminal-map u) H ``` -### A type `B` is `f`-local if and only if the terminal projection `B → unit` is `f`-local +### A map is `f`-local if and only if it is `f`-orthogonal -This remains to be formalized. +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {X' : UU l5} {Y' : UU l6} + (f : A → B) (g : X → Y) + where -### A map is `f`-local if and only if it is `f`-orthogonal + is-local-map-is-orthogonal-pullback-condition : + is-orthogonal-pullback-condition f g → is-local-map f g + is-local-map-is-orthogonal-pullback-condition G y = + is-local-is-orthogonal-pullback-condition-terminal-map f + ( is-orthogonal-pullback-condition-right-base-change f g + ( terminal-map (fiber g y)) + ( fiber-cartesian-hom-arrow g y) + ( G)) + + is-local-map-is-orthogonal : is-orthogonal f g → is-local-map f g + is-local-map-is-orthogonal G y = + is-local-is-orthogonal-terminal-map f + ( is-orthogonal-right-base-change f g + ( terminal-map (fiber g y)) + ( fiber-cartesian-hom-arrow g y) + ( G)) +``` -This remains to be formalized. +The converse remains to be formalized. ## See also diff --git a/src/orthogonal-factorization-systems/local-types.lagda.md b/src/orthogonal-factorization-systems/local-types.lagda.md index 6fef761549..07c0e3ca8c 100644 --- a/src/orthogonal-factorization-systems/local-types.lagda.md +++ b/src/orthogonal-factorization-systems/local-types.lagda.md @@ -15,9 +15,11 @@ open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences open import foundation.empty-types open import foundation.equivalences +open import foundation.families-of-equivalences open import foundation.function-extensionality open import foundation.function-types open import foundation.functoriality-dependent-function-types +open import foundation.homotopies open import foundation.identity-types open import foundation.postcomposition-functions open import foundation.precomposition-dependent-functions @@ -96,22 +98,73 @@ module _ {l1 l2 : Level} {Y : UU l1} {X : UU l2} (f : Y → X) where - map-distributive-Π-is-local-dependent-type : + distributive-Π-is-local-dependent-type : {l3 l4 : Level} {A : UU l3} (B : A → X → UU l4) → ((a : A) → is-local-dependent-type f (B a)) → is-local-dependent-type f (λ x → (a : A) → B a x) - map-distributive-Π-is-local-dependent-type B f-loc = + distributive-Π-is-local-dependent-type B f-loc = is-equiv-map-equiv ( ( equiv-swap-Π) ∘e ( equiv-Π-equiv-family (λ a → precomp-Π f (B a) , (f-loc a))) ∘e ( equiv-swap-Π)) - map-distributive-Π-is-local : + distributive-Π-is-local : {l3 l4 : Level} {A : UU l3} (B : A → UU l4) → ((a : A) → is-local f (B a)) → is-local f ((a : A) → B a) - map-distributive-Π-is-local B = - map-distributive-Π-is-local-dependent-type (λ a _ → B a) + distributive-Π-is-local B = + distributive-Π-is-local-dependent-type (λ a _ → B a) +``` + +### Local types are closed under equivalences + +```agda +module _ + {l1 l2 l3 l4 : Level} + {Y : UU l1} {X : UU l2} {A : X → UU l3} {B : X → UU l4} + (f : Y → X) + where + + is-local-dependent-type-fam-equiv : + fam-equiv A B → is-local-dependent-type f B → is-local-dependent-type f A + is-local-dependent-type-fam-equiv e is-local-B = + is-equiv-htpy-equiv + ( ( equiv-Π-equiv-family (inv-equiv ∘ e ∘ f)) ∘e + ( precomp-Π f B , is-local-B) ∘e + ( equiv-Π-equiv-family e)) + ( λ g → + eq-htpy (λ y → inv (is-retraction-map-inv-equiv (e (f y)) (g (f y))))) + + is-local-dependent-type-inv-fam-equiv : + fam-equiv B A → is-local-dependent-type f B → is-local-dependent-type f A + is-local-dependent-type-inv-fam-equiv e = + is-local-dependent-type-fam-equiv (inv-equiv ∘ e) + +module _ + {l1 l2 l3 l4 : Level} + {Y : UU l1} {X : UU l2} {A : UU l3} {B : UU l4} + (f : Y → X) + where + + is-local-equiv : A ≃ B → is-local f B → is-local f A + is-local-equiv e = is-local-dependent-type-fam-equiv f (λ _ → e) + + is-local-inv-equiv : B ≃ A → is-local f B → is-local f A + is-local-inv-equiv e = is-local-dependent-type-inv-fam-equiv f (λ _ → e) +``` + +### Locality is preserved by homotopies + +```agda +module _ + {l1 l2 l3 : Level} {Y : UU l1} {X : UU l2} {A : UU l3} {f f' : Y → X} + where + + is-local-htpy : (H : f ~ f') → is-local f' A → is-local f A + is-local-htpy H = is-equiv-htpy (precomp f' A) (htpy-precomp H A) + + is-local-htpy' : (H : f ~ f') → is-local f A → is-local f' A + is-local-htpy' H = is-equiv-htpy' (precomp f A) (htpy-precomp H A) ``` ### If `S` is `f`-local then `S` is local at every retract of `f` @@ -253,7 +306,7 @@ module _ is-local-is-equiv is-equiv-f = is-equiv-precomp-is-equiv f is-equiv-f ``` -### Families of contractible types are local at any map +### Contractible types are local at any map ```agda module _ diff --git a/src/orthogonal-factorization-systems/null-types.lagda.md b/src/orthogonal-factorization-systems/null-types.lagda.md index 8983cf73b1..7b75cf5b95 100644 --- a/src/orthogonal-factorization-systems/null-types.lagda.md +++ b/src/orthogonal-factorization-systems/null-types.lagda.md @@ -8,11 +8,14 @@ module orthogonal-factorization-systems.null-types where ```agda open import foundation.constant-maps +open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.fibers-of-maps open import foundation.precomposition-functions open import foundation.propositions open import foundation.type-arithmetic-unit-type open import foundation.unit-type +open import foundation.universal-property-family-of-fibers-of-maps open import foundation.universe-levels open import orthogonal-factorization-systems.local-types @@ -40,6 +43,13 @@ module _ is-null : UU (l1 ⊔ l2) is-null = is-equiv (const Y A) + + is-prop-is-null : is-prop is-null + is-prop-is-null = is-property-is-equiv (const Y A) + + is-null-Prop : Prop (l1 ⊔ l2) + pr1 is-null-Prop = is-null + pr2 is-null-Prop = is-prop-is-null ``` ## Properties @@ -76,3 +86,20 @@ module _ equiv-is-null-is-local : is-local (λ y → star) A ≃ is-null Y A equiv-is-null-is-local = inv-equiv equiv-is-local-is-null ``` + +### A type is `f`-local if it is null at every fiber of `f` + +```agda +module _ + {l1 l2 l3 : Level} {Y : UU l1} {X : UU l2} (f : Y → X) + where + + is-local-dependent-type-is-null-fiber : + (A : X → UU l3) → + ((x : X) → is-null (fiber f x) (A x)) → is-local-dependent-type f A + is-local-dependent-type-is-null-fiber A = is-equiv-precomp-Π-fiber-condition + + is-local-is-null-fiber : + (A : UU l3) → ((x : X) → is-null (fiber f x) A) → is-local f A + is-local-is-null-fiber A = is-local-dependent-type-is-null-fiber (λ _ → A) +``` diff --git a/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md b/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md index 556d452c8b..adf460f2b2 100644 --- a/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md +++ b/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md @@ -7,11 +7,39 @@ module orthogonal-factorization-systems.orthogonal-maps where
Imports ```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-morphisms-arrows +open import foundation.cartesian-product-types +open import foundation.composition-algebra +open import foundation.contractible-maps +open import foundation.contractible-types +open import foundation.coproduct-types open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.fibered-maps +open import foundation.fibers-of-maps +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-coproduct-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.postcomposition-functions +open import foundation.precomposition-functions open import foundation.propositions +open import foundation.pullbacks +open import foundation.type-arithmetic-dependent-function-types +open import foundation.unit-type +open import foundation.universal-property-cartesian-product-types +open import foundation.universal-property-coproduct-types +open import foundation.universal-property-dependent-pair-types +open import foundation.universal-property-equivalences +open import foundation.universal-property-pullbacks open import foundation.universe-levels +open import foundation.whiskering-homotopies +open import orthogonal-factorization-systems.lifting-squares +open import orthogonal-factorization-systems.local-types open import orthogonal-factorization-systems.pullback-hom ``` @@ -19,60 +47,1028 @@ open import orthogonal-factorization-systems.pullback-hom ## Idea -The map `f : A → X` is said to be **orthogonal** to `g : B → Y` if their -[pullback-hom](orthogonal-factorization-systems.pullback-hom.md) is an -equivalence. This means that there is a unique -[lifting operation](orthogonal-factorization-systems.lifting-operations.md) -between `f` and `g`. +The map `f : A → B` is said to be +{{#concept "orthogonal" Disambiguation="maps of types" Agda=is-orthogonal}} to +`g : X → Y` if any of the following equivalent definitions hold -In this case we say that `f` is **left orthogonal** to `g` and `g` is **right -orthogonal** to `f`. +1. Their [pullback-hom](orthogonal-factorization-systems.pullback-hom.md) is an + equivalence. -## Definition +2. There is a [unique](foundation-core.contractible-types.md) + [lifting operation](orthogonal-factorization-systems.lifting-operations.md) + between `f` and `g`. + +3. The following is a [pullback square](foundation.pullback-squares.md): + + ```text + - ∘ f + B → X -------> A → X + | | + g ∘ - | | g ∘ - + V V + B → Y -------> A → Y. + - ∘ f + ``` + +4. The [fibers](foundation-core.fibers-of-maps.md) of `g` are + [`f`-local](orthogonal-factorization-systems.local-types.md), i.e., `g` is an + [`f`-local map](orthogonal-factorization-systems.local-maps.md). + +If `f` is orthogonal to `g`, we say that `f` is +{{#concept "left orthogonal" Disambiguation="maps of types" Agda=is-left-orthogonal}} +to `g`, and `g` is +{{#concept "right orthogonal" Disambiguation="maps of types" Agda=is-right-orthogonal}} +to `f`, and may write `f ⊥ g`. + +## Definitions + +### The orthogonality predicate ```agda module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) where - is-orthogonal : (A → X) → (B → Y) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) - is-orthogonal f g = is-equiv (pullback-hom f g) + is-orthogonal : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-orthogonal = is-equiv (pullback-hom f g) _⊥_ = is-orthogonal -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where + is-prop-is-orthogonal : is-prop is-orthogonal + is-prop-is-orthogonal = is-property-is-equiv (pullback-hom f g) + + is-orthogonal-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + pr1 is-orthogonal-Prop = is-orthogonal + pr2 is-orthogonal-Prop = is-prop-is-orthogonal ``` A term of `is-right-orthogonal f g` asserts that `g` is right orthogonal to `f`. ```agda - is-right-orthogonal : (A → X) → (B → Y) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-right-orthogonal : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) is-right-orthogonal = is-orthogonal ``` A term of `is-left-orthogonal f g` asserts that `g` is left orthogonal to `f`. ```agda - is-left-orthogonal : (A → X) → (B → Y) → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) - is-left-orthogonal g f = is-orthogonal f g +module _ + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-left-orthogonal : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-left-orthogonal = is-orthogonal g f +``` + +### The pullback condition for orthogonal maps + +The maps `f` and `g` are orthogonal if and only if the square + +```text + - ∘ f + B → X -------> A → X + | | + g ∘ - | | g ∘ - + V V + B → Y -------> A → Y. + - ∘ f +``` + +is a pullback. + +```agda +module _ + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-orthogonal-pullback-condition = + is-pullback (precomp f Y) (postcomp A g) (cone-pullback-hom' f g) + + is-orthogonal-pullback-condition-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-orthogonal-pullback-condition-Prop = + is-pullback-Prop (precomp f Y) (postcomp A g) (cone-pullback-hom' f g) + + is-prop-is-orthogonal-pullback-condition : + is-prop is-orthogonal-pullback-condition + is-prop-is-orthogonal-pullback-condition = + is-prop-is-pullback (precomp f Y) (postcomp A g) (cone-pullback-hom' f g) +``` + +### The universal property of orthogonal maps + +The universal property of orthogonal maps is the universal property associated +to the pullback condition, which, as opposed to the pullback condition itself, +is a large proposition. + +```agda +module _ + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + universal-property-orthogonal-maps : UUω + universal-property-orthogonal-maps = + universal-property-pullback + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) ``` ## Properties -### Orthogonality is a property +### Being orthogonal means that the type of lifting squares is contractible ```agda module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - (f : A → X) (g : B → Y) + (f : A → B) (g : X → Y) where - is-property-is-orthogonal : is-prop (is-orthogonal f g) - is-property-is-orthogonal = is-property-is-equiv (pullback-hom f g) + unique-lifting-squares-is-orthogonal : + is-orthogonal f g → + (h : fibered-map f g) → + is-contr (lifting-square-fibered-map f g h) + unique-lifting-squares-is-orthogonal H h = + is-contr-equiv + ( fiber (pullback-hom f g) h) + ( compute-fiber-pullback-hom f g h) + ( is-contr-map-is-equiv H h) - is-orthogonal-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) - pr1 is-orthogonal-Prop = is-orthogonal f g - pr2 is-orthogonal-Prop = is-property-is-orthogonal + is-orthogonal-unique-lifting-squares : + ( (h : fibered-map f g) → is-contr (lifting-square-fibered-map f g h)) → + is-orthogonal f g + is-orthogonal-unique-lifting-squares H = + is-equiv-is-contr-map + ( λ h → + is-contr-equiv' + ( lifting-square-fibered-map f g h) + ( compute-fiber-pullback-hom f g h) + ( H h)) +``` + +### Being orthogonal means satisfying the pullback condition of being orthogonal maps + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition-is-orthogonal : + is-orthogonal f g → is-orthogonal-pullback-condition f g + is-orthogonal-pullback-condition-is-orthogonal = + is-equiv-right-factor + ( map-fibered-map-type-standard-pullback-hom f g) + ( gap (precomp f Y) (postcomp A g) (cone-pullback-hom' f g)) + ( is-equiv-map-equiv (equiv-fibered-map-type-standard-pullback-hom f g)) + + is-orthogonal-is-orthogonal-pullback-condition : + is-orthogonal-pullback-condition f g → is-orthogonal f g + is-orthogonal-is-orthogonal-pullback-condition H = + is-equiv-comp + ( map-fibered-map-type-standard-pullback-hom f g) + ( gap (precomp f Y) (postcomp A g) (cone-pullback-hom' f g)) + ( H) + ( is-equiv-map-equiv (equiv-fibered-map-type-standard-pullback-hom f g)) +``` + +### Being orthogonal means satisfying the universal property of being orthogonal + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + universal-property-orthogonal-maps-is-orthogonal : + is-orthogonal f g → universal-property-orthogonal-maps f g + universal-property-orthogonal-maps-is-orthogonal H = + universal-property-pullback-is-pullback + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( is-orthogonal-pullback-condition-is-orthogonal f g H) + + is-orthogonal-universal-property-orthogonal-maps : + universal-property-orthogonal-maps f g → is-orthogonal f g + is-orthogonal-universal-property-orthogonal-maps H = + is-orthogonal-is-orthogonal-pullback-condition f g + ( is-pullback-universal-property-pullback + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( H)) ``` + +### Orthogonality is preserved by homotopies + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition-htpy : + {f' : A → B} {g' : X → Y} → f ~ f' → g ~ g' → + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f' g' + is-orthogonal-pullback-condition-htpy F G = + is-pullback-htpy' + ( htpy-precomp F Y) + ( htpy-postcomp A G) + ( cone-pullback-hom' f g) + ( ( htpy-postcomp B G) , + ( htpy-precomp F X) , + ( ( commutative-htpy-postcomp-htpy-precomp F G) ∙h + ( inv-htpy-right-unit-htpy))) + + is-orthogonal-pullback-condition-htpy-left : + {f' : A → B} → f ~ f' → + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f' g + is-orthogonal-pullback-condition-htpy-left F = + is-orthogonal-pullback-condition-htpy F refl-htpy + + is-orthogonal-pullback-condition-htpy-right : + {g' : X → Y} → g ~ g' → + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f g' + is-orthogonal-pullback-condition-htpy-right = + is-orthogonal-pullback-condition-htpy refl-htpy +``` + +### Equivalences are orthogonal to every map + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition-is-equiv-left : + is-equiv f → is-orthogonal-pullback-condition f g + is-orthogonal-pullback-condition-is-equiv-left is-equiv-f = + is-pullback-is-equiv-horizontal-maps + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( is-equiv-precomp-is-equiv f is-equiv-f Y) + ( is-equiv-precomp-is-equiv f is-equiv-f X) + + is-orthogonal-is-equiv-left : is-equiv f → is-orthogonal f g + is-orthogonal-is-equiv-left is-equiv-f = + is-orthogonal-is-orthogonal-pullback-condition f g + ( is-orthogonal-pullback-condition-is-equiv-left is-equiv-f) + + is-orthogonal-pullback-condition-is-equiv-right : + is-equiv g → is-orthogonal-pullback-condition f g + is-orthogonal-pullback-condition-is-equiv-right is-equiv-g = + is-pullback-is-equiv-vertical-maps + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( is-equiv-postcomp-is-equiv g is-equiv-g A) + ( is-equiv-postcomp-is-equiv g is-equiv-g B) + + is-orthogonal-is-equiv-right : is-equiv g → is-orthogonal f g + is-orthogonal-is-equiv-right is-equiv-g = + is-orthogonal-is-orthogonal-pullback-condition f g + ( is-orthogonal-pullback-condition-is-equiv-right is-equiv-g) + +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-orthogonal-pullback-condition-equiv-left : + (f : A ≃ B) (g : X → Y) → is-orthogonal-pullback-condition (map-equiv f) g + is-orthogonal-pullback-condition-equiv-left f g = + is-orthogonal-pullback-condition-is-equiv-left + ( map-equiv f) + ( g) + ( is-equiv-map-equiv f) + + is-orthogonal-equiv-left : + (f : A ≃ B) (g : X → Y) → is-orthogonal (map-equiv f) g + is-orthogonal-equiv-left f g = + is-orthogonal-is-equiv-left (map-equiv f) g (is-equiv-map-equiv f) + + is-orthogonal-pullback-condition-equiv-right : + (f : A → B) (g : X ≃ Y) → is-orthogonal-pullback-condition f (map-equiv g) + is-orthogonal-pullback-condition-equiv-right f g = + is-orthogonal-pullback-condition-is-equiv-right + ( f) + ( map-equiv g) + ( is-equiv-map-equiv g) + + is-orthogonal-equiv-right : + (f : A → B) (g : X ≃ Y) → is-orthogonal f (map-equiv g) + is-orthogonal-equiv-right f g = + is-orthogonal-is-equiv-right f (map-equiv g) (is-equiv-map-equiv g) +``` + +### Right orthogonal maps are closed under composition and have the left cancellation property + +Given two composable maps `h` after `g`, if `f ⊥ h`, then `f ⊥ g` if and only if +`f ⊥ (h ∘ g)`. + +**Proof:** By the pullback condition of orthogonal maps, the top square in the +below diagram is a pullback precisely when `g` is right orthogonal to `f`: + +```text + - ∘ f + B → X -------> A → X + | | + g ∘ - | | g ∘ - + V V + B → Y -------> A → Y + | ⌟ | + h ∘ - | | h ∘ - + V V + B → Z -------> A → Z. + - ∘ f +``` + +so the result is an instance of the vertical pasting property for pullbacks. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {Z : UU l5} + (f : A → B) (g : X → Y) (h : Y → Z) + where + + is-orthogonal-pullback-condition-right-comp : + is-orthogonal-pullback-condition f h → + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f (h ∘ g) + is-orthogonal-pullback-condition-right-comp = + is-pullback-rectangle-is-pullback-top + ( precomp f Z) + ( postcomp A h) + ( postcomp A g) + ( cone-pullback-hom' f h) + ( cone-pullback-hom' f g) + + up-orthogonal-right-comp : + universal-property-orthogonal-maps f h → + universal-property-orthogonal-maps f g → + universal-property-orthogonal-maps f (h ∘ g) + up-orthogonal-right-comp = + up-pullback-rectangle-up-pullback-top + ( precomp f Z) + ( postcomp A h) + ( postcomp A g) + ( cone-pullback-hom' f h) + ( cone-pullback-hom' f g) + + is-orthogonal-right-comp : + is-orthogonal f h → is-orthogonal f g → is-orthogonal f (h ∘ g) + is-orthogonal-right-comp H G = + is-orthogonal-is-orthogonal-pullback-condition f (h ∘ g) + ( is-orthogonal-pullback-condition-right-comp + ( is-orthogonal-pullback-condition-is-orthogonal f h H) + ( is-orthogonal-pullback-condition-is-orthogonal f g G)) + + is-orthogonal-pullback-condition-right-right-factor : + is-orthogonal-pullback-condition f h → + is-orthogonal-pullback-condition f (h ∘ g) → + is-orthogonal-pullback-condition f g + is-orthogonal-pullback-condition-right-right-factor = + is-pullback-top-is-pullback-rectangle + ( precomp f Z) + ( postcomp A h) + ( postcomp A g) + ( cone-pullback-hom' f h) + ( cone-pullback-hom' f g) + + up-orthogonal-right-right-factor : + universal-property-orthogonal-maps f h → + universal-property-orthogonal-maps f (h ∘ g) → + universal-property-orthogonal-maps f g + up-orthogonal-right-right-factor = + up-pullback-top-up-pullback-rectangle + ( precomp f Z) + ( postcomp A h) + ( postcomp A g) + ( cone-pullback-hom' f h) + ( cone-pullback-hom' f g) + + is-orthogonal-right-right-factor : + is-orthogonal f h → is-orthogonal f (h ∘ g) → is-orthogonal f g + is-orthogonal-right-right-factor H HG = + is-orthogonal-is-orthogonal-pullback-condition f g + ( is-orthogonal-pullback-condition-right-right-factor + ( is-orthogonal-pullback-condition-is-orthogonal f h H) + ( is-orthogonal-pullback-condition-is-orthogonal f (h ∘ g) HG)) +``` + +### Left orthogonal maps are closed under composition and have the right cancellation property + +Given two composable maps `h` after `f`, if `f ⊥ g`, then `h ⊥ g` if and only if +`(h ∘ f) ⊥ g`. + +**Proof:** By the universal property of orthogonal maps, the right square in the +below diagram is a pullback precisely when `f` is left orthogonal to `g`: + +```text + - ∘ h - ∘ f + C → X -------> B → X -------> A → X + | | ⌟ | + g ∘ - | | | g ∘ - + V V V + C → Y -------> B → Y -------> A → Y + - ∘ h - ∘ f +``` + +so the result is an instance of the horizontal pasting property for pullbacks. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {C : UU l5} + (f : A → B) (h : B → C) (g : X → Y) + where + + is-orthogonal-pullback-condition-left-comp : + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition h g → + is-orthogonal-pullback-condition (h ∘ f) g + is-orthogonal-pullback-condition-left-comp = + is-pullback-rectangle-is-pullback-left-square + ( precomp h Y) + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( cone-pullback-hom' h g) + + up-orthogonal-left-comp : + universal-property-orthogonal-maps f g → + universal-property-orthogonal-maps h g → + universal-property-orthogonal-maps (h ∘ f) g + up-orthogonal-left-comp = + up-pullback-rectangle-up-pullback-left-square + ( precomp h Y) + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( cone-pullback-hom' h g) + + is-orthogonal-left-comp : + is-orthogonal f g → is-orthogonal h g → is-orthogonal (h ∘ f) g + is-orthogonal-left-comp F H = + is-orthogonal-is-orthogonal-pullback-condition (h ∘ f) g + ( is-orthogonal-pullback-condition-left-comp + ( is-orthogonal-pullback-condition-is-orthogonal f g F) + ( is-orthogonal-pullback-condition-is-orthogonal h g H)) + + is-orthogonal-pullback-condition-left-left-factor : + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition (h ∘ f) g → + is-orthogonal-pullback-condition h g + is-orthogonal-pullback-condition-left-left-factor = + is-pullback-left-square-is-pullback-rectangle + ( precomp h Y) + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( cone-pullback-hom' h g) + + up-orthogonal-left-left-factor : + universal-property-orthogonal-maps f g → + universal-property-orthogonal-maps (h ∘ f) g → + universal-property-orthogonal-maps h g + up-orthogonal-left-left-factor = + up-pullback-left-square-up-pullback-rectangle + ( precomp h Y) + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + ( cone-pullback-hom' h g) + + is-orthogonal-left-left-factor : + is-orthogonal f g → is-orthogonal (h ∘ f) g → is-orthogonal h g + is-orthogonal-left-left-factor F HF = + is-orthogonal-is-orthogonal-pullback-condition h g + ( is-orthogonal-pullback-condition-left-left-factor + ( is-orthogonal-pullback-condition-is-orthogonal f g F) + ( is-orthogonal-pullback-condition-is-orthogonal (h ∘ f) g HF)) +``` + +### Right orthogonality is preserved by dependent products + +If `f ⊥ gᵢ`, for each `i : I`, then `f ⊥ (map-Π g)`. + +**Proof:** We need to show that the square + +```text + - ∘ f + (B → Πᵢ Xᵢ) ---------------> (A → Πᵢ Xᵢ) + | | + | | + map-Π g ∘ - | | map-Π g ∘ - + | | + v v + (B → Πᵢ Yᵢ) ---------------> (A → Πᵢ Yᵢ) + - ∘ f +``` + +is a pullback. By swapping the arguments at each vertex, this square is +equivalent to + +```text + map-Π (- ∘ f) + (Πᵢ B → Xᵢ) ---------------> (Πᵢ A → Xᵢ) + | | + | | + map-Π (gᵢ ∘ -) | | map-Π (gᵢ ∘ -) + | | + v v + (Πᵢ B → Yᵢ) ---------------> (Πᵢ A → Yᵢ) + map-Π (- ∘ f) +``` + +which is a pullback by assumption since pullbacks are preserved by dependent +products. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {I : UU l1} {A : UU l2} {B : UU l3} {X : I → UU l4} {Y : I → UU l5} + (f : A → B) (g : (i : I) → X i → Y i) + where + + is-orthogonal-pullback-condition-right-Π : + ((i : I) → is-orthogonal-pullback-condition f (g i)) → + is-orthogonal-pullback-condition f (map-Π g) + is-orthogonal-pullback-condition-right-Π G = + is-pullback-bottom-is-pullback-top-cube-is-equiv + ( postcomp B (map-Π g)) + ( precomp f ((i : I) → X i)) + ( precomp f ((i : I) → Y i)) + ( postcomp A (map-Π g)) + ( map-Π (λ i → postcomp B (g i))) + ( map-Π (λ i → precomp f (X i))) + ( map-Π (λ i → precomp f (Y i))) + ( map-Π (λ i → postcomp A (g i))) + ( swap-Π) + ( swap-Π) + ( swap-Π) + ( swap-Π) + ( λ _ → eq-htpy refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( ( ap swap-Π) ·l + ( eq-htpy-refl-htpy ∘ map-Π (λ i → precomp f (Y i) ∘ postcomp B (g i)))) + ( is-equiv-swap-Π) + ( is-equiv-swap-Π) + ( is-equiv-swap-Π) + ( is-equiv-swap-Π) + ( is-pullback-Π + ( λ i → precomp f (Y i)) + ( λ i → postcomp A (g i)) + ( λ i → cone-pullback-hom' f (g i)) + ( G)) + + is-orthogonal-right-Π : + ((i : I) → is-orthogonal f (g i)) → is-orthogonal f (map-Π g) + is-orthogonal-right-Π G = + is-orthogonal-is-orthogonal-pullback-condition f (map-Π g) + ( is-orthogonal-pullback-condition-right-Π + ( λ i → is-orthogonal-pullback-condition-is-orthogonal f (g i) (G i))) +``` + +### Any map that is left orthogonal to a map `g` is also left orthogonal to postcomposing by `g` + +If `f ⊥ g` then `f ⊥ postcomp S g` for every type `S`. + +**Proof:** This is a special case of the previous result by taking `g` to be +constant over `S`. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} (S : UU l5) + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition-right-postcomp : + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f (postcomp S g) + is-orthogonal-pullback-condition-right-postcomp G = + is-orthogonal-pullback-condition-right-Π f (λ _ → g) (λ _ → G) + + is-orthogonal-right-postcomp : + is-orthogonal f g → is-orthogonal f (postcomp S g) + is-orthogonal-right-postcomp G = + is-orthogonal-right-Π f (λ _ → g) (λ _ → G) +``` + +### Right orthogonality is preserved by products + +If `f ⊥ g` and `f ⊥ g'`, then `f ⊥ (g × g')`. + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {X' : UU l5} {Y' : UU l6} + (f : A → B) (g : X → Y) (g' : X' → Y') + where + + is-orthogonal-pullback-condition-right-prod : + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f g' → + is-orthogonal-pullback-condition f (map-prod g g') + is-orthogonal-pullback-condition-right-prod G G' = + is-pullback-top-is-pullback-bottom-cube-is-equiv + ( map-prod (postcomp B g) (postcomp B g')) + ( map-prod (precomp f X) (precomp f X')) + ( map-prod (precomp f Y) (precomp f Y')) + ( map-prod (postcomp A g) (postcomp A g')) + ( postcomp B (map-prod g g')) + ( precomp f (X × X')) + ( precomp f (Y × Y')) + ( postcomp A (map-prod g g')) + ( map-up-product) + ( map-up-product) + ( map-up-product) + ( map-up-product) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( up-product) + ( up-product) + ( up-product) + ( up-product) + ( is-pullback-prod-is-pullback-pair + ( precomp f Y) + ( postcomp A g) + ( precomp f Y') + ( postcomp A g') + ( cone-pullback-hom' f g) + ( cone-pullback-hom' f g') + ( G) + ( G')) + + is-orthogonal-right-prod : + is-orthogonal f g → is-orthogonal f g' → is-orthogonal f (map-prod g g') + is-orthogonal-right-prod G G' = + is-orthogonal-is-orthogonal-pullback-condition f (map-prod g g') + ( is-orthogonal-pullback-condition-right-prod + ( is-orthogonal-pullback-condition-is-orthogonal f g G) + ( is-orthogonal-pullback-condition-is-orthogonal f g' G')) +``` + +### Left orthogonality is preserved by dependent sums + +If `fᵢ ⊥ g` for every `i`, then `(tot f) ⊥ g`. + +**Proof:** We need to show that the square + +```text + - ∘ (tot f) + ((Σ I B) → X) ---------------> ((Σ I A) → X) + | | + | | + g ∘ - | | g ∘ - + | | + v v + ((Σ I B) → Y) ---------------> ((Σ I A) → Y) + - ∘ (tot f) +``` + +is a pullback. However, by the universal property of dependent pair types this +square is equivalent to + +```text + Πᵢ (- ∘ fᵢ) + Πᵢ (Bᵢ → X) -----------> Πᵢ (Aᵢ → X) + | | + | | + Πᵢ (g ∘ -) | | Πᵢ (g ∘ -) + | | + v v + Πᵢ (Bᵢ → Y) -----------> Πᵢ (Aᵢ → Y), + Πᵢ (- ∘ fᵢ) +``` + +which is a pullback by assumption and the fact that pullbacks are preserved +under dependent products. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {I : UU l1} {A : I → UU l2} {B : I → UU l3} {X : UU l4} {Y : UU l5} + (f : (i : I) → A i → B i) (g : X → Y) + where + + is-orthogonal-pullback-condition-left-Σ : + ((i : I) → is-orthogonal-pullback-condition (f i) g) → + is-orthogonal-pullback-condition (tot f) g + is-orthogonal-pullback-condition-left-Σ F = + is-pullback-top-is-pullback-bottom-cube-is-equiv + ( map-Π (λ i → postcomp (B i) g)) + ( map-Π (λ i → precomp (f i) X)) + ( map-Π (λ i → precomp (f i) Y)) + ( map-Π (λ i → postcomp (A i) g)) + ( postcomp (Σ I B) g) + ( precomp (tot f) X) + ( precomp (tot f) Y) + ( postcomp (Σ I A) g) + ( ev-pair) + ( ev-pair) + ( ev-pair) + ( ev-pair) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( λ _ → eq-htpy refl-htpy) + ( inv-htpy + ( ( right-unit-htpy) ∙h + ( ( eq-htpy-refl-htpy) ·r + ( ( map-Π (λ i → precomp (f i) Y)) ∘ + ( map-Π (λ i → postcomp (B i) g)) ∘ + ( ev-pair))))) + ( is-equiv-ev-pair) + ( is-equiv-ev-pair) + ( is-equiv-ev-pair) + ( is-equiv-ev-pair) + ( is-pullback-Π + ( λ i → precomp (f i) Y) + ( λ i → postcomp (A i) g) + ( λ i → cone-pullback-hom' (f i) g) + ( F)) + + is-orthogonal-left-Σ : + ((i : I) → is-orthogonal (f i) g) → is-orthogonal (tot f) g + is-orthogonal-left-Σ F = + is-orthogonal-is-orthogonal-pullback-condition (tot f) g + ( is-orthogonal-pullback-condition-left-Σ + ( λ i → is-orthogonal-pullback-condition-is-orthogonal (f i) g (F i))) +``` + +### Left orthogonality is preserved by coproducts + +If `f ⊥ g` and `f' ⊥ g`, then `(f + f') ⊥ g`. + +**Proof:** We need to show that the square + +```text + - ∘ (f + f') + ((B + B') → X) ---------------> ((A + A') → X) + | | + | | + g ∘ - | | g ∘ - + | | + v v + ((B + B') → Y) ---------------> ((A + A') → Y) + - ∘ (f + f') +``` + +is a pullback. However, by the universal property of coproducts this square is +equivalent to + +```text + (- ∘ f) × (- ∘ f') + (B → X) × (B' → X) -----------> (A → X) × (A' → X) + | | + | | + (g ∘ -) × (g ∘ -) | | (g ∘ -) × (g ∘ -) + | | + v v + (B → Y) × (B' → Y) -----------> (A → Y) × (A' → Y), + (- ∘ f) × (- ∘ f') +``` + +which is a pullback by assumption and the fact that pullbacks are preserved +under products. + +**Note:** This result can also be obtained as a special case of the previous one +by taking the indexing type to be the +[two-element type](foundation.booleans.md). + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {A' : UU l3} {B' : UU l4} {X : UU l5} {Y : UU l6} + (f : A → B) (f' : A' → B') (g : X → Y) + where + + is-orthogonal-pullback-condition-left-coprod : + is-orthogonal-pullback-condition f g → + is-orthogonal-pullback-condition f' g → + is-orthogonal-pullback-condition (map-coprod f f') g + is-orthogonal-pullback-condition-left-coprod F F' = + is-pullback-top-is-pullback-bottom-cube-is-equiv + ( map-prod (postcomp B g) (postcomp B' g)) + ( map-prod (precomp f X) (precomp f' X)) + ( map-prod (precomp f Y) (precomp f' Y)) + ( map-prod (postcomp A g) (postcomp A' g)) + ( postcomp (B + B') g) + ( precomp (map-coprod f f') X) + ( precomp (map-coprod f f') Y) + ( postcomp (A + A') g) + ( ev-inl-inr (λ _ → X)) + ( ev-inl-inr (λ _ → Y)) + ( ev-inl-inr (λ _ → X)) + ( ev-inl-inr (λ _ → Y)) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( refl-htpy) + ( universal-property-coprod X) + ( universal-property-coprod Y) + ( universal-property-coprod X) + ( universal-property-coprod Y) + ( is-pullback-prod-is-pullback-pair + ( precomp f Y) + ( postcomp A g) + ( precomp f' Y) + ( postcomp A' g) + ( cone-pullback-hom' f g) + ( cone-pullback-hom' f' g) + ( F) + ( F')) + + is-orthogonal-left-coprod : + is-orthogonal f g → is-orthogonal f' g → is-orthogonal (map-coprod f f') g + is-orthogonal-left-coprod F F' = + is-orthogonal-is-orthogonal-pullback-condition (map-coprod f f') g + ( is-orthogonal-pullback-condition-left-coprod + ( is-orthogonal-pullback-condition-is-orthogonal f g F) + ( is-orthogonal-pullback-condition-is-orthogonal f' g F')) +``` + +### Right orthogonality is preserved under base change + +Given a pullback square + +```text + X' -----> X + | ⌟ | + g'| | g + v v + Y' -----> Y, +``` + +if `f ⊥ g`, then `f ⊥ g'`. + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {X' : UU l5} {Y' : UU l6} + (f : A → B) (g : X → Y) (g' : X' → Y') (α : cartesian-hom-arrow g' g) + where + + is-orthogonal-pullback-condition-right-base-change : + is-orthogonal-pullback-condition f g → is-orthogonal-pullback-condition f g' + is-orthogonal-pullback-condition-right-base-change G = + is-pullback-back-right-is-pullback-back-left-cube + ( refl-htpy) + ( htpy-postcomp B (coh-cartesian-hom-arrow g' g α)) + ( refl-htpy) + ( refl-htpy) + ( htpy-postcomp A (coh-cartesian-hom-arrow g' g α)) + ( refl-htpy) + ( ( right-unit-htpy) ∙h + ( right-unit-htpy) ∙h + ( inv-htpy + ( commutative-precomp-htpy-postcomp + ( f) + ( coh-cartesian-hom-arrow g' g α)))) + ( G) + ( is-pullback-postcomp-is-pullback + ( map-codomain-cartesian-hom-arrow g' g α) + ( g) + ( cone-cartesian-hom-arrow g' g α) + ( is-cartesian-cartesian-hom-arrow g' g α) + ( A)) + ( is-pullback-postcomp-is-pullback + ( map-codomain-cartesian-hom-arrow g' g α) + ( g) + ( cone-cartesian-hom-arrow g' g α) + ( is-cartesian-cartesian-hom-arrow g' g α) + ( B)) + + is-orthogonal-right-base-change : is-orthogonal f g → is-orthogonal f g' + is-orthogonal-right-base-change G = + is-orthogonal-is-orthogonal-pullback-condition f g' + ( is-orthogonal-pullback-condition-right-base-change + ( is-orthogonal-pullback-condition-is-orthogonal f g G)) +``` + +### A type is `f`-local if and only if the terminal map is `f`-orthogonal + +**Proof (forward direction):** If the terminal map is right orthogonal to `f`, +that means we have a pullback square + +```text + - ∘ f + B → X -----> A → X + | ⌟ | + ! ∘ - | | ! ∘ - + v v + B → 1 -----> A → 1. + - ∘ f +``` + +which displays `precomp f X` as a pullback of an equivalence, hence it is itself +an equivalence. + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {X : UU l3} (f : A → B) + where + + is-local-is-orthogonal-pullback-condition-terminal-map : + is-orthogonal-pullback-condition f (terminal-map X) → is-local f X + is-local-is-orthogonal-pullback-condition-terminal-map = + is-equiv-horizontal-map-is-pullback + ( precomp f unit) + ( postcomp A (terminal-map X)) + ( cone-pullback-hom' f (terminal-map X)) + ( is-local-is-contr f unit is-contr-unit) + + is-local-is-orthogonal-terminal-map : + is-orthogonal f (terminal-map X) → is-local f X + is-local-is-orthogonal-terminal-map F = + is-local-is-orthogonal-pullback-condition-terminal-map + ( is-orthogonal-pullback-condition-is-orthogonal f (terminal-map X) F) + + is-orthogonal-pullback-condition-terminal-map-is-local : + is-local f X → is-orthogonal-pullback-condition f (terminal-map X) + is-orthogonal-pullback-condition-terminal-map-is-local = + is-pullback-is-equiv-horizontal-maps + ( precomp f unit) + ( postcomp A (terminal-map X)) + ( cone-pullback-hom' f (terminal-map X)) + ( is-local-is-contr f unit is-contr-unit) + + is-orthogonal-terminal-map-is-local : + is-local f X → is-orthogonal f (terminal-map X) + is-orthogonal-terminal-map-is-local F = + is-orthogonal-is-orthogonal-pullback-condition f (terminal-map X) + ( is-orthogonal-pullback-condition-terminal-map-is-local F) +``` + +### If the codomain of `g` is `f`-local, then `g` is `f`-orthogonal if and only if the domain of `g` is `f`-local + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + is-orthogonal-pullback-condition-is-local-domain-is-local-codomain : + is-local f Y → is-local f X → is-orthogonal-pullback-condition f g + is-orthogonal-pullback-condition-is-local-domain-is-local-codomain = + is-pullback-is-equiv-horizontal-maps + ( precomp f Y) + ( postcomp A g) + ( cone-pullback-hom' f g) + + is-orthogonal-is-local-domain-is-local-codomain : + is-local f Y → is-local f X → is-orthogonal f g + is-orthogonal-is-local-domain-is-local-codomain H K = + is-orthogonal-is-orthogonal-pullback-condition f g + ( is-orthogonal-pullback-condition-is-local-domain-is-local-codomain H K) + + is-local-domain-is-orthogonal-pullback-condition-is-local-codomain : + is-local f Y → is-orthogonal-pullback-condition f g → is-local f X + is-local-domain-is-orthogonal-pullback-condition-is-local-codomain H G = + is-local-is-orthogonal-pullback-condition-terminal-map f + ( is-orthogonal-pullback-condition-right-comp f g (terminal-map Y) + ( is-orthogonal-pullback-condition-terminal-map-is-local f H) + ( G)) + + is-local-domain-is-orthogonal-is-local-codomain : + is-local f Y → is-orthogonal f g → is-local f X + is-local-domain-is-orthogonal-is-local-codomain H G = + is-local-domain-is-orthogonal-pullback-condition-is-local-codomain H + ( is-orthogonal-pullback-condition-is-orthogonal f g G) +``` + +## References + +- Ulrik Buchholtz and Jonathan Weinberger, _Synthetic fibered (∞, 1)-category + theory_ (2023) ([arXiv:2105.01724](https://arxiv.org/abs/2105.01724), + [DOI:10.21136/HS.2023.04](https://articles.math.cas.cz/10.21136/HS.2023.04)) diff --git a/src/orthogonal-factorization-systems/pullback-hom.lagda.md b/src/orthogonal-factorization-systems/pullback-hom.lagda.md index c902b1ebf7..073c337dc7 100644 --- a/src/orthogonal-factorization-systems/pullback-hom.lagda.md +++ b/src/orthogonal-factorization-systems/pullback-hom.lagda.md @@ -155,8 +155,7 @@ fibered maps. is-equiv-map-equiv equiv-type-standard-pullback-hom-fibered-map universal-property-pullback-fibered-map : - universal-property-pullback - ( precomp f Y) (postcomp A g) (cone-pullback-hom) + universal-property-pullback (precomp f Y) (postcomp A g) (cone-pullback-hom) universal-property-pullback-fibered-map = universal-property-pullback-is-pullback ( precomp f Y) @@ -187,8 +186,13 @@ module _ (f : A → B) (g : X → Y) where + cone-pullback-hom' : cone (precomp f Y) (postcomp A g) (B → X) + pr1 cone-pullback-hom' = postcomp B g + pr1 (pr2 cone-pullback-hom') = precomp f X + pr2 (pr2 cone-pullback-hom') = refl-htpy + pullback-hom : (B → X) → fibered-map f g - pullback-hom = gap-pullback-hom f g (postcomp B g , precomp f X , refl-htpy) + pullback-hom = gap-pullback-hom f g cone-pullback-hom' infix 30 _⋔_ _⋔_ = pullback-hom @@ -240,6 +244,44 @@ module _ compute-fiber-pullback-hom = inv-equiv inv-compute-fiber-pullback-hom ``` +### Computing the pullback-hom of a composite + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {S : UU l5} + (f : A → B) (g : X → Y) (h : Y → S) + where + + map-fibered-map-comp-right-fibered-map : + fibered-map f g → fibered-map f (h ∘ g) + pr1 (map-fibered-map-comp-right-fibered-map (j , i , H)) = h ∘ j + pr1 (pr2 (map-fibered-map-comp-right-fibered-map (j , i , H))) = i + pr2 (pr2 (map-fibered-map-comp-right-fibered-map (j , i , H))) = h ·l H + + compute-pullback-hom-comp-right : + pullback-hom f (h ∘ g) ~ + map-fibered-map-comp-right-fibered-map ∘ pullback-hom f g + compute-pullback-hom-comp-right = refl-htpy + +module _ + {l1 l2 l3 l4 l5 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {S : UU l5} + (f : A → B) (g : X → Y) (h : S → A) + where + + map-fibered-map-comp-left-fibered-map : + fibered-map f g → fibered-map (f ∘ h) g + pr1 (map-fibered-map-comp-left-fibered-map (j , i , H)) = j + pr1 (pr2 (map-fibered-map-comp-left-fibered-map (j , i , H))) = i ∘ h + pr2 (pr2 (map-fibered-map-comp-left-fibered-map (j , i , H))) = H ·r h + + compute-pullback-hom-comp-left : + pullback-hom (f ∘ h) g ~ + map-fibered-map-comp-left-fibered-map ∘ pullback-hom f g + compute-pullback-hom-comp-left = refl-htpy +``` + ## Table of files about pullbacks The following table lists files that are about pullbacks as a general concept. diff --git a/src/synthetic-homotopy-theory/acyclic-maps.lagda.md b/src/synthetic-homotopy-theory/acyclic-maps.lagda.md index eba32753db..16b4fef745 100644 --- a/src/synthetic-homotopy-theory/acyclic-maps.lagda.md +++ b/src/synthetic-homotopy-theory/acyclic-maps.lagda.md @@ -436,8 +436,8 @@ module _ is-acyclic-map (vertical-map-cone f g c) is-acyclic-map-vertical-map-cone-is-pullback pb ac a = is-acyclic-equiv - ( map-fiber-cone f g c a , - is-fiberwise-equiv-map-fiber-cone-is-pullback f g c pb a) + ( map-fiber-vertical-map-cone f g c a , + is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback f g c pb a) ( ac (f a)) module _ diff --git a/src/synthetic-homotopy-theory/cocartesian-morphisms-arrows.lagda.md b/src/synthetic-homotopy-theory/cocartesian-morphisms-arrows.lagda.md index bab25532e8..8648a7cc45 100644 --- a/src/synthetic-homotopy-theory/cocartesian-morphisms-arrows.lagda.md +++ b/src/synthetic-homotopy-theory/cocartesian-morphisms-arrows.lagda.md @@ -7,11 +7,15 @@ module synthetic-homotopy-theory.cocartesian-morphisms-arrows where
Imports ```agda +open import foundation.commuting-squares-of-maps +open import foundation.dependent-pair-types open import foundation.morphisms-arrows +open import foundation.propositions open import foundation.universe-levels open import synthetic-homotopy-theory.cocones-under-spans open import synthetic-homotopy-theory.pushouts +open import synthetic-homotopy-theory.universal-property-pushouts ```
@@ -19,8 +23,8 @@ open import synthetic-homotopy-theory.pushouts ## Idea A [morphism of arrows](foundation.morphisms-arrows.md) `h : f → g` is said to be -{{#concept "cocartesian" Disambiguation="morphism of arrows"}} if the -[commuting square](foundation-core.commuting-squares-of-maps.md) +{{#concept "cocartesian" Disambiguation="morphism of arrows" Agda=is-cocartesian-hom-arrow}} +if the [commuting square](foundation-core.commuting-squares-of-maps.md) ```text i @@ -32,7 +36,7 @@ A [morphism of arrows](foundation.morphisms-arrows.md) `h : f → g` is said to j ``` -is a [pushout square](synthetic-homotopy-theory.pushouts.md). +is a [pushout](synthetic-homotopy-theory.pushouts.md) square. ## Definitions @@ -47,4 +51,76 @@ module _ is-cocartesian-hom-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) is-cocartesian-hom-arrow = is-pushout f (map-domain-hom-arrow f g h) (cocone-hom-arrow f g h) + + is-prop-is-cocartesian-hom-arrow : is-prop is-cocartesian-hom-arrow + is-prop-is-cocartesian-hom-arrow = + is-prop-is-pushout f (map-domain-hom-arrow f g h) (cocone-hom-arrow f g h) + + is-cocartesian-hom-arrow-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + pr1 is-cocartesian-hom-arrow-Prop = is-cocartesian-hom-arrow + pr2 is-cocartesian-hom-arrow-Prop = is-prop-is-cocartesian-hom-arrow ``` + +### The type of cocartesian morphisms of arrows + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + cocartesian-hom-arrow : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + cocartesian-hom-arrow = Σ (hom-arrow f g) (is-cocartesian-hom-arrow f g) + +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) (h : cocartesian-hom-arrow f g) + where + + hom-arrow-cocartesian-hom-arrow : hom-arrow f g + hom-arrow-cocartesian-hom-arrow = pr1 h + + is-cocartesian-cocartesian-hom-arrow : + is-cocartesian-hom-arrow f g hom-arrow-cocartesian-hom-arrow + is-cocartesian-cocartesian-hom-arrow = pr2 h + + map-domain-cocartesian-hom-arrow : A → X + map-domain-cocartesian-hom-arrow = + map-domain-hom-arrow f g hom-arrow-cocartesian-hom-arrow + + map-codomain-cocartesian-hom-arrow : B → Y + map-codomain-cocartesian-hom-arrow = + map-codomain-hom-arrow f g hom-arrow-cocartesian-hom-arrow + + coh-cocartesian-hom-arrow : + coherence-square-maps + ( map-domain-cocartesian-hom-arrow) + ( f) + ( g) + ( map-codomain-cocartesian-hom-arrow) + coh-cocartesian-hom-arrow = + coh-hom-arrow f g hom-arrow-cocartesian-hom-arrow + + cocone-cocartesian-hom-arrow : + cocone f map-domain-cocartesian-hom-arrow Y + cocone-cocartesian-hom-arrow = + cocone-hom-arrow f g hom-arrow-cocartesian-hom-arrow + + universal-property-cocartesian-hom-arrow : + {l : Level} → + universal-property-pushout l + ( f) + ( map-domain-cocartesian-hom-arrow) + ( cocone-cocartesian-hom-arrow) + universal-property-cocartesian-hom-arrow = + universal-property-pushout-is-pushout + ( f) + ( map-domain-cocartesian-hom-arrow) + ( cocone-cocartesian-hom-arrow) + ( is-cocartesian-cocartesian-hom-arrow) +``` + +## See also + +- [Cartesian morphisms of arrows](foundation.cartesian-morphisms-arrows.md) for + the dual. diff --git a/src/synthetic-homotopy-theory/dependent-universal-property-pushouts.lagda.md b/src/synthetic-homotopy-theory/dependent-universal-property-pushouts.lagda.md index 2c9c57bcbe..0278952857 100644 --- a/src/synthetic-homotopy-theory/dependent-universal-property-pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/dependent-universal-property-pushouts.lagda.md @@ -136,14 +136,15 @@ dependent-universal-property-pushout-induction-principle-pushout ( ind-induction-principle-pushout f g c (ind-c l) P) ( pr2 (ind-c l P)) ( λ h → - eq-htpy (htpy-eq-dependent-cocone-map f g c - ( ind-c l) - ( ind-induction-principle-pushout f g c + eq-htpy + ( htpy-eq-dependent-cocone-map f g c ( ind-c l) - ( P) - ( dependent-cocone-map f g c P h)) - ( h) - ( pr2 (ind-c l P) (dependent-cocone-map f g c P h)))) + ( ind-induction-principle-pushout f g c + ( ind-c l) + ( P) + ( dependent-cocone-map f g c P h)) + ( h) + ( pr2 (ind-c l P) (dependent-cocone-map f g c P h)))) ``` #### The dependent universal property of pushouts implies the induction principle of pushouts diff --git a/src/synthetic-homotopy-theory/pushouts.lagda.md b/src/synthetic-homotopy-theory/pushouts.lagda.md index 8d9e69a856..599d95be59 100644 --- a/src/synthetic-homotopy-theory/pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/pushouts.lagda.md @@ -16,6 +16,7 @@ open import foundation.function-types open import foundation.functoriality-dependent-pair-types open import foundation.homotopies open import foundation.identity-types +open import foundation.propositions open import foundation.retractions open import foundation.sections open import foundation.universe-levels @@ -150,14 +151,23 @@ module _ is-retraction-map-inv-equiv (equiv-up-pushout f g X) ``` -### The predicate of being a pushout cocone +### The small predicate of being a pushout cocone ```agda -is-pushout : - { l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} - ( f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) → - UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) -is-pushout f g c = is-equiv (cogap f g c) +module _ + {l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} + (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) + where + + is-pushout : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + is-pushout = is-equiv (cogap f g c) + + is-prop-is-pushout : is-prop is-pushout + is-prop-is-pushout = is-property-is-equiv (cogap f g c) + + is-pushout-Prop : Prop (l1 ⊔ l2 ⊔ l3 ⊔ l4) + pr1 is-pushout-Prop = is-pushout + pr2 is-pushout-Prop = is-prop-is-pushout ``` ## Properties @@ -166,8 +176,8 @@ is-pushout f g c = is-equiv (cogap f g c) ```agda module _ - { l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} - ( f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) + {l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} + (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) where universal-property-pushout-is-pushout : diff --git a/src/synthetic-homotopy-theory/truncated-acyclic-maps.lagda.md b/src/synthetic-homotopy-theory/truncated-acyclic-maps.lagda.md index 187e53b696..f9e9f7385a 100644 --- a/src/synthetic-homotopy-theory/truncated-acyclic-maps.lagda.md +++ b/src/synthetic-homotopy-theory/truncated-acyclic-maps.lagda.md @@ -408,8 +408,8 @@ module _ is-truncated-acyclic-map k (vertical-map-cone f g c) is-truncated-acyclic-map-vertical-map-cone-is-pullback pb ac a = is-truncated-acyclic-equiv - ( map-fiber-cone f g c a , - is-fiberwise-equiv-map-fiber-cone-is-pullback f g c pb a) + ( map-fiber-vertical-map-cone f g c a , + is-fiberwise-equiv-map-fiber-vertical-map-cone-is-pullback f g c pb a) ( ac (f a)) module _ diff --git a/src/synthetic-homotopy-theory/universal-property-pushouts.lagda.md b/src/synthetic-homotopy-theory/universal-property-pushouts.lagda.md index 03cdadd271..5c419dde39 100644 --- a/src/synthetic-homotopy-theory/universal-property-pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/universal-property-pushouts.lagda.md @@ -304,11 +304,12 @@ is-equiv-universal-property-pushout : {l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} {C : UU l4} (f : S → A) (g : S → B) (c : cocone f g C) → is-equiv f → - ({l : Level} → universal-property-pushout l f g c) → is-equiv (pr1 (pr2 c)) + ({l : Level} → universal-property-pushout l f g c) → + is-equiv (vertical-map-cocone f g c) is-equiv-universal-property-pushout f g (i , j , H) is-equiv-f up-c = is-equiv-is-equiv-precomp j ( λ T → - is-equiv-is-pullback' + is-equiv-horizontal-map-is-pullback ( _∘ f) ( _∘ g) ( cone-pullback-property-pushout f g (i , j , H) T) @@ -336,14 +337,14 @@ pr2 (equiv-universal-property-pushout e g c up-c) = universal-property-pushout-is-equiv : {l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} {C : UU l4} (f : S → A) (g : S → B) (c : cocone f g C) → - is-equiv f → is-equiv (pr1 (pr2 c)) → + is-equiv f → is-equiv (vertical-map-cocone f g c) → ({l : Level} → universal-property-pushout l f g c) universal-property-pushout-is-equiv f g (i , j , H) is-equiv-f is-equiv-j {l} = let c = (i , j , H) in universal-property-pushout-pullback-property-pushout l f g c ( λ T → - is-pullback-is-equiv' + is-pullback-is-equiv-horizontal-maps ( _∘ f) ( _∘ g) ( cone-pullback-property-pushout f g c T) @@ -364,7 +365,7 @@ is-equiv-universal-property-pushout' f g c is-equiv-g up-c = is-equiv-is-equiv-precomp ( horizontal-map-cocone f g c) ( λ T → - is-equiv-is-pullback + is-equiv-vertical-map-is-pullback ( precomp f T) ( precomp g T) ( cone-pullback-property-pushout f g c T) @@ -394,7 +395,7 @@ universal-property-pushout-is-equiv' f g (i , j , H) is-equiv-g is-equiv-i {l} = let c = (i , j , H) in universal-property-pushout-pullback-property-pushout l f g c ( λ T → - is-pullback-is-equiv + is-pullback-is-equiv-vertical-maps ( precomp f T) ( precomp g T) ( cone-pullback-property-pushout f g c T) diff --git a/src/trees/w-types.lagda.md b/src/trees/w-types.lagda.md index 63d2d3dad3..31d1a29151 100644 --- a/src/trees/w-types.lagda.md +++ b/src/trees/w-types.lagda.md @@ -18,6 +18,7 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.homotopy-induction open import foundation.identity-types +open import foundation.postcomposition-functions open import foundation.propositional-truncations open import foundation.sets open import foundation.torsorial-type-families @@ -294,7 +295,7 @@ compute-structure-htpy-hom-𝕎-Alg : ( htpy-polynomial-endofunctor A B H (pair x α))) = ( ap ( λ t → structure-algebra-polynomial-endofunctor X (pair x t)) - ( eq-htpy (H ·r α))) + ( htpy-postcomp (B x) H α)) compute-structure-htpy-hom-𝕎-Alg {A = A} {B} X x α = ind-htpy ( map-hom-𝕎-Alg X) @@ -304,7 +305,7 @@ compute-structure-htpy-hom-𝕎-Alg {A = A} {B} X x α = ( htpy-polynomial-endofunctor A B H (pair x α))) = ( ap ( λ t → structure-algebra-polynomial-endofunctor X (pair x t)) - ( eq-htpy (H ·r α)))) + ( htpy-postcomp (B x) H α))) ( ap ( ap (pr2 X)) ( coh-refl-htpy-polynomial-endofunctor A B diff --git a/src/univalent-combinatorics/decidable-subtypes.lagda.md b/src/univalent-combinatorics/decidable-subtypes.lagda.md index 66b79ec69f..3717f68b97 100644 --- a/src/univalent-combinatorics/decidable-subtypes.lagda.md +++ b/src/univalent-combinatorics/decidable-subtypes.lagda.md @@ -189,8 +189,7 @@ is-decidable-subtype-is-finite-has-decidable-eq S dec-A fin-S a = ( fin-S) ( is-decidable-Prop (S a)) ( λ count-S → - ind-coprod - ( λ _ → type-Prop (is-decidable-Prop (S a))) + rec-coprod ( λ x → inl (tr (type-Prop ∘ S) (inv (pr2 x)) (pr2 (pr1 x)))) ( λ x → inr λ S-a → x (( (a , S-a) , refl))) ( is-decidable-Σ-count count-S λ s → dec-A a (pr1 s))) diff --git a/src/univalent-combinatorics/pi-finite-types.lagda.md b/src/univalent-combinatorics/pi-finite-types.lagda.md index e8c6f6a032..199c63f019 100644 --- a/src/univalent-combinatorics/pi-finite-types.lagda.md +++ b/src/univalent-combinatorics/pi-finite-types.lagda.md @@ -703,7 +703,7 @@ module _ where map-is-coprod-codomain : (im (f ∘ inl) + im (f ∘ inr)) → B - map-is-coprod-codomain = ind-coprod (λ x → B) pr1 pr1 + map-is-coprod-codomain = rec-coprod pr1 pr1 triangle-is-coprod-codomain : ( ( map-is-coprod-codomain) ∘ @@ -864,7 +864,7 @@ abstract ( right-distributive-Σ-coprod ( im (f ∘ inl)) ( im (f ∘ inr)) - ( ind-coprod (λ x → UU l2) (B ∘ pr1) (B ∘ pr1)))) + ( rec-coprod (B ∘ pr1) (B ∘ pr1)))) i : Fin k → type-trunc-Set (im (f ∘ inl)) i = unit-trunc-Set ∘ map-unit-im (f ∘ inl) is-surjective-i : is-surjective i diff --git a/tables/pullbacks.md b/tables/pullbacks.md index 257eb4584c..f5951ec94c 100644 --- a/tables/pullbacks.md +++ b/tables/pullbacks.md @@ -10,6 +10,7 @@ | Pullbacks (foundation) | [`foundation.pullbacks`](foundation.pullbacks.md) | | Functoriality of pullbacks | [`foundation.functoriality-pullbacks`](foundation.functoriality-pullbacks.md) | | Pullback squares | [`foundation.pullback-squares`](foundation.pullback-squares.md) | +| Cartesian morphisms of arrows | [`foundation.cartesian-morphisms-arrows`](foundation.cartesian-morphisms-arrows.md) | | The pullback-hom | [`orthogonal-factorization-systems.pullback-hom`](orthogonal-factorization-systems.pullback-hom.md) | | Functoriality of the pullback-hom | [`orthogonal-factorization-systems.functoriality-pullback-hom`](orthogonal-factorization-systems.functoriality-pullback-hom.md) | | Pullbacks in precategories | [`category-theory.pullbacks-in-precategories`](category-theory.pullbacks-in-precategories.md) |