diff --git a/agda-unimath.agda-lib b/agda-unimath.agda-lib index d57349cb40..5b2757d7fb 100644 --- a/agda-unimath.agda-lib +++ b/agda-unimath.agda-lib @@ -1,3 +1,3 @@ name: agda-unimath include: src -flags: --without-K --exact-split --no-import-sorts --auto-inline -WnoWithoutKFlagPrimEraseEquality +flags: --without-K --exact-split --no-import-sorts --auto-inline --confluence-check -WnoWithoutKFlagPrimEraseEquality diff --git a/references.bib b/references.bib index c2d2c7db82..475c40dddc 100644 --- a/references.bib +++ b/references.bib @@ -338,7 +338,7 @@ @article{MP87 doi = {10.1016/0021-8693(87)90046-9} } -@online{MR23, +@online{MR23a, title = {Delooping the Sign Homomorphism in Univalent Mathematics}, author = {Mangel, Éléonore and Rijke, Egbert}, date = {2023-01-24}, @@ -352,6 +352,19 @@ @online{MR23 keywords = {20B30 03B15,Mathematics - Group Theory,Mathematics - Logic} } +@online{MR23b, + title = {Commuting {{Cohesions}}}, + author = {Myers, David Jaz and Riley, Mitchell}, + date = {2023-01-31}, + month = {01}, + year = {2023}, + eprint = {2301.13780}, + eprinttype = {arxiv}, + eprintclass = {math}, + abstract = {Shulman's spatial type theory internalizes the modalities of Lawvere's axiomatic cohesion in a homotopy type theory, enabling many of the constructions from Schreiber's modal approach to differential cohomology to be carried out synthetically. In spatial type theory, every type carries a spatial cohesion among its points and every function is continuous with respect to this. But in mathematical practice, objects may be spatial in more than one way at the same time; for example, a simplicial space has both topological and simplicial structures. In this paper, we put forward a type theory with "commuting focuses" which allows for types to carry multiple kinds of spatial structure. The theory is a relatively painless extension of spatial type theory, and enables us to give a synthetic account of simplicial, differential, equivariant, and other cohesions carried by the same types. We demonstrate the theory by showing that the homotopy type of any differential stack may be computed from a discrete simplicial set derived from the \textbackslash v\{C\}ech nerve of any good cover. We also give other examples of commuting cohesions, such as differential equivariant types and supergeometric types, laying the groundwork for a synthetic account of Schreiber and Sati's proper orbifold cohomology.}, + langid = {english} +} + @book{MRR88, title = {A {{Course}} in {{Constructive Algebra}}}, author = {Mines, Ray and Richman, Fred and Ruitenburg, Wim}, @@ -452,6 +465,25 @@ @online{Rij22draft pubstate = {draft} } +@misc{RS17, + title = {A type theory for synthetic $\infty$-categories}, + author = {Riehl, Emily and Shulman, Michael}, + year = {2017}, + shortjournal = {High. Struct.}, + journal = {Higher Structures}, + volume = {1}, + year = {2017}, + number = {1}, + pages = {147--224}, + issn = {2209-0606}, + eprint = {1705.07442}, + eprinttype = {arxiv}, + eprintclass = {math}, + archiveprefix = {arXiv}, + primaryclass = {math.CT}, + url = {https://emilyriehl.github.io/files/synthetic.pdf} +} + @article{RSS20, title = {Modalities in Homotopy Type Theory}, author = {Rijke, Egbert and Shulman, Michael and Spitters, Bas}, diff --git a/src/finite-group-theory/cartier-delooping-sign-homomorphism.lagda.md b/src/finite-group-theory/cartier-delooping-sign-homomorphism.lagda.md index 78b0a6dcbc..9d2a1a90e4 100644 --- a/src/finite-group-theory/cartier-delooping-sign-homomorphism.lagda.md +++ b/src/finite-group-theory/cartier-delooping-sign-homomorphism.lagda.md @@ -202,4 +202,4 @@ module _ ## References -{{#bibliography}} {{#reference MR23}} +{{#bibliography}} {{#reference MR23a}} diff --git a/src/finite-group-theory/delooping-sign-homomorphism.lagda.md b/src/finite-group-theory/delooping-sign-homomorphism.lagda.md index f9be0f8f1b..13be992da0 100644 --- a/src/finite-group-theory/delooping-sign-homomorphism.lagda.md +++ b/src/finite-group-theory/delooping-sign-homomorphism.lagda.md @@ -1598,4 +1598,4 @@ module _ ## References -{{#bibliography}} {{#reference MR23}} +{{#bibliography}} {{#reference MR23a}} diff --git a/src/finite-group-theory/simpson-delooping-sign-homomorphism.lagda.md b/src/finite-group-theory/simpson-delooping-sign-homomorphism.lagda.md index 479553f8ba..d3c80b441c 100644 --- a/src/finite-group-theory/simpson-delooping-sign-homomorphism.lagda.md +++ b/src/finite-group-theory/simpson-delooping-sign-homomorphism.lagda.md @@ -800,4 +800,4 @@ module _ ## References -{{#bibliography}} {{#reference MR23}} +{{#bibliography}} {{#reference MR23a}} diff --git a/src/foundation-core/truncated-types.lagda.md b/src/foundation-core/truncated-types.lagda.md index ef0e87f4a0..463283fbaa 100644 --- a/src/foundation-core/truncated-types.lagda.md +++ b/src/foundation-core/truncated-types.lagda.md @@ -20,7 +20,9 @@ open import foundation-core.equality-dependent-pair-types open import foundation-core.equivalences open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.invertible-maps open import foundation-core.propositions +open import foundation-core.retractions open import foundation-core.retracts-of-types open import foundation-core.transport-along-identifications open import foundation-core.truncation-levels @@ -75,8 +77,7 @@ module _ abstract is-trunc-succ-is-trunc : (k : 𝕋) {l : Level} {A : UU l} → is-trunc k A → is-trunc (succ-𝕋 k) A - pr1 (is-trunc-succ-is-trunc neg-two-𝕋 H x y) = eq-is-contr H - pr2 (is-trunc-succ-is-trunc neg-two-𝕋 H x .x) refl = left-inv (pr2 H x) + is-trunc-succ-is-trunc neg-two-𝕋 = is-prop-is-contr is-trunc-succ-is-trunc (succ-𝕋 k) H x y = is-trunc-succ-is-trunc k (H x y) truncated-type-succ-Truncated-Type : @@ -121,11 +122,11 @@ module _ where is-trunc-retract-of : - {k : 𝕋} {A : UU l1} {B : UU l2} → + (k : 𝕋) {A : UU l1} {B : UU l2} → A retract-of B → is-trunc k B → is-trunc k A - is-trunc-retract-of {neg-two-𝕋} = is-contr-retract-of _ - is-trunc-retract-of {succ-𝕋 k} R H x y = - is-trunc-retract-of (retract-eq R x y) (H (pr1 R x) (pr1 R y)) + is-trunc-retract-of neg-two-𝕋 = is-contr-retract-of _ + is-trunc-retract-of (succ-𝕋 k) R H x y = + is-trunc-retract-of k (retract-eq R x y) (H (pr1 R x) (pr1 R y)) ``` ### `k`-truncated types are closed under equivalences @@ -136,7 +137,7 @@ abstract {l1 l2 : Level} (k : 𝕋) {A : UU l1} (B : UU l2) (f : A → B) → is-equiv f → is-trunc k B → is-trunc k A is-trunc-is-equiv k B f is-equiv-f = - is-trunc-retract-of (pair f (pr2 is-equiv-f)) + is-trunc-retract-of k (pair f (pr2 is-equiv-f)) abstract is-trunc-equiv : @@ -180,6 +181,18 @@ abstract is-trunc-emb k f = is-trunc-is-emb k (map-emb f) (is-emb-map-emb f) ``` +In fact, it suffices that the map's action on identifications has a retraction. + +```agda +abstract + is-trunc-retraction-ap : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (f : A → B) → + ((x y : A) → retraction (ap f {x} {y})) → + is-trunc (succ-𝕋 k) B → is-trunc (succ-𝕋 k) A + is-trunc-retraction-ap k f Ef H x y = + is-trunc-retract-of k (ap f , Ef x y) (H (f x) (f y)) +``` + ### Truncated types are closed under dependent pair types ```agda @@ -225,7 +238,7 @@ abstract {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → is-trunc k A → is-trunc k B → is-trunc k (A × B) is-trunc-product k is-trunc-A is-trunc-B = - is-trunc-Σ is-trunc-A (λ x → is-trunc-B) + is-trunc-Σ is-trunc-A (λ _ → is-trunc-B) product-Truncated-Type : {l1 l2 : Level} (k : 𝕋) → @@ -236,37 +249,43 @@ pr2 (product-Truncated-Type k A B) = is-trunc-product k ( is-trunc-type-Truncated-Type A) ( is-trunc-type-Truncated-Type B) +``` + +We need only show that each factor is `k`-truncated given that the opposite +factor has an element when `k ≥ -1`. +```agda is-trunc-product' : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → - (B → is-trunc (succ-𝕋 k) A) → (A → is-trunc (succ-𝕋 k) B) → + (B → is-trunc (succ-𝕋 k) A) → + (A → is-trunc (succ-𝕋 k) B) → is-trunc (succ-𝕋 k) (A × B) is-trunc-product' k f g (pair a b) (pair a' b') = is-trunc-equiv k - ( Eq-product (pair a b) (pair a' b')) + ( Eq-product (a , b) (pair a' b')) ( equiv-pair-eq (pair a b) (pair a' b')) ( is-trunc-product k (f b a a') (g a b b')) -is-trunc-left-factor-product : +is-trunc-left-factor-product' : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → is-trunc k (A × B) → B → is-trunc k A -is-trunc-left-factor-product neg-two-𝕋 {A} {B} H b = +is-trunc-left-factor-product' neg-two-𝕋 {A} {B} H b = is-contr-left-factor-product A B H -is-trunc-left-factor-product (succ-𝕋 k) H b a a' = - is-trunc-left-factor-product k {A = (a = a')} {B = (b = b)} +is-trunc-left-factor-product' (succ-𝕋 k) H b a a' = + is-trunc-left-factor-product' k {A = (a = a')} {B = (b = b)} ( is-trunc-equiv' k ( pair a b = pair a' b) ( equiv-pair-eq (pair a b) (pair a' b)) ( H (pair a b) (pair a' b))) ( refl) -is-trunc-right-factor-product : +is-trunc-right-factor-product' : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → is-trunc k (A × B) → A → is-trunc k B -is-trunc-right-factor-product neg-two-𝕋 {A} {B} H a = +is-trunc-right-factor-product' neg-two-𝕋 {A} {B} H a = is-contr-right-factor-product A B H -is-trunc-right-factor-product (succ-𝕋 k) {A} {B} H a b b' = - is-trunc-right-factor-product k {A = (a = a)} {B = (b = b')} +is-trunc-right-factor-product' (succ-𝕋 k) {A} {B} H a b b' = + is-trunc-right-factor-product' k {A = (a = a)} {B = (b = b')} ( is-trunc-equiv' k ( pair a b = pair a b') ( equiv-pair-eq (pair a b) (pair a b')) @@ -334,7 +353,7 @@ abstract {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → is-trunc k B → is-trunc k (A → B) is-trunc-function-type k {A} {B} is-trunc-B = - is-trunc-Π k {B = λ (x : A) → B} (λ x → is-trunc-B) + is-trunc-Π k {B = λ (x : A) → B} (λ _ → is-trunc-B) function-type-Truncated-Type : {l1 l2 : Level} {k : 𝕋} (A : UU l1) (B : Truncated-Type l2 k) → @@ -402,7 +421,12 @@ module _ ( is-trunc-function-type k H) ( λ h → is-trunc-Π k (λ x → is-trunc-Id H (h (f x)) x)))) +``` + +Alternatively, this follows from the fact that equivalences embed into function +types, and function types between `k`-truncated types are `k`-truncated. +```agda type-equiv-Truncated-Type : {l1 l2 : Level} {k : 𝕋} (A : Truncated-Type l1 k) (B : Truncated-Type l2 k) → UU (l1 ⊔ l2) diff --git a/src/foundation-core/univalence.lagda.md b/src/foundation-core/univalence.lagda.md index 5ac8ad328f..294e0f3faa 100644 --- a/src/foundation-core/univalence.lagda.md +++ b/src/foundation-core/univalence.lagda.md @@ -105,8 +105,8 @@ abstract is-torsorial-equiv-based-univalence : {l : Level} (A : UU l) → based-univalence-axiom A → is-torsorial (λ (B : UU l) → A ≃ B) - is-torsorial-equiv-based-univalence A UA = - fundamental-theorem-id' (λ B → equiv-eq) UA + is-torsorial-equiv-based-univalence A = + fundamental-theorem-id' (λ B → equiv-eq) ``` ### Contractibility of the total space of equivalences implies univalence diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index c67a0410cf..2bd56638e8 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -39,8 +39,15 @@ open import foundation-core.truncation-levels ## Idea -A type is said to be connected if its type of connected components, i.e., its -set truncation, is contractible. +A type is said to be +{{#concept "0-connected" Disambiguation="type" Agda=is-0-connected}} if its type +of [connected components](foundation.connected-components.md), i.e., its +[set truncation](foundation.set-truncations.md), is +[contractible](foundation-core.contractible-types.md). + +## Definitions + +### The predicate on types of being 0-connected ```agda is-0-connected-Prop : {l : Level} → UU l → Prop l @@ -66,7 +73,13 @@ abstract {l : Level} {A : UU l} → is-0-connected A → (x y : A) → mere-eq x y mere-eq-is-0-connected {A = A} H x y = apply-effectiveness-unit-trunc-Set (eq-is-contr H) +``` +## Properties + +### A type is 0-connected if there is an element of that type such that every element is merely equal to it + +```agda abstract is-0-connected-mere-eq : {l : Level} {A : UU l} (a : A) → @@ -77,7 +90,11 @@ abstract ( apply-dependent-universal-property-trunc-Set' ( λ x → set-Prop (Id-Prop (trunc-Set A) (unit-trunc-Set a) x)) ( λ x → apply-effectiveness-unit-trunc-Set' (e x))) +``` + +### A type is 0-connected if it is inhabited and all elements are merely equal +```agda abstract is-0-connected-mere-eq-is-inhabited : {l : Level} {A : UU l} → @@ -86,7 +103,11 @@ abstract apply-universal-property-trunc-Prop H ( is-0-connected-Prop _) ( λ a → is-0-connected-mere-eq a (K a)) +``` + +### A type is is 0-connected iff there is a point inclusion which is surjective +```agda is-0-connected-is-surjective-point : {l1 : Level} {A : UU l1} (a : A) → is-surjective (point a) → is-0-connected A @@ -107,10 +128,15 @@ abstract ( mere-eq-is-0-connected H a x) ( trunc-Prop (fiber (point a) x)) ( λ where refl → unit-trunc-Prop (star , refl)) +``` + +### The evaluation map at a point of a 0-connected type into a `k+1`-truncated type is `k`-truncated +```agda is-trunc-map-ev-point-is-connected : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (a : A) → - is-0-connected A → is-trunc (succ-𝕋 k) B → + is-0-connected A → + is-trunc (succ-𝕋 k) B → is-trunc-map k (ev-point' a {B}) is-trunc-map-ev-point-is-connected k {A} {B} a H K = is-trunc-map-comp k @@ -121,11 +147,14 @@ is-trunc-map-ev-point-is-connected k {A} {B} a H K = ( is-trunc-map-precomp-Π-is-surjective k ( is-surjective-point-is-0-connected a H) ( λ _ → (B , K))) +``` +### 0-connected types satisfy the dependent universal property of 0-connected types + +```agda equiv-dependent-universal-property-is-0-connected : {l1 : Level} {A : UU l1} (a : A) → is-0-connected A → - ( {l : Level} (P : A → Prop l) → - ((x : A) → type-Prop (P x)) ≃ type-Prop (P a)) + {l : Level} (P : A → Prop l) → ((x : A) → type-Prop (P x)) ≃ type-Prop (P a) equiv-dependent-universal-property-is-0-connected a H P = ( equiv-universal-property-unit (type-Prop (P a))) ∘e ( equiv-dependent-universal-property-surjection-is-surjective @@ -167,19 +196,24 @@ abstract is-0-connected A is-0-connected-is-surjective-fiber-inclusion a H = is-0-connected-mere-eq a (mere-eq-is-surjective-fiber-inclusion a H) +``` +### 0-connected types are closed under equivalences + +```agda is-0-connected-equiv : {l1 l2 : Level} {A : UU l1} {B : UU l2} → - (A ≃ B) → is-0-connected B → is-0-connected A -is-0-connected-equiv e = is-contr-equiv _ (equiv-trunc-Set e) + A ≃ B → is-0-connected B → is-0-connected A +is-0-connected-equiv {B = B} e = + is-contr-equiv (type-trunc-Set B) (equiv-trunc-Set e) is-0-connected-equiv' : {l1 l2 : Level} {A : UU l1} {B : UU l2} → - (A ≃ B) → is-0-connected A → is-0-connected B + A ≃ B → is-0-connected A → is-0-connected B is-0-connected-equiv' e = is-0-connected-equiv (inv-equiv e) ``` -### `0-connected` types are closed under cartesian products +### 0-connected types are closed under cartesian products ```agda module _ @@ -195,12 +229,11 @@ module _ ( is-contr-product p1 p2) ``` -### A contractible type is `0`-connected +### Contractible types are 0-connected ```agda is-0-connected-is-contr : - {l : Level} (X : UU l) → - is-contr X → is-0-connected X + {l : Level} (X : UU l) → is-contr X → is-0-connected X is-0-connected-is-contr X p = is-contr-equiv X (inv-equiv (equiv-unit-trunc-Set (X , is-set-is-contr p))) p ``` diff --git a/src/foundation/action-on-identifications-functions.lagda.md b/src/foundation/action-on-identifications-functions.lagda.md index 5e4162ad9d..215e151c74 100644 --- a/src/foundation/action-on-identifications-functions.lagda.md +++ b/src/foundation/action-on-identifications-functions.lagda.md @@ -100,7 +100,7 @@ ap-inv f refl = refl ```agda ap-const : {l1 l2 : Level} {A : UU l1} {B : UU l2} (b : B) {x y : A} - (p : x = y) → (ap (const A b) p) = refl + (p : x = y) → ap (const A b) p = refl ap-const b refl = refl ``` diff --git a/src/foundation/booleans.lagda.md b/src/foundation/booleans.lagda.md index 9b69465c3e..1f13b05d6a 100644 --- a/src/foundation/booleans.lagda.md +++ b/src/foundation/booleans.lagda.md @@ -65,6 +65,17 @@ module _ ind-bool pt pf false = pf ``` +### The recursion principle of the booleans + +```agda +module _ + {l : Level} {A : UU l} + where + + rec-bool : A → A → bool → A + rec-bool = ind-bool (λ _ → A) +``` + ### The `if_then_else` function ```agda @@ -73,8 +84,7 @@ module _ where if_then_else_ : bool → A → A → A - if true then x else y = x - if false then x else y = y + if b then x else y = rec-bool x y b ``` ### Raising universe levels of the booleans diff --git a/src/foundation/cartesian-morphisms-arrows.lagda.md b/src/foundation/cartesian-morphisms-arrows.lagda.md index 8750d63496..21b04fdb0b 100644 --- a/src/foundation/cartesian-morphisms-arrows.lagda.md +++ b/src/foundation/cartesian-morphisms-arrows.lagda.md @@ -37,6 +37,7 @@ open import foundation.whiskering-homotopies-composition open import foundation-core.commuting-squares-of-maps open import foundation-core.homotopies open import foundation-core.propositions +open import foundation-core.sections open import foundation-core.universal-property-pullbacks ``` @@ -1039,6 +1040,79 @@ module _ ( is-cartesian-cartesian-hom-arrow-lift-map-codomain-cartesian-hom-arrow) ``` +### Base change of sections + +Given a cartesian morphism of arrows + +```text + A ------> X + | ⌟ | + f | | g + ∨ ∨ + B ------> Y + j +``` + +then if `g` has a section so does `f`. More generally, for every map `s` such +that `g ∘ s ∘ j ~ j`, there exists a section of `f`. + +```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} + (α : cartesian-hom-arrow f g) + (s : Y → X) + (H : + map-codomain-cartesian-hom-arrow f g α ~ + g ∘ s ∘ map-codomain-cartesian-hom-arrow f g α) + where + + cone-section-base-change' : + cone (map-codomain-cartesian-hom-arrow f g α) g B + cone-section-base-change' = + ( id , s ∘ map-codomain-cartesian-hom-arrow f g α , H) + + map-section-base-change' : + B → A + map-section-base-change' = + gap-is-pullback + ( map-codomain-cartesian-hom-arrow f g α) + ( g) + ( cone-cartesian-hom-arrow f g α) + ( is-cartesian-cartesian-hom-arrow f g α) + ( cone-section-base-change') + + is-section-map-section-base-change' : + is-section f map-section-base-change' + is-section-map-section-base-change' = + htpy-vertical-map-gap-is-pullback + ( map-codomain-cartesian-hom-arrow f g α) + ( g) + ( cone-cartesian-hom-arrow f g α) + ( is-cartesian-cartesian-hom-arrow f g α) + ( cone-section-base-change') + + section-base-change' : section f + section-base-change' = + ( map-section-base-change' , is-section-map-section-base-change') + +module _ + {l1 l2 l3 l4 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + {f : A → B} {g : X → Y} + (α : cartesian-hom-arrow f g) + (s : section g) + where + + section-base-change : section f + section-base-change = + section-base-change' α + ( map-section g s) + ( ( inv-htpy (is-section-map-section g s)) ·r + ( map-codomain-cartesian-hom-arrow f g α)) +``` + ## See also - [Cocartesian morphisms of arrows](synthetic-homotopy-theory.cocartesian-morphisms-arrows.md) diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index 193097e706..967334a223 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -34,7 +34,10 @@ open import foundation-core.truncation-levels ## Idea -A type is said to be **`k`-connected** if its `k`-truncation is contractible. +A type is said to be +{{#concept "`k`-connected" Disambiguation="type" Agda=is-connected}} if its +`k`-[truncation](foundation.truncations.md) is +[contractible](foundation-core.contractible-types.md). ## Definition @@ -191,8 +194,7 @@ module _ ( map-trunc k f) ( is-equiv-map-equiv-trunc k (f , e)) - is-connected-equiv : - A ≃ B → is-connected k B → is-connected k A + is-connected-equiv : A ≃ B → is-connected k B → is-connected k A is-connected-equiv f = is-connected-is-equiv (map-equiv f) (is-equiv-map-equiv f) @@ -217,9 +219,7 @@ module _ where is-connected-retract-of : - A retract-of B → - is-connected k B → - is-connected k A - is-connected-retract-of R c = - is-contr-retract-of (type-trunc k B) (retract-of-trunc-retract-of R) c + A retract-of B → is-connected k B → is-connected k A + is-connected-retract-of R = + is-contr-retract-of (type-trunc k B) (retract-of-trunc-retract-of R) ``` diff --git a/src/foundation/constant-type-families.lagda.md b/src/foundation/constant-type-families.lagda.md index 58fd987d4a..15db228460 100644 --- a/src/foundation/constant-type-families.lagda.md +++ b/src/foundation/constant-type-families.lagda.md @@ -10,13 +10,13 @@ module foundation.constant-type-families where open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types +open import foundation.identity-types open import foundation.transport-along-identifications open import foundation.universe-levels open import foundation-core.commuting-squares-of-identifications open import foundation-core.dependent-identifications open import foundation-core.equivalences -open import foundation-core.identity-types ``` @@ -76,6 +76,26 @@ tr-constant-type-family : tr-constant-type-family refl b = refl ``` +### Computing dependent identifications in constant type families + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + map-compute-dependent-identification-constant-type-family : + {x y : A} (p : x = y) {x' y' : B} → + x' = y' → dependent-identification (λ _ → B) p x' y' + map-compute-dependent-identification-constant-type-family p {x'} q = + tr-constant-type-family p x' ∙ q + + compute-dependent-identification-constant-type-family : + {x y : A} (p : x = y) {x' y' : B} → + (x' = y') ≃ dependent-identification (λ _ → B) p x' y' + compute-dependent-identification-constant-type-family p {x'} {y'} = + equiv-concat (tr-constant-type-family p x') y' +``` + ### Dependent action on paths of sections of standard constant type families ```agda diff --git a/src/foundation/dependent-identifications.lagda.md b/src/foundation/dependent-identifications.lagda.md index fc4b128c5d..e0406f4211 100644 --- a/src/foundation/dependent-identifications.lagda.md +++ b/src/foundation/dependent-identifications.lagda.md @@ -9,8 +9,10 @@ open import foundation-core.dependent-identifications public
Imports ```agda +open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types +open import foundation.strictly-right-unital-concatenation-identifications open import foundation.transport-along-higher-identifications open import foundation.universe-levels @@ -26,7 +28,7 @@ open import foundation-core.transport-along-identifications ## Idea We characterize dependent paths in the family of depedent paths; define the -groupoidal operators on dependent paths; define the cohrences paths: prove the +groupoidal operators on dependent paths; define the coherence paths; prove the operators are equivalences. ## Properites @@ -120,12 +122,34 @@ module _ dependent-identification B (p ∙ q) x' z' concat-dependent-identification refl q refl q' = q' - compute-concat-dependent-identification-refl : + compute-concat-dependent-identification-left-base-refl : { y z : A} (q : y = z) → { x' y' : B y} {z' : B z} (p' : x' = y') → ( q' : dependent-identification B q y' z') → - ( concat-dependent-identification refl q p' q') = ap (tr B q) p' ∙ q' - compute-concat-dependent-identification-refl refl refl q' = refl + concat-dependent-identification refl q p' q' = ap (tr B q) p' ∙ q' + compute-concat-dependent-identification-left-base-refl q refl q' = refl +``` + +#### Strictly right unital concatenation of dependent identifications + +```agda +module _ + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) + where + + right-strict-concat-dependent-identification : + {x y z : A} (p : x = y) (q : y = z) {x' : B x} {y' : B y} {z' : B z} → + dependent-identification B p x' y' → + dependent-identification B q y' z' → + dependent-identification B (p ∙ᵣ q) x' z' + right-strict-concat-dependent-identification p refl p' q' = p' ∙ᵣ q' + + compute-right-strict-concat-dependent-identification-refl : + { x y : A} (p : x = y) → + { x' : B x} {y' z' : B y} (p' : dependent-identification B p x' y') → + ( q' : y' = z') → + right-strict-concat-dependent-identification p refl p' q' = p' ∙ᵣ q' + compute-right-strict-concat-dependent-identification-refl q p' q' = refl ``` #### Inverses of dependent identifications diff --git a/src/foundation/disjunction.lagda.md b/src/foundation/disjunction.lagda.md index 21a764be49..996a1f345c 100644 --- a/src/foundation/disjunction.lagda.md +++ b/src/foundation/disjunction.lagda.md @@ -7,6 +7,7 @@ module foundation.disjunction where
Imports ```agda +open import foundation.coproduct-types open import foundation.decidable-types open import foundation.dependent-pair-types open import foundation.functoriality-coproduct-types @@ -16,11 +17,11 @@ open import foundation.propositional-truncations open import foundation.universe-levels open import foundation-core.cartesian-product-types -open import foundation-core.coproduct-types open import foundation-core.decidable-propositions open import foundation-core.empty-types open import foundation-core.equivalences open import foundation-core.function-types +open import foundation-core.negation open import foundation-core.propositions ``` @@ -335,6 +336,84 @@ module _ ( is-decidable-Decidable-Prop Q)) ``` +### The disjunction of mutually exclusive types + +If two propositions are mutually exclusive, then their disjunction is equivalent +to the coproduct of their underlying types + +```text + P ∨ Q = P + Q. +``` + +```agda +module _ + {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) + (f : type-Prop P → ¬ (type-Prop Q)) + where + + map-compute-disjunction-mutually-exclusive : + type-disjunction-Prop P Q → type-Prop P + type-Prop Q + map-compute-disjunction-mutually-exclusive = + elim-disjunction (coproduct-Prop P Q f) inl inr + + compute-disjunction-mutually-exclusive : + type-disjunction-Prop P Q ↔ type-Prop P + type-Prop Q + compute-disjunction-mutually-exclusive = + ( map-compute-disjunction-mutually-exclusive , unit-trunc-Prop) + + inv-compute-disjunction-mutually-exclusive : + type-Prop P + type-Prop Q ↔ type-disjunction-Prop P Q + inv-compute-disjunction-mutually-exclusive = + inv-iff compute-disjunction-mutually-exclusive + + equiv-compute-disjunction-mutually-exclusive : + type-disjunction-Prop P Q ≃ type-Prop P + type-Prop Q + equiv-compute-disjunction-mutually-exclusive = + equiv-iff' + ( disjunction-Prop P Q) + ( coproduct-Prop P Q f) + ( compute-disjunction-mutually-exclusive) + +module _ + {l1 l2 : Level} (A : UU l1) (B : UU l2) + (f : ║ A ║₋₁ → ¬ ║ B ║₋₁) + where + + map-compute-disjunction-type-mutually-exclusive : + disjunction-type A B → ║ A ║₋₁ + ║ B ║₋₁ + map-compute-disjunction-type-mutually-exclusive = + elim-disjunction + ( coproduct-Prop (trunc-Prop A) (trunc-Prop B) f) + ( inl ∘ unit-trunc-Prop) + ( inr ∘ unit-trunc-Prop) + + map-inv-compute-disjunction-type-mutually-exclusive : + ║ A ║₋₁ + ║ B ║₋₁ → disjunction-type A B + map-inv-compute-disjunction-type-mutually-exclusive (inl a) = + rec-trunc-Prop (disjunction-type-Prop A B) inl-disjunction a + map-inv-compute-disjunction-type-mutually-exclusive (inr b) = + rec-trunc-Prop (disjunction-type-Prop A B) inr-disjunction b + + compute-disjunction-type-mutually-exclusive : + disjunction-type A B ↔ (║ A ║₋₁ + ║ B ║₋₁) + compute-disjunction-type-mutually-exclusive = + ( map-compute-disjunction-type-mutually-exclusive , + map-inv-compute-disjunction-type-mutually-exclusive) + + inv-compute-disjunction-type-mutually-exclusive : + (║ A ║₋₁ + ║ B ║₋₁) ↔ disjunction-type A B + inv-compute-disjunction-type-mutually-exclusive = + inv-iff compute-disjunction-type-mutually-exclusive + + equiv-compute-disjunction-type-mutually-exclusive : + disjunction-type A B ≃ (║ A ║₋₁ + ║ B ║₋₁) + equiv-compute-disjunction-type-mutually-exclusive = + equiv-iff' + ( disjunction-type-Prop A B) + ( coproduct-Prop (trunc-Prop A) (trunc-Prop B) f) + ( compute-disjunction-type-mutually-exclusive) +``` + ## See also - The indexed counterpart to disjunction is diff --git a/src/foundation/fixed-points-endofunctions.lagda.md b/src/foundation/fixed-points-endofunctions.lagda.md index 58d4618a9e..40c2520930 100644 --- a/src/foundation/fixed-points-endofunctions.lagda.md +++ b/src/foundation/fixed-points-endofunctions.lagda.md @@ -18,8 +18,13 @@ open import foundation-core.identity-types ## Idea Given an [endofunction](foundation-core.endomorphisms.md) `f : A → A`, the type -of {{#concept "fixed points"}} is the type of elements `x : A` such that -`f x = x`. +of +{{#concept "fixed points" Disambiguation="of an endofunction" Agda=fixed-point}} +is the type of elements `x : A` such that `f x = x` + +```text + fixed-point f := Σ (x : A), (f x = x). +``` ## Definitions diff --git a/src/foundation/identity-types.lagda.md b/src/foundation/identity-types.lagda.md index 8bbd6bea4c..0346f2da7e 100644 --- a/src/foundation/identity-types.lagda.md +++ b/src/foundation/identity-types.lagda.md @@ -21,6 +21,8 @@ open import foundation-core.equivalences open import foundation-core.fibers-of-maps open import foundation-core.function-types open import foundation-core.homotopies +open import foundation-core.retractions +open import foundation-core.sections ```
@@ -362,3 +364,50 @@ module _ pr2 compute-fiber-map-out-of-identity-type = is-equiv-map-compute-fiber-map-out-of-identity-type ``` + +### Computing the total type of identifications + +```text + (Σ (x y : A), (x = y)) ≃ A +``` + +```agda +module _ + {l : Level} {A : UU l} + where + + map-compute-total-Id : A → Σ A (λ x → Σ A (λ y → x = y)) + map-compute-total-Id x = (x , x , refl) + + map-inv-compute-total-Id : Σ A (λ x → Σ A (λ y → x = y)) → A + map-inv-compute-total-Id = pr1 + + is-section-map-inv-compute-total-Id : + is-section map-compute-total-Id map-inv-compute-total-Id + is-section-map-inv-compute-total-Id (x , .x , refl) = refl + + is-retraction-map-inv-compute-total-Id : + is-retraction map-compute-total-Id map-inv-compute-total-Id + is-retraction-map-inv-compute-total-Id = refl-htpy + + is-equiv-map-compute-total-Id : is-equiv map-compute-total-Id + is-equiv-map-compute-total-Id = + is-equiv-is-invertible + ( map-inv-compute-total-Id) + ( is-section-map-inv-compute-total-Id) + ( is-retraction-map-inv-compute-total-Id) + + compute-total-Id : A ≃ Σ A (λ x → Σ A (λ y → x = y)) + compute-total-Id = (map-compute-total-Id , is-equiv-map-compute-total-Id) + + is-equiv-map-inv-compute-total-Id : is-equiv map-inv-compute-total-Id + is-equiv-map-inv-compute-total-Id = + is-equiv-is-invertible + ( map-compute-total-Id) + ( is-retraction-map-inv-compute-total-Id) + ( is-section-map-inv-compute-total-Id) + + inv-compute-total-Id : Σ A (λ x → Σ A (λ y → x = y)) ≃ A + inv-compute-total-Id = + ( map-inv-compute-total-Id , is-equiv-map-inv-compute-total-Id) +``` diff --git a/src/foundation/monomorphisms.lagda.md b/src/foundation/monomorphisms.lagda.md index 1908560489..82a42b94ff 100644 --- a/src/foundation/monomorphisms.lagda.md +++ b/src/foundation/monomorphisms.lagda.md @@ -29,10 +29,11 @@ open import foundation-core.truncation-levels ## Idea -A function `f : A → B` is a monomorphism if whenever we have two functions -`g h : X → A` such that `f ∘ g = f ∘ h`, then in fact `g = h`. The way to state -this in Homotopy Type Theory is to say that postcomposition by `f` is an -embedding. +A function `f : A → B` is a +{{#concept "monomorphism" Disambiguation="of types" Agda=is-mono}} if whenever +we have two functions `g h : X → A` such that `f ∘ g = f ∘ h`, then in fact +`g = h`. The way to state this in Homotopy Type Theory is to say that +postcomposition by `f` is an embedding. ## Definition @@ -43,7 +44,7 @@ module _ where is-mono-Prop : Prop (l1 ⊔ l2 ⊔ lsuc l3) - is-mono-Prop = Π-Prop (UU l3) λ X → is-emb-Prop (postcomp X f) + is-mono-Prop = Π-Prop (UU l3) (λ X → is-emb-Prop (postcomp X f)) is-mono : UU (l1 ⊔ l2 ⊔ lsuc l3) is-mono = type-Prop is-mono-Prop diff --git a/src/foundation/postcomposition-pullbacks.lagda.md b/src/foundation/postcomposition-pullbacks.lagda.md index b24c28bb74..acf493df36 100644 --- a/src/foundation/postcomposition-pullbacks.lagda.md +++ b/src/foundation/postcomposition-pullbacks.lagda.md @@ -184,3 +184,9 @@ module _ The following table lists files that are about pullbacks as a general concept. {{#include tables/pullbacks.md}} + +## See also + +- For the dual property for [pushouts](synthetic-homotopy-theory.pushouts.md), + see + [the pullback property of pushouts](synthetic-homotopy-theory.pullback-property-pushouts.md). diff --git a/src/foundation/propositional-truncations.lagda.md b/src/foundation/propositional-truncations.lagda.md index 1181cc2572..88c7b93f96 100644 --- a/src/foundation/propositional-truncations.lagda.md +++ b/src/foundation/propositional-truncations.lagda.md @@ -284,7 +284,7 @@ module _ abstract dependent-universal-property-trunc-Prop : {l : Level} {A : UU l} → - dependent-universal-property-propositional-truncation + dependent-universal-property-propositional-truncation ( trunc-Prop A) ( unit-trunc-Prop) dependent-universal-property-trunc-Prop {A = A} = diff --git a/src/foundation/propositions.lagda.md b/src/foundation/propositions.lagda.md index 5a9092e007..835e2ada3b 100644 --- a/src/foundation/propositions.lagda.md +++ b/src/foundation/propositions.lagda.md @@ -46,7 +46,7 @@ module _ where is-prop-retract-of : A retract-of B → is-prop B → is-prop A - is-prop-retract-of = is-trunc-retract-of + is-prop-retract-of = is-trunc-retract-of neg-one-𝕋 ``` ### If a type embeds into a proposition, then it is a proposition diff --git a/src/foundation/raising-universe-levels.lagda.md b/src/foundation/raising-universe-levels.lagda.md index 75fcc7d91a..ecfd5103d0 100644 --- a/src/foundation/raising-universe-levels.lagda.md +++ b/src/foundation/raising-universe-levels.lagda.md @@ -71,6 +71,9 @@ compute-raise : (l : Level) {l1 : Level} (A : UU l1) → A ≃ raise l A pr1 (compute-raise l A) = map-raise pr2 (compute-raise l A) = is-equiv-map-raise +inv-compute-raise : (l : Level) {l1 : Level} (A : UU l1) → raise l A ≃ A +inv-compute-raise l A = inv-equiv (compute-raise l A) + Raise : (l : Level) {l1 : Level} (A : UU l1) → Σ (UU (l1 ⊔ l)) (λ X → A ≃ X) pr1 (Raise l A) = raise l A pr2 (Raise l A) = compute-raise l A diff --git a/src/foundation/type-arithmetic-booleans.lagda.md b/src/foundation/type-arithmetic-booleans.lagda.md index 5d7dfaad12..9bbb29dcff 100644 --- a/src/foundation/type-arithmetic-booleans.lagda.md +++ b/src/foundation/type-arithmetic-booleans.lagda.md @@ -9,20 +9,24 @@ module foundation.type-arithmetic-booleans where ```agda open import foundation.booleans open import foundation.dependent-pair-types +open import foundation.function-extensionality open import foundation.universe-levels +open import foundation-core.cartesian-product-types open import foundation-core.coproduct-types 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 ```
## Idea -We prove arithmetical laws involving the booleans +We prove arithmetical laws involving the booleans. ## Laws @@ -34,12 +38,12 @@ module _ where map-Σ-bool-coproduct : Σ bool A → A true + A false - map-Σ-bool-coproduct (pair true a) = inl a - map-Σ-bool-coproduct (pair false a) = inr a + map-Σ-bool-coproduct (true , a) = inl a + map-Σ-bool-coproduct (false , a) = inr a map-inv-Σ-bool-coproduct : A true + A false → Σ bool A - map-inv-Σ-bool-coproduct (inl a) = pair true a - map-inv-Σ-bool-coproduct (inr a) = pair false a + map-inv-Σ-bool-coproduct (inl a) = (true , a) + map-inv-Σ-bool-coproduct (inr a) = (false , a) is-section-map-inv-Σ-bool-coproduct : ( map-Σ-bool-coproduct ∘ map-inv-Σ-bool-coproduct) ~ id @@ -48,8 +52,8 @@ module _ is-retraction-map-inv-Σ-bool-coproduct : ( map-inv-Σ-bool-coproduct ∘ map-Σ-bool-coproduct) ~ id - is-retraction-map-inv-Σ-bool-coproduct (pair true a) = refl - is-retraction-map-inv-Σ-bool-coproduct (pair false a) = refl + is-retraction-map-inv-Σ-bool-coproduct (true , a) = refl + is-retraction-map-inv-Σ-bool-coproduct (false , a) = refl is-equiv-map-Σ-bool-coproduct : is-equiv map-Σ-bool-coproduct is-equiv-map-Σ-bool-coproduct = @@ -73,3 +77,55 @@ module _ pr1 inv-equiv-Σ-bool-coproduct = map-inv-Σ-bool-coproduct pr2 inv-equiv-Σ-bool-coproduct = is-equiv-map-inv-Σ-bool-coproduct ``` + +### Dependent products over the booleans are cartesian products + +This is also the dependent +[universal property of the booleans](foundation.universal-property-booleans.md). + +```agda +module _ + {l : Level} (A : bool → UU l) + where + + map-Π-bool-product : ((b : bool) → A b) → A true × A false + map-Π-bool-product f = (f true , f false) + + map-inv-Π-bool-product : A true × A false → ((b : bool) → A b) + map-inv-Π-bool-product (ft , ff) = ind-bool A ft ff + + is-section-map-inv-Π-bool-product : + is-section map-Π-bool-product map-inv-Π-bool-product + is-section-map-inv-Π-bool-product _ = refl + + abstract + is-retraction-map-inv-Π-bool-product : + is-retraction map-Π-bool-product map-inv-Π-bool-product + is-retraction-map-inv-Π-bool-product f = + eq-htpy + ( λ where + true → refl + false → refl) + + is-equiv-map-Π-bool-product : is-equiv map-Π-bool-product + is-equiv-map-Π-bool-product = + is-equiv-is-invertible + map-inv-Π-bool-product + is-section-map-inv-Π-bool-product + is-retraction-map-inv-Π-bool-product + + equiv-Π-bool-product : ((b : bool) → A b) ≃ (A true × A false) + pr1 equiv-Π-bool-product = map-Π-bool-product + pr2 equiv-Π-bool-product = is-equiv-map-Π-bool-product + + is-equiv-map-inv-Π-bool-product : is-equiv map-inv-Π-bool-product + is-equiv-map-inv-Π-bool-product = + is-equiv-is-invertible + map-Π-bool-product + is-retraction-map-inv-Π-bool-product + is-section-map-inv-Π-bool-product + + inv-equiv-Π-bool-product : (A true × A false) ≃ ((b : bool) → A b) + pr1 inv-equiv-Π-bool-product = map-inv-Π-bool-product + pr2 inv-equiv-Π-bool-product = is-equiv-map-inv-Π-bool-product +``` diff --git a/src/foundation/universal-property-booleans.lagda.md b/src/foundation/universal-property-booleans.lagda.md index b976f49d78..33d77cb4df 100644 --- a/src/foundation/universal-property-booleans.lagda.md +++ b/src/foundation/universal-property-booleans.lagda.md @@ -14,10 +14,13 @@ open import foundation.function-extensionality open import foundation.universe-levels open import foundation-core.cartesian-product-types +open import foundation-core.coproduct-types 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 ``` @@ -36,9 +39,8 @@ map-universal-property-bool (pair x y) false = y abstract is-section-map-universal-property-bool : {l : Level} {A : UU l} → - ((ev-true-false A) ∘ map-universal-property-bool) ~ id - is-section-map-universal-property-bool (pair x y) = - eq-pair refl refl + is-section (ev-true-false A) (map-universal-property-bool) + is-section-map-universal-property-bool (pair x y) = refl abstract is-retraction-map-universal-property-bool' : @@ -50,14 +52,14 @@ abstract abstract is-retraction-map-universal-property-bool : {l : Level} {A : UU l} → - (map-universal-property-bool ∘ (ev-true-false A)) ~ id + is-retraction (ev-true-false A) (map-universal-property-bool) is-retraction-map-universal-property-bool f = eq-htpy (is-retraction-map-universal-property-bool' f) abstract universal-property-bool : {l : Level} (A : UU l) → - is-equiv (λ (f : bool → A) → pair (f true) (f false)) + is-equiv (λ (f : bool → A) → (f true , f false)) universal-property-bool A = is-equiv-is-invertible map-universal-property-bool @@ -115,3 +117,15 @@ eq-false-equiv' e p (inr x) = ( pair false (eq-true (map-equiv e false) x))))) -} ``` + +### The canonical projection from a coproduct to the booleans + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + projection-bool-coproduct : A + B → bool + projection-bool-coproduct (inl _) = true + projection-bool-coproduct (inr _) = false +``` diff --git a/src/foundation/universal-property-unit-type.lagda.md b/src/foundation/universal-property-unit-type.lagda.md index 6d066ead4f..c93cee6932 100644 --- a/src/foundation/universal-property-unit-type.lagda.md +++ b/src/foundation/universal-property-unit-type.lagda.md @@ -123,3 +123,15 @@ abstract ( universal-property-unit-is-equiv-point x is-equiv-point Y) ( refl-htpy) ``` + +### The unit type is terminal + +```agda +module _ + {l : Level} {X : UU l} + where + + is-equiv-terminal-map-Π-unit : is-equiv (terminal-map (X → unit)) + is-equiv-terminal-map-Π-unit = + is-equiv-is-invertible (const X) refl-htpy refl-htpy +``` diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 95c940b7c6..cd81f656fd 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -12,6 +12,7 @@ module orthogonal-factorization-systems where open import orthogonal-factorization-systems.cd-structures public open import orthogonal-factorization-systems.cellular-maps public open import orthogonal-factorization-systems.closed-modalities public +open import orthogonal-factorization-systems.coproducts-null-types public open import orthogonal-factorization-systems.double-lifts-families-of-elements public open import orthogonal-factorization-systems.extensions-double-lifts-families-of-elements public open import orthogonal-factorization-systems.extensions-lifts-families-of-elements public diff --git a/src/orthogonal-factorization-systems/coproducts-null-types.lagda.md b/src/orthogonal-factorization-systems/coproducts-null-types.lagda.md new file mode 100644 index 0000000000..1217e799ee --- /dev/null +++ b/src/orthogonal-factorization-systems/coproducts-null-types.lagda.md @@ -0,0 +1,116 @@ +# Coproducts of null types + +```agda +module orthogonal-factorization-systems.coproducts-null-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.diagonal-maps-of-types +open import foundation.equivalences +open import foundation.equivalences-arrows +open import foundation.fibers-of-maps +open import foundation.function-extensionality +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-coproduct-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.inhabited-types +open import foundation.logical-equivalences +open import foundation.postcomposition-dependent-functions +open import foundation.postcomposition-functions +open import foundation.precomposition-dependent-functions +open import foundation.precomposition-functions +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.raising-universe-levels +open import foundation.retracts-of-maps +open import foundation.retracts-of-types +open import foundation.type-arithmetic-booleans +open import foundation.type-arithmetic-unit-type +open import foundation.unit-type +open import foundation.universal-property-booleans +open import foundation.universal-property-equivalences +open import foundation.universal-property-family-of-fibers-of-maps +open import foundation.universal-property-unit-type +open import foundation.universe-levels + +open import orthogonal-factorization-systems.local-maps +open import orthogonal-factorization-systems.local-types +open import orthogonal-factorization-systems.null-maps +open import orthogonal-factorization-systems.null-types +open import orthogonal-factorization-systems.orthogonal-maps +``` + +
+ +## Idea + +The universe of `Y`-[null](orthogonal-factorization-systems.null-types.md) types +are closed under [coproducts](foundation.coproduct-types.md) if and only if the +[booleans](foundation.booleans.md) are `Y`-null. + +## Properties + +### The canonical map `A + B → bool` is `Y`-null if `A` and `B` are + +```agda +module _ + {l1 l2 l3 : Level} + (Y : UU l1) {A : UU l2} {B : UU l3} + (is-null-A : is-null Y A) + (is-null-B : is-null Y B) + where + + abstract + is-null-projection-bool-coproduct : + is-null-map Y (projection-bool-coproduct {A = A} {B}) + is-null-projection-bool-coproduct = + is-null-map-left-map-triangle Y + ( λ where + (inl _) → refl + (inr _) → refl) + ( is-null-map-pr1-is-null-family Y + ( rec-bool (raise (l2 ⊔ l3) A) (raise (l2 ⊔ l3) B)) + ( λ where + true → + is-null-equiv-base (inv-compute-raise (l2 ⊔ l3) A) is-null-A + false → + is-null-equiv-base (inv-compute-raise (l2 ⊔ l3) B) is-null-B)) + ( is-null-map-map-equiv Y + ( ( inv-equiv-Σ-bool-coproduct + ( rec-bool (raise (l2 ⊔ l3) A) (raise (l2 ⊔ l3) B))) ∘e + ( equiv-coproduct + ( compute-raise (l2 ⊔ l3) A) + ( compute-raise (l2 ⊔ l3) B)))) +``` + +### If the booleans are `Y`-null then coproducts of `Y`-null types are `Y`-null + +```agda +module _ + {l1 l2 l3 : Level} + (Y : UU l1) {A : UU l2} {B : UU l3} + (is-null-bool : is-null Y bool) + (is-null-A : is-null Y A) + (is-null-B : is-null Y B) + where + + is-null-coproduct-is-null-bool : is-null Y (A + B) + is-null-coproduct-is-null-bool = + is-null-is-orthogonal-terminal-maps + ( is-orthogonal-right-comp + ( terminal-map Y) + ( projection-bool-coproduct) + ( terminal-map bool) + ( is-orthogonal-terminal-maps-is-null is-null-bool) + ( is-orthogonal-terminal-map-is-null-map Y + ( projection-bool-coproduct) + ( is-null-projection-bool-coproduct Y is-null-A is-null-B))) +``` diff --git a/src/orthogonal-factorization-systems/local-types.lagda.md b/src/orthogonal-factorization-systems/local-types.lagda.md index 0f957651be..0a8c3191ec 100644 --- a/src/orthogonal-factorization-systems/local-types.lagda.md +++ b/src/orthogonal-factorization-systems/local-types.lagda.md @@ -21,10 +21,12 @@ open import foundation.function-types open import foundation.functoriality-dependent-function-types open import foundation.homotopies open import foundation.identity-types +open import foundation.inhabited-types open import foundation.logical-equivalences open import foundation.postcomposition-functions open import foundation.precomposition-dependent-functions open import foundation.precomposition-functions +open import foundation.propositional-truncations open import foundation.propositions open import foundation.retracts-of-maps open import foundation.retracts-of-types @@ -299,6 +301,28 @@ module _ is-local-dependent-type-is-prop (λ _ → A) (λ _ → is-prop-A) ``` +### Propositions are `f`-local if the domain of `f` is inhabited + +```agda +module _ + {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) + where + + is-local-has-element-domain-is-prop : + {l : Level} (A : UU l) → + is-prop A → X → is-local f A + is-local-has-element-domain-is-prop A is-prop-A x = + is-local-is-prop f A is-prop-A (λ h y → h x) + + is-local-is-inhabited-domain-is-prop : + {l : Level} (A : UU l) → + is-prop A → is-inhabited X → is-local f A + is-local-is-inhabited-domain-is-prop A is-prop-A = + rec-trunc-Prop + ( is-local-Prop f A) + ( is-local-has-element-domain-is-prop A is-prop-A) +``` + ### All type families are local at equivalences ```agda @@ -342,7 +366,7 @@ module _ ```agda is-contr-is-local : - {l : Level} (A : UU l) → is-local (λ (_ : empty) → star) A → is-contr A + {l : Level} (A : UU l) → is-local (terminal-map empty) A → is-contr A is-contr-is-local A is-local-A = is-contr-is-equiv ( empty → A) diff --git a/src/orthogonal-factorization-systems/null-maps.lagda.md b/src/orthogonal-factorization-systems/null-maps.lagda.md index 2863c0ff12..04e8dca0ee 100644 --- a/src/orthogonal-factorization-systems/null-maps.lagda.md +++ b/src/orthogonal-factorization-systems/null-maps.lagda.md @@ -15,6 +15,7 @@ open import foundation.equivalences-arrows open import foundation.families-of-equivalences open import foundation.fibers-of-maps open import foundation.function-types +open import foundation.functoriality-cartesian-product-types open import foundation.functoriality-fibers-of-maps open import foundation.homotopies open import foundation.identity-types @@ -219,7 +220,7 @@ module _ inv-equiv equiv-is-local-terminal-map-is-null-map ``` -### A map is `Y`-null if and only if the terminal projection of `Y` is orthogonal to `f` +### A map is `Y`-null if and only if it is `Y`-orthogonal ```agda module _ @@ -278,3 +279,182 @@ module _ is-orthogonal-is-orthogonal-fiber-condition-right-map (terminal-map Y) f ( is-orthogonal-fiber-condition-terminal-map-is-null-map H) ``` + +### Equivalences are null at any type + +```agda +module _ + {l1 l2 l3 : Level} (Y : UU l1) {A : UU l2} {B : UU l3} (e : A ≃ B) + where + + is-null-map-map-equiv : is-null-map Y (map-equiv e) + is-null-map-map-equiv = + is-null-map-is-orthogonal-terminal-map Y + ( map-equiv e) + ( is-orthogonal-equiv-right (terminal-map Y) e) +``` + +### Null maps are closed under homotopies + +```agda +module _ + {l1 l2 l3 : Level} (Y : UU l1) {A : UU l2} {B : UU l3} + {f g : A → B} + where + + is-null-map-htpy : f ~ g → is-null-map Y f → is-null-map Y g + is-null-map-htpy H is-null-f = + is-null-map-is-orthogonal-terminal-map Y g + ( is-orthogonal-htpy-right (terminal-map Y) f H + ( is-orthogonal-terminal-map-is-null-map Y f is-null-f)) + + is-null-map-htpy' : g ~ f → is-null-map Y f → is-null-map Y g + is-null-map-htpy' H = is-null-map-htpy (inv-htpy H) +``` + +### A family is `Y`-null if and only if it is `Y`-orthogonal + +```agda +module _ + {l1 l2 l3 : Level} (Y : UU l1) {A : UU l2} {B : A → UU l3} + where + + is-null-family-is-orthogonal-terminal-map : + is-orthogonal (terminal-map Y) (pr1 {B = B}) → is-null-family Y B + is-null-family-is-orthogonal-terminal-map H = + is-null-family-is-null-map-pr1 Y B + ( is-null-map-is-orthogonal-terminal-map Y pr1 H) + + is-orthogonal-terminal-map-is-null-family : + is-null-family Y B → is-orthogonal (terminal-map Y) (pr1 {B = B}) + is-orthogonal-terminal-map-is-null-family H = + is-orthogonal-terminal-map-is-null-map Y pr1 + ( is-null-map-pr1-is-null-family Y B H) +``` + +### The dependent sum of a `Y`-null type family over a `Y`-null base is `Y`-null + +```agda +module _ + {l1 l2 l3 : Level} (Y : UU l1) {A : UU l2} {B : A → UU l3} + where + + abstract + is-null-Σ : is-null Y A → is-null-family Y B → is-null Y (Σ A B) + is-null-Σ is-null-A is-null-B = + is-null-is-orthogonal-terminal-maps + ( is-orthogonal-right-comp + ( terminal-map Y) + ( pr1) + ( terminal-map A) + ( is-orthogonal-terminal-maps-is-null is-null-A) + ( is-orthogonal-terminal-map-is-null-family Y is-null-B)) + + abstract + is-null-family-is-null-Σ : + is-null Y A → is-null Y (Σ A B) → is-null-family Y B + is-null-family-is-null-Σ is-null-A is-null-ΣAB = + is-null-family-is-orthogonal-terminal-map Y + ( is-orthogonal-right-right-factor + ( terminal-map Y) + ( pr1) + ( terminal-map A) + ( is-orthogonal-terminal-maps-is-null is-null-A) + ( is-orthogonal-terminal-maps-is-null is-null-ΣAB)) +``` + +### Composition and right cancellation of null maps + +```agda +module _ + {l1 l2 l3 l4 : Level} + (Y : UU l1) {A : UU l2} {B : UU l3} {C : UU l4} + {f : A → B} {g : B → C} + where + + abstract + is-null-map-comp : + is-null-map Y g → is-null-map Y f → is-null-map Y (g ∘ f) + is-null-map-comp is-null-g is-null-f = + is-null-map-is-orthogonal-terminal-map Y (g ∘ f) + ( is-orthogonal-right-comp (terminal-map Y) f g + ( is-orthogonal-terminal-map-is-null-map Y g is-null-g) + ( is-orthogonal-terminal-map-is-null-map Y f is-null-f)) + + abstract + is-null-map-right-factor : + is-null-map Y g → is-null-map Y (g ∘ f) → is-null-map Y f + is-null-map-right-factor is-null-g is-null-gf = + is-null-map-is-orthogonal-terminal-map Y f + ( is-orthogonal-right-right-factor (terminal-map Y) f g + ( is-orthogonal-terminal-map-is-null-map Y g is-null-g) + ( is-orthogonal-terminal-map-is-null-map Y (g ∘ f) is-null-gf)) + +module _ + {l1 l2 l3 l4 : Level} + (Y : UU l1) {A : UU l2} {B : UU l3} {C : UU l4} + {f : A → B} {g : B → C} {h : A → C} (H : h ~ g ∘ f) + where + + is-null-map-left-map-triangle : + is-null-map Y g → is-null-map Y f → is-null-map Y h + is-null-map-left-map-triangle is-null-g is-null-f = + is-null-map-htpy' Y H (is-null-map-comp Y is-null-g is-null-f) + + is-null-map-top-map-triangle : + is-null-map Y g → is-null-map Y h → is-null-map Y f + is-null-map-top-map-triangle is-null-g is-null-h = + is-null-map-right-factor Y is-null-g (is-null-map-htpy Y H is-null-h) +``` + +### Null maps are closed under dependent products + +```agda +module _ + {l1 l2 l3 l4 : Level} + (Y : UU l1) {I : UU l2} {A : I → UU l3} {B : I → UU l4} + {f : (i : I) → A i → B i} + where + + abstract + is-null-map-Π : ((i : I) → is-null-map Y (f i)) → is-null-map Y (map-Π f) + is-null-map-Π is-null-f = + is-null-map-is-orthogonal-terminal-map Y (map-Π f) + ( is-orthogonal-right-Π (terminal-map Y) f + ( λ i → is-orthogonal-terminal-map-is-null-map Y (f i) (is-null-f i))) +``` + +### Null maps are closed under exponentiation + +```agda +module _ + {l1 l2 l3 l4 : Level} + (Y : UU l1) (I : UU l2) {A : UU l3} {B : UU l4} {f : A → B} + where + + abstract + is-null-map-postcomp : is-null-map Y f → is-null-map Y (postcomp I f) + is-null-map-postcomp is-null-f = + is-null-map-is-orthogonal-terminal-map Y (postcomp I f) + ( is-orthogonal-right-postcomp I (terminal-map Y) f + ( is-orthogonal-terminal-map-is-null-map Y f is-null-f)) +``` + +### Null maps are closed under cartesian products + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} + (Y : UU l1) {A : UU l2} {B : UU l3} {C : UU l4} {D : UU l5} + {f : A → B} {g : C → D} + where + + abstract + is-null-map-product : + is-null-map Y f → is-null-map Y g → is-null-map Y (map-product f g) + is-null-map-product is-null-f is-null-g = + is-null-map-is-orthogonal-terminal-map Y (map-product f g) + ( is-orthogonal-right-product (terminal-map Y) f g + ( is-orthogonal-terminal-map-is-null-map Y f is-null-f) + ( is-orthogonal-terminal-map-is-null-map Y g is-null-g)) +``` diff --git a/src/orthogonal-factorization-systems/null-types.lagda.md b/src/orthogonal-factorization-systems/null-types.lagda.md index 6e50e36ba6..6eb9ee88c3 100644 --- a/src/orthogonal-factorization-systems/null-types.lagda.md +++ b/src/orthogonal-factorization-systems/null-types.lagda.md @@ -7,18 +7,24 @@ module orthogonal-factorization-systems.null-types where
Imports ```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types open import foundation.dependent-pair-types open import foundation.diagonal-maps-of-types open import foundation.equivalences open import foundation.equivalences-arrows open import foundation.fibers-of-maps open import foundation.function-extensionality +open import foundation.functoriality-cartesian-product-types open import foundation.homotopies open import foundation.identity-types +open import foundation.inhabited-types open import foundation.logical-equivalences +open import foundation.postcomposition-dependent-functions open import foundation.postcomposition-functions open import foundation.precomposition-dependent-functions open import foundation.precomposition-functions +open import foundation.propositional-truncations open import foundation.propositions open import foundation.retracts-of-maps open import foundation.retracts-of-types @@ -227,3 +233,109 @@ module _ (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) ``` + +### Propositions are null at inhabited types + +```agda +module _ + {l1 l2 : Level} {Y : UU l1} + where + + is-null-is-prop-is-inhabited' : + {P : UU l2} → Y → is-prop P → is-null Y P + is-null-is-prop-is-inhabited' {P} y is-prop-P = + is-equiv-has-converse-is-prop + ( is-prop-P) + ( is-prop-function-type is-prop-P) + ( λ f → f y) + + is-null-is-prop-is-inhabited : + {P : UU l2} → is-inhabited Y → is-prop P → is-null Y P + is-null-is-prop-is-inhabited {P} is-inhabited-Y is-prop-P = + is-equiv-has-converse-is-prop + ( is-prop-P) + ( is-prop-function-type is-prop-P) + ( λ f → rec-trunc-Prop (P , is-prop-P) f is-inhabited-Y) + + is-null-prop-is-inhabited : + is-inhabited Y → (P : Prop l2) → is-null Y (type-Prop P) + is-null-prop-is-inhabited is-inhabited-Y P = + is-null-is-prop-is-inhabited is-inhabited-Y (is-prop-type-Prop P) +``` + +### Null types are closed under dependent products + +```agda +module _ + {l1 l2 l3 : Level} {Y : UU l1} {I : UU l2} {B : I → UU l3} + where + + abstract + is-null-Π : ((i : I) → is-null Y (B i)) → is-null Y ((i : I) → B i) + is-null-Π is-null-B = + is-null-is-orthogonal-terminal-maps + ( is-orthogonal-right-comp + ( terminal-map Y) + ( postcomp-Π I (λ {i : I} → terminal-map (B i))) + ( terminal-map (I → unit)) + ( is-orthogonal-is-equiv-right + ( terminal-map Y) + ( terminal-map (I → unit)) + ( is-equiv-terminal-map-Π-unit)) + ( is-orthogonal-right-Π + ( terminal-map Y) + ( λ i → terminal-map (B i)) + ( λ i → (is-orthogonal-terminal-maps-is-null (is-null-B i))))) +``` + +### Null types are closed under exponentiation + +```agda +module _ + {l1 l2 l3 : Level} {Y : UU l1} {I : UU l2} {B : UU l3} + where + + is-null-function-type : is-null Y B → is-null Y (I → B) + is-null-function-type is-null-B = is-null-Π (λ _ → is-null-B) +``` + +### Null types are closed under cartesian products + +```agda +module _ + {l1 l2 l3 : Level} {Y : UU l1} {A : UU l2} {B : UU l3} + where + + abstract + is-null-product : is-null Y A → is-null Y B → is-null Y (A × B) + is-null-product is-null-A is-null-B = + is-null-is-orthogonal-terminal-maps + ( is-orthogonal-right-comp + ( terminal-map Y) + ( map-product (terminal-map A) (terminal-map B)) + ( terminal-map (unit × unit)) + ( is-orthogonal-is-equiv-right + ( terminal-map Y) + ( terminal-map (unit × unit)) + ( is-equiv-map-right-unit-law-product)) + ( is-orthogonal-right-product + ( terminal-map Y) + ( terminal-map A) + ( terminal-map B) + ( is-orthogonal-terminal-maps-is-null is-null-A) + ( is-orthogonal-terminal-maps-is-null is-null-B))) +``` + +### Null types are closed under identity types + +This remains to be formalized. + +## See also + +- We show that null types are closed under dependent sums in + [`null-maps`](orthogonal-factorization-systems.null-maps.md). +- In + [`coproducts-null-types`](orthogonal-factorization-systems.coproducts-null-types.md) + we show that `Y`-null types are closed under + [coproducts](foundation.coproduct-types.md) if and only if the + [booleans](foundation.booleans.md) are `Y`-null. diff --git a/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md b/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md index 1f47496951..0acd0dac27 100644 --- a/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md +++ b/src/orthogonal-factorization-systems/orthogonal-maps.lagda.md @@ -419,8 +419,37 @@ module _ is-orthogonal-pullback-condition f g' is-orthogonal-pullback-condition-htpy-right = is-orthogonal-pullback-condition-htpy refl-htpy + + is-orthogonal-htpy : + {f' : A → B} {g' : X → Y} → f ~ f' → g ~ g' → + is-orthogonal f g → is-orthogonal f' g' + is-orthogonal-htpy {f'} {g'} F G H = + is-orthogonal-is-orthogonal-pullback-condition f' g' + ( is-orthogonal-pullback-condition-htpy F G + ( is-orthogonal-pullback-condition-is-orthogonal f g H)) + + is-orthogonal-htpy-left : + {f' : A → B} → f ~ f' → is-orthogonal f g → is-orthogonal f' g + is-orthogonal-htpy-left F = is-orthogonal-htpy F refl-htpy + + is-orthogonal-htpy-right : + {g' : X → Y} → g ~ g' → is-orthogonal f g → is-orthogonal f g' + is-orthogonal-htpy-right G = is-orthogonal-htpy refl-htpy G ``` +### Orthogonality is preserved under equivalences of maps + +Given equivalences of arrows `f' ≃ f` and `g' ≃ g`, then `f ⊥ g` if and only if +`f' ⊥ g'`. + +This remains to be formalized. + +### Orthogonality is preserved under retracts of maps + +Given retracts of maps `f' → f` and `g → g'`, then `f ⊥ g` implies `f' ⊥ g'`. + +This remains to be formalized. + ### Equivalences are orthogonal to every map ```agda @@ -946,7 +975,7 @@ If `f ⊥ g` and `f' ⊥ g`, then `(f + f') ⊥ g`. **Proof:** We need to show that the square ```text - - ∘ (f + f') + - ∘ (f + f') ((B + B') → X) ---------------> ((A + A') → X) | | | | @@ -1094,6 +1123,41 @@ module _ ( is-orthogonal-pullback-condition-is-orthogonal f g G)) ``` +### Right orthogonality is preserved under pullback-homs with arbitrary maps + +Given an `f`-orthogonal map `g` and an arbitrary map `k`, then the pullback-hom +`⟨k , g⟩` is also `f`-orthogonal. This is Corollary 3.1.13 of {{#cite BW23}}. + +This remains to be formalized. + +### Left orthogonality is preserved under pushout-products with arbitrary maps + +Given maps `f ⊥ g` and an arbitrary map `k`, then `(f □ k) ⊥ g`. + +**Proof.** Follows from the previous result using the adjoint relation to the +pullback-hom: + +```text + ⟨f □ k , g⟩ ~ ⟨f , ⟨k , g⟩⟩. +``` + +This remains to be formalized. + +### Right orthogonality is preserved under sequential limits + +Given a [tower](foundation.inverse-sequential-diagrams.md) of `f`-orthogonal +maps + +```text + g₃ g₂ g₁ g₀ + ⋯ ---> X₃ ---> X₂ ---> X₁ ---> X₀. +``` + +then the [sequential limit](foundation.sequential-limits.md) `g∞ : X∞ → X₀` is +also `f`-orthogonal. This is Proposition 3.1.14 of {{#cite BW23}}. + +This remains to be formalized. + ### A type is `f`-local if and only if its terminal map is `f`-orthogonal **Proof (forward direction):** If the terminal map is right orthogonal to `f`, diff --git a/src/reflection/erasing-equality.lagda.md b/src/reflection/erasing-equality.lagda.md index 3f3a0bec8a..ce95bb592a 100644 --- a/src/reflection/erasing-equality.lagda.md +++ b/src/reflection/erasing-equality.lagda.md @@ -18,11 +18,11 @@ open import foundation-core.identity-types Agda's builtin primitive `primEraseEquality` is a special construct on [identifications](foundation-core.identity-types.md) that for every -identification `x = y` gives an identification `x = y` with the following +identification `x = y` gives another identification `x = y` with the following reduction behaviour: -- If the two end points `x = y` normalize to the same term, `primEraseEquality` - reduces to `refl`. +- If the two end points of `p : x = y` normalize to the same term, + `primEraseEquality p` reduces to `refl`. For example, `primEraseEquality` applied to the loop of the [circle](synthetic-homotopy-theory.circle.md) will compute to `refl`, while diff --git a/src/simplicial-type-theory.lagda.md b/src/simplicial-type-theory.lagda.md new file mode 100644 index 0000000000..b3fec39887 --- /dev/null +++ b/src/simplicial-type-theory.lagda.md @@ -0,0 +1,54 @@ +# Simplicial type theory + +```agda +{-# OPTIONS --rewriting #-} +``` + +## Files in the simplicial type theory folder + +```agda +module simplicial-type-theory where + +open import simplicial-type-theory.2-simplices public +open import simplicial-type-theory.action-on-directed-edges-dependent-functions public +open import simplicial-type-theory.action-on-directed-edges-functions public +open import simplicial-type-theory.boundary-standard-simplices public +open import simplicial-type-theory.comma-types public +open import simplicial-type-theory.compositions-of-directed-edges public +open import simplicial-type-theory.conormed-maps-between-types public +open import simplicial-type-theory.dependent-directed-edges public +open import simplicial-type-theory.directed-circle public +open import simplicial-type-theory.directed-edges public +open import simplicial-type-theory.directed-edges-booleans public +open import simplicial-type-theory.directed-edges-cartesian-product-types public +open import simplicial-type-theory.directed-edges-dependent-pair-types public +open import simplicial-type-theory.directed-interval-type public +open import simplicial-type-theory.free-directed-loops public +open import simplicial-type-theory.fully-faithful-maps public +open import simplicial-type-theory.globularly-coskeletal-types public +open import simplicial-type-theory.horizontal-composition-directed-edges-functions public +open import simplicial-type-theory.horizontal-composition-natural-transformations public +open import simplicial-type-theory.horizontal-composition-simplicial-arrows-functions public +open import simplicial-type-theory.inequality-directed-interval-type public +open import simplicial-type-theory.inner-2-horn public +open import simplicial-type-theory.lax-suspension public +open import simplicial-type-theory.natural-transformations public +open import simplicial-type-theory.normed-maps-between-types public +open import simplicial-type-theory.representing-biinvertible-simplicial-arrow public +open import simplicial-type-theory.rewriting-directed-circle public +open import simplicial-type-theory.simplicial-arrows public +open import simplicial-type-theory.simplicial-cones public +open import simplicial-type-theory.simplicial-cubes public +open import simplicial-type-theory.simplicial-mapping-cylinders public +open import simplicial-type-theory.simplicial-spines public +open import simplicial-type-theory.simplicially-covariant-type-families public +open import simplicial-type-theory.simplicially-discrete-types public +open import simplicial-type-theory.standard-simplices public +open import simplicial-type-theory.standard-simplicial-joins public +open import simplicial-type-theory.transposing-adjunctions-between-types public +open import simplicial-type-theory.transposing-biadjunctions-between-types public +open import simplicial-type-theory.universal-property-directed-circle public +open import simplicial-type-theory.whiskering-directed-edges public +open import simplicial-type-theory.whiskering-simplicial-arrows-functions public +open import simplicial-type-theory.whiskering-simplicial-edges-functions public +``` diff --git a/src/simplicial-type-theory/2-simplices.lagda.md b/src/simplicial-type-theory/2-simplices.lagda.md new file mode 100644 index 0000000000..036a83e8a3 --- /dev/null +++ b/src/simplicial-type-theory/2-simplices.lagda.md @@ -0,0 +1,193 @@ +# 2-simplices + +```agda +module simplicial-type-theory.2-simplices where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-cones + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Definitions + + + +```agda +eq-image-eq-point-is-prop : + {l1 l2 : Level} {A : UU l1} {B : UU l2} → is-prop A → + (f : A → B) → (b : B) (x : A) → b = f x → {y : A} → b = f y +eq-image-eq-point-is-prop is-prop-A f b x p = p ∙ ap f (eq-is-prop is-prop-A) +``` + +### The lower simplicial triangle + +```agda +subtype-lower-simplicial-triangle : subtype lzero (𝟚 × 𝟚) +subtype-lower-simplicial-triangle (t , s) = leq-𝟚-Prop s t + +lower-simplicial-triangle = type-subtype subtype-lower-simplicial-triangle + +subtype-boundary-lower-simplicial-triangle : subtype lzero (𝟚 × 𝟚) +subtype-boundary-lower-simplicial-triangle (t , s) = + join-Prop (Id-𝟚-Prop 0₂ s) (join-Prop (Id-𝟚-Prop s t) (Id-𝟚-Prop t 1₂)) +boundary-lower-simplicial-triangle = + type-subtype subtype-boundary-lower-simplicial-triangle + +◿⊆◢ : + subtype-boundary-lower-simplicial-triangle ⊆ subtype-lower-simplicial-triangle +◿⊆◢ (x , y) = + rec-join-Prop + ( leq-𝟚-Prop y x) + ( min-leq-eq-𝟚 ∘ inv) + ( rec-join-Prop (leq-𝟚-Prop y x) (leq-eq-𝟚) (max-leq-eq-𝟚)) + +inclusion-boundary-lower-simplicial-triangle : + boundary-lower-simplicial-triangle → lower-simplicial-triangle +inclusion-boundary-lower-simplicial-triangle = tot ◿⊆◢ +``` + +### The upper simplicial triangle + +```agda +subtype-◤ : subtype lzero (𝟚 × 𝟚) +subtype-◤ (t , s) = leq-𝟚-Prop t s + +◤ = type-subtype subtype-◤ + +subtype-◸ : subtype lzero (𝟚 × 𝟚) +subtype-◸ (t , s) = + join-Prop + ( Id-𝟚-Prop 0₂ t) + ( join-Prop + ( Id-𝟚-Prop t s) + ( Id-𝟚-Prop s 1₂)) + +◸ = type-subtype subtype-◸ + +◸⊆◤ : subtype-◸ ⊆ subtype-◤ +◸⊆◤ (x , y) = + rec-join-Prop + ( leq-𝟚-Prop x y) + ( min-leq-eq-𝟚 ∘ inv) + ( rec-join-Prop (leq-𝟚-Prop x y) (leq-eq-𝟚) (max-leq-eq-𝟚)) + +◸→◤ : ◸ → ◤ +◸→◤ = tot ◸⊆◤ +``` + +### The standard 2-simplex + +We define the standard 2-simplex as the lower simplicial triangle. + +```agda +subtype-Δ² = subtype-lower-simplicial-triangle +Δ² = lower-simplicial-triangle +``` + +### The boundary of the standard 2-simplex + +```agda +subtype-∂Δ² : subtype lzero (𝟚 × 𝟚) +subtype-∂Δ² = subtype-boundary-lower-simplicial-triangle + +∂Δ² : UU lzero +∂Δ² = boundary-lower-simplicial-triangle + +∂Δ²⊆Δ² = ◿⊆◢ +inclusion-∂Δ² = inclusion-boundary-lower-simplicial-triangle + +rec-simplicial-arrow-∂Δ² : + {l : Level} {A : UU l} + (f g h : simplicial-arrow A) → + f 0₂ = h 0₂ → + f 1₂ = g 0₂ → + g 1₂ = h 1₂ → + ∂Δ² → A +rec-simplicial-arrow-∂Δ² {A = A} f g h f0=h0 f1=g0 g1=h1 ((t , s) , u) = + cogap-join A + ( ( λ _ → f t) , + ( C) , + ( λ (0=s , vw) → + cogap-join _ + ( ( λ s=t → + eq-image-eq-point-is-prop + ( is-prop-join-is-prop (is-set-𝟚 s t) (is-set-𝟚 t 1₂)) + ( C) + ( f t) + ( inl-join s=t) + ( ( ap f (inv (0=s ∙ s=t)) ∙ f0=h0 ∙ ap h (0=s ∙ s=t)) ∙ + ( inv (compute-inl-cogap-join _ s=t)))) , + ( ( λ t=1 → + eq-image-eq-point-is-prop + ( is-prop-join-is-prop (is-set-𝟚 s t) (is-set-𝟚 t 1₂)) + ( C) + ( f t) + ( inr-join t=1) + ( ( ap f t=1 ∙ f1=g0 ∙ ap g 0=s) ∙ + ( inv (compute-inr-cogap-join _ t=1))))) , + ( λ (s=t , t=1) → ex-falso (is-nontrivial-𝟚 (0=s ∙ s=t ∙ t=1)))) + ( vw))) + ( u) + where + C = + cogap-join A + ( ( λ _ → h t) , + ( λ _ → g s) , + ( λ (s=t , t=1) → ap h t=1 ∙ inv (ap g (s=t ∙ t=1) ∙ g1=h1))) + +rec-simplicial-hom-∂Δ² : + {l : Level} {A : UU l} + {x y z : A} → + simplicial-hom x y → simplicial-hom y z → simplicial-hom x z → + ∂Δ² → A +rec-simplicial-hom-∂Δ² f g h = + rec-simplicial-arrow-∂Δ² + ( simplicial-arrow-simplicial-hom f) + ( simplicial-arrow-simplicial-hom g) + ( simplicial-arrow-simplicial-hom h) + ( eq-source-source-simplicial-hom f h) + ( eq-source-target-simplicial-hom f g) + ( eq-target-target-simplicial-hom g h) +``` + +### The 2-simplex as a simplicial cone + +```agda +simplicial-cone-Δ² : UU lzero +simplicial-cone-Δ² = simplicial-cone 𝟚 +``` diff --git a/src/simplicial-type-theory/action-on-directed-edges-dependent-functions.lagda.md b/src/simplicial-type-theory/action-on-directed-edges-dependent-functions.lagda.md new file mode 100644 index 0000000000..40f963fc69 --- /dev/null +++ b/src/simplicial-type-theory/action-on-directed-edges-dependent-functions.lagda.md @@ -0,0 +1,46 @@ +# The action on directed edges of dependent functions + +```agda +module simplicial-type-theory.action-on-directed-edges-dependent-functions where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.action-on-identifications-functions +open import foundation.constant-maps +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.universe-levels + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +``` + +
+ +## Idea + +Any dependent function `f : (x : A) → B x` preserves +[directed edges](simplicial-type-theory.directed-edges.md), in the sense that it +maps any edge `p : x →₂ y` in `A` to a +[dependent edge](simplicial-type-theory.dependent-directed-edges.md) +`action-simplicial-hom-dependent-function f p` from `f x` to `f y` over `p` in +`B`. This action on directed edges can be thought of as functoriality of +dependent functions in simplicial type theory. + +## Definition + +### The functorial action of dependent functions on directed edges + +```agda +action-simplicial-hom-dependent-function : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} (f : (x : A) → B x) {x y : A} → + (α : x →₂ y) → dependent-simplicial-hom B α (f x) (f y) +action-simplicial-hom-dependent-function f (α , s , t) = + ( f ∘ α , apd f s , apd f t) +``` diff --git a/src/simplicial-type-theory/action-on-directed-edges-functions.lagda.md b/src/simplicial-type-theory/action-on-directed-edges-functions.lagda.md new file mode 100644 index 0000000000..06cf7381ab --- /dev/null +++ b/src/simplicial-type-theory/action-on-directed-edges-functions.lagda.md @@ -0,0 +1,95 @@ +# The action on directed edges of functions + +```agda +module simplicial-type-theory.action-on-directed-edges-functions where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.constant-maps +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +``` + +
+ +## Idea + +Any function `f : A → B` preserves +[directed edges](simplicial-type-theory.directed-edges.md), in the sense that it +maps any edge `p : x →₂ y` in `A` to an edge +`action-simplicial-hom f p : f x →₂ f y` in `B`. This action on directed edges +can be thought of as a functorial action of functions in simplicial type theory. + +## Definition + +### The functorial action of functions on directed edges + +```agda +action-simplicial-hom-function : + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) {x y : A} → + x →₂ y → f x →₂ f y +action-simplicial-hom-function f (α , s , t) = (f ∘ α , ap f s , ap f t) +``` + +## Properties + +### The identity function acts trivially on directed edges + +```agda +compute-action-simplicial-hom-id-function : + {l : Level} {A : UU l} {x y : A} (p : x →₂ y) → + (action-simplicial-hom-function id p) = p +compute-action-simplicial-hom-id-function (α , s , t) = + eq-pair-eq-fiber (eq-pair (ap-id s) (ap-id t)) +``` + +### The action on directed edges of a composite function is the composite of the actions + +```agda +compute-action-simplicial-hom-comp-function : + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} (g : B → C) + (f : A → B) {x y : A} (p : x →₂ y) → + (action-simplicial-hom-function (g ∘ f) p) = + ((action-simplicial-hom-function g ∘ action-simplicial-hom-function f) p) +compute-action-simplicial-hom-comp-function g f (α , s , t) = + eq-pair-eq-fiber (eq-pair (ap-comp g f s) (ap-comp g f t)) + +associative-action-simplicial-hom-comp-function : + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + (h : C → D) (g : B → C) (f : A → B) {x y : A} (p : x →₂ y) → + action-simplicial-hom-function (h ∘ g) (action-simplicial-hom-function f p) = + action-simplicial-hom-function h (action-simplicial-hom-function (g ∘ f) p) +associative-action-simplicial-hom-comp-function h g f (α , s , t) = + eq-pair-eq-fiber (eq-pair (ap-comp-assoc h g f s) (ap-comp-assoc h g f t)) +``` + +### The action on directed edges of any map preserves the identity edges + +In fact, the identity edges are preserved strictly. + +```agda +compute-action-id-simplicial-hom-function : + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) (x : A) → + action-simplicial-hom-function f (id-simplicial-hom x) = + id-simplicial-hom (f x) +compute-action-id-simplicial-hom-function f x = refl +``` + +### The action on identifications of a constant map is constant + +```agda +compute-action-simplicial-hom-const-function : + {l1 l2 : Level} {A : UU l1} {B : UU l2} (b : B) {x y : A} (p : x →₂ y) → + action-simplicial-hom-function (const A b) p = id-simplicial-hom b +compute-action-simplicial-hom-const-function b (α , s , t) = + eq-pair-eq-fiber (eq-pair (ap-const b s) (ap-const b t)) +``` diff --git a/src/simplicial-type-theory/boundary-standard-simplices.lagda.md b/src/simplicial-type-theory/boundary-standard-simplices.lagda.md new file mode 100644 index 0000000000..b1d2ce44cc --- /dev/null +++ b/src/simplicial-type-theory/boundary-standard-simplices.lagda.md @@ -0,0 +1,179 @@ +# The boundary of the standard simplices + +```agda +module simplicial-type-theory.boundary-standard-simplices where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.embeddings +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-cubes +open import simplicial-type-theory.standard-simplices + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Definitions + +### The boundary of the standard simplices + +```agda +subtype-auxillary-face-standard-simplex : + (n r : ℕ) → subtype lzero (simplicial-cube n) +subtype-auxillary-face-standard-simplex 0 _ _ = + empty-Prop +subtype-auxillary-face-standard-simplex 1 0 x = + Id-Prop 𝟚-Set x 0₂ +subtype-auxillary-face-standard-simplex 1 (succ-ℕ _) x = + empty-Prop +subtype-auxillary-face-standard-simplex 2 zero-ℕ (x , y) = + Id-Prop 𝟚-Set x y +subtype-auxillary-face-standard-simplex 2 (succ-ℕ r) (x , y) = + subtype-auxillary-face-standard-simplex 1 r y +subtype-auxillary-face-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) 0 (x , y , u) = + ( Id-Prop 𝟚-Set x y) ∧ + ( subtype-standard-simplex (succ-ℕ (succ-ℕ n)) (y , u)) +subtype-auxillary-face-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) (succ-ℕ r) (x , y , u) = + ( leq-𝟚-Prop y x) ∧ + ( subtype-auxillary-face-standard-simplex (succ-ℕ (succ-ℕ n)) r (y , u)) + +subtype-first-face-standard-simplex : + (n : ℕ) → subtype lzero (simplicial-cube n) +subtype-first-face-standard-simplex 0 _ = empty-Prop +subtype-first-face-standard-simplex 1 x = Id-Prop 𝟚-Set 1₂ x +subtype-first-face-standard-simplex 2 (x , _) = Id-Prop 𝟚-Set 1₂ x +subtype-first-face-standard-simplex (succ-ℕ (succ-ℕ (succ-ℕ n))) (x , u) = + Id-Prop 𝟚-Set 1₂ x ∧ subtype-standard-simplex (succ-ℕ (succ-ℕ n)) u + +subtype-face-standard-simplex : (n r : ℕ) → subtype lzero (simplicial-cube n) +subtype-face-standard-simplex n 0 = + subtype-first-face-standard-simplex n +subtype-face-standard-simplex n (succ-ℕ r) = + subtype-auxillary-face-standard-simplex n r +``` + +> These definitions would've been a bit simpler if we used list induction. + +```agda +subtype-faces-up-to-standard-simplex : + (n k : ℕ) → subtype lzero (simplicial-cube n) +subtype-faces-up-to-standard-simplex n 0 = subtype-face-standard-simplex n 0 +subtype-faces-up-to-standard-simplex n (succ-ℕ k) = + union-subtype + ( subtype-face-standard-simplex n (succ-ℕ k)) + ( subtype-faces-up-to-standard-simplex n k) + +subtype-boundary-standard-simplex : (n : ℕ) → subtype lzero (simplicial-cube n) +subtype-boundary-standard-simplex n = subtype-faces-up-to-standard-simplex n n + +∂Δ : ℕ → UU +∂Δ = type-subtype ∘ subtype-boundary-standard-simplex + +-- TODO: add alternative definition +``` + +### The boundary of the standard simplex is included in the standard simplex + +```agda +leq-subtype-auxillary-face-standard-simplex-standard-simplex : + (n r : ℕ) → + subtype-auxillary-face-standard-simplex n r ⊆ subtype-standard-simplex n +leq-subtype-auxillary-face-standard-simplex-standard-simplex 1 0 _ _ = + star +leq-subtype-auxillary-face-standard-simplex-standard-simplex 2 0 (x , y) = + leq-inv-eq-𝟚 +leq-subtype-auxillary-face-standard-simplex-standard-simplex 2 1 _ = + min-leq-eq-𝟚 +leq-subtype-auxillary-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) 0 _ (x=y , d) = + ( leq-inv-eq-𝟚 x=y , d) +leq-subtype-auxillary-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) (succ-ℕ r) (_ , u) (x≥y , f) = + ( x≥y , + leq-subtype-auxillary-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ n)) r u f) + +leq-subtype-face-standard-simplex-standard-simplex : + (n r : ℕ) → subtype-face-standard-simplex n r ⊆ subtype-standard-simplex n +leq-subtype-face-standard-simplex-standard-simplex 1 0 _ _ = + star +leq-subtype-face-standard-simplex-standard-simplex 1 (succ-ℕ r) _ _ = + star +leq-subtype-face-standard-simplex-standard-simplex 2 0 _ = + max-leq-eq-𝟚 ∘ inv +leq-subtype-face-standard-simplex-standard-simplex 2 (succ-ℕ r) = + leq-subtype-auxillary-face-standard-simplex-standard-simplex 2 r +leq-subtype-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) 0 _ (1=x , p) = + ( max-leq-eq-𝟚 (inv 1=x) , p) +leq-subtype-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) (succ-ℕ r) = + leq-subtype-auxillary-face-standard-simplex-standard-simplex + ( succ-ℕ (succ-ℕ (succ-ℕ n))) + ( r) + +leq-subtype-faces-up-to-standard-simplex-standard-simplex : + (n k : ℕ) → + subtype-faces-up-to-standard-simplex n k ⊆ subtype-standard-simplex n +leq-subtype-faces-up-to-standard-simplex-standard-simplex n 0 = + leq-subtype-face-standard-simplex-standard-simplex n 0 +leq-subtype-faces-up-to-standard-simplex-standard-simplex n (succ-ℕ k) x = + elim-disjunction + ( subtype-standard-simplex n x) + ( leq-subtype-face-standard-simplex-standard-simplex n (succ-ℕ k) x) + ( leq-subtype-faces-up-to-standard-simplex-standard-simplex n k x) + +leq-subtype-boundary-standard-simplex-standard-simplex : + (n : ℕ) → subtype-boundary-standard-simplex n ⊆ subtype-standard-simplex n +leq-subtype-boundary-standard-simplex-standard-simplex n = + leq-subtype-faces-up-to-standard-simplex-standard-simplex n n + +inclusion-boundary-standard-simplex : (n : ℕ) → ∂Δ n → Δ n +inclusion-boundary-standard-simplex n = + tot (leq-subtype-boundary-standard-simplex-standard-simplex n) +``` + +## Properties + +### The standard 𝑛-simplex is a retract of the directed 𝑛-cube + +This remains to be formalized. Lemma 4.2.2 {{#cite MR23b}} + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/comma-types.lagda.md b/src/simplicial-type-theory/comma-types.lagda.md new file mode 100644 index 0000000000..ac1e7ba0b3 --- /dev/null +++ b/src/simplicial-type-theory/comma-types.lagda.md @@ -0,0 +1,126 @@ +# Comma types + +```agda +module simplicial-type-theory.comma-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.cones-over-cospan-diagrams +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.pullbacks +open import foundation.standard-pullbacks +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + + + +## Definitions + +### The standard simplicial comma type + +```agda +simplicial-comma-type : + {l1 l2 l3 : Level} + {A : UU l1} {B : UU l2} {C : UU l3} → + (B → A) → (C → A) → UU (l1 ⊔ l2 ⊔ l3) +simplicial-comma-type {B = B} {C} f g = + Σ B (λ b → Σ C (λ c → simplicial-hom (f b) (g c))) + +_↓▵_ = simplicial-comma-type +``` + +## Properties + +### The universal property of the simplicial comma type + +The comma type `f ↓▵ g` is the pullback in the following diagram + +```text + f ↓▵ g --------> A^𝟚 + | ⌟ | + | | (source , target) + ∨ ∨ + B × C --------> A × A + f × g +``` + +```agda +module _ + {l1 l2 l3 : Level} + {A : UU l1} {B : UU l2} {C : UU l3} + (f : B → A) (g : C → A) + where + + cone-simplicial-comma-type : + cone + {A = B × C} {𝟚 → A} {A × A} + ( λ (b , c) → (f b , g c)) + ( λ α → (α 0₂ , α 1₂)) + ( f ↓▵ g) + pr1 (cone-simplicial-comma-type) (b , c , _) = (b , c) + pr1 (pr2 (cone-simplicial-comma-type)) (_ , _ , α , _) = α + pr2 (pr2 (cone-simplicial-comma-type)) (_ , _ , _ , α0=fb , α1=gc) = + inv (eq-pair α0=fb α1=gc) + + gap-simplicial-comma-type : + f ↓▵ g → standard-pullback (λ (b , c) → (f b , g c)) (λ α → α 0₂ , α 1₂) + gap-simplicial-comma-type = + gap + ( λ (b , c) → (f b , g c)) + ( λ α → (α 0₂ , α 1₂)) + ( cone-simplicial-comma-type) + + map-inv-gap-simplicial-comma-type : + ( standard-pullback + {A = B × C} {𝟚 → A} {A × A} + ( λ (b , c) → (f b , g c)) + ( λ α → α 0₂ , α 1₂)) → + f ↓▵ g + map-inv-gap-simplicial-comma-type ((b , c) , α , coh) = + ( b , c , α , pair-eq (inv coh)) + + is-section-gap-simplicial-comma-type : + map-inv-gap-simplicial-comma-type ∘ gap-simplicial-comma-type ~ id + is-section-gap-simplicial-comma-type (b , c , α , coh) = + eq-pair-eq-fiber + ( eq-pair-eq-fiber + ( eq-pair-eq-fiber + ( ap pair-eq (inv-inv (eq-pair' coh)) ∙ is-retraction-pair-eq coh))) + + is-retraction-gap-simplicial-comma-type : + gap-simplicial-comma-type ∘ map-inv-gap-simplicial-comma-type ~ id + is-retraction-gap-simplicial-comma-type ((b , c) , α , coh) = + eq-pair-eq-fiber + ( eq-pair-eq-fiber + ( ap inv (is-section-pair-eq (inv coh)) ∙ inv-inv coh)) + + is-pullback-simplicial-comma-type : + is-pullback + ( λ (b , c) → (f b , g c)) + ( λ α → (α 0₂ , α 1₂)) + ( cone-simplicial-comma-type) + is-pullback-simplicial-comma-type = + is-equiv-is-invertible + ( map-inv-gap-simplicial-comma-type) + ( is-retraction-gap-simplicial-comma-type) + ( is-section-gap-simplicial-comma-type) +``` diff --git a/src/simplicial-type-theory/compositions-of-directed-edges.lagda.md b/src/simplicial-type-theory/compositions-of-directed-edges.lagda.md new file mode 100644 index 0000000000..02123c6c3b --- /dev/null +++ b/src/simplicial-type-theory/compositions-of-directed-edges.lagda.md @@ -0,0 +1,366 @@ +# Compositions of directed edges + +```agda +module simplicial-type-theory.compositions-of-directed-edges where +``` + +
Imports + +```agda +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-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.retractions +open import foundation.sections +open import foundation.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.truncated-types +open import foundation.truncation-levels +open import foundation.type-arithmetic-booleans +open import foundation.type-arithmetic-cartesian-product-types +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a pair of consecutive +[directed edges](simplicial-type-theory.directed-edges.md) + +```text + f g + x ----> y ----> z +``` + +in a type `A`, then a +{{#concept "composite" Disambiguation="of directed edges in a simplicial type"}} +is a 2-simplex + +```text + σ : Δ² → A +``` + +such that the restriction along the first axis is `f` and the restriction along +the other axis if `g`. + +## Definition + +```agda +module _ + {l : Level} {A : UU l} {x y z : A} + where + + composition-simplicial-hom : hom▵ y z → hom▵ x y → UU l + composition-simplicial-hom g f = {! Σ Δ² ? !} +``` + +A composition of two arrows `f : x → y` and `g: y → z` in a type `A` is a +2-simplex that restricts on the boundary to `f` and `g` as follows. + +```md + z + ^^ + /..| + /....g + /......| + x - f -> y +``` + +The diagonal arrow is then a composite of `g` after `f`. + +## Definitions + +### Compositions + +```agda +dependent-composition-horn : + {l : Level} (A : Δ 2 → UU l) → ((u : Λ²₁) → A (Λ²₁→Δ² u)) → UU l +dependent-composition-horn A = extension-dependent-type Λ²₁→Δ² A + +module _ + {l : Level} {A : UU l} + where + + composition-horn : (Λ²₁ → A) → UU l + composition-horn = dependent-composition-horn (λ _ → A) + + composition-arr : (f g : arr A) → f 1₂ = g 0₂ → UU l + composition-arr f g p = composition-horn (rec-arr-Λ²₁ f g p) + + composition : {x y z : A} → hom x y → hom y z → UU l + composition f g = composition-horn (rec-hom-Λ²₁ f g) +``` + +### Composition witnesses + +```agda +module _ + {l : Level} {A : UU l} + where + + witness-composition-horn : + {fg : Λ²₁ → A} → composition-horn fg → Δ 2 → A + witness-composition-horn = pr1 + + witness-composition-arr : + {f g : arr A} {p : f 1₂ = g 0₂} → composition-arr f g p → Δ 2 → A + witness-composition-arr = pr1 + + witness-composition : + {x y z : A} (f : hom x y) (g : hom y z) → composition f g → Δ 2 → A + witness-composition f g = pr1 +``` + +### Extension witnesses compositions + +```agda +module _ + {l : Level} {A : UU l} + where + + htpy-composition-horn : + {fg : Λ²₁ → A} (c : composition-horn fg) → + fg ~ witness-composition-horn c ∘ Λ²₁→Δ² + htpy-composition-horn = pr2 + + htpy-composition-arr : + {f g : arr A} {p : f 1₂ = g 0₂} (c : composition-arr f g p) → + rec-arr-Λ²₁ f g p ~ witness-composition-arr c ∘ Λ²₁→Δ² + htpy-composition-arr = pr2 + + htpy-composition : + {x y z : A} (f : hom x y) (g : hom y z) (c : composition f g) → + rec-hom-Λ²₁ f g ~ witness-composition f g c ∘ Λ²₁→Δ² + htpy-composition f g = pr2 +``` + +### Composites + +```agda +module _ + {l : Level} {A : UU l} + where + + arr-composite-composition-horn : + (fg : Λ²₁ → A) → composition-horn fg → arr A + arr-composite-composition-horn fg c t = + witness-composition-horn c ((t , t) , refl-≤) + + arr-composite-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) → composition-arr f g p → arr A + arr-composite-composition-arr f g p = + arr-composite-composition-horn (rec-arr-Λ²₁ f g p) + + eq-source-arr-composite-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) (c : composition-arr f g p) → + arr-composite-composition-arr f g p c 0₂ = f 0₂ + eq-source-arr-composite-composition-arr f g p c = + ( ap (witness-composition-arr c) (eq-type-subtype (subtype-Δ 2) refl)) ∙ + ( inv (htpy-composition-arr c ((0₂ , 0₂) , inl-join refl))) ∙ + ( compute-inl-cogap-join _ refl) + + eq-target-arr-composite-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) (c : composition-arr f g p) → + arr-composite-composition-arr f g p c 1₂ = g 1₂ + eq-target-arr-composite-composition-arr f g p c = + ( ap (witness-composition-arr c) (eq-type-subtype (subtype-Δ 2) refl)) ∙ + ( inv (htpy-composition-arr c ((1₂ , 1₂) , inr-join refl))) ∙ + ( compute-inr-cogap-join _ refl) + + composite-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) → composition-arr f g p → hom (f 0₂) (g 1₂) + pr1 (composite-composition-arr f g p c) = + arr-composite-composition-arr f g p c + pr1 (pr2 (composite-composition-arr f g p c)) = + eq-source-arr-composite-composition-arr f g p c + pr2 (pr2 (composite-composition-arr f g p c)) = + eq-target-arr-composite-composition-arr f g p c +``` + +```agda + arr-composite-composition : + {x y z : A} (f : hom x y) (g : hom y z) → composition f g → arr A + arr-composite-composition f g = + arr-composite-composition-horn (rec-hom-Λ²₁ f g) + + eq-source-arr-composite-composition : + {x y z : A} (f : hom x y) (g : hom y z) (c : composition f g) → + arr-composite-composition f g c 0₂ = x + eq-source-arr-composite-composition f g c = + ( eq-source-arr-composite-composition-arr + ( arr-hom f) (arr-hom g) (eq-source-target-hom f g) c) ∙ + ( eq-source-hom f) + + eq-target-arr-composite-composition : + {x y z : A} (f : hom x y) (g : hom y z) (c : composition f g) → + arr-composite-composition f g c 1₂ = z + eq-target-arr-composite-composition f g c = + ( eq-target-arr-composite-composition-arr + ( arr-hom f) (arr-hom g) (eq-source-target-hom f g) c) ∙ + ( eq-target-hom g) + + composite-composition : + {x y z : A} (f : hom x y) (g : hom y z) → composition f g → hom x z + pr1 (composite-composition f g c) = + arr-composite-composition-horn (rec-hom-Λ²₁ f g) c + pr1 (pr2 (composite-composition f g c)) = + eq-source-arr-composite-composition f g c + pr2 (pr2 (composite-composition f g c)) = + eq-target-arr-composite-composition f g c +``` + +## Computations + +### Extensionality of compositions + +```agda +module _ + {l : Level} {A : Δ 2 → UU l} + where + + extensionality-composition-horn : + (i : (u : Λ²₁) → A (Λ²₁→Δ² u)) + (c d : dependent-composition-horn A i) → + (c = d) ≃ htpy-extension Λ²₁→Δ² i c d + extensionality-composition-horn = extensionality-extension Λ²₁→Δ² + + eq-htpy-composition-horn : + (i : (u : Λ²₁) → A (Λ²₁→Δ² u)) + (c d : dependent-composition-horn A i) + (H : map-extension c ~ map-extension d) → + coherence-htpy-extension Λ²₁→Δ² i c d H → c = d + eq-htpy-composition-horn = eq-htpy-extension Λ²₁→Δ² + + htpy-eq-composition-horn : + (i : (u : Λ²₁) → A (Λ²₁→Δ² u)) + (c d : dependent-composition-horn A i) → + c = d → htpy-extension Λ²₁→Δ² i c d + htpy-eq-composition-horn = htpy-eq-extension Λ²₁→Δ² +``` + +### Computing with composition witnesses + +```agda +module _ + {l : Level} {A : UU l} + where + + compute-first-witness-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) (c : composition-arr f g p) → + (t : 𝟚) {r : predicate-Δ 2 (t , 0₂)} → + witness-composition-arr c ((t , 0₂) , r) = f t + compute-first-witness-composition-arr f g p c t = + ( ap + ( λ r → witness-composition-arr c ((t , 0₂) , r)) + ( eq-is-in-subtype (subtype-Δ 2))) ∙ + ( inv (pr2 c ((t , 0₂) , inl-join refl))) ∙ + ( compute-first-rec-arr-Λ²₁ f g p t) + + compute-second-witness-composition-arr : + (f g : arr A) (p : f 1₂ = g 0₂) (c : composition-arr f g p) → + (t : 𝟚) {r : predicate-Δ 2 (1₂ , t)} → + witness-composition-arr c ((1₂ , t) , r) = g t + compute-second-witness-composition-arr f g p c t = + ( ap + ( λ r → witness-composition-arr c ((1₂ , t) , r)) + ( eq-is-in-subtype (subtype-Δ 2))) ∙ + ( inv (pr2 c ((1₂ , t) , inr-join refl))) ∙ + ( compute-second-rec-arr-Λ²₁ f g p t) + + compute-first-witness-composition : + {x y z : A} (f : hom x y) (g : hom y z) (c : composition f g) → + (t : 𝟚) {r : predicate-Δ 2 (t , 0₂)} → + witness-composition f g c ((t , 0₂) , r) = arr-hom f t + compute-first-witness-composition f g = + compute-first-witness-composition-arr + (arr-hom f) (arr-hom g) (eq-source-target-hom f g) + + compute-second-witness-composition : + {x y z : A} (f : hom x y) (g : hom y z) (c : composition f g) → + (t : 𝟚) {r : predicate-Δ 2 (1₂ , t)} → + witness-composition f g c ((1₂ , t) , r) = arr-hom g t + compute-second-witness-composition f g = + compute-second-witness-composition-arr + (arr-hom f) (arr-hom g) (eq-source-target-hom f g) +``` + +TODO: move part below + +### The `is-composite` family + +An arrow `h` is the **composite** of `f` and `g` if there is a composition of +`f` and `g` such that their composite is equal to `h`. + +```agda +module _ + {l : Level} {A : UU l} (fg : Λ²₁ → A) (h : 𝟚 → A) + where + + is-composite-horn : UU l + is-composite-horn = + Σ (composition-horn fg) (λ c → arr-composite-composition-horn fg c = h) + + triangle-horn : + fg ((0₂ , 0₂) , inl-join refl) = h 0₂ → + fg ((1₂ , 1₂) , inr-join refl) = h 1₂ → + ∂Δ² → A + triangle-horn h0 h1 = + rec-arr-∂Δ² + ( λ t → fg ((t , 0₂) , inl-join refl)) + ( λ s → fg ((1₂ , s) , inr-join refl)) + ( h) + ( h0) + ( ap (λ p → fg ((1₂ , 0₂) , p)) (glue-join (refl , refl))) + ( h1) + + is-composite-horn' : UU l + is-composite-horn' = + Σ ( ( fg ((0₂ , 0₂) , inl-join refl) = h 0₂) × + ( fg ((1₂ , 1₂) , inr-join refl) = h 1₂)) + ( λ (h0 , h1) → extension ∂Δ²→Δ² (triangle-horn h0 h1)) + +hom² : + {l : Level} {A : UU l} {x y z : A} → + hom x y → hom y z → hom x z → UU l +hom² f g h = extension ∂Δ²→Δ² (rec-hom-∂Δ² f g h) + +-- hom²-composition : +-- {l : Level} {A : UU l} {x y z : A} +-- (f : hom x y) (g : hom y z) (c : composition f g) → hom² f g (composite-composition f g c) +-- pr1 (hom²-composition f g c) = witness-composition f g c +-- pr2 (hom²-composition f g c) x = {! !} +``` + +```agda +is-composite-arr : + {l : Level} {A : UU l} → (f g : arr A) → f 1₂ = g 0₂ → arr A → UU l +is-composite-arr f g p h = is-composite-horn (rec-arr-Λ²₁ f g p) h +``` + +These definitions are not compatible in the same way as the previous ones, as +the second formulation also requires coherence at the end points. + +```agda +is-composite : + {l : Level} {A : UU l} {x y z : A} → hom x y → hom y z → hom x z → UU l +is-composite f g h = Σ (composition f g) (λ c → composite-composition f g c = h) +``` diff --git a/src/simplicial-type-theory/conormed-maps-between-types.lagda.md b/src/simplicial-type-theory/conormed-maps-between-types.lagda.md new file mode 100644 index 0000000000..3ff60350e5 --- /dev/null +++ b/src/simplicial-type-theory/conormed-maps-between-types.lagda.md @@ -0,0 +1,230 @@ +# Conormed maps between types + +```agda +module simplicial-type-theory.conormed-maps-between-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.commuting-triangles-of-identifications +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.functoriality-dependent-function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.postcomposition-functions +open import foundation.precomposition-functions +open import foundation.univalence +open import foundation.universal-property-equivalences +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition +open import foundation.whiskering-identifications-concatenation + +open import foundation-core.equivalences +open import foundation-core.homotopies + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-edges-dependent-pair-types +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.horizontal-composition-natural-transformations +open import simplicial-type-theory.natural-transformations +open import simplicial-type-theory.transposing-adjunctions-between-types +open import simplicial-type-theory.transposing-biadjunctions-between-types +``` + +
+ +## Idea + +Given a map of types `q : A → B`, that has a left and a right adjoint + +```text + q! ⊣ q ⊣ q∗ +``` + +then we say `q` is +{{#concept "conormed" Disambiguation="map of simplicial types"}} if there's a +natural transformation + +```text + Nm_q : q∗ ⇒ q!. +``` + +## Definitions + +### Conormed maps between types + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-conormed-map : (A → B) → UU (l1 ⊔ l2) + is-conormed-map q = + Σ ( is-transposing-biadjoint q) + ( λ Q → + ( map-right-adjoint-is-transposing-biadjoint Q) ⇒▵ + ( map-left-adjoint-is-transposing-biadjoint Q)) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-conormed-map q) + where + + is-transposing-biadjoint-is-conormed-map : is-transposing-biadjoint q + is-transposing-biadjoint-is-conormed-map = pr1 H + + is-transposing-right-adjoint-is-conormed-map : + is-transposing-right-adjoint q + is-transposing-right-adjoint-is-conormed-map = + is-transposing-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + is-transposing-left-adjoint-is-conormed-map : + is-transposing-left-adjoint q + is-transposing-left-adjoint-is-conormed-map = + is-transposing-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + map-left-adjoint-is-conormed-map : B → A + map-left-adjoint-is-conormed-map = + map-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + is-transposing-adjoint-map-left-adjoint-is-conormed-map : + is-transposing-adjunction map-left-adjoint-is-conormed-map q + is-transposing-adjoint-map-left-adjoint-is-conormed-map = + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + unit-left-adjoint-is-conormed-map : + id ⇒▵ q ∘ map-left-adjoint-is-conormed-map + unit-left-adjoint-is-conormed-map = + unit-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + counit-left-adjoint-is-conormed-map : + map-left-adjoint-is-conormed-map ∘ q ⇒▵ id + counit-left-adjoint-is-conormed-map = + counit-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + map-right-adjoint-is-conormed-map : B → A + map-right-adjoint-is-conormed-map = + map-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + is-transposing-adjoint-map-right-adjoint-is-conormed-map : + is-transposing-adjunction q map-right-adjoint-is-conormed-map + is-transposing-adjoint-map-right-adjoint-is-conormed-map = + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + unit-right-adjoint-is-conormed-map : + id ⇒▵ map-right-adjoint-is-conormed-map ∘ q + unit-right-adjoint-is-conormed-map = + unit-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + counit-right-adjoint-is-conormed-map : + q ∘ map-right-adjoint-is-conormed-map ⇒▵ id + counit-right-adjoint-is-conormed-map = + counit-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-conormed-map + + conorm-map-is-conormed-map : + map-right-adjoint-is-conormed-map ⇒▵ map-left-adjoint-is-conormed-map + conorm-map-is-conormed-map = pr2 H +``` + +### The wrong way unit associated to a conormed map + +Given a transposing biadjoint `q! ⊣ q* ⊣ q∗` and a conorm map +`Conm_q : q∗ ⇒ q!`, we have the _wrong way unit_ + +```text + η∗ Conm_q + μ : id =====> q∗q* =====> q!q*. +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} + (H : is-conormed-map q) + where + + -- TODO +``` + +### The integral map associated to a conormed map + +Here we need composition. + +```agda + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-conormed-map q) + where + + integral-map-is-conormed-map : (x y : A) → q x →₂ q y → x →₂ y + integral-map-is-conormed-map x y f = + map-equiv + {! is-transposing-adjoint-map-left-adjoint-is-conormed-map H (q x) y ∘e ? !} + {! action-simplicial-hom-simplicial-natural-transformation (conorm-map-is-conormed-map H) f !} + -- λ t → conorm-map-is-conormed-map H (simplicial-arrow-simplicial-hom f t) -- action-simplicial-hom-function (map-right-adjoint-is-conormed-map H) f + -- TODO +``` + +### The cardinality operator associated to a conormed map + +Here we need composition. + +```agda + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-conormed-map q) + where + + -- TODO +``` + +## Properties + +### Composition of conormed maps + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {p : A → B} {q : B → C} + where + + is-conormed-map-comp : + is-conormed-map q → is-conormed-map p → is-conormed-map (q ∘ p) + is-conormed-map-comp (Q , Conm-q) (P , Conm-p) = + ( is-transposing-biadjoint-comp Q P) , + ( horizontal-comp-simplicial-natural-transformation Conm-p Conm-q) +``` + +## Examples + +### The identity conormed map + +```agda +is-conormed-map-id : {l : Level} {A : UU l} → is-conormed-map (id {A = A}) +pr1 is-conormed-map-id = is-transposing-biadjoint-id +pr2 is-conormed-map-id = id-simplicial-hom +``` + +### Equivalences are conormed maps + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} (H : is-equiv f) + where + + is-conormed-map-is-equiv : is-conormed-map f + pr1 is-conormed-map-is-equiv = is-transposing-biadjoint-is-equiv H + pr2 is-conormed-map-is-equiv = id-simplicial-hom ∘ map-inv-is-equiv H +``` + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/dependent-directed-edges.lagda.md b/src/simplicial-type-theory/dependent-directed-edges.lagda.md new file mode 100644 index 0000000000..1925940018 --- /dev/null +++ b/src/simplicial-type-theory/dependent-directed-edges.lagda.md @@ -0,0 +1,239 @@ +# Dependent directed edges + +```agda +module simplicial-type-theory.dependent-directed-edges where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-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.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a type family `B : A → 𝒰` and a +[directed edge](simplicial-type-theory.directed-edges.md) `α : x →₂ y` in `A`, a +{{#concept "dependent directed edge" Disambiguation="simplicial type theory" Agda=dependent-simplicial-hom}} +_over_ `α` from `x'` to `y'` is a simplicial arrow `β` in `B ∘ α : 𝟚 → 𝒰`. such +that `β 0₂ = x'` over the identification `α 0₂ = x` and `β 1₂ = y'` over the +identification `α 1₂ = y`. + +Assuming for simplicity that the endpoints are strict, i.e., `α 0₂ ≐ x` and +`α 1₂ ≐ y`, the situation can be drawn as in the following diagram. The +dependent arrow `β` lives in the dependent type `B` varying over `A`, pointing +from an element in the fiber `B x` to an element in the fiber `B y`. On each end +of the dependent arrow, there is a correcting identification within the +respective fiber of `B` + +```text + B x B y + ~~~~~~~~~~~~~~~~~~~~~~~~~ + │ │ │ │ + │ x' ∙ ··········> │ │ B + │ │ β ∙ y' │ + │ │ │ │ + ~~~~~~~~~~~~~~~~~~~~~~~~~ + ↧ ↧ + ───── ∙ ──────────> ∙ ───── A. + x α y +``` + +## Definitions + +### Dependent directed edges + +```agda +module _ + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) + {x y : A} (α : x →₂ y) + where + + dependent-simplicial-hom : B x → B y → UU l2 + dependent-simplicial-hom x' y' = + Σ ( simplicial-arrow' (B ∘ simplicial-arrow-simplicial-hom α)) + ( λ β → + ( dependent-identification B (eq-source-simplicial-hom α) (β 0₂) x') × + ( dependent-identification B (eq-target-simplicial-hom α) (β 1₂) y')) + +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} + {x y : A} (α : x →₂ y) {x' : B x} {y' : B y} + (β : dependent-simplicial-hom B α x' y') + where + + simplicial-arrow-dependent-simplicial-hom : + simplicial-arrow' (B ∘ simplicial-arrow-simplicial-hom α) + simplicial-arrow-dependent-simplicial-hom = pr1 β + + eq-source-dependent-simplicial-hom : + dependent-identification B + ( eq-source-simplicial-hom α) + ( simplicial-arrow-dependent-simplicial-hom 0₂) + ( x') + eq-source-dependent-simplicial-hom = pr1 (pr2 β) + + eq-target-dependent-simplicial-hom : + dependent-identification B + ( eq-target-simplicial-hom α) + ( simplicial-arrow-dependent-simplicial-hom 1₂) + ( y') + eq-target-dependent-simplicial-hom = pr2 (pr2 β) +``` + +### The identity/constant dependent directed edges + +```agda +id-dependent-simplicial-hom : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} {x : A} (x' : B x) → + dependent-simplicial-hom B (id-simplicial-hom x) x' x' +id-dependent-simplicial-hom = id-simplicial-hom +``` + +### Dependent directed edges arising from identifications + +```agda +dependent-simplicial-hom-eq : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} {x y : A} (p : x = y) + {x' : B x} {y' : B y} → dependent-identification B p x' y' → + dependent-simplicial-hom B (simplicial-hom-eq p) x' y' +dependent-simplicial-hom-eq refl = simplicial-hom-eq +``` + +### Characterizing equality of dependent directed edges over a fixed edge + +Given a directed edge `α : x →₂ y` in `A` and elements `x' : B x` and +`y' : B y`, we want to characterize identifications of dependent directed edges +from `x'` to `y'` over `α`. The situation is as follows + +```text + B x B y + ~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~ + │ │ │ │ │ │ + │ │ │ β │ │ │ + │ │ ══════ ∙ ··················> ∙ ══════ │ │ + │ x' ∙ │ │ ∙ y' │ B + │ │ ══════ ∙ ··················> ∙ ══════ │ │ + │ │ │ β' │ │ │ + │ │ │ │ │ │ + ~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~~ + ↧ ↧ ↧ ↧ + ───── ∙ ══════ ∙ ──────────────────> ∙ ====== ∙ ───── A. + x α 0₂ α α 1₂ y +``` + +We are looking for coherence data to fill the area between `x'` and `y'`, and +the area is naturally subdivided into three sections. To fill the middle, we ask +for a homotopy of the underlying dependent arrows `H : β ~ β'`. Now, to also +fill the triangles at the end points, we require further coherences of the +dependent triangles depicted below + +```text + ──── ∙ β 0₂ β 1₂ ∙ ──── + ╱ │ │ ╲ + x' ∙ │ H 0₂ H 1₂ │ ∙ y' + ╲ │ │ ╱ + ──── ∙ β' 0₂ β' 1₂ ∙ ──── + + ↧ ↧ ↧ ↧ + + x ∙ ────── ∙ α 0₂ α 1₂ ∙ ────── ∙ y. +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} + {x y : A} (α : x →₂ y) {x' : B x} {y' : B y} + where + + coherence-htpy-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → + simplicial-arrow-dependent-simplicial-hom α β ~ + simplicial-arrow-dependent-simplicial-hom α β' → + UU l2 + coherence-htpy-dependent-simplicial-hom-over β β' H = + ( ( eq-source-dependent-simplicial-hom α β) = + ( concat-dependent-identification B refl + ( eq-source-simplicial-hom α) + ( H 0₂) + ( eq-source-dependent-simplicial-hom α β'))) × + ( ( eq-target-dependent-simplicial-hom α β) = + ( concat-dependent-identification B refl + ( eq-target-simplicial-hom α) + ( H 1₂) + ( eq-target-dependent-simplicial-hom α β'))) + + htpy-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → UU l2 + htpy-dependent-simplicial-hom-over β β' = + Σ ( simplicial-arrow-dependent-simplicial-hom α β ~ + simplicial-arrow-dependent-simplicial-hom α β') + ( coherence-htpy-dependent-simplicial-hom-over β β') + + refl-htpy-dependent-simplicial-hom-over : + (β : dependent-simplicial-hom B α x' y') → + htpy-dependent-simplicial-hom-over β β + refl-htpy-dependent-simplicial-hom-over β = (refl-htpy , refl , refl) + + htpy-eq-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → + β = β' → htpy-dependent-simplicial-hom-over β β' + htpy-eq-dependent-simplicial-hom-over β .β refl = + refl-htpy-dependent-simplicial-hom-over β + + is-torsorial-htpy-dependent-simplicial-hom-over : + (β : dependent-simplicial-hom B α x' y') → + is-torsorial (htpy-dependent-simplicial-hom-over β) + is-torsorial-htpy-dependent-simplicial-hom-over β = + is-torsorial-Eq-structure + ( is-torsorial-htpy (simplicial-arrow-dependent-simplicial-hom α β)) + ( simplicial-arrow-dependent-simplicial-hom α β , refl-htpy) + ( is-torsorial-Eq-structure + ( is-torsorial-Id (eq-source-dependent-simplicial-hom α β)) + ( eq-source-dependent-simplicial-hom α β , refl) + ( is-torsorial-Id (eq-target-dependent-simplicial-hom α β))) + + is-equiv-htpy-eq-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → + is-equiv (htpy-eq-dependent-simplicial-hom-over β β') + is-equiv-htpy-eq-dependent-simplicial-hom-over β = + fundamental-theorem-id + ( is-torsorial-htpy-dependent-simplicial-hom-over β) + ( htpy-eq-dependent-simplicial-hom-over β) + + extensionality-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → + (β = β') ≃ htpy-dependent-simplicial-hom-over β β' + extensionality-dependent-simplicial-hom-over β β' = + ( htpy-eq-dependent-simplicial-hom-over β β' , + is-equiv-htpy-eq-dependent-simplicial-hom-over β β') + + eq-htpy-dependent-simplicial-hom-over : + (β β' : dependent-simplicial-hom B α x' y') → + htpy-dependent-simplicial-hom-over β β' → β = β' + eq-htpy-dependent-simplicial-hom-over β β' = + map-inv-equiv (extensionality-dependent-simplicial-hom-over β β') +``` diff --git a/src/simplicial-type-theory/directed-circle.lagda.md b/src/simplicial-type-theory/directed-circle.lagda.md new file mode 100644 index 0000000000..4a359fae0f --- /dev/null +++ b/src/simplicial-type-theory/directed-circle.lagda.md @@ -0,0 +1,286 @@ +# The directed circle + +```agda +module simplicial-type-theory.directed-circle where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-dependent-functions +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.injective-maps +open import foundation.negated-equality +open import foundation.propositions +open import foundation.retractions +open import foundation.sections +open import foundation.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import reflection.erasing-equality + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.free-directed-loops +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicially-discrete-types +open import simplicial-type-theory.universal-property-directed-circle + +open import synthetic-homotopy-theory.circle +``` + +
+ +## Idea + +The {{#concept "directed circle"}} is the type consisting of a point `*` and a +nontrivial [directed edge](simplicial-type-theory.directed-edges.md) `* →₂ *`. +The directed circle classifies +[free directed loops](simplicial-type-theory.free-directed-loops.md), meaning +that maps `directed-circle → X` are in correspondence with free directed loops +of `X`. + +The directed circle also satisfies the +[universal property of the coequalizer](synthetic-homotopy-theory.universal-property-coequalizers.md) +of the diagram + +```text + 0₂ + -------> + 1 𝟚. + -------> + 1₂ +``` + +In +[`rewriting-directed-circle`](simplicial-type-theory.rewriting-directed-circle.md) +we make the directed circle the strict coequalizer of this diagram. + +## Postulates + +### The standard directed circle type + +```agda +postulate + directed-circle : UU lzero + +postulate + arrow-directed-circle : simplicial-arrow directed-circle + +base-directed-circle : directed-circle +base-directed-circle = arrow-directed-circle 0₂ + +postulate + compute-target-arrow-directed-circle : + arrow-directed-circle 1₂ = base-directed-circle + +compute-source-arrow-directed-circle : + arrow-directed-circle 0₂ = base-directed-circle +compute-source-arrow-directed-circle = refl + +eq-source-target-arrow-directed-circle : + arrow-directed-circle 1₂ = arrow-directed-circle 0₂ +eq-source-target-arrow-directed-circle = + compute-target-arrow-directed-circle + +eq-target-source-arrow-directed-circle : + arrow-directed-circle 0₂ = arrow-directed-circle 1₂ +eq-target-source-arrow-directed-circle = + inv compute-target-arrow-directed-circle + +free-loop-directed-circle : free-directed-loop directed-circle +free-loop-directed-circle = + ( arrow-directed-circle , eq-source-target-arrow-directed-circle) + +loop-directed-circle : base-directed-circle →₂ base-directed-circle +loop-directed-circle = + ( arrow-directed-circle , + compute-source-arrow-directed-circle , + compute-target-arrow-directed-circle) +``` + +### The induction principle of the standard directed circle + +```agda +module _ + {l : Level} (P : directed-circle → UU l) + where + + postulate + ind-directed-circle : + free-dependent-directed-loop free-loop-directed-circle P → + (x : directed-circle) → P x + + compute-arrow-ind-directed-circle : + (α : free-dependent-directed-loop free-loop-directed-circle P) (t : 𝟚) → + ind-directed-circle α (arrow-directed-circle t) = + arrow-free-dependent-directed-loop free-loop-directed-circle α t + compute-arrow-ind-directed-circle α t = + primEraseEquality compute-arrow-ind-directed-circle' + where postulate + compute-arrow-ind-directed-circle' : + ind-directed-circle α (arrow-directed-circle t) = + arrow-free-dependent-directed-loop free-loop-directed-circle α t + + postulate + coherence-ind-directed-circle : + (α : free-dependent-directed-loop free-loop-directed-circle P) → + coherence-htpy-free-dependent-directed-loop + ( free-loop-directed-circle) + ( P) + ( ev-free-directed-loop-Π free-loop-directed-circle P + ( ind-directed-circle α)) + ( α) + ( compute-arrow-ind-directed-circle α) + + compute-ind-directed-circle : + is-section + ( ev-free-directed-loop-Π free-loop-directed-circle P) + ( ind-directed-circle) + compute-ind-directed-circle α = + eq-htpy-free-dependent-directed-loop free-loop-directed-circle P + ( ev-free-directed-loop-Π free-loop-directed-circle P + ( ind-directed-circle α)) + ( α) + ( compute-arrow-ind-directed-circle α , coherence-ind-directed-circle α) + +induction-principle-directed-circle' : + induction-principle-directed-circle free-loop-directed-circle +induction-principle-directed-circle' P = + ( ind-directed-circle P , compute-ind-directed-circle P) +``` + +## Definitions + +### The dependent universal property of the directed circle + +```agda +dup-directed-circle : + dependent-universal-property-directed-circle free-loop-directed-circle +dup-directed-circle = + dependent-universal-property-induction-principle-directed-circle + ( free-loop-directed-circle) + ( induction-principle-directed-circle') +``` + +### The universal property of the directed circle + +```agda +up-directed-circle : + universal-property-directed-circle free-loop-directed-circle +up-directed-circle = + universal-property-induction-principle-directed-circle + ( free-loop-directed-circle) + ( induction-principle-directed-circle') +``` + +### The recursion principle of the directed circle + +```agda +module _ + {l : Level} {X : UU l} + where + + rec-directed-circle : free-directed-loop X → directed-circle → X + rec-directed-circle α = + ind-directed-circle + ( λ _ → X) + ( map-compute-free-dependent-directed-loop-constant-type-family + ( free-loop-directed-circle) + ( X) + ( α)) + + compute-arrow-rec-directed-circle : + (α : free-directed-loop X) → + rec-directed-circle α ∘ arrow-directed-circle ~ + arrow-free-directed-loop α + compute-arrow-rec-directed-circle α = + compute-arrow-ind-directed-circle + ( λ _ → X) + ( map-compute-free-dependent-directed-loop-constant-type-family + ( free-loop-directed-circle) + ( X) + ( α)) +``` + +## Properties + +### The directed circle as a coequalizer + +The directed circle satisfies the universal property of the coequalizer of the +diagram + +```text + 0₂ + -------> + 1 𝟚. + -------> + 1₂ +``` + +This remains to be formalized. + +### The canonical comparison map to the homotopical circle + +```agda +map-directed-circle-circle : directed-circle → 𝕊¹ +map-directed-circle-circle = + rec-directed-circle (free-directed-loop-free-loop free-loop-𝕊¹) + +compute-map-directed-circle-circle-id-arrow : + (x : directed-circle) → + map-directed-circle-circle ∘ id-simplicial-arrow x ~ + id-simplicial-arrow (map-directed-circle-circle x) +compute-map-directed-circle-circle-id-arrow x = refl-htpy +``` + +### The loop of the directed circle is nontrivial + +```agda +module _ + (is-discrete-𝕊¹ : is-simplicially-discrete 𝕊¹) + where + + is-nontrivial-loop-simplicial-hom-𝕊¹ : + simplicial-hom-eq loop-𝕊¹ ≠ id-simplicial-hom base-𝕊¹ + is-nontrivial-loop-simplicial-hom-𝕊¹ p = + is-nontrivial-loop-𝕊¹ + ( is-injective-is-equiv (is-discrete-𝕊¹ base-𝕊¹ base-𝕊¹) p) +``` + +Steps: + +- construct computation on edges of the recursor of the directed circle +- show that the loop of the directed circle maps to `simplicial-hom-eq loop-𝕊¹` + +```agda + -- is-nontrivial-loop-directed-circle : + -- loop-directed-circle ≠ id-simplicial-hom base-directed-circle + -- is-nontrivial-loop-directed-circle p = + -- is-nontrivial-loop-simplicial-hom-𝕊¹ + -- {! ? ∙ ap (action-simplicial-hom-function map-directed-circle-circle) p ∙ ? !} +``` + +It remains to formalize that the circle is simplicially discrete. Note that the +proof only uses that `simplicial-hom-eq` is injective. diff --git a/src/simplicial-type-theory/directed-edges-booleans.lagda.md b/src/simplicial-type-theory/directed-edges-booleans.lagda.md new file mode 100644 index 0000000000..a02cd2f173 --- /dev/null +++ b/src/simplicial-type-theory/directed-edges-booleans.lagda.md @@ -0,0 +1,138 @@ +# Directed edges in the booleans + +```agda +module simplicial-type-theory.directed-edges-booleans where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.coproduct-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-coproduct-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.negated-equality +open import foundation.negation +open import foundation.raising-universe-levels +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.type-arithmetic-booleans +open import foundation.unit-type +open import foundation.universal-property-booleans +open import foundation.universe-levels + +open import orthogonal-factorization-systems.coproducts-null-types +open import orthogonal-factorization-systems.extensions-of-maps +open import orthogonal-factorization-systems.null-families-of-types +open import orthogonal-factorization-systems.null-maps +open import orthogonal-factorization-systems.null-types +open import orthogonal-factorization-systems.orthogonal-maps + +open import simplicial-type-theory.action-on-directed-edges-dependent-functions +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicially-discrete-types +``` + +
+ +## Idea + +We postulate that the [booleans](foundation.booleans.md) are +[simplicially discrete](simplicial-type-theory.simplicially-discrete-types.md) +and hence that all its +[directed edges](simplicial-type-theory.directed-edges.md) are constant. This +refutes the models of the +[directed interval](simplicial-type-theory.directed-interval.md) in the booleans +and is a distinguishing property between the two. + +As a corollary we have that the universe of simplicially discrete types is +closed under finite coproducts and contains the finite types. + +## Postulate + +```agda +postulate + is-simplicially-discrete-bool : is-simplicially-discrete bool +``` + +## Properties + +### The booleans are 𝟚-null + +```agda +is-𝟚-null-bool : is-null 𝟚 bool +is-𝟚-null-bool = + is-𝟚-null-is-simplicially-discrete is-simplicially-discrete-bool +``` + +### The booleans are not a directed interval + +```agda +is-not-directed-interval-bool' : 𝟚 ≠ bool +is-not-directed-interval-bool' = + nonequal-leibniz' + ( is-simplicially-discrete) + ( 𝟚) + ( bool) + ( is-not-simplicially-discrete-𝟚) + ( is-simplicially-discrete-bool) + +is-not-directed-interval-bool : ¬ (𝟚 ≃ bool) +is-not-directed-interval-bool e = + is-not-simplicially-discrete-𝟚 + ( is-simplicially-discrete-equiv e is-simplicially-discrete-bool) + +is-not-retract-of-directed-interval-bool : ¬ (𝟚 retract-of bool) +is-not-retract-of-directed-interval-bool r = + is-not-simplicially-discrete-𝟚 + ( is-simplicially-discrete-retract r is-simplicially-discrete-bool) +``` + +### Coproducts of simplicially discrete types are simplicially discrete + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + (is-disc-A : is-simplicially-discrete A) + (is-disc-B : is-simplicially-discrete B) + where + + is-simplicially-discrete-coproduct : + is-simplicially-discrete A → + is-simplicially-discrete B → + is-simplicially-discrete (A + B) + is-simplicially-discrete-coproduct is-disc-A is-disc-B = + is-simplicially-discrete-is-𝟚-null + ( is-null-coproduct-is-null-bool 𝟚 + ( is-𝟚-null-bool) + ( is-𝟚-null-is-simplicially-discrete is-disc-A) + ( is-𝟚-null-is-simplicially-discrete is-disc-B)) +``` + +### Finite types are simplicially discrete + +> This remains to be formalized. + +### The directed interval is not finite + +> This remains to be formalized. diff --git a/src/simplicial-type-theory/directed-edges-cartesian-product-types.lagda.md b/src/simplicial-type-theory/directed-edges-cartesian-product-types.lagda.md new file mode 100644 index 0000000000..8fb4c95498 --- /dev/null +++ b/src/simplicial-type-theory/directed-edges-cartesian-product-types.lagda.md @@ -0,0 +1,105 @@ +# Directed edges in cartesian product types + +```agda +module simplicial-type-theory.directed-edges-cartesian-product-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-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.retractions +open import foundation.sections +open import foundation.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.action-on-directed-edges-dependent-functions +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +A [directed edge](simplicial-type-theory.directed-edges.md) +`f : (x , x') →₂ (y , y')` in a +[cartesian product type](foundation.dependent-pair-types.md) `A × B` consists of +a directed edge `α : x →₂ y` in `A` and a directed edge `β : x' →₂ y'` in `B`. + +## Properties + +### Characterizing directed edges in cartesian product types + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + simplicial-hom-product : (x y : A × B) → UU (l1 ⊔ l2) + simplicial-hom-product (x , x') (y , y') = (x →₂ y) × (x' →₂ y') + + map-compute-simplicial-hom-product : + {x y : A × B} → x →₂ y → simplicial-hom-product x y + map-compute-simplicial-hom-product α = + ( action-simplicial-hom-function pr1 α , + action-simplicial-hom-function pr2 α) + + map-inv-compute-simplicial-hom-product : + {x y : A × B} → simplicial-hom-product x y → x →₂ y + map-inv-compute-simplicial-hom-product ((α , p , q) , (β , p' , q')) = + ((λ t → (α t , β t)) , eq-pair p p' , eq-pair q q') + + is-section-map-inv-compute-simplicial-hom-product : + {x y : A × B} → + is-section + ( map-compute-simplicial-hom-product {x} {y}) + ( map-inv-compute-simplicial-hom-product) + is-section-map-inv-compute-simplicial-hom-product + ( ( α , refl , refl) , (β , refl , refl)) = + refl + + is-retraction-map-inv-compute-simplicial-hom-product : + {x y : A × B} → + is-retraction + ( map-compute-simplicial-hom-product {x} {y}) + ( map-inv-compute-simplicial-hom-product) + is-retraction-map-inv-compute-simplicial-hom-product (α , refl , refl) = refl + + is-equiv-map-compute-simplicial-hom-product : + {x y : A × B} → is-equiv (map-compute-simplicial-hom-product {x} {y}) + is-equiv-map-compute-simplicial-hom-product = + is-equiv-is-invertible + ( map-inv-compute-simplicial-hom-product) + ( is-section-map-inv-compute-simplicial-hom-product) + ( is-retraction-map-inv-compute-simplicial-hom-product) + + compute-simplicial-hom-product : + {x y : A × B} → (x →₂ y) ≃ simplicial-hom-product x y + compute-simplicial-hom-product = + ( map-compute-simplicial-hom-product , + is-equiv-map-compute-simplicial-hom-product) + + inv-compute-simplicial-hom-product : + {x y : A × B} → simplicial-hom-product x y ≃ (x →₂ y) + inv-compute-simplicial-hom-product = inv-equiv compute-simplicial-hom-product +``` diff --git a/src/simplicial-type-theory/directed-edges-dependent-pair-types.lagda.md b/src/simplicial-type-theory/directed-edges-dependent-pair-types.lagda.md new file mode 100644 index 0000000000..c4141438c7 --- /dev/null +++ b/src/simplicial-type-theory/directed-edges-dependent-pair-types.lagda.md @@ -0,0 +1,117 @@ +# Directed edges in dependent pair types + +```agda +module simplicial-type-theory.directed-edges-dependent-pair-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-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.retractions +open import foundation.sections +open import foundation.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.action-on-directed-edges-dependent-functions +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +A [directed edge](simplicial-type-theory.directed-edges.md) +`f : (x , x') →₂ (y , y')` in a +[dependent pair type](foundation.dependent-pair-types.md) `Σ A B` consists of a +directed edge `α : x →₂ y` in the base `A` and a +[dependent directed edge](simplicial-type-theory.dependent-directed-edges.md) +`β` in `B` over `α` from `x'` to `y'`. + +## Properties + +### Characterizing directed edges in dependent pair types + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} + where + + simplicial-hom-Σ : (x y : Σ A B) → UU (l1 ⊔ l2) + simplicial-hom-Σ (x , x') (y , y') = + Σ (x →₂ y) (λ α → dependent-simplicial-hom B α x' y') + + action-simplicial-hom-dependent-function-pr2' : + {x y : Σ A B} (α : x →₂ y) → + dependent-simplicial-hom B + ( action-simplicial-hom-function pr1 α) + ( pr2 x) + ( pr2 y) + pr1 (action-simplicial-hom-dependent-function-pr2' (α , p , q)) = pr2 ∘ α + pr1 (pr2 (action-simplicial-hom-dependent-function-pr2' (α , refl , q))) = + refl + pr2 (pr2 (action-simplicial-hom-dependent-function-pr2' (α , p , refl))) = + refl + + map-compute-simplicial-hom-Σ : + {x y : Σ A B} → x →₂ y → simplicial-hom-Σ x y + map-compute-simplicial-hom-Σ α = + ( action-simplicial-hom-function pr1 α , + action-simplicial-hom-dependent-function-pr2' α) + + map-inv-compute-simplicial-hom-Σ : + {x y : Σ A B} → simplicial-hom-Σ x y → x →₂ y + map-inv-compute-simplicial-hom-Σ ((α , p , q) , (β , p' , q')) = + ((λ t → (α t , β t)) , eq-pair-Σ p p' , eq-pair-Σ q q') + + is-section-map-inv-compute-simplicial-hom-Σ : + {x y : Σ A B} → + is-section + ( map-compute-simplicial-hom-Σ {x} {y}) + ( map-inv-compute-simplicial-hom-Σ) + is-section-map-inv-compute-simplicial-hom-Σ + ( (α , refl , refl) , (β , refl , refl)) = + refl + + is-retraction-map-inv-compute-simplicial-hom-Σ : + {x y : Σ A B} → + is-retraction + ( map-compute-simplicial-hom-Σ {x} {y}) + ( map-inv-compute-simplicial-hom-Σ) + is-retraction-map-inv-compute-simplicial-hom-Σ (α , refl , refl) = refl + + is-equiv-map-compute-simplicial-hom-Σ : + {x y : Σ A B} → is-equiv (map-compute-simplicial-hom-Σ {x} {y}) + is-equiv-map-compute-simplicial-hom-Σ = + is-equiv-is-invertible + ( map-inv-compute-simplicial-hom-Σ) + ( is-section-map-inv-compute-simplicial-hom-Σ) + ( is-retraction-map-inv-compute-simplicial-hom-Σ) + + compute-simplicial-hom-Σ : {x y : Σ A B} → (x →₂ y) ≃ simplicial-hom-Σ x y + compute-simplicial-hom-Σ = + ( map-compute-simplicial-hom-Σ , is-equiv-map-compute-simplicial-hom-Σ) + + inv-compute-simplicial-hom-Σ : + {x y : Σ A B} → simplicial-hom-Σ x y ≃ (x →₂ y) + inv-compute-simplicial-hom-Σ = inv-equiv compute-simplicial-hom-Σ +``` diff --git a/src/simplicial-type-theory/directed-edges.lagda.md b/src/simplicial-type-theory/directed-edges.lagda.md new file mode 100644 index 0000000000..2d9ea43b6d --- /dev/null +++ b/src/simplicial-type-theory/directed-edges.lagda.md @@ -0,0 +1,411 @@ +# Directed edges + +```agda +module simplicial-type-theory.directed-edges where +``` + +
Imports + +```agda +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-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.retractions +open import foundation.sections +open import foundation.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.truncated-types +open import foundation.truncation-levels +open import foundation.type-arithmetic-booleans +open import foundation.type-arithmetic-cartesian-product-types +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +A +{{#concept "directed edge" Disambiguation="simplicial type theory" Agda=simplicial-hom}} +in a type `A` from `x : A` to `y : A` is a +[simplicial arrow](simplicial-type-theory.simplicial-arrows.md) `α` in `A` +together with [identifications](foundation-core.identity-types.md) `α 0₂ = x` +and `α 1₂ = y`. We call `x` the _source_, and `y` the _target_ of the edge. + +We introduce the notation `x →₂ y` for the type of directed edges from `x` to +`y`. + +## Definitions + +### Directed edges in types dependent over the directed interval + +```agda +module _ + {l : Level} {A : 𝟚 → UU l} + where + + simplicial-hom' : A 0₂ → A 1₂ → UU l + simplicial-hom' x y = + Σ (simplicial-arrow' A) (λ α → (α 0₂ = x) × (α 1₂ = y)) + + simplicial-arrow-simplicial-hom : + {x : A 0₂} {y : A 1₂} → + simplicial-hom' x y → + simplicial-arrow' A + simplicial-arrow-simplicial-hom = pr1 + + simplicial-hom-simplicial-arrow : + (α : simplicial-arrow' A) → simplicial-hom' (α 0₂) (α 1₂) + simplicial-hom-simplicial-arrow α = (α , refl , refl) +``` + +### Directed edges + +```agda +module _ + {l : Level} {A : UU l} + where + + _→₂_ : A → A → UU l + _→₂_ = simplicial-hom' {A = λ _ → A} + + infix 7 _→₂_ + + simplicial-hom : A → A → UU l + simplicial-hom = _→₂_ + + hom▵ : A → A → UU l + hom▵ = _→₂_ + + eq-source-simplicial-hom : + {x y : A} (f : x →₂ y) → simplicial-arrow-simplicial-hom f 0₂ = x + eq-source-simplicial-hom f = pr1 (pr2 f) + + inv-eq-source-simplicial-hom : + {x y : A} (f : x →₂ y) → x = simplicial-arrow-simplicial-hom f 0₂ + inv-eq-source-simplicial-hom f = inv (eq-source-simplicial-hom f) + + eq-target-simplicial-hom : + {x y : A} (f : x →₂ y) → simplicial-arrow-simplicial-hom f 1₂ = y + eq-target-simplicial-hom f = pr2 (pr2 f) + + eq-source-target-simplicial-hom : + {x y z : A} (f : x →₂ y) (g : y →₂ z) → + simplicial-arrow-simplicial-hom f 1₂ = simplicial-arrow-simplicial-hom g 0₂ + eq-source-target-simplicial-hom f g = + eq-target-simplicial-hom f ∙ inv-eq-source-simplicial-hom g + + eq-source-source-simplicial-hom : + {x y z : A} (f : x →₂ y) (g : x →₂ z) → + simplicial-arrow-simplicial-hom f 0₂ = simplicial-arrow-simplicial-hom g 0₂ + eq-source-source-simplicial-hom f g = + eq-source-simplicial-hom f ∙ inv-eq-source-simplicial-hom g + + eq-target-target-simplicial-hom : + {x y z : A} (f : x →₂ y) (g : z →₂ y) → + simplicial-arrow-simplicial-hom f 1₂ = simplicial-arrow-simplicial-hom g 1₂ + eq-target-target-simplicial-hom f g = + eq-target-simplicial-hom f ∙ inv (eq-target-simplicial-hom g) +``` + +### The identity/constant directed edges + +```agda +id-simplicial-hom : {l : Level} {A : UU l} (x : A) → x →₂ x +id-simplicial-hom = simplicial-hom-simplicial-arrow ∘ id-simplicial-arrow +``` + +### The representing edge of the directed interval + +```agda +representing-hom-𝟚 : 0₂ →₂ 1₂ +representing-hom-𝟚 = (id , refl , refl) +``` + +### Directed edges arising from equalities + +```agda +simplicial-hom-eq : + {l : Level} {A : UU l} {x y : A} → x = y → x →₂ y +simplicial-hom-eq p = + ( simplicial-arrow-eq p , + compute-source-simplicial-arrow-eq p , + compute-target-simplicial-arrow-eq p) +``` + +## Properties + +### Characterizing equality of homomorphisms + +```agda +module _ + {l : Level} {A : UU l} {x y : A} + where + + coherence-htpy-simplicial-hom : + (f g : x →₂ y) → + simplicial-arrow-simplicial-hom f ~ simplicial-arrow-simplicial-hom g → + UU l + coherence-htpy-simplicial-hom f g H = + ( eq-source-simplicial-hom f = H 0₂ ∙ eq-source-simplicial-hom g) × + ( eq-target-simplicial-hom f = H 1₂ ∙ eq-target-simplicial-hom g) + + htpy-simplicial-hom : (f g : x →₂ y) → UU l + htpy-simplicial-hom f g = + Σ ( simplicial-arrow-simplicial-hom f ~ simplicial-arrow-simplicial-hom g) + ( coherence-htpy-simplicial-hom f g) + + refl-htpy-simplicial-hom : + (f : x →₂ y) → htpy-simplicial-hom f f + refl-htpy-simplicial-hom f = (refl-htpy , refl , refl) + + htpy-eq-simplicial-hom : + (f g : x →₂ y) → f = g → htpy-simplicial-hom f g + htpy-eq-simplicial-hom f .f refl = refl-htpy-simplicial-hom f + + abstract + is-torsorial-htpy-simplicial-hom : + (f : x →₂ y) → + is-contr (Σ (x →₂ y) (htpy-simplicial-hom f)) + is-torsorial-htpy-simplicial-hom f = + is-torsorial-Eq-structure + ( is-torsorial-htpy (simplicial-arrow-simplicial-hom f)) + ( simplicial-arrow-simplicial-hom f , refl-htpy) + ( is-torsorial-Eq-structure + ( is-torsorial-Id (eq-source-simplicial-hom f)) + ( eq-source-simplicial-hom f , refl) + ( is-torsorial-Id (eq-target-simplicial-hom f))) + + is-equiv-htpy-eq-simplicial-hom : + (f g : x →₂ y) → is-equiv (htpy-eq-simplicial-hom f g) + is-equiv-htpy-eq-simplicial-hom f = + fundamental-theorem-id + ( is-torsorial-htpy-simplicial-hom f) + ( htpy-eq-simplicial-hom f) + + extensionality-simplicial-hom : + (f g : x →₂ y) → (f = g) ≃ (htpy-simplicial-hom f g) + extensionality-simplicial-hom f g = + ( htpy-eq-simplicial-hom f g , is-equiv-htpy-eq-simplicial-hom f g) + + eq-htpy-simplicial-hom : + (f g : x →₂ y) → htpy-simplicial-hom f g → f = g + eq-htpy-simplicial-hom f g = + map-inv-equiv (extensionality-simplicial-hom f g) +``` + +### The homotopy of directed edges associated to a homotopy of simplicial arrows + +```agda +module _ + {l : Level} {A : UU l} {x y : A} + {f : simplicial-arrow A} (g : x →₂ y) + (H : f ~ simplicial-arrow-simplicial-hom g) + where + + simplicial-hom-htpy-simplicial-arrow : x →₂ y + simplicial-hom-htpy-simplicial-arrow = + ( f , + H 0₂ ∙ eq-source-simplicial-hom g , + H 1₂ ∙ eq-target-simplicial-hom g) + + htpy-simplicial-hom-htpy-simplicial-arrow : + htpy-simplicial-hom simplicial-hom-htpy-simplicial-arrow g + htpy-simplicial-hom-htpy-simplicial-arrow = (H , refl , refl) + + eq-simplicial-hom-htpy-simplicial-arrow : + simplicial-hom-htpy-simplicial-arrow = g + eq-simplicial-hom-htpy-simplicial-arrow = + eq-htpy-simplicial-hom + ( simplicial-hom-htpy-simplicial-arrow) + ( g) + ( htpy-simplicial-hom-htpy-simplicial-arrow) +``` + +### Computing the based total type of directed edges + +```text + Σ (𝟚 → A) (λ α → α 0₂ = x) ≃ Σ (y : A), (x →₂ y) +``` + +```agda +module _ + {l : Level} {A : UU l} (x : A) + where + + based-simplicial-hom : UU l + based-simplicial-hom = Σ A (λ y → (x →₂ y)) + + map-compute-based-simplicial-hom : + Σ (𝟚 → A) (λ α → α 0₂ = x) → based-simplicial-hom + map-compute-based-simplicial-hom (α , p) = (α 1₂ , α , p , refl) + + map-inv-compute-based-simplicial-hom : + based-simplicial-hom → Σ (𝟚 → A) (λ α → α 0₂ = x) + map-inv-compute-based-simplicial-hom (y , α , p , q) = (α , p) + + is-section-map-inv-compute-based-simplicial-hom : + is-section + ( map-compute-based-simplicial-hom) + ( map-inv-compute-based-simplicial-hom) + is-section-map-inv-compute-based-simplicial-hom + (.(α 1₂) , α , p , refl) = refl + + is-retraction-map-inv-compute-based-simplicial-hom : + is-retraction + ( map-compute-based-simplicial-hom) + ( map-inv-compute-based-simplicial-hom) + is-retraction-map-inv-compute-based-simplicial-hom = refl-htpy + + is-equiv-map-compute-based-simplicial-hom : + is-equiv map-compute-based-simplicial-hom + is-equiv-map-compute-based-simplicial-hom = + is-equiv-is-invertible + ( map-inv-compute-based-simplicial-hom) + ( is-section-map-inv-compute-based-simplicial-hom) + ( is-retraction-map-inv-compute-based-simplicial-hom) + + is-equiv-map-inv-compute-based-simplicial-hom : + is-equiv map-inv-compute-based-simplicial-hom + is-equiv-map-inv-compute-based-simplicial-hom = + is-equiv-is-invertible + ( map-compute-based-simplicial-hom) + ( is-retraction-map-inv-compute-based-simplicial-hom) + ( is-section-map-inv-compute-based-simplicial-hom) + + compute-based-simplicial-hom : + Σ (𝟚 → A) (λ α → α 0₂ = x) ≃ based-simplicial-hom + compute-based-simplicial-hom = + ( map-compute-based-simplicial-hom , + is-equiv-map-compute-based-simplicial-hom) + + inv-compute-based-simplicial-hom : + based-simplicial-hom ≃ Σ (𝟚 → A) (λ α → α 0₂ = x) + inv-compute-based-simplicial-hom = + ( map-inv-compute-based-simplicial-hom , + is-equiv-map-inv-compute-based-simplicial-hom) +``` + +### Computing the total type of directed edges + +The directed interval type classifies the total type of directed edges in a +type. + +```text + (𝟚 → A) ≃ Σ (x y : A), (x →₂ y) +``` + +```agda +module _ + {l : Level} {A : UU l} + where + + total-simplicial-hom : UU l + total-simplicial-hom = Σ A based-simplicial-hom + + map-compute-total-simplicial-hom : + (𝟚 → A) → total-simplicial-hom + map-compute-total-simplicial-hom α = (α 0₂ , α 1₂ , α , refl , refl) + + map-inv-compute-total-simplicial-hom : + total-simplicial-hom → 𝟚 → A + map-inv-compute-total-simplicial-hom (x , y , α , p , q) = α + + is-section-map-inv-compute-total-simplicial-hom : + is-section + ( map-compute-total-simplicial-hom) + ( map-inv-compute-total-simplicial-hom) + is-section-map-inv-compute-total-simplicial-hom + (.(α 0₂) , .(α 1₂) , α , refl , refl) = refl + + is-retraction-map-inv-compute-total-simplicial-hom : + is-retraction + ( map-compute-total-simplicial-hom) + ( map-inv-compute-total-simplicial-hom) + is-retraction-map-inv-compute-total-simplicial-hom = refl-htpy + + is-equiv-map-compute-total-simplicial-hom : + is-equiv map-compute-total-simplicial-hom + is-equiv-map-compute-total-simplicial-hom = + is-equiv-is-invertible + ( map-inv-compute-total-simplicial-hom) + ( is-section-map-inv-compute-total-simplicial-hom) + ( is-retraction-map-inv-compute-total-simplicial-hom) + + is-equiv-map-inv-compute-total-simplicial-hom : + is-equiv map-inv-compute-total-simplicial-hom + is-equiv-map-inv-compute-total-simplicial-hom = + is-equiv-is-invertible + ( map-compute-total-simplicial-hom) + ( is-retraction-map-inv-compute-total-simplicial-hom) + ( is-section-map-inv-compute-total-simplicial-hom) + + compute-total-simplicial-hom : + (𝟚 → A) ≃ total-simplicial-hom + compute-total-simplicial-hom = + ( map-compute-total-simplicial-hom , + is-equiv-map-compute-total-simplicial-hom) + + inv-compute-total-simplicial-hom : + total-simplicial-hom ≃ (𝟚 → A) + inv-compute-total-simplicial-hom = + ( map-inv-compute-total-simplicial-hom , + is-equiv-map-inv-compute-total-simplicial-hom) +``` + +### The hom type is an extension type + +The hom-type `x →₂ y` is equivalent to the +[type of extensions](orthogonal-factorization-systems.extensions-of-maps.md) of +`[x , y] : ∂𝟚 → A` along the inclusion `∂𝟚 ↪ 𝟚`. + +```agda +module _ + {l : Level} {A : UU l} (x y : A) + where + + compute-extension-type-simplicial-hom : + (x →₂ y) ≃ extension map-directed-interval-bool (rec-bool y x) + compute-extension-type-simplicial-hom = + equiv-tot + ( λ α → + ( inv-equiv-Π-bool-product + ( λ b → rec-bool y x b = α (map-directed-interval-bool b))) ∘e + ( commutative-product) ∘e + ( equiv-product (equiv-inv (α 0₂) x) (equiv-inv (α 1₂) y))) +``` + +### The hom-types of a truncated type are truncated + +```agda +module _ + {l : Level} (k : 𝕋) {A : UU l} (x y : A) + where + + is-trunc-simplicial-hom : is-trunc k A → is-trunc k (x →₂ y) + is-trunc-simplicial-hom is-trunc-A = + is-trunc-equiv k + ( extension map-directed-interval-bool (rec-bool y x)) + ( compute-extension-type-simplicial-hom x y) + ( is-trunc-extension-dependent-type k + ( map-directed-interval-bool) + ( rec-bool y x) + ( λ _ → is-trunc-A)) +``` diff --git a/src/simplicial-type-theory/directed-interval-type.lagda.md b/src/simplicial-type-theory/directed-interval-type.lagda.md new file mode 100644 index 0000000000..a35a051d78 --- /dev/null +++ b/src/simplicial-type-theory/directed-interval-type.lagda.md @@ -0,0 +1,208 @@ +# The directed interval type + +```agda +module simplicial-type-theory.directed-interval-type where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.0-connected-types +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.fibers-of-maps +open import foundation.function-types +open import foundation.identity-types +open import foundation.injective-maps +open import foundation.mere-equality +open import foundation.negated-equality +open import foundation.negation +open import foundation.noncontractible-types +open import foundation.propositional-truncations +open import foundation.propositions +open import foundation.retractions +open import foundation.sections +open import foundation.sets +open import foundation.subtypes +open import foundation.surjective-maps +open import foundation.unit-type +open import foundation.universe-levels +``` + +
+ +## Idea + +The +{{#concept "directed interval type" Disambiguation="simplicial type theory" Agda=𝟚}} +`𝟚` is the representing type for the simplicial structure on types. It is a type +consisting of a distinct source and target element, `0₂` and `1₂`, and comes +[equipped](foundation.structure.md) with a directed relation which defines a +[total order](order-theory.total-orders.md) with `0₂` as a +[bottom element](order-theory.bottom-elements-posets.md), and `1₂` as a +[top element](order-theory.top-elements-posets.md). + +In this file, we postulate the existence of the directed interval type together +with its distinct source and target elements. In the module +[`inequality-directed-interval-type`](simplicial-type-theory.inequality-directed-interval-type.md), +we postulate the existence of the directed relation on the directed interval +type. + +## Postulates + +### The directed interval type + +```agda +postulate + 𝟚 : UU lzero + + 0₂ 1₂ : 𝟚 + + is-nontrivial-𝟚 : 0₂ ≠ 1₂ +``` + +## Properties + +### The directed interval type is not a proposition + +```agda +is-not-prop-𝟚 : ¬ (is-prop 𝟚) +is-not-prop-𝟚 H = is-nontrivial-𝟚 (eq-is-prop H) +``` + +### The directed interval type is not contractible + +```agda +is-not-contractible-𝟚 : is-not-contractible 𝟚 +is-not-contractible-𝟚 H = is-nontrivial-𝟚 (eq-is-contr H) +``` + +### The directed interval type is noncontractible + +```agda +is-noncontractible-𝟚' : is-noncontractible' 𝟚 1 +is-noncontractible-𝟚' = (0₂ , 1₂ , is-nontrivial-𝟚) + +is-noncontractible-𝟚 : is-noncontractible 𝟚 +is-noncontractible-𝟚 = (1 , is-noncontractible-𝟚') +``` + +## Definitions + +### The boundary of the directed interval + +```agda +subtype-∂𝟚' : subtype lzero 𝟚 +subtype-∂𝟚' t = + coproduct-Prop + ( mere-eq-Prop t 0₂) + ( mere-eq-Prop t 1₂) + ( λ |t=0| → + rec-trunc-Prop empty-Prop + ( λ t=1 → + rec-trunc-Prop empty-Prop + ( λ t=0 → is-nontrivial-𝟚 (inv t=0 ∙ t=1)) + ( |t=0|))) + +∂𝟚 : UU lzero +∂𝟚 = type-subtype subtype-∂𝟚' +``` + +### The canonical inclusion of the booleans into the directed interval + +The canonical inclusion of the booleans into the directed interval is the map +that sends `false` to `0₂` and `true` to `1₂`. We call the +[image](foundation.images.md) of this map the boundary of the directed interval, +`∂𝟚`, and we show that `bool` is a [retract](foundation.retracts-of-types.md) of +`∂𝟚`. + +```agda +map-boundary-directed-interval-bool : bool → ∂𝟚 +map-boundary-directed-interval-bool true = (1₂ , inr (refl-mere-eq 1₂)) +map-boundary-directed-interval-bool false = (0₂ , inl (refl-mere-eq 0₂)) + +map-bool-boundary-directed-interval : ∂𝟚 → bool +map-bool-boundary-directed-interval (t , inl x) = false +map-bool-boundary-directed-interval (t , inr x) = true + +is-retraction-map-bool-boundary-directed-interval : + is-retraction + ( map-boundary-directed-interval-bool) + ( map-bool-boundary-directed-interval) +is-retraction-map-bool-boundary-directed-interval true = refl +is-retraction-map-bool-boundary-directed-interval false = refl + +is-surjective-map-bool-boundary-directed-interval : + is-surjective map-bool-boundary-directed-interval +is-surjective-map-bool-boundary-directed-interval = + is-surjective-has-section + ( map-boundary-directed-interval-bool , + is-retraction-map-bool-boundary-directed-interval) + +is-surjective-map-boundary-directed-interval-bool : + is-surjective map-boundary-directed-interval-bool +is-surjective-map-boundary-directed-interval-bool (t , inl |p|) = + rec-trunc-Prop + ( trunc-Prop (fiber map-boundary-directed-interval-bool (t , inl |p|))) + ( λ t=0 → unit-trunc-Prop (false , eq-type-subtype subtype-∂𝟚' (inv t=0))) + ( |p|) +is-surjective-map-boundary-directed-interval-bool (t , inr |p|) = + rec-trunc-Prop + ( trunc-Prop (fiber map-boundary-directed-interval-bool (t , inr |p|))) + ( λ t=1 → unit-trunc-Prop (true , eq-type-subtype subtype-∂𝟚' (inv t=1))) + ( |p|) + +map-directed-interval-bool : bool → 𝟚 +map-directed-interval-bool true = 1₂ +map-directed-interval-bool false = 0₂ + +is-injective-map-directed-interval-bool : + is-injective map-directed-interval-bool +is-injective-map-directed-interval-bool {true} {true} p = + refl +is-injective-map-directed-interval-bool {true} {false} p = + ex-falso (is-nontrivial-𝟚 (inv p)) +is-injective-map-directed-interval-bool {false} {true} p = + ex-falso (is-nontrivial-𝟚 p) +is-injective-map-directed-interval-bool {false} {false} p = + refl + +is-retraction-is-injective-map-directed-interval-bool : + {x y : bool} → + is-retraction + ( ap map-directed-interval-bool {x} {y}) + ( is-injective-map-directed-interval-bool) +is-retraction-is-injective-map-directed-interval-bool {true} refl = refl +is-retraction-is-injective-map-directed-interval-bool {false} refl = refl + +retraction-ap-map-directed-interval-bool : + {x y : bool} → retraction (ap map-directed-interval-bool {x} {y}) +retraction-ap-map-directed-interval-bool = + ( is-injective-map-directed-interval-bool , + is-retraction-is-injective-map-directed-interval-bool) +``` + +We show that `map-directed-interval-bool` is an +[embedding](foundation-core.embeddings.md) in +[`inequality-directed-interval-type`](simplicial-type-theory.inequality-directed-interval-type.md). + +### The directed interval is not connected + +**Proof.** A type is 0-connected only if all pairs of elements are +[merely equal](foundation.mere-equality.md), and since we are attempting to +deduce a contradiction we may assume we have that all elements are equal, but +`0₂` and `1₂` are not. + +```agda +is-not-0-connected-𝟚 : ¬ (is-0-connected 𝟚) +is-not-0-connected-𝟚 H = + rec-trunc-Prop empty-Prop is-nontrivial-𝟚 (mere-eq-is-0-connected H 0₂ 1₂) +``` diff --git a/src/simplicial-type-theory/free-directed-loops.lagda.md b/src/simplicial-type-theory/free-directed-loops.lagda.md new file mode 100644 index 0000000000..992d4f07d1 --- /dev/null +++ b/src/simplicial-type-theory/free-directed-loops.lagda.md @@ -0,0 +1,321 @@ +# Free directed loops + +```agda +module simplicial-type-theory.free-directed-loops where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.commuting-squares-of-identifications +open import foundation.constant-type-families +open import foundation.contractible-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.function-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.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.transport-along-identifications +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.free-loops +``` + +
+ +## Idea + +A +{{#concept "free directed loop" Disambiguation="in a simplicial type" Agda=free-directed-loop}} +in a type `X` consists of a +[directed arrow](simplicial-type-theory.simplicial-arrows.md) `α : 𝟚 → X` and an +[identification](foundation-core.identity-types.md) `α 1₂ = α 0₂`. Free +directed loops are classified by the +[directed circle](simplicial-type-theory.directed-circle.md), meaning that the +type of free directed loops in `X` is +[equivalent](foundation-core.equivalences.md) to the type of maps +`directed-circle → X`. + +## Definitions + +### Free directed loops + +```agda +free-directed-loop : {l : Level} → UU l → UU l +free-directed-loop X = Σ (𝟚 → X) (λ α → α 1₂ = α 0₂) + +module _ + {l1 : Level} {X : UU l1} + where + + arrow-free-directed-loop : free-directed-loop X → 𝟚 → X + arrow-free-directed-loop = pr1 + + base-free-directed-loop : free-directed-loop X → X + base-free-directed-loop α = arrow-free-directed-loop α 0₂ + + compute-target-free-directed-loop : + (α : free-directed-loop X) → + arrow-free-directed-loop α 1₂ = base-free-directed-loop α + compute-target-free-directed-loop = pr2 + + loop-free-directed-loop : + (α : free-directed-loop X) → + base-free-directed-loop α →₂ base-free-directed-loop α + loop-free-directed-loop α = + ( arrow-free-directed-loop α , refl , compute-target-free-directed-loop α) +``` + +### Free dependent directed loops + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (P : X → UU l2) + where + + free-dependent-directed-loop : UU l2 + free-dependent-directed-loop = + Σ ( simplicial-arrow' (P ∘ arrow-free-directed-loop α)) + ( λ β → + dependent-identification P + ( compute-target-free-directed-loop α) + ( β 1₂) + ( β 0₂)) + +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) {P : X → UU l2} + ( β : free-dependent-directed-loop α P) + where + + arrow-free-dependent-directed-loop : + simplicial-arrow' (P ∘ arrow-free-directed-loop α) + arrow-free-dependent-directed-loop = pr1 β + + base-free-dependent-directed-loop : P (base-free-directed-loop α) + base-free-dependent-directed-loop = arrow-free-dependent-directed-loop 0₂ + + compute-target-free-dependent-directed-loop : + dependent-identification P + ( compute-target-free-directed-loop α) + ( arrow-free-dependent-directed-loop 1₂) + ( arrow-free-dependent-directed-loop 0₂) + compute-target-free-dependent-directed-loop = pr2 β + + loop-free-dependent-directed-loop : + dependent-simplicial-hom P + ( loop-free-directed-loop α) + ( base-free-dependent-directed-loop) + ( base-free-dependent-directed-loop) + loop-free-dependent-directed-loop = + ( arrow-free-dependent-directed-loop , + refl , + compute-target-free-dependent-directed-loop) +``` + +### Free directed loops from free loops + +```agda +module _ + {l1 : Level} {X : UU l1} + where + + free-directed-loop-free-loop : free-loop X → free-directed-loop X + pr1 (free-directed-loop-free-loop α) = + simplicial-arrow-eq (loop-free-loop α) + pr2 (free-directed-loop-free-loop α) = + ( compute-target-simplicial-arrow-eq (loop-free-loop α) ∙ + inv (compute-source-simplicial-arrow-eq (loop-free-loop α))) +``` + +### Constant free directed loops + +```agda +module _ + {l1 : Level} {X : UU l1} + where + + constant-free-directed-loop : X → free-directed-loop X + constant-free-directed-loop x = (id-simplicial-arrow x , refl) +``` + +## Properties + +### Characterization of the identity type of the type of free directed loops + +```agda +module _ + {l1 : Level} {X : UU l1} + where + + coherence-htpy-free-directed-loop : + (α α' : free-directed-loop X) → + arrow-free-directed-loop α ~ arrow-free-directed-loop α' → + UU l1 + coherence-htpy-free-directed-loop α α' H = + coherence-square-identifications + ( compute-target-free-directed-loop α) + ( H 1₂) + ( H 0₂) + ( compute-target-free-directed-loop α') + + htpy-free-directed-loop : (α α' : free-directed-loop X) → UU l1 + htpy-free-directed-loop α α' = + Σ ( arrow-free-directed-loop α ~ arrow-free-directed-loop α') + ( coherence-htpy-free-directed-loop α α') + + refl-htpy-free-directed-loop : + (α : free-directed-loop X) → htpy-free-directed-loop α α + refl-htpy-free-directed-loop α = (refl-htpy , inv right-unit) + + htpy-eq-free-directed-loop : + (α α' : free-directed-loop X) → α = α' → htpy-free-directed-loop α α' + htpy-eq-free-directed-loop α .α refl = refl-htpy-free-directed-loop α + + is-torsorial-htpy-free-directed-loop : + (α : free-directed-loop X) → is-torsorial (htpy-free-directed-loop α) + is-torsorial-htpy-free-directed-loop α = + is-torsorial-Eq-structure + ( is-torsorial-htpy (arrow-free-directed-loop α)) + ( arrow-free-directed-loop α , refl-htpy) + ( is-torsorial-Id' (compute-target-free-directed-loop α ∙ refl)) + + is-equiv-htpy-eq-free-directed-loop : + (α α' : free-directed-loop X) → is-equiv (htpy-eq-free-directed-loop α α') + is-equiv-htpy-eq-free-directed-loop α = + fundamental-theorem-id + ( is-torsorial-htpy-free-directed-loop α) + ( htpy-eq-free-directed-loop α) + + extensionality-free-directed-loop : + (α α' : free-directed-loop X) → (α = α') ≃ htpy-free-directed-loop α α' + extensionality-free-directed-loop α α' = + ( htpy-eq-free-directed-loop α α' , + is-equiv-htpy-eq-free-directed-loop α α') + + eq-htpy-free-directed-loop : + (α α' : free-directed-loop X) → htpy-free-directed-loop α α' → α = α' + eq-htpy-free-directed-loop α α' = + map-inv-equiv (extensionality-free-directed-loop α α') +``` + +### Characterization of the identity type of free dependent loops + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (P : X → UU l2) + where + + coherence-htpy-free-dependent-directed-loop : + (β β' : free-dependent-directed-loop α P) → + arrow-free-dependent-directed-loop α β ~ + arrow-free-dependent-directed-loop α β' → + UU l2 + coherence-htpy-free-dependent-directed-loop β β' H = + ( right-strict-concat-dependent-identification P + ( compute-target-free-directed-loop α) + ( refl) + ( compute-target-free-dependent-directed-loop α β) + ( H 0₂)) = + ( concat-dependent-identification P + ( refl) + ( compute-target-free-directed-loop α) + ( H 1₂) + ( compute-target-free-dependent-directed-loop α β')) + + htpy-free-dependent-directed-loop : + (β β' : free-dependent-directed-loop α P) → UU l2 + htpy-free-dependent-directed-loop β β' = + Σ ( arrow-free-dependent-directed-loop α β ~ + arrow-free-dependent-directed-loop α β') + ( coherence-htpy-free-dependent-directed-loop β β') + + refl-htpy-free-dependent-directed-loop : + (β : free-dependent-directed-loop α P) → + htpy-free-dependent-directed-loop β β + refl-htpy-free-dependent-directed-loop β = (refl-htpy , refl) + + htpy-free-dependent-directed-loop-eq : + ( β β' : free-dependent-directed-loop α P) → + β = β' → htpy-free-dependent-directed-loop β β' + htpy-free-dependent-directed-loop-eq β .β refl = + refl-htpy-free-dependent-directed-loop β + + is-torsorial-htpy-free-dependent-directed-loop : + ( β : free-dependent-directed-loop α P) → + is-torsorial (htpy-free-dependent-directed-loop β) + is-torsorial-htpy-free-dependent-directed-loop β = + is-torsorial-Eq-structure + ( is-torsorial-htpy (arrow-free-dependent-directed-loop α β)) + ( arrow-free-dependent-directed-loop α β , refl-htpy) + (is-torsorial-Id (compute-target-free-dependent-directed-loop α β)) + + is-equiv-htpy-free-dependent-directed-loop-eq : + (β β' : free-dependent-directed-loop α P) → + is-equiv (htpy-free-dependent-directed-loop-eq β β') + is-equiv-htpy-free-dependent-directed-loop-eq β = + fundamental-theorem-id + ( is-torsorial-htpy-free-dependent-directed-loop β) + ( htpy-free-dependent-directed-loop-eq β) + + eq-htpy-free-dependent-directed-loop : + (β β' : free-dependent-directed-loop α P) → + htpy-free-dependent-directed-loop β β' → β = β' + eq-htpy-free-dependent-directed-loop β β' = + map-inv-is-equiv (is-equiv-htpy-free-dependent-directed-loop-eq β β') +``` + +### The type of free dependent loops in a constant family of types is equivalent to the type of ordinary free directed loops + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (Y : UU l2) + where + + compute-free-dependent-directed-loop-constant-type-family : + free-directed-loop Y ≃ free-dependent-directed-loop α (λ _ → Y) + compute-free-dependent-directed-loop-constant-type-family = + equiv-tot + ( λ _ → + compute-dependent-identification-constant-type-family + ( compute-target-free-directed-loop α)) + + map-compute-free-dependent-directed-loop-constant-type-family : + free-directed-loop Y → free-dependent-directed-loop α (λ _ → Y) + map-compute-free-dependent-directed-loop-constant-type-family = + map-equiv compute-free-dependent-directed-loop-constant-type-family + + map-inv-compute-free-dependent-directed-loop-constant-type-family : + free-dependent-directed-loop α (λ _ → Y) → free-directed-loop Y + map-inv-compute-free-dependent-directed-loop-constant-type-family = + map-inv-equiv compute-free-dependent-directed-loop-constant-type-family +``` + +### The type of free directed loops is equivalent to the type of directed edges with common source and target + +```agda +module _ + {l : Level} {X : UU l} + where + + compute-free-dependent-directed-loop : + Σ X (λ x → x →₂ x) ≃ free-directed-loop X + compute-free-dependent-directed-loop = + ( equiv-tot + ( λ α → + ( left-unit-law-Σ-is-contr (is-torsorial-Id (α 0₂)) (α 0₂ , refl)) ∘e + ( inv-associative-Σ X (α 0₂ =_) (λ z → α 1₂ = pr1 z)))) ∘e + ( equiv-left-swap-Σ) +``` diff --git a/src/simplicial-type-theory/fully-faithful-maps.lagda.md b/src/simplicial-type-theory/fully-faithful-maps.lagda.md new file mode 100644 index 0000000000..03772a99a5 --- /dev/null +++ b/src/simplicial-type-theory/fully-faithful-maps.lagda.md @@ -0,0 +1,561 @@ +# Fully-faithful maps + +```agda +module simplicial-type-theory.fully-faithful-maps where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.commuting-squares-of-maps +open import foundation.commuting-triangles-of-maps +open import foundation.cones-over-cospan-diagrams +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.embeddings +open import foundation.equivalences +open import foundation.fibers-of-maps +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopies +open import foundation.postcomposition-functions +open import foundation.pullbacks +open import foundation.retractions +open import foundation.torsorial-type-families +open import foundation.transport-along-identifications +open import foundation.universe-levels + +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.propositions +open import foundation-core.sections + +open import orthogonal-factorization-systems.orthogonal-maps + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.whiskering-directed-edges +``` + +
+ +## Idea + +A +{{#concept "simplicially fully-faithful map" Disambiguation="of simplicial types" Agda=is-simplicially-fully-faithful}} +from one type into another is a map that induces +[equivalences](foundation-core.equivalences.md) on +[hom-types](simplicial-type-theory.directed-edges.md). In other words, the +directed edges `f x →₂ f y` for a simplicially fully-faithful map `f : A → B` +are in one-to-one correspondence with the directed edges `x →₂ y`. + +## Definitions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-simplicially-fully-faithful : (A → B) → UU (l1 ⊔ l2) + is-simplicially-fully-faithful f = + (x y : A) → is-equiv (action-simplicial-hom-function f {x} {y}) + + equiv-action-is-simplicially-fully-faithful : + {f : A → B} (e : is-simplicially-fully-faithful f) + {x y : A} → (x →₂ y) ≃ (f x →₂ f y) + equiv-action-is-simplicially-fully-faithful {f} e {x} {y} = + ( action-simplicial-hom-function f , e x y) + + inv-equiv-action-is-simplicially-fully-faithful : + {f : A → B} (e : is-simplicially-fully-faithful f) + {x y : A} → (f x →₂ f y) ≃ (x →₂ y) + inv-equiv-action-is-simplicially-fully-faithful e = + inv-equiv (equiv-action-is-simplicially-fully-faithful e) + +infix 5 _↪▵_ +_↪▵_ : {l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2) +A ↪▵ B = Σ (A → B) (is-simplicially-fully-faithful) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + map-simplicially-fully-faithful-map : A ↪▵ B → A → B + map-simplicially-fully-faithful-map = pr1 + + is-simplicially-fully-faithful-map-simplicially-fully-faithful-map : + (f : A ↪▵ B) → + is-simplicially-fully-faithful (map-simplicially-fully-faithful-map f) + is-simplicially-fully-faithful-map-simplicially-fully-faithful-map = pr2 + + equiv-action-simplicially-fully-faithful-map : + (e : A ↪▵ B) {x y : A} → + ( x →₂ y) ≃ + ( map-simplicially-fully-faithful-map e x →₂ + map-simplicially-fully-faithful-map e y) + equiv-action-simplicially-fully-faithful-map e = + equiv-action-is-simplicially-fully-faithful + ( is-simplicially-fully-faithful-map-simplicially-fully-faithful-map e) + + inv-equiv-action-simplicially-fully-faithful-map : + (e : A ↪▵ B) + {x y : A} → + ( map-simplicially-fully-faithful-map e x →₂ + map-simplicially-fully-faithful-map e y) ≃ + ( x →₂ y) + inv-equiv-action-simplicially-fully-faithful-map e = + inv-equiv (equiv-action-simplicially-fully-faithful-map e) +``` + +## Properties + +### Being simplicially fully faithful is a property + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-property-is-simplicially-fully-faithful : + (f : A → B) → is-prop (is-simplicially-fully-faithful f) + is-property-is-simplicially-fully-faithful f = + is-prop-Π + ( λ x → + is-prop-Π + ( λ y → is-property-is-equiv (action-simplicial-hom-function f))) + + is-simplicially-fully-faithful-Prop : (A → B) → Prop (l1 ⊔ l2) + is-simplicially-fully-faithful-Prop f = + ( is-simplicially-fully-faithful f , + is-property-is-simplicially-fully-faithful f) +``` + +### The identity map is simplicially fully faithful + +```agda +module _ + {l : Level} {A : UU l} + where + + is-simplicially-fully-faithful-id : + is-simplicially-fully-faithful (id {A = A}) + is-simplicially-fully-faithful-id x y = + is-equiv-htpy id compute-action-simplicial-hom-id-function is-equiv-id + + id-simplicially-fully-faithful-map : A ↪▵ A + id-simplicially-fully-faithful-map = + ( id , is-simplicially-fully-faithful-id) +``` + +### Equivalences are simplicially fully faithful + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-simplicially-fully-faithful-is-equiv : + {f : A → B} → is-equiv f → is-simplicially-fully-faithful f + is-simplicially-fully-faithful-is-equiv {f} H x y = + is-equiv-map-Σ + ( λ α → (α 0₂ = f x) × (α 1₂ = f y)) + ( is-equiv-postcomp-is-equiv f H 𝟚) + ( λ α → + is-equiv-map-product + ( ap f) + ( ap f) + ( is-emb-is-equiv H (α 0₂) x) + ( is-emb-is-equiv H (α 1₂) y)) + + equiv-action-simplicial-hom : + (e : A ≃ B) → (x y : A) → hom▵ x y ≃ hom▵ (map-equiv e x) (map-equiv e y) + equiv-action-simplicial-hom e x y = + ( action-simplicial-hom-function (map-equiv e) , + is-simplicially-fully-faithful-is-equiv (is-equiv-map-equiv e) x y) +``` + +### A map is simplicially fully faithful if and only if it is `(∂𝟚 → 𝟚)`-orthogonal + +> This remains to be formalized. + +### To prove that a map is simplicially fully faithful, a point in the domain may be assumed + +```agda +module _ + {l : Level} {A : UU l} {l2 : Level} {B : UU l2} {f : A → B} + where + + is-simplicially-fully-faithful-is-simplicially-fully-faithful : + (A → is-simplicially-fully-faithful f) → is-simplicially-fully-faithful f + is-simplicially-fully-faithful-is-simplicially-fully-faithful H x y = H x x y +``` + +### Simplicially fully faithful maps are closed under homotopies + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + abstract + is-simplicially-fully-faithful-htpy : + {f g : A → B} (H : f ~ g) → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful f + is-simplicially-fully-faithful-htpy {f} {g} H is-ff-g x y = + is-equiv-top-map-triangle + ( action-simplicial-hom-function g) + ( double-whisker-simplicial-hom (H x) (H y)) + ( action-simplicial-hom-function f) + ( nat-htpy▵ H) + ( is-equiv-double-whisker-simplicial-hom (H x) (H y)) + ( is-ff-g x y) + + is-simplicially-fully-faithful-htpy-simplicially-fully-faithful-map : + {f : A → B} (e : A ↪▵ B) → + f ~ map-simplicially-fully-faithful-map e → + is-simplicially-fully-faithful f + is-simplicially-fully-faithful-htpy-simplicially-fully-faithful-map e H = + is-simplicially-fully-faithful-htpy H + ( is-simplicially-fully-faithful-map-simplicially-fully-faithful-map e) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + abstract + is-simplicially-fully-faithful-htpy' : + {f g : A → B} (H : f ~ g) → + is-simplicially-fully-faithful f → + is-simplicially-fully-faithful g + is-simplicially-fully-faithful-htpy' H is-ff-f = + is-simplicially-fully-faithful-htpy (inv-htpy H) is-ff-f + + is-simplicially-fully-faithful-htpy-simplicially-fully-faithful-map' : + (e : A ↪▵ B) {g : A → B} → + map-simplicially-fully-faithful-map e ~ g → + is-simplicially-fully-faithful g + is-simplicially-fully-faithful-htpy-simplicially-fully-faithful-map' e H = + is-simplicially-fully-faithful-htpy' H + ( is-simplicially-fully-faithful-map-simplicially-fully-faithful-map e) +``` + +### Any map between propositions is simplicially fully faithful + +**Proof:** Propositions are simplicially discrete, so a simplicially +fully-faithful map between them is an embedding. + +```text +is-simplicially-fully-faithful-is-prop : + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} → + is-prop A → is-prop B → is-simplicially-fully-faithful f +is-simplicially-fully-faithful-is-prop H K = + is-simplicially-fully-faithful-is-prop-map (is-trunc-map-is-trunc-domain-codomain neg-one-𝕋 H K) +``` + +### Any map between `-1`-coskeletal types is simplicially fully faithful + +> This remains to be formalized. + +### Simplicially fully faithful maps are closed under retracts of maps + +> This remains to be formalized. + +### Simplicially fully faithful maps are closed under composition + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-simplicially-fully-faithful-comp : + (g : B → C) (h : A → B) → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful h → + is-simplicially-fully-faithful (g ∘ h) + is-simplicially-fully-faithful-comp g h is-ff-g is-ff-h x y = + is-equiv-left-map-triangle + ( action-simplicial-hom-function (g ∘ h)) + ( action-simplicial-hom-function g) + ( action-simplicial-hom-function h) + ( compute-action-simplicial-hom-comp-function g h) + ( is-ff-h x y) + ( is-ff-g (h x) (h y)) + + abstract + is-simplicially-fully-faithful-left-map-triangle : + (f : A → C) (g : B → C) (h : A → B) (H : coherence-triangle-maps f g h) → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful h → is-simplicially-fully-faithful f + is-simplicially-fully-faithful-left-map-triangle f g h H is-ff-g is-ff-h = + is-simplicially-fully-faithful-htpy H + ( is-simplicially-fully-faithful-comp g h is-ff-g is-ff-h) + + comp-simplicially-fully-faithful-map : + (B ↪▵ C) → (A ↪▵ B) → (A ↪▵ C) + comp-simplicially-fully-faithful-map (g , H) (f , K) = + ( g ∘ f , is-simplicially-fully-faithful-comp g f H K) +``` + +### The right factor of a composed embedding is an embedding + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-simplicially-fully-faithful-right-factor : + (g : B → C) (h : A → B) → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful (g ∘ h) → + is-simplicially-fully-faithful h + is-simplicially-fully-faithful-right-factor g h is-ff-g is-ff-gh x y = + is-equiv-top-map-triangle + ( action-simplicial-hom-function (g ∘ h)) + ( action-simplicial-hom-function g) + ( action-simplicial-hom-function h) + ( compute-action-simplicial-hom-comp-function g h) + ( is-ff-g (h x) (h y)) + ( is-ff-gh x y) + + abstract + is-simplicially-fully-faithful-top-map-triangle : + (f : A → C) (g : B → C) (h : A → B) + (H : coherence-triangle-maps f g h) → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful f → + is-simplicially-fully-faithful h + is-simplicially-fully-faithful-top-map-triangle + f g h H is-ff-g is-ff-f x y = + is-equiv-top-map-triangle + ( action-simplicial-hom-function (g ∘ h)) + ( action-simplicial-hom-function g) + ( action-simplicial-hom-function h) + ( compute-action-simplicial-hom-comp-function g h) + ( is-ff-g (h x) (h y)) + ( is-simplicially-fully-faithful-htpy (inv-htpy H) is-ff-f x y) + + abstract + is-simplicially-fully-faithful-triangle-is-equiv : + (f : A → C) (g : B → C) (e : A → B) (H : coherence-triangle-maps f g e) → + is-equiv e → + is-simplicially-fully-faithful g → + is-simplicially-fully-faithful f + is-simplicially-fully-faithful-triangle-is-equiv + f g e H is-equiv-e is-ff-g = + is-simplicially-fully-faithful-left-map-triangle f g e H is-ff-g + ( is-simplicially-fully-faithful-is-equiv is-equiv-e) + +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + abstract + is-simplicially-fully-faithful-triangle-is-equiv' : + (f : A → C) (g : B → C) (e : A → B) (H : coherence-triangle-maps f g e) → + is-equiv e → + is-simplicially-fully-faithful f → + is-simplicially-fully-faithful g + is-simplicially-fully-faithful-triangle-is-equiv' + f g e H is-equiv-e is-ff-f = + is-simplicially-fully-faithful-triangle-is-equiv g f + ( map-inv-is-equiv is-equiv-e) + ( triangle-section f g e H + ( pair + ( map-inv-is-equiv is-equiv-e) + ( is-section-map-inv-is-equiv is-equiv-e))) + ( is-equiv-map-inv-is-equiv is-equiv-e) + ( is-ff-f) +``` + +### The map on total spaces induced by a family of simplicially fully faithful maps is simplicially fully faithful + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3} + where + + -- is-simplicially-fully-faithful-tot : + -- {f : (x : A) → B x → C x} → ((x : A) → is-simplicially-fully-faithful (f x)) → is-simplicially-fully-faithful (tot f) + -- is-simplicially-fully-faithful-tot H = + -- is-simplicially-fully-faithful-is-prop-map (is-prop-map-tot (λ x → is-prop-map-is-simplicially-fully-faithful (H x))) + + -- simplicially-fully-faithful-map-tot : ((x : A) → B x ↪▵ C x) → Σ A B ↪▵ Σ A C + -- pr1 (simplicially-fully-faithful-map-tot f) = tot (λ x → map-simplicially-fully-faithful-map (f x)) + -- pr2 (simplicially-fully-faithful-map-tot f) = is-simplicially-fully-faithful-tot (λ x → is-simplicially-fully-faithful-map-simplicially-fully-faithful-map (f x)) +``` + +### The functoriality of dependent pair types preserves simplicially fully faithful maps + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} + where + + -- abstract + -- is-simplicially-fully-faithful-map-Σ-map-base : + -- {f : A → B} (C : B → UU l3) → is-simplicially-fully-faithful f → is-simplicially-fully-faithful (map-Σ-map-base f C) + -- is-simplicially-fully-faithful-map-Σ-map-base C H = + -- is-simplicially-fully-faithful-is-prop-map (is-prop-map-map-Σ-map-base C (is-prop-map-is-simplicially-fully-faithful H)) + + -- simplicially-fully-faithful-map-Σ-simplicially-fully-faithful-map-base : + -- (f : A ↪▵ B) (C : B → UU l3) → Σ A (λ a → C (map-simplicially-fully-faithful-map f a)) ↪▵ Σ B C + -- pr1 (simplicially-fully-faithful-map-Σ-simplicially-fully-faithful-map-base f C) = map-Σ-map-base (map-simplicially-fully-faithful-map f) C + -- pr2 (simplicially-fully-faithful-map-Σ-simplicially-fully-faithful-map-base f C) = + -- is-simplicially-fully-faithful-map-Σ-map-base C (is-simplicially-fully-faithful-map-simplicially-fully-faithful-map f) + +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} + where + + -- is-simplicially-fully-faithful-map-Σ : + -- (D : B → UU l4) {f : A → B} {g : (x : A) → C x → D (f x)} → + -- is-simplicially-fully-faithful f → ((x : A) → is-simplicially-fully-faithful (g x)) → is-simplicially-fully-faithful (map-Σ D f g) + -- is-simplicially-fully-faithful-map-Σ D H K = + -- is-simplicially-fully-faithful-is-prop-map + -- ( is-prop-map-map-Σ D + -- ( is-prop-map-is-simplicially-fully-faithful H) + -- ( λ x → is-prop-map-is-simplicially-fully-faithful (K x))) + + -- simplicially-fully-faithful-map-Σ : + -- (D : B → UU l4) (f : A ↪▵ B) (g : (x : A) → C x ↪▵ D (map-simplicially-fully-faithful-map f x)) → + -- Σ A C ↪▵ Σ B D + -- pr1 (simplicially-fully-faithful-map-Σ D f g) = map-Σ D (map-simplicially-fully-faithful-map f) (λ x → map-simplicially-fully-faithful-map (g x)) + -- pr2 (simplicially-fully-faithful-map-Σ D f g) = + -- is-simplicially-fully-faithful-map-Σ D (is-simplicially-fully-faithful-map-simplicially-fully-faithful-map f) (λ x → is-simplicially-fully-faithful-map-simplicially-fully-faithful-map (g x)) +``` + +### Equivalence on total spaces induced by simplicially fully faithful maps on the base types + +We saw above that given an embedding `f : A ↪▵ B` and a type family `C` over `B` +we obtain an embedding + +```text + Σ A (C ∘ f) ↪▵ Σ B C. +``` + +This embedding can be upgraded to an equivalence if we furthermore know that the +support of `C` is contained in the image of `f`. More precisely, if we are given +a section `((b , c) : Σ B C) → fiber f b`, then it follows that + +```text + Σ A (C ∘ f) ≃ Σ B C. +``` + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : B → UU l3} (f : A ↪▵ B) + (H : ((b , c) : Σ B C) → fiber (map-simplicially-fully-faithful-map f) b) + where + + map-inv-Σ-simplicially-fully-faithful-map-base : + Σ B C → Σ A (C ∘ map-simplicially-fully-faithful-map f) + pr1 (map-inv-Σ-simplicially-fully-faithful-map-base u) = pr1 (H u) + pr2 (map-inv-Σ-simplicially-fully-faithful-map-base u) = + inv-tr C (pr2 (H u)) (pr2 u) + + is-section-map-inv-Σ-simplicially-fully-faithful-map-base : + is-section + ( map-Σ-map-base (map-simplicially-fully-faithful-map f) C) + ( map-inv-Σ-simplicially-fully-faithful-map-base) + is-section-map-inv-Σ-simplicially-fully-faithful-map-base (b , c) = + ap + ( λ s → (pr1 s , inv-tr C (pr2 s) c)) + ( eq-is-contr (is-torsorial-Id' b)) + + -- is-retraction-map-inv-Σ-simplicially-fully-faithful-map-base : + -- is-retraction (map-Σ-map-base (map-simplicially-fully-faithful-map f) C) map-inv-Σ-simplicially-fully-faithful-map-base + -- is-retraction-map-inv-Σ-simplicially-fully-faithful-map-base (a , c) = + -- ap + -- ( λ s → (pr1 s , inv-tr C (pr2 s) c)) + -- ( eq-is-prop (is-prop-map-is-simplicially-fully-faithful (pr2 f) (map-simplicially-fully-faithful-map f a))) + + -- equiv-Σ-simplicially-fully-faithful-map-base : Σ A (C ∘ map-simplicially-fully-faithful-map f) ≃ Σ B C + -- pr1 equiv-Σ-simplicially-fully-faithful-map-base = map-Σ-map-base (map-simplicially-fully-faithful-map f) C + -- pr2 equiv-Σ-simplicially-fully-faithful-map-base = + -- is-equiv-is-invertible + -- map-inv-Σ-simplicially-fully-faithful-map-base + -- is-section-map-inv-Σ-simplicially-fully-faithful-map-base + -- is-retraction-map-inv-Σ-simplicially-fully-faithful-map-base +``` + +### The product of two simplicially fully faithful maps is simplicially fully faithful + +```agda +-- module _ +-- {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} +-- where + +-- simplicially-fully-faithful-map-product : (A ↪▵ C) → (B ↪▵ D) → ((A × B) ↪▵ (C × D)) +-- simplicially-fully-faithful-map-product f g = simplicially-fully-faithful-map-Σ (λ _ → D) f (λ _ → g) + +-- is-simplicially-fully-faithful-map-product : +-- (f : A → C) (g : B → D) → is-simplicially-fully-faithful f → is-simplicially-fully-faithful g → (is-simplicially-fully-faithful (map-product f g)) +-- is-simplicially-fully-faithful-map-product f g is-ff-f is-ff-g = +-- is-simplicially-fully-faithful-map-simplicially-fully-faithful-map (simplicially-fully-faithful-map-product (f , is-ff-f) (g , is-ff-g)) +``` + +### Simplicially fully-faithful maps are closed under pullback + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + {X : UU l4} (f : A → X) (g : B → X) (c : cone f g C) + where + + -- abstract + -- is-simplicially-fully-faithful-vertical-map-cone-is-pullback : + -- is-pullback f g c → is-simplicially-fully-faithful g → is-simplicially-fully-faithful (vertical-map-cone f g c) + -- is-simplicially-fully-faithful-vertical-map-cone-is-pullback pb is-ff-g = + -- is-simplicially-fully-faithful-is-prop-map + -- ( is-trunc-vertical-map-is-pullback neg-one-𝕋 f g c pb + -- ( is-prop-map-is-simplicially-fully-faithful is-ff-g)) + + -- abstract + -- is-simplicially-fully-faithful-horizontal-map-cone-is-pullback : + -- is-pullback f g c → is-simplicially-fully-faithful f → is-simplicially-fully-faithful (horizontal-map-cone f g c) + -- is-simplicially-fully-faithful-horizontal-map-cone-is-pullback pb is-ff-f = + -- is-simplicially-fully-faithful-is-prop-map + -- ( is-trunc-horizontal-map-is-pullback neg-one-𝕋 f g c pb + -- ( is-prop-map-is-simplicially-fully-faithful is-ff-f)) +``` + +### A map is an embedding if and only if it has contractible fibers at values + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where + + -- is-simplicially-fully-faithful-is-contr-fibers-values' : + -- ((a : A) → is-contr (fiber' f (f a))) → is-simplicially-fully-faithful f + -- is-simplicially-fully-faithful-is-contr-fibers-values' c a = + -- fundamental-theorem-id (c a) (λ x → ap f {a} {x}) + + -- is-simplicially-fully-faithful-is-contr-fibers-values : + -- ((a : A) → is-contr (fiber f (f a))) → is-simplicially-fully-faithful f + -- is-simplicially-fully-faithful-is-contr-fibers-values c = + -- is-simplicially-fully-faithful-is-contr-fibers-values' + -- ( λ a → + -- is-contr-equiv' + -- ( fiber f (f a)) + -- ( equiv-fiber f (f a)) + -- ( c a)) + + -- is-contr-fibers-values-is-simplicially-fully-faithful' : + -- is-simplicially-fully-faithful f → ((a : A) → is-contr (fiber' f (f a))) + -- is-contr-fibers-values-is-simplicially-fully-faithful' e a = + -- fundamental-theorem-id' (λ x → ap f {a} {x}) (e a) + + -- is-contr-fibers-values-is-simplicially-fully-faithful : + -- is-simplicially-fully-faithful f → ((a : A) → is-contr (fiber f (f a))) + -- is-contr-fibers-values-is-simplicially-fully-faithful e a = + -- is-contr-equiv + -- ( fiber' f (f a)) + -- ( equiv-fiber f (f a)) + -- ( is-contr-fibers-values-is-simplicially-fully-faithful' e a) +``` diff --git a/src/simplicial-type-theory/globularly-coskeletal-types.lagda.md b/src/simplicial-type-theory/globularly-coskeletal-types.lagda.md new file mode 100644 index 0000000000..9e3b3e3d3b --- /dev/null +++ b/src/simplicial-type-theory/globularly-coskeletal-types.lagda.md @@ -0,0 +1,597 @@ +# Globularly coskeletal types + +```agda +module simplicial-type-theory.globularly-coskeletal-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.function-extensionality +open import foundation.functoriality-cartesian-product-types +open import foundation.inhabited-types +open import foundation.propositional-truncations +open import foundation.raising-universe-levels +open import foundation.type-arithmetic-booleans +open import foundation.universe-levels + +open import foundation-core.cartesian-product-types +open import foundation-core.contractible-types +open import foundation-core.embeddings +open import foundation-core.equality-dependent-pair-types +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.propositions +open import foundation-core.retractions +open import foundation-core.retracts-of-types +open import foundation-core.transport-along-identifications +open import foundation-core.truncation-levels + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-edges-cartesian-product-types +open import simplicial-type-theory.directed-edges-dependent-pair-types +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.natural-transformations +open import simplicial-type-theory.simplicially-discrete-types +open import simplicial-type-theory.whiskering-directed-edges +``` + +
+ +## Idea + +The globular coskeletality of a type is a measure of the complexity of its +hom-types. The simplest case is a contractible type. This is the base case of +the inductive definition of globular coskeletality for types. A type is +`k+1`-globularly coskeletal if its hom-types are `k`-globularly coskeletal. + +**Note.** This is not coskeletality in simplicial spaces, but coskeletality in +globular spaces. However, the two agree in many cases we care about, i.e. when +`k = 0` or the types are Segal. + +**Question.** Should `k+1`-globular coskeletality also require the identity +types to be `k`-globular coskeletal? Probably yes, c.f. higher modalities. + +## Definition + +### The condition of coskeletality + +```agda +is-globularly-coskeletal : {l : Level} (k : 𝕋) → UU l → UU l +is-globularly-coskeletal neg-two-𝕋 A = is-contr A +is-globularly-coskeletal (succ-𝕋 k) A = + (x y : A) → is-globularly-coskeletal k (x →₂ y) + +is-globularly-coskeletal-eq : + {l : Level} {k k' : 𝕋} {A : UU l} → + k = k' → is-globularly-coskeletal k A → is-globularly-coskeletal k' A +is-globularly-coskeletal-eq refl H = H +``` + +### The universe of globularly coskeletal types + +```agda +Globularly-Coskeletal-Type : (l : Level) → 𝕋 → UU (lsuc l) +Globularly-Coskeletal-Type l k = Σ (UU l) (is-globularly-coskeletal k) + +module _ + {k : 𝕋} {l : Level} + where + + type-Globularly-Coskeletal-Type : Globularly-Coskeletal-Type l k → UU l + type-Globularly-Coskeletal-Type = pr1 + + is-globularly-coskeletal-type-Globularly-Coskeletal-Type : + (A : Globularly-Coskeletal-Type l k) → + is-globularly-coskeletal k (type-Globularly-Coskeletal-Type A) + is-globularly-coskeletal-type-Globularly-Coskeletal-Type = pr2 +``` + +## Properties + +### Being globularly coskeletal is a property + +```agda +abstract + is-prop-is-globularly-coskeletal : + {l : Level} (k : 𝕋) (A : UU l) → is-prop (is-globularly-coskeletal k A) + is-prop-is-globularly-coskeletal neg-two-𝕋 A = is-property-is-contr + is-prop-is-globularly-coskeletal (succ-𝕋 k) A = + is-prop-Π + ( λ x → is-prop-Π (λ y → is-prop-is-globularly-coskeletal k (x →₂ y))) + +is-globularly-coskeletal-Prop : {l : Level} (k : 𝕋) (A : UU l) → Prop l +pr1 (is-globularly-coskeletal-Prop k A) = is-globularly-coskeletal k A +pr2 (is-globularly-coskeletal-Prop k A) = is-prop-is-globularly-coskeletal k A +``` + +### A type is `-1`-globularly coskeletal if and only if it is `∂𝟚 ↪ 𝟚`-local + +This remains to be formalized. + +### A type is `k`-globularly coskeletal if and only if it is local at the `k`'th directed suspension of `0 → 1` + +This remains to be formalized. + +### If a type is `k`-globularly coskeletal then it is `k+1`-globularly coskeletal + +```agda +is-contr-hom-is-contr : + {l : Level} {A : UU l} → is-contr A → (x y : A) → is-contr (x →₂ y) +is-contr-hom-is-contr H x y = + is-contr-is-equiv' + ( x = y) + ( simplicial-hom-eq) + ( is-simplicially-discrete-is-contr H x y) + ( is-prop-is-contr H x y) + +abstract + is-globularly-coskeletal-succ-is-globularly-coskeletal : + (k : 𝕋) {l : Level} {A : UU l} → + is-globularly-coskeletal k A → is-globularly-coskeletal (succ-𝕋 k) A + is-globularly-coskeletal-succ-is-globularly-coskeletal neg-two-𝕋 = + is-contr-hom-is-contr + is-globularly-coskeletal-succ-is-globularly-coskeletal (succ-𝕋 k) H x y = + is-globularly-coskeletal-succ-is-globularly-coskeletal k (H x y) + +truncated-type-succ-Globularly-Coskeletal-Type : + (k : 𝕋) {l : Level} → + Globularly-Coskeletal-Type l k → Globularly-Coskeletal-Type l (succ-𝕋 k) +pr1 (truncated-type-succ-Globularly-Coskeletal-Type k A) = + type-Globularly-Coskeletal-Type A +pr2 (truncated-type-succ-Globularly-Coskeletal-Type k A) = + is-globularly-coskeletal-succ-is-globularly-coskeletal k + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) +``` + +### The hom-types of a `k`-globularly coskeletal type are `k`-globularly coskeletal + +```agda +abstract + is-globularly-coskeletal-hom : + {l : Level} {k : 𝕋} {A : UU l} → + is-globularly-coskeletal k A → + (x y : A) → is-globularly-coskeletal k (x →₂ y) + is-globularly-coskeletal-hom {k = k} = + is-globularly-coskeletal-succ-is-globularly-coskeletal k + +hom-Globularly-Coskeletal-Type : + {l : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l (succ-𝕋 k)) → + (x y : type-Globularly-Coskeletal-Type A) → Globularly-Coskeletal-Type l k +hom-Globularly-Coskeletal-Type A x y = + ( (x →₂ y) , is-globularly-coskeletal-type-Globularly-Coskeletal-Type A x y) + +hom-Globularly-Coskeletal-Type' : + {l : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l k) → + (x y : type-Globularly-Coskeletal-Type A) → Globularly-Coskeletal-Type l k +hom-Globularly-Coskeletal-Type' A x y = + ( (x →₂ y) , + is-globularly-coskeletal-hom + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) x y) +``` + +### The identity types of a `k`-globularly coskeletal type are `k`-globularly coskeletal + +This should be true for coskeletality to be a modality. + +```text +-- abstract +-- is-globularly-coskeletal-Id : +-- {l : Level} {k : 𝕋} {A : UU l} → +-- is-globularly-coskeletal k A → (x y : A) → is-globularly-coskeletal k (x = y) +-- is-globularly-coskeletal-Id {k = neg-two-𝕋} = is-prop-is-contr +-- is-globularly-coskeletal-Id {k = succ-𝕋 k} H x y p q = {! !} + +-- Id-Globularly-Coskeletal-Type' : +-- {l : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l k) → +-- (x y : type-Globularly-Coskeletal-Type A) → Globularly-Coskeletal-Type l k +-- pr1 (Id-Globularly-Coskeletal-Type' A x y) = (x = y) +-- pr2 (Id-Globularly-Coskeletal-Type' A x y) = +-- is-globularly-coskeletal-Id (is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) x y +``` + +### `k`-globularly coskeletal types are closed under retracts + +```agda +module _ + {l1 l2 : Level} + where + + is-globularly-coskeletal-retract-of : + (k : 𝕋) {A : UU l1} {B : UU l2} → + A retract-of B → is-globularly-coskeletal k B → is-globularly-coskeletal k A + is-globularly-coskeletal-retract-of neg-two-𝕋 = is-contr-retract-of _ + is-globularly-coskeletal-retract-of (succ-𝕋 k) R H x y = + is-globularly-coskeletal-retract-of k + ( retract-simplicial-hom R x y) + ( H (pr1 R x) (pr1 R y)) +``` + +### `k`-globularly coskeletal types are closed under equivalences + +```agda +abstract + is-globularly-coskeletal-is-equiv : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} (B : UU l2) (f : A → B) → is-equiv f → + is-globularly-coskeletal k B → is-globularly-coskeletal k A + is-globularly-coskeletal-is-equiv k B f is-equiv-f = + is-globularly-coskeletal-retract-of k (f , pr2 is-equiv-f) + +abstract + is-globularly-coskeletal-equiv : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} (B : UU l2) (e : A ≃ B) → + is-globularly-coskeletal k B → is-globularly-coskeletal k A + is-globularly-coskeletal-equiv k B (f , is-equiv-f) = + is-globularly-coskeletal-is-equiv k B f is-equiv-f + +abstract + is-globularly-coskeletal-is-equiv' : + {l1 l2 : Level} (k : 𝕋) (A : UU l1) {B : UU l2} (f : A → B) → + is-equiv f → is-globularly-coskeletal k A → is-globularly-coskeletal k B + is-globularly-coskeletal-is-equiv' + k A f is-equiv-f is-globularly-coskeletal-A = + is-globularly-coskeletal-is-equiv k A + ( map-inv-is-equiv is-equiv-f) + ( is-equiv-map-inv-is-equiv is-equiv-f) + ( is-globularly-coskeletal-A) + +abstract + is-globularly-coskeletal-equiv' : + {l1 l2 : Level} (k : 𝕋) (A : UU l1) {B : UU l2} (e : A ≃ B) → + is-globularly-coskeletal k A → is-globularly-coskeletal k B + is-globularly-coskeletal-equiv' k A (f , is-equiv-f) = + is-globularly-coskeletal-is-equiv' k A f is-equiv-f +``` + +### If a type simplicially embeds into a `k+1`-globularly coskeletal type, then it is `k+1`-globularly coskeletal + +```agda +abstract + is-globularly-coskeletal-is-simplicially-fully-faithful : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (f : A → B) → + is-simplicially-fully-faithful f → + is-globularly-coskeletal (succ-𝕋 k) B → + is-globularly-coskeletal (succ-𝕋 k) A + is-globularly-coskeletal-is-simplicially-fully-faithful k f Ef H x y = + is-globularly-coskeletal-is-equiv k (f x →₂ f y) + ( action-simplicial-hom-function f {x} {y}) + ( Ef x y) + ( H (f x) (f y)) + +abstract + is-globularly-coskeletal-simplicially-fully-faithful-map : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (f : A ↪▵ B) → + is-globularly-coskeletal (succ-𝕋 k) B → + is-globularly-coskeletal (succ-𝕋 k) A + is-globularly-coskeletal-simplicially-fully-faithful-map k f = + is-globularly-coskeletal-is-simplicially-fully-faithful k + ( map-simplicially-fully-faithful-map f) + ( is-simplicially-fully-faithful-map-simplicially-fully-faithful-map f) +``` + +In fact, it suffices that the action on homs has a retraction. + +```agda +abstract + is-globularly-coskeletal-retraction-ap : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (f : A → B) → + ((x y : A) → retraction (action-simplicial-hom-function f {x} {y})) → + is-globularly-coskeletal (succ-𝕋 k) B → + is-globularly-coskeletal (succ-𝕋 k) A + is-globularly-coskeletal-retraction-ap k f Ef H x y = + is-globularly-coskeletal-retract-of k + ( action-simplicial-hom-function f {x} {y} , Ef x y) + ( H (f x) (f y)) +``` + +### Globularly coskeletal types are closed under dependent pair types + +```text +-- abstract +-- is-globularly-coskeletal-Σ : +-- {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : A → UU l2} → +-- is-globularly-coskeletal k A → ((x : A) → is-globularly-coskeletal k (B x)) → is-globularly-coskeletal k (Σ A B) +-- is-globularly-coskeletal-Σ {k = neg-two-𝕋} is-globularly-coskeletal-A is-globularly-coskeletal-B = +-- is-contr-Σ' is-globularly-coskeletal-A is-globularly-coskeletal-B +-- is-globularly-coskeletal-Σ {k = succ-𝕋 k} {B = B} is-globularly-coskeletal-A is-globularly-coskeletal-B s t = +-- is-globularly-coskeletal-equiv k +-- ( simplicial-hom-Σ s t) +-- ( compute-simplicial-hom-Σ) +-- ( is-globularly-coskeletal-Σ +-- ( is-globularly-coskeletal-A (pr1 s) (pr1 t)) +-- {! is-globularly-coskeletal-B ? ? ? !}) + + -- is-globularly-coskeletal-equiv k + -- ( Σ (pr1 s = pr1 t) (λ p → tr B p (pr2 s) = pr2 t)) + -- ( equiv-pair-eq-Σ s t) + -- ( is-globularly-coskeletal-Σ + -- ( is-globularly-coskeletal-A (pr1 s) (pr1 t)) + -- ( λ p → is-globularly-coskeletal-B (pr1 t) (tr B p (pr2 s)) (pr2 t))) + +-- Σ-Globularly-Coskeletal-Type : +-- {l1 l2 : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l1 k) +-- (B : type-Globularly-Coskeletal-Type A → Globularly-Coskeletal-Type l2 k) → +-- Globularly-Coskeletal-Type (l1 ⊔ l2) k +-- pr1 (Σ-Globularly-Coskeletal-Type A B) = +-- Σ (type-Globularly-Coskeletal-Type A) (λ a → type-Globularly-Coskeletal-Type (B a)) +-- pr2 (Σ-Globularly-Coskeletal-Type A B) = +-- is-globularly-coskeletal-Σ +-- ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) +-- ( λ a → is-globularly-coskeletal-type-Globularly-Coskeletal-Type (B a)) + +-- fiber-Globularly-Coskeletal-Type : +-- {l1 l2 : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l1 k) +-- (B : Globularly-Coskeletal-Type l2 k) +-- (f : type-Globularly-Coskeletal-Type A → type-Globularly-Coskeletal-Type B) → +-- type-Globularly-Coskeletal-Type B → Globularly-Coskeletal-Type (l1 ⊔ l2) k +-- fiber-Globularly-Coskeletal-Type A B f b = +-- Σ-Globularly-Coskeletal-Type A (λ a → Id-Globularly-Coskeletal-Type' B (f a) b) +``` + +### Products of families of globularly coskeletal types are globularly coskeletal + +```agda +abstract + is-globularly-coskeletal-Π : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} → + ((x : A) → is-globularly-coskeletal k (B x)) → + is-globularly-coskeletal k ((x : A) → B x) + is-globularly-coskeletal-Π neg-two-𝕋 is-globularly-coskeletal-B = + is-contr-Π is-globularly-coskeletal-B + is-globularly-coskeletal-Π (succ-𝕋 k) is-globularly-coskeletal-B f g = + is-globularly-coskeletal-is-equiv k (f ⇒▵ g) + ( simplicial-natural-transformation-simplicial-edge-of-dependent-functions) + ( is-equiv-simplicial-natural-transformation-simplicial-edge-of-dependent-functions) + ( is-globularly-coskeletal-Π k + ( λ x → is-globularly-coskeletal-B x (f x) (g x))) + +type-Π-Globularly-Coskeletal-Type' : + (k : 𝕋) {l1 l2 : Level} + (A : UU l1) (B : A → Globularly-Coskeletal-Type l2 k) → + UU (l1 ⊔ l2) +type-Π-Globularly-Coskeletal-Type' k A B = + (x : A) → type-Globularly-Coskeletal-Type (B x) + +is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type' : + (k : 𝕋) {l1 l2 : Level} + (A : UU l1) (B : A → Globularly-Coskeletal-Type l2 k) → + is-globularly-coskeletal k (type-Π-Globularly-Coskeletal-Type' k A B) +is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type' k A B = + is-globularly-coskeletal-Π k + ( λ x → is-globularly-coskeletal-type-Globularly-Coskeletal-Type (B x)) + +Π-Globularly-Coskeletal-Type' : + (k : 𝕋) {l1 l2 : Level} + (A : UU l1) (B : A → Globularly-Coskeletal-Type l2 k) → + Globularly-Coskeletal-Type (l1 ⊔ l2) k +pr1 (Π-Globularly-Coskeletal-Type' k A B) = + type-Π-Globularly-Coskeletal-Type' k A B +pr2 (Π-Globularly-Coskeletal-Type' k A B) = + is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type' k A B + +type-Π-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : type-Globularly-Coskeletal-Type A → Globularly-Coskeletal-Type l2 k) → + UU (l1 ⊔ l2) +type-Π-Globularly-Coskeletal-Type k A B = + type-Π-Globularly-Coskeletal-Type' k (type-Globularly-Coskeletal-Type A) B + +is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : type-Globularly-Coskeletal-Type A → Globularly-Coskeletal-Type l2 k) → + is-globularly-coskeletal k (type-Π-Globularly-Coskeletal-Type k A B) +is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type k A B = + is-globularly-coskeletal-type-Π-Globularly-Coskeletal-Type' k + ( type-Globularly-Coskeletal-Type A) B + +Π-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : type-Globularly-Coskeletal-Type A → Globularly-Coskeletal-Type l2 k) → + Globularly-Coskeletal-Type (l1 ⊔ l2) k +Π-Globularly-Coskeletal-Type k A B = + Π-Globularly-Coskeletal-Type' k (type-Globularly-Coskeletal-Type A) B +``` + +### The type of functions into a globularly coskeletal type is globularly coskeletal + +```agda +abstract + is-globularly-coskeletal-function-type : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k B → is-globularly-coskeletal k (A → B) + is-globularly-coskeletal-function-type k is-globularly-coskeletal-B = + is-globularly-coskeletal-Π k (λ _ → is-globularly-coskeletal-B) + +function-type-Globularly-Coskeletal-Type : + {l1 l2 : Level} {k : 𝕋} (A : UU l1) (B : Globularly-Coskeletal-Type l2 k) → + Globularly-Coskeletal-Type (l1 ⊔ l2) k +pr1 (function-type-Globularly-Coskeletal-Type A B) = + A → type-Globularly-Coskeletal-Type B +pr2 (function-type-Globularly-Coskeletal-Type A B) = + is-globularly-coskeletal-function-type _ + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type B) + +type-function-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : Globularly-Coskeletal-Type l2 k) → UU (l1 ⊔ l2) +type-function-Globularly-Coskeletal-Type k A B = + type-Globularly-Coskeletal-Type A → type-Globularly-Coskeletal-Type B + +is-globularly-coskeletal-type-function-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : Globularly-Coskeletal-Type l2 k) → + is-globularly-coskeletal k (type-function-Globularly-Coskeletal-Type k A B) +is-globularly-coskeletal-type-function-Globularly-Coskeletal-Type k A B = + is-globularly-coskeletal-function-type k + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type B) + +function-Globularly-Coskeletal-Type : + (k : 𝕋) {l1 l2 : Level} (A : Globularly-Coskeletal-Type l1 k) + (B : Globularly-Coskeletal-Type l2 k) → Globularly-Coskeletal-Type (l1 ⊔ l2) k +pr1 (function-Globularly-Coskeletal-Type k A B) = + type-function-Globularly-Coskeletal-Type k A B +pr2 (function-Globularly-Coskeletal-Type k A B) = + is-globularly-coskeletal-type-function-Globularly-Coskeletal-Type k A B +``` + +### Products of globularly coskeletal types are globularly coskeletal + +```agda +abstract + is-globularly-coskeletal-product-Level : + {l : Level} (k : 𝕋) {A B : UU l} → + is-globularly-coskeletal k A → + is-globularly-coskeletal k B → + is-globularly-coskeletal k (A × B) + is-globularly-coskeletal-product-Level + k {A} {B} is-globularly-coskeletal-A is-globularly-coskeletal-B = + is-globularly-coskeletal-equiv' k + ( (b : bool) → rec-bool A B b) + ( equiv-Π-bool-product (rec-bool A B)) + ( is-globularly-coskeletal-Π k + ( ind-bool + ( is-globularly-coskeletal k ∘ rec-bool A B) + ( is-globularly-coskeletal-A) + ( is-globularly-coskeletal-B))) + +abstract + is-globularly-coskeletal-product : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k A → + is-globularly-coskeletal k B → + is-globularly-coskeletal k (A × B) + is-globularly-coskeletal-product + {l1} {l2} k {A} {B} is-globularly-coskeletal-A is-globularly-coskeletal-B = + is-globularly-coskeletal-equiv k + ( raise (l1 ⊔ l2) A × raise (l1 ⊔ l2) B) + ( equiv-product (compute-raise (l1 ⊔ l2) A) (compute-raise (l1 ⊔ l2) B)) + ( is-globularly-coskeletal-product-Level k + ( is-globularly-coskeletal-equiv' k A + ( compute-raise (l1 ⊔ l2) A) + ( is-globularly-coskeletal-A)) + ( is-globularly-coskeletal-equiv' k B + ( compute-raise (l1 ⊔ l2) B) + ( is-globularly-coskeletal-B))) + +product-Globularly-Coskeletal-Type : + {l1 l2 : Level} (k : 𝕋) → + Globularly-Coskeletal-Type l1 k → + Globularly-Coskeletal-Type l2 k → + Globularly-Coskeletal-Type (l1 ⊔ l2) k +pr1 (product-Globularly-Coskeletal-Type k A B) = + type-Globularly-Coskeletal-Type A × type-Globularly-Coskeletal-Type B +pr2 (product-Globularly-Coskeletal-Type k A B) = + is-globularly-coskeletal-product k + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) + ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type B) + +is-globularly-coskeletal-product' : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + (B → is-globularly-coskeletal (succ-𝕋 k) A) → + (A → is-globularly-coskeletal (succ-𝕋 k) B) → + is-globularly-coskeletal (succ-𝕋 k) (A × B) +is-globularly-coskeletal-product' k f g (a , b) (a' , b') = + is-globularly-coskeletal-equiv k + ( simplicial-hom-product (a , b) (a' , b')) + ( compute-simplicial-hom-product) + ( is-globularly-coskeletal-product k (f b a a') (g a b b')) + +is-globularly-coskeletal-left-factor-product' : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k (A × B) → B → is-globularly-coskeletal k A +is-globularly-coskeletal-left-factor-product' neg-two-𝕋 {A} {B} H b = + is-contr-left-factor-product A B H +is-globularly-coskeletal-left-factor-product' (succ-𝕋 k) H b a a' = + is-globularly-coskeletal-left-factor-product' k + ( is-globularly-coskeletal-equiv' k + ( (a , b) →₂ (a' , b)) + ( compute-simplicial-hom-product) + ( H (a , b) (a' , b))) + ( id-simplicial-hom b) + +is-globularly-coskeletal-left-factor-product : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k (A × B) → + is-inhabited B → + is-globularly-coskeletal k A +is-globularly-coskeletal-left-factor-product k {A} {B} H = + rec-trunc-Prop + ( is-globularly-coskeletal-Prop k A) + ( is-globularly-coskeletal-left-factor-product' k H) + +is-globularly-coskeletal-right-factor-product' : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k (A × B) → A → is-globularly-coskeletal k B +is-globularly-coskeletal-right-factor-product' neg-two-𝕋 {A} {B} H a = + is-contr-right-factor-product A B H +is-globularly-coskeletal-right-factor-product' (succ-𝕋 k) H a b b' = + is-globularly-coskeletal-right-factor-product' k + ( is-globularly-coskeletal-equiv' k + ( (a , b) →₂ (a , b')) + ( compute-simplicial-hom-product) + ( H (a , b) (a , b'))) + ( id-simplicial-hom a) + +is-globularly-coskeletal-right-factor-product : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} → + is-globularly-coskeletal k (A × B) → + is-inhabited A → + is-globularly-coskeletal k B +is-globularly-coskeletal-right-factor-product k {A} {B} H = + rec-trunc-Prop + ( is-globularly-coskeletal-Prop k B) + ( is-globularly-coskeletal-right-factor-product' k H) +``` + +### The type of equivalences between globularly coskeletal types is globularly coskeletal + +**Proof.** The type of equivalences `A ≃ B` is a pullback + +```text + A ≃ B ---> (B → A) × (A → B) × (B → A) + | ⌟ | + | | + ∨ ∨ + 1 ----------> (A → A) × (B → B) + (id , id) +``` + +so the result follows by pullback stability. + +```text +-- module _ +-- {l1 l2 : Level} {A : UU l1} {B : UU l2} +-- where + +-- is-globularly-coskeletal-equiv-is-globularly-coskeletal : +-- (k : 𝕋) → is-globularly-coskeletal k A → is-globularly-coskeletal k B → is-globularly-coskeletal k (A ≃ B) +-- is-globularly-coskeletal-equiv-is-globularly-coskeletal k H K = {! !} + +-- type-equiv-Globularly-Coskeletal-Type : +-- {l1 l2 : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l1 k) (B : Globularly-Coskeletal-Type l2 k) → +-- UU (l1 ⊔ l2) +-- type-equiv-Globularly-Coskeletal-Type A B = +-- type-Globularly-Coskeletal-Type A ≃ type-Globularly-Coskeletal-Type B + +-- is-globularly-coskeletal-type-equiv-Globularly-Coskeletal-Type : +-- {l1 l2 : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l1 k) (B : Globularly-Coskeletal-Type l2 k) → +-- is-globularly-coskeletal k (type-equiv-Globularly-Coskeletal-Type A B) +-- is-globularly-coskeletal-type-equiv-Globularly-Coskeletal-Type A B = +-- is-globularly-coskeletal-equiv-is-globularly-coskeletal _ +-- ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type A) +-- ( is-globularly-coskeletal-type-Globularly-Coskeletal-Type B) + +-- equiv-Globularly-Coskeletal-Type : +-- {l1 l2 : Level} {k : 𝕋} (A : Globularly-Coskeletal-Type l1 k) (B : Globularly-Coskeletal-Type l2 k) → +-- Globularly-Coskeletal-Type (l1 ⊔ l2) k +-- pr1 (equiv-Globularly-Coskeletal-Type A B) = type-equiv-Globularly-Coskeletal-Type A B +-- pr2 (equiv-Globularly-Coskeletal-Type A B) = is-globularly-coskeletal-type-equiv-Globularly-Coskeletal-Type A B +``` diff --git a/src/simplicial-type-theory/horizontal-composition-directed-edges-functions.lagda.md b/src/simplicial-type-theory/horizontal-composition-directed-edges-functions.lagda.md new file mode 100644 index 0000000000..5f046e4685 --- /dev/null +++ b/src/simplicial-type-theory/horizontal-composition-directed-edges-functions.lagda.md @@ -0,0 +1,105 @@ +# Horizontal composition of directed edges between functions + +```agda +module simplicial-type-theory.horizontal-composition-directed-edges-functions where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.horizontal-composition-simplicial-arrows-functions +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a directed edge `α` between functions `f g : A → B` and a directed edge +`β` of functions `f' g' : B → C`, we may +{{#concept "horizontally compose" Disambiguation="directed edges of functions" Agda=horizontal-comp-simplicial-hom}} +them to obtain a directed edge of functions `f' ∘ f →▵ g' ∘ g`. The horizontal +composite is constructed by "synchronously traversing `α` and `β`", defined on +the underlying [simplicial arrows](simplicial-type-theory.simplicial-arrows.md) +as: + +```text + β □ α := (t ↦ x ↦ β t (α t x)). +``` + +## Definitions + +### Horizontal composition of directed edges between functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + {f g : A → B} {f' g' : B → C} + where + + horizontal-comp-simplicial-hom : f' →₂ g' → f →₂ g → (f' ∘ f) →₂ (g' ∘ g) + horizontal-comp-simplicial-hom (β , s , t) (α , s' , t') = + ( ( horizontal-comp-simplicial-arrow β α) , + ( ap (β 0₂ ∘_) s' ∙ ap (_∘ f) s) , + ( ap (β 1₂ ∘_) t' ∙ ap (_∘ g) t)) +``` + +## Properties + +### Unit laws for horizontal composition of directed edges of functions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} + where + + left-unit-law-horizontal-comp-simplicial-hom : + (α : f →₂ g) → + horizontal-comp-simplicial-hom (id-simplicial-hom id) α = α + left-unit-law-horizontal-comp-simplicial-hom (α , s , t) = + eq-pair-eq-fiber (eq-pair (right-unit ∙ ap-id s) (right-unit ∙ ap-id t)) + + right-unit-law-horizontal-comp-simplicial-hom : + (α : f →₂ g) → + horizontal-comp-simplicial-hom α (id-simplicial-hom id) = α + right-unit-law-horizontal-comp-simplicial-hom (α , s , t) = + eq-pair-eq-fiber (eq-pair (ap-id s) (ap-id t)) +``` + +### Associativity of horizontal composition of directed edges of functions + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + {h h' : C → D} {g g' : B → C} {f f' : A → B} + where + + associative-horizontal-comp-simplicial-hom : + (γ : h →₂ h') (β : g →₂ g') (α : f →₂ f') → + horizontal-comp-simplicial-hom (horizontal-comp-simplicial-hom γ β) α = + horizontal-comp-simplicial-hom γ (horizontal-comp-simplicial-hom β α) + associative-horizontal-comp-simplicial-hom + ( γ , refl , refl) (β , refl , refl) (α , refl , refl) = refl +``` diff --git a/src/simplicial-type-theory/horizontal-composition-natural-transformations.lagda.md b/src/simplicial-type-theory/horizontal-composition-natural-transformations.lagda.md new file mode 100644 index 0000000000..5089e9d27a --- /dev/null +++ b/src/simplicial-type-theory/horizontal-composition-natural-transformations.lagda.md @@ -0,0 +1,160 @@ +# Horizontal composition of natural transformations between functions + +```agda +module simplicial-type-theory.horizontal-composition-natural-transformations where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.horizontal-composition-directed-edges-functions +open import simplicial-type-theory.horizontal-composition-simplicial-arrows-functions +open import simplicial-type-theory.natural-transformations +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a natural transformation `α` between functions `f g : A → B` and a natural +transformation `β` of functions `f' g' : B → C`, we may +{{#concept "horizontally compose" Disambiguation="natural transformations of functions" Agda=horizontal-comp-simplicial-natural-transformation}} +them to obtain a natural transformation of functions `f' ∘ f ⇒▵ g' ∘ g`. The +horizontal composite is constructed by "synchronously traversing `α` and `β`", +defined on the underlying +[simplicial arrows](simplicial-type-theory.simplicial-arrows.md) as: + +```text + β □ α := (t ↦ x ↦ β t (α t x)). +``` + +## Definitions + +### Horizontal composition of natural transformations between functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + {f g : A → B} {f' g' : B → C} + where + + horizontal-comp-simplicial-natural-transformation : + f' ⇒▵ g' → f ⇒▵ g → (f' ∘ f) ⇒▵ (g' ∘ g) + horizontal-comp-simplicial-natural-transformation β α x = + ( λ t → + simplicial-arrow-simplicial-hom + ( β (simplicial-arrow-simplicial-hom (α x) t)) + ( t)) , + ( ( ap + ( λ u → simplicial-arrow-simplicial-hom (β u) 0₂) + ( eq-source-simplicial-natural-transformation α x)) ∙ + ( eq-source-simplicial-natural-transformation β (f x))) , + ( ( ap + ( λ u → simplicial-arrow-simplicial-hom (β u) 1₂) + ( eq-target-simplicial-natural-transformation α x)) ∙ + ( eq-target-simplicial-natural-transformation β (g x))) +``` + +### The action of a natural transformation on a directed edge + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} + where + + action-simplicial-hom-simplicial-natural-transformation : + f ⇒▵ g → {x y : A} → x →₂ y → f x →₂ g y + action-simplicial-hom-simplicial-natural-transformation α a = + ( λ t → + family-of-simplicial-arrows-simplicial-natural-transformation α + ( simplicial-arrow-simplicial-hom a t) + ( t)) , + ( ( eq-source-simplicial-natural-transformation α + ( simplicial-arrow-simplicial-hom a 0₂)) ∙ + ( ap f (eq-source-simplicial-hom a))) , + ( ( eq-target-simplicial-natural-transformation α + ( simplicial-arrow-simplicial-hom a 1₂)) ∙ + ( ap g (eq-target-simplicial-hom a))) +``` + +## Properties + +### Unit laws for horizontal composition of directed edges of functions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} + where + + left-unit-law-horizontal-comp-simplicial-natural-transformation : + (α : f ⇒▵ g) → + ( horizontal-comp-simplicial-natural-transformation + ( id-simplicial-natural-transformation id) + ( α)) = + ( α) + left-unit-law-horizontal-comp-simplicial-natural-transformation α = + eq-htpy + ( λ x → + eq-pair-eq-fiber + ( eq-pair + ( right-unit ∙ + ap-id (eq-source-simplicial-natural-transformation α x)) + ( right-unit ∙ + ap-id (eq-target-simplicial-natural-transformation α x)))) + + right-unit-law-horizontal-comp-simplicial-natural-transformation : + (α : f ⇒▵ g) → + ( horizontal-comp-simplicial-natural-transformation + ( α) + ( id-simplicial-natural-transformation id)) = + ( α) + right-unit-law-horizontal-comp-simplicial-natural-transformation α = refl +``` + +### Associativity of horizontal composition of directed edges of functions + +```text +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + {h h' : C → D} {g g' : B → C} {f f' : A → B} + where + + associative-horizontal-comp-simplicial-natural-transformation : + (γ : h ⇒▵ h') (β : g ⇒▵ g') (α : f ⇒▵ f') → + ( horizontal-comp-simplicial-natural-transformation + ( horizontal-comp-simplicial-natural-transformation γ β) + ( α)) = + ( horizontal-comp-simplicial-natural-transformation + ( γ) + ( horizontal-comp-simplicial-natural-transformation β α)) + associative-horizontal-comp-simplicial-natural-transformation γ β α = + eq-htpy + ( λ x → + eq-pair-eq-fiber + ( eq-pair + ( equational-reasoning {! !} = {! !} by {! !}) + {! !})) +``` diff --git a/src/simplicial-type-theory/horizontal-composition-simplicial-arrows-functions.lagda.md b/src/simplicial-type-theory/horizontal-composition-simplicial-arrows-functions.lagda.md new file mode 100644 index 0000000000..cab9cae5bb --- /dev/null +++ b/src/simplicial-type-theory/horizontal-composition-simplicial-arrows-functions.lagda.md @@ -0,0 +1,99 @@ +# Horizontal composition of simplicial arrows of functions + +```agda +module simplicial-type-theory.horizontal-composition-simplicial-arrows-functions where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a [simplicial arrow](simplicial-type-theory.simplicial-arrows.md) `α` of +functions `A → B` and a simplicial arrow `β` of functions `B → C`, we may +{{#concept "horizontally compose" Disambiguation="simplicial arrows of functions" Agda=horizontal-comp-simplicial-arrow}} +them to obtain a simplicial arrow of functions `A → C`. The horizontal composite +is constructed by "synchronously traversing `α` and `β`": + +```text + β □ α := (t ↦ x ↦ β t (α t x)), +``` + +and thus satisfies unit laws and associativity strictly. + +## Definitions + +### Horizontal composition of simplicial arrows of functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + horizontal-comp-simplicial-arrow : + simplicial-arrow (B → C) → + simplicial-arrow (A → B) → + simplicial-arrow (A → C) + horizontal-comp-simplicial-arrow β α t x = β t (α t x) +``` + +## Properties + +### Unit laws for horizontal composition of simplicial arrows of functions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + left-unit-law-horizontal-comp-simplicial-arrow : + (α : simplicial-arrow (A → B)) → + horizontal-comp-simplicial-arrow (id-simplicial-arrow id) α = α + left-unit-law-horizontal-comp-simplicial-arrow α = refl + + right-unit-law-horizontal-comp-simplicial-arrow : + (α : simplicial-arrow (A → B)) → + horizontal-comp-simplicial-arrow α (id-simplicial-arrow id) = α + right-unit-law-horizontal-comp-simplicial-arrow α = refl +``` + +### Associativity of horizontal composition of simplicial arrows of functions + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + where + + associative-horizontal-comp-simplicial-arrow : + (γ : simplicial-arrow (C → D)) + (β : simplicial-arrow (B → C)) + (α : simplicial-arrow (A → B)) → + horizontal-comp-simplicial-arrow (horizontal-comp-simplicial-arrow γ β) α = + horizontal-comp-simplicial-arrow γ (horizontal-comp-simplicial-arrow β α) + associative-horizontal-comp-simplicial-arrow γ β α = refl +``` diff --git a/src/simplicial-type-theory/inequality-directed-interval-type.lagda.md b/src/simplicial-type-theory/inequality-directed-interval-type.lagda.md new file mode 100644 index 0000000000..bbdc575986 --- /dev/null +++ b/src/simplicial-type-theory/inequality-directed-interval-type.lagda.md @@ -0,0 +1,240 @@ +# The inequality on the directed interval type + +```agda +module simplicial-type-theory.inequality-directed-interval-type where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.embeddings +open import foundation.function-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.injective-maps +open import foundation.negated-equality +open import foundation.negation +open import foundation.propositions +open import foundation.sets +open import foundation.universe-levels + +open import order-theory.posets +open import order-theory.preorders +open import order-theory.total-orders + +open import simplicial-type-theory.directed-interval-type + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +``` + +
+ +## Idea + +The +{{#concept "directed relation" Disambiguation="on the directed interval type" Agda=_≤-𝟚_}} +on the +[directed interval type](simplicial-type-theory.directed-interval-type.md) `𝟚` +is a [total order](order-theory.total-orders.md) with `0₂` as the +[bottom element](order-theory.bottom-elements-posets.md), and `1₂` as the +[top element](order-theory.top-elements-posets.md). + +## Postulates + +### The directed relation on the directed interval + +```agda +postulate + _≤-𝟚_ : 𝟚 → 𝟚 → UU lzero + + is-prop-leq-𝟚 : {x y : 𝟚} → is-prop (x ≤-𝟚 y) + +infix 30 _≤-𝟚_ + +leq-𝟚 : 𝟚 → 𝟚 → UU lzero +leq-𝟚 = _≤-𝟚_ +``` + +### The directed relation is a total order + +```agda +postulate + refl-leq-𝟚 : {x : 𝟚} → x ≤-𝟚 x + + transitive-leq-𝟚 : {x y z : 𝟚} → y ≤-𝟚 z → x ≤-𝟚 y → x ≤-𝟚 z + + antisymmetric-leq-𝟚 : {x y : 𝟚} → x ≤-𝟚 y → y ≤-𝟚 x → x = y + + total-leq-𝟚 : {x y : 𝟚} → (x ≤-𝟚 y) * (y ≤-𝟚 x) +``` + +**Note.** We define totality using the +[join operation](synthetic-homotopy-theory.joins-of-types.md) on types as we may +commonly want to eliminate the directed relation into types, and not just +propositions. Moreover, the join operation is defined in terms of the +[standard pushout](synthetic-homotopy-theory.pushouts.md) which we can enable +rewrite rules for by importing the module +[`synthetic-homotopy-theory.rewriting-pushouts`](synthetic-homotopy-theory.rewriting-pushouts.md). + +### The source and target of the directed interval are bottom and top elements + +```agda +postulate + min-leq-𝟚 : {x : 𝟚} → 0₂ ≤-𝟚 x + + max-leq-𝟚 : {x : 𝟚} → x ≤-𝟚 1₂ +``` + +## Operations + +### The binary max function on the directed interval + +```agda +cocone-max-𝟚 : (t s : 𝟚) → cocone pr1 pr2 𝟚 +cocone-max-𝟚 t s = + ( (λ _ → s) , (λ _ → t) , (λ (p , q) → antisymmetric-leq-𝟚 q p)) + +max-𝟚 : 𝟚 → 𝟚 → 𝟚 +max-𝟚 t s = cogap-join 𝟚 (cocone-max-𝟚 t s) total-leq-𝟚 + +abstract + compute-left-max-𝟚 : {t s : 𝟚} → s ≤-𝟚 t → max-𝟚 t s = t + compute-left-max-𝟚 {t} {s} p = + ( ap + ( cogap-join 𝟚 (cocone-max-𝟚 t s)) + ( eq-is-prop (is-prop-join-is-prop is-prop-leq-𝟚 is-prop-leq-𝟚))) ∙ + ( compute-inr-cogap-join (cocone-max-𝟚 t s) p) + +abstract + compute-right-max-𝟚 : {t s : 𝟚} → t ≤-𝟚 s → max-𝟚 t s = s + compute-right-max-𝟚 {t} {s} p = + ( ap + ( cogap-join 𝟚 (cocone-max-𝟚 t s)) + ( eq-is-prop (is-prop-join-is-prop is-prop-leq-𝟚 is-prop-leq-𝟚))) ∙ + ( compute-inl-cogap-join (cocone-max-𝟚 t s) p) +``` + +### The binary minimum function on the directed interval + +```agda +cocone-min-𝟚 : (t s : 𝟚) → cocone pr1 pr2 𝟚 +cocone-min-𝟚 t s = + ( (λ _ → t) , (λ _ → s) , (λ (p , q) → antisymmetric-leq-𝟚 p q)) + +min-𝟚 : 𝟚 → 𝟚 → 𝟚 +min-𝟚 t s = cogap-join 𝟚 (cocone-min-𝟚 t s) total-leq-𝟚 + +abstract + compute-left-min-𝟚 : {t s : 𝟚} (p : t ≤-𝟚 s) → min-𝟚 t s = t + compute-left-min-𝟚 {t} {s} p = + ( ap + ( cogap-join 𝟚 (cocone-min-𝟚 t s)) + ( eq-is-prop (is-prop-join-is-prop is-prop-leq-𝟚 is-prop-leq-𝟚))) ∙ + ( compute-inl-cogap-join (cocone-min-𝟚 t s) p) + +abstract + compute-right-min-𝟚 : {t s : 𝟚} (p : s ≤-𝟚 t) → min-𝟚 t s = s + compute-right-min-𝟚 {t} {s} p = + ( ap + ( cogap-join 𝟚 (cocone-min-𝟚 t s)) + ( eq-is-prop (is-prop-join-is-prop is-prop-leq-𝟚 is-prop-leq-𝟚))) ∙ + ( compute-inr-cogap-join (cocone-min-𝟚 t s) p) +``` + +## Definitions + +### The inequality binary propositional relation on the directed interval + +```agda +leq-𝟚-Prop : (x y : 𝟚) → Prop lzero +leq-𝟚-Prop x y = (x ≤-𝟚 y , is-prop-leq-𝟚) + +{-# INLINE leq-𝟚-Prop #-} + +min-leq-eq-𝟚 : {x y : 𝟚} → x = 0₂ → x ≤-𝟚 y +min-leq-eq-𝟚 refl = min-leq-𝟚 + +max-leq-eq-𝟚 : {x y : 𝟚} → y = 1₂ → x ≤-𝟚 y +max-leq-eq-𝟚 refl = max-leq-𝟚 + +leq-eq-𝟚 : {x y : 𝟚} → x = y → x ≤-𝟚 y +leq-eq-𝟚 refl = refl-leq-𝟚 + +leq-inv-eq-𝟚 : {x y : 𝟚} → x = y → y ≤-𝟚 x +leq-inv-eq-𝟚 = leq-eq-𝟚 ∘ inv +``` + +### The poset structure on the directed interval + +```agda +𝟚-Preorder : Preorder lzero lzero +pr1 𝟚-Preorder = 𝟚 +pr1 (pr2 𝟚-Preorder) = leq-𝟚-Prop +pr1 (pr2 (pr2 𝟚-Preorder)) x = refl-leq-𝟚 {x} +pr2 (pr2 (pr2 𝟚-Preorder)) x y z = transitive-leq-𝟚 + +𝟚-Poset : Poset lzero lzero +pr1 𝟚-Poset = 𝟚-Preorder +pr2 𝟚-Poset x y = antisymmetric-leq-𝟚 +``` + +### The total order on the directed interval + +```agda +is-total-leq-𝟚 : is-total-Poset 𝟚-Poset +is-total-leq-𝟚 x y = + map-disjunction-join-Prop (leq-𝟚-Prop x y) (leq-𝟚-Prop y x) (total-leq-𝟚) + +𝟚-Total-Order : Total-Order lzero lzero +𝟚-Total-Order = (𝟚-Poset , is-total-leq-𝟚) +``` + +## Properties + +### The directed interval is a set + +```agda +is-set-𝟚 : is-set 𝟚 +is-set-𝟚 = is-set-type-Poset 𝟚-Poset + +𝟚-Set : Set lzero +𝟚-Set = (𝟚 , is-set-𝟚) + +Id-𝟚-Prop : 𝟚 → 𝟚 → Prop lzero +Id-𝟚-Prop = Id-Prop 𝟚-Set + +{-# INLINE Id-𝟚-Prop #-} +``` + +### The maximal element is not less than or equal to the bottom element of the directed interval + +```agda +not-leq-target-source-𝟚 : ¬ (1₂ ≤-𝟚 0₂) +not-leq-target-source-𝟚 leq-1-0 = + is-nontrivial-𝟚 (antisymmetric-leq-𝟚 min-leq-𝟚 leq-1-0) +``` + +### The directed relation `t ≤-𝟚 s` is equivalent to the relation `max-𝟚 t s = s` + +> This remains to be formalized. + +### The directed relation `t ≤-𝟚 s` is equivalent to the relation `min-𝟚 t s = t` + +> This remains to be formalized. + +### The canonical inclusion of the booleans into the directed interval is an embedding + +```agda +is-emb-map-directed-interval-bool : is-emb map-directed-interval-bool +is-emb-map-directed-interval-bool = + is-emb-is-injective is-set-𝟚 is-injective-map-directed-interval-bool +``` + +### The canonical inclusion of the booleans into the directed interval preserves and reflects its ordering + +> This remains to be formalized. diff --git a/src/simplicial-type-theory/inner-2-horn.lagda.md b/src/simplicial-type-theory/inner-2-horn.lagda.md new file mode 100644 index 0000000000..bc7301e3fc --- /dev/null +++ b/src/simplicial-type-theory/inner-2-horn.lagda.md @@ -0,0 +1,207 @@ +# The inner 2-horn + +```agda +module simplicial-type-theory.inner-2-horn where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.2-simplices +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-spines +open import simplicial-type-theory.standard-simplices + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Idea + +The {{#concept "inner 2-horn" Disambiguation="simplicial type"}} $Λ²₁$, also +called the _2-1-horn_, is the universal type generated from the data of two +directed edges such that the source of the first is the target of the second, +pictorially: + +```text + 0 ----> 1 ----> 2. +``` + +The inner 2-horn has multiple defining properties: + +1. The inner 2-horn is the subtype of the standard + [2-simplex](simplicial-type-theory.2-simplices.md) defined by + + ```text + Λ²₁ = {(x , y) ∈ 𝟚 × 𝟚 | (y = 0₂) ∨ (x = 1₂)} ⊆ Δ². + ``` + +2. The inner 2-horn is the pushout + + ```text + 0₂ + 1 -----> 𝟚 + | | + 1₂ | | + ∨ ⌜ ∨ + 𝟚 -----> Λ²₁. + ``` + +3. The inner 2-horn is the 2-[spine](simplicial-type-theory.spines.md). + +## Definitions + +### The inner 2-horn as a subtype of the lower simplicial triangle + +> TODO: replace with `subtype-spine 2` + +```agda +subtype-inner-two-horn : subtype lzero (𝟚 × 𝟚) +subtype-inner-two-horn (x , y) = + join-Prop (Id-Prop 𝟚-Set y 0₂) (Id-Prop 𝟚-Set x 1₂) + +inner-two-horn : UU +inner-two-horn = type-subtype subtype-inner-two-horn + +inl-inner-two-horn : 𝟚 → inner-two-horn +inl-inner-two-horn t = ((t , 0₂) , inl-join refl) + +inr-inner-two-horn : 𝟚 → inner-two-horn +inr-inner-two-horn s = ((1₂ , s) , inr-join refl) +``` + +```agda +Λ²₁ : UU +Λ²₁ = inner-two-horn +``` + +### The cogap map of the inner 2-horn as a subtype of the lower simplicial triangle + +```agda +module _ + {l : Level} {A : UU l} (f g : 𝟚 → A) (p : f 1₂ = g 0₂) + where + + cogap-inner-two-horn : inner-two-horn → A + cogap-inner-two-horn ((x , y) , H) = + cogap-join A + ( ( λ y=0 → f x) , + ( λ x=1 → g y) , + ( λ (y=0 , x=1) → ap f x=1 ∙ p ∙ ap g (inv y=0))) + ( H) +``` + +### The inner 2-horn as a pushout + +```text + 0₂ + 1 -----> 𝟚 + | | + 1₂ | | + ∨ ⌜ ∨ + 𝟚 -----> Λ²₁. +``` + +```agda +pushout-inner-two-horn : UU +pushout-inner-two-horn = pushout (point 1₂) (point 0₂) + +inl-pushout-inner-two-horn : 𝟚 → pushout-inner-two-horn +inl-pushout-inner-two-horn = inl-pushout (point 1₂) (point 0₂) + +inr-pushout-inner-two-horn : 𝟚 → pushout-inner-two-horn +inr-pushout-inner-two-horn = inr-pushout (point 1₂) (point 0₂) +``` + +## Properties + +### The inner 2-horn is a set + +```agda +is-set-inner-two-horn : is-set inner-two-horn +is-set-inner-two-horn = + is-set-type-subtype subtype-inner-two-horn (is-set-product is-set-𝟚 is-set-𝟚) +``` + +### The canonical map from the inner 2-horn as a pushout to the inner 2-horn as a subtype of the square + +```agda +map-inner-two-horn-pushout-inner-two-horn : + pushout-inner-two-horn → inner-two-horn +map-inner-two-horn-pushout-inner-two-horn = + cogap + ( point 1₂) + ( point 0₂) + ( inl-inner-two-horn , + inr-inner-two-horn , + λ _ → eq-type-subtype subtype-inner-two-horn refl) + +map-pushout-inner-two-horn-inner-two-horn : + inner-two-horn → pushout-inner-two-horn +map-pushout-inner-two-horn-inner-two-horn = + cogap-inner-two-horn + ( inl-pushout-inner-two-horn) + ( inr-pushout-inner-two-horn) + ( glue-pushout (point 1₂) (point 0₂) star) + +-- is-retraction-map-inner-two-horn-pushout-inner-two-horn : +-- is-retraction +-- ( map-pushout-inner-two-horn-inner-two-horn) +-- ( map-inner-two-horn-pushout-inner-two-horn) +-- is-retraction-map-inner-two-horn-pushout-inner-two-horn ((x , y) , H) = +-- dependent-cogap-join +-- { P = +-- λ H → +-- map-inner-two-horn-pushout-inner-two-horn +-- ( map-pushout-inner-two-horn-inner-two-horn ((x , y) , H)) = +-- ((x , y) , H)} +-- ( ( λ y=0 → {! compute-inl-cogap ? ? ? ? !}) , +-- ( λ x=1 → {! !}) , +-- {! !}) +-- ( H) +``` + +### The inclusion of the 2-horn into the 2-simplex + +```agda +leq-subtype-two-simplex-inner-two-horn : + subtype-inner-two-horn ⊆ subtype-lower-simplicial-triangle +leq-subtype-two-simplex-inner-two-horn (x , y) = + cogap-join + ( y ≤-𝟚 x) + ( min-leq-eq-𝟚 , max-leq-eq-𝟚 , λ _ → eq-is-prop is-prop-leq-𝟚) + +inclusion-two-simplex-inner-two-horn : Λ²₁ → Δ² +inclusion-two-simplex-inner-two-horn = + tot leq-subtype-two-simplex-inner-two-horn +``` diff --git a/src/simplicial-type-theory/lax-suspension.lagda.md b/src/simplicial-type-theory/lax-suspension.lagda.md new file mode 100644 index 0000000000..e89da333ec --- /dev/null +++ b/src/simplicial-type-theory/lax-suspension.lagda.md @@ -0,0 +1,74 @@ +# The simplicial suspension + +```agda +module simplicial-type-theory.lax-suspension where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Idea + +Given a type `X`, we define the {{#concept "lax suspension"}} of `X` as the +[pushout](synthetic-homotopy-theory.pushouts.md) + +```text + id × 0₂ + id × 1₂ + X + X ------------------> X × 𝟚 + | | + | | + ∨ ⌜ ∨ + ∂𝟚 ---------------------> Σ▵X. +``` + +or in other words the oriented pushout + +```text + X ------> 1 + | | + | ⇗ | + ∨ ⌜ ∨ + 1 -----> Σ▵X. +``` + +Intuitively, the lax suspension of `X` can be understood as a type consisting of +a point at the north and south pole, and a +[directed edge](simplicial-type-theory.directed-edges.md) `north →₂ south` for +every element `x : X`. It is constructed by taking the +[cartesian product](foundation-core.cartesian-product-types.md) `X × 𝟚`, and +"pinching" it together at the north and south pole. diff --git a/src/simplicial-type-theory/natural-transformations.lagda.md b/src/simplicial-type-theory/natural-transformations.lagda.md new file mode 100644 index 0000000000..bc93c70416 --- /dev/null +++ b/src/simplicial-type-theory/natural-transformations.lagda.md @@ -0,0 +1,191 @@ +# Natural transformations + +```agda +module simplicial-type-theory.natural-transformations where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given two dependent functions `f g : (x : A) → B x`, a +{{#concept "simplicial natural transformation" Disambiguation="simplicial type theory" Agda=simplicial-natural-transformation}} +`α` from `f` to `g` is a family of directed edges + +```text + α : (x : A) → (f x →₂ g x). +``` + +Naturality follows automatically from the fact that every section is natural in +the base. I.e., for every edge `x →₂ y` in `A`, we have a dependent edge +`α x →₂ α y` over it, giving us a commuting dependent square of simplicial +arrows + +```text + α x + f x ------> g x + | | + f p | α p | g p + ∨ ∨ + f y ------> g y. + α y +``` + +## Definitions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} + where + + _⇒▵_ : ((x : A) → B x) → ((x : A) → B x) → UU (l1 ⊔ l2) + f ⇒▵ g = (x : A) → f x →₂ g x + + infix 7 _⇒▵_ + + simplicial-natural-transformation : + ((x : A) → B x) → ((x : A) → B x) → UU (l1 ⊔ l2) + simplicial-natural-transformation = _⇒▵_ + +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} {f g : (x : A) → B x} (α : f ⇒▵ g) + where + + family-of-simplicial-arrows-simplicial-natural-transformation : + (x : A) → simplicial-arrow (B x) + family-of-simplicial-arrows-simplicial-natural-transformation x t = + simplicial-arrow-simplicial-hom (α x) t + + eq-source-simplicial-natural-transformation : + (x : A) → + family-of-simplicial-arrows-simplicial-natural-transformation x 0₂ = f x + eq-source-simplicial-natural-transformation x = + eq-source-simplicial-hom (α x) + + eq-target-simplicial-natural-transformation : + (x : A) → + family-of-simplicial-arrows-simplicial-natural-transformation x 1₂ = g x + eq-target-simplicial-natural-transformation x = + eq-target-simplicial-hom (α x) +``` + +## Properties + +### Families of simplicial arrows are simplicial arrows of dependent functions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → 𝟚 → UU l2} + where + + family-of-simplicial-arrows-simplicial-arrow-of-dependent-functions : + simplicial-arrow' (λ t → (x : A) → B x t) → + (x : A) → simplicial-arrow' (B x) + family-of-simplicial-arrows-simplicial-arrow-of-dependent-functions = swap-Π + + simplicial-arrow-of-dependent-functions-family-of-simplicial-arrows : + ((x : A) → simplicial-arrow' (B x)) → + simplicial-arrow' (λ t → (x : A) → B x t) + simplicial-arrow-of-dependent-functions-family-of-simplicial-arrows = swap-Π + + equiv-family-of-simplicial-arrows-simplicial-arrow-of-dependent-functions : + ( simplicial-arrow' (λ t → (x : A) → B x t)) ≃ + ( (x : A) → simplicial-arrow' (B x)) + equiv-family-of-simplicial-arrows-simplicial-arrow-of-dependent-functions = + equiv-swap-Π +``` + +### Extensionality for simplicial natural transformations + +A simplicial natural transformation between functions `f` and `g` is the same as +a directed edge from `f` to `g`. + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} {f g : (x : A) → B x} + where + + simplicial-natural-transformation-simplicial-edge-of-dependent-functions : + f →₂ g → f ⇒▵ g + simplicial-natural-transformation-simplicial-edge-of-dependent-functions + ( α , p , q) x = + ( ( λ t → α t x) , htpy-eq p x , htpy-eq q x) + + simplicial-edge-of-dependent-functions-simplicial-natural-transformation : + f ⇒▵ g → f →₂ g + simplicial-edge-of-dependent-functions-simplicial-natural-transformation α = + ( (λ t x → pr1 (α x) t) , eq-htpy (pr1 ∘ pr2 ∘ α) , eq-htpy (pr2 ∘ pr2 ∘ α)) + + is-section-simplicial-edge-of-dependent-functions-simplicial-natural-transformation : + is-section + ( simplicial-natural-transformation-simplicial-edge-of-dependent-functions) + ( simplicial-edge-of-dependent-functions-simplicial-natural-transformation) + is-section-simplicial-edge-of-dependent-functions-simplicial-natural-transformation + α = + eq-htpy + ( λ x → + eq-pair-eq-fiber + ( eq-pair + ( htpy-eq (is-section-eq-htpy (pr1 ∘ pr2 ∘ α)) x) + ( htpy-eq (is-section-eq-htpy (pr2 ∘ pr2 ∘ α)) x))) + + is-retraction-simplicial-edge-of-dependent-functions-simplicial-natural-transformation : + is-retraction + ( simplicial-natural-transformation-simplicial-edge-of-dependent-functions) + ( simplicial-edge-of-dependent-functions-simplicial-natural-transformation) + is-retraction-simplicial-edge-of-dependent-functions-simplicial-natural-transformation + ( α , p , q) = + eq-pair-eq-fiber + ( eq-pair (is-retraction-eq-htpy p) (is-retraction-eq-htpy q)) + + is-equiv-simplicial-natural-transformation-simplicial-edge-of-dependent-functions : + is-equiv + ( simplicial-natural-transformation-simplicial-edge-of-dependent-functions) + is-equiv-simplicial-natural-transformation-simplicial-edge-of-dependent-functions = + is-equiv-is-invertible + ( simplicial-edge-of-dependent-functions-simplicial-natural-transformation) + ( is-section-simplicial-edge-of-dependent-functions-simplicial-natural-transformation) + ( is-retraction-simplicial-edge-of-dependent-functions-simplicial-natural-transformation) + + extensionality-simplicial-natural-transformation : (f →₂ g) ≃ (f ⇒▵ g) + extensionality-simplicial-natural-transformation = + ( simplicial-natural-transformation-simplicial-edge-of-dependent-functions , + is-equiv-simplicial-natural-transformation-simplicial-edge-of-dependent-functions) +``` + +## The identity natural transformation + +```agda +id-simplicial-natural-transformation : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} (f : (x : A) → B x) → f ⇒▵ f +id-simplicial-natural-transformation f x = id-simplicial-hom (f x) +``` diff --git a/src/simplicial-type-theory/normed-maps-between-types.lagda.md b/src/simplicial-type-theory/normed-maps-between-types.lagda.md new file mode 100644 index 0000000000..38fe991f05 --- /dev/null +++ b/src/simplicial-type-theory/normed-maps-between-types.lagda.md @@ -0,0 +1,312 @@ +# Normed maps between types + +```agda +module simplicial-type-theory.normed-maps-between-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.commuting-triangles-of-identifications +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.functoriality-dependent-function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.postcomposition-functions +open import foundation.precomposition-functions +open import foundation.univalence +open import foundation.universal-property-equivalences +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition +open import foundation.whiskering-identifications-concatenation + +open import foundation-core.equivalences +open import foundation-core.homotopies + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-edges-dependent-pair-types +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.horizontal-composition-natural-transformations +open import simplicial-type-theory.natural-transformations +open import simplicial-type-theory.transposing-adjunctions-between-types +open import simplicial-type-theory.transposing-biadjunctions-between-types +``` + +
+ +## Idea + +Given a map of types `q : A → B`, that has a left and a right adjoint + +```text + q! ⊣ q ⊣ q∗ +``` + +then we say `q` is +{{#concept "normed" Disambiguation="map of simplicial types"}} if there's a +natural transformation + +```text + Nm_q : q! ⇒ q∗. +``` + +## Definitions + +### Normed maps between types + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-normed-map : (A → B) → UU (l1 ⊔ l2) + is-normed-map q = + Σ ( is-transposing-biadjoint q) + ( λ Q → + ( map-left-adjoint-is-transposing-biadjoint Q) ⇒▵ + ( map-right-adjoint-is-transposing-biadjoint Q)) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-normed-map q) + where + + is-transposing-biadjoint-is-normed-map : is-transposing-biadjoint q + is-transposing-biadjoint-is-normed-map = pr1 H + + is-transposing-right-adjoint-is-normed-map : + is-transposing-right-adjoint q + is-transposing-right-adjoint-is-normed-map = + is-transposing-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + is-transposing-left-adjoint-is-normed-map : + is-transposing-left-adjoint q + is-transposing-left-adjoint-is-normed-map = + is-transposing-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + map-left-adjoint-is-normed-map : B → A + map-left-adjoint-is-normed-map = + map-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + is-transposing-adjoint-map-left-adjoint-is-normed-map : + is-transposing-adjunction map-left-adjoint-is-normed-map q + is-transposing-adjoint-map-left-adjoint-is-normed-map = + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + unit-left-adjoint-is-normed-map : + id ⇒▵ q ∘ map-left-adjoint-is-normed-map + unit-left-adjoint-is-normed-map = + unit-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + counit-left-adjoint-is-normed-map : + map-left-adjoint-is-normed-map ∘ q ⇒▵ id + counit-left-adjoint-is-normed-map = + counit-left-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + map-right-adjoint-is-normed-map : B → A + map-right-adjoint-is-normed-map = + map-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + is-transposing-adjoint-map-right-adjoint-is-normed-map : + is-transposing-adjunction q map-right-adjoint-is-normed-map + is-transposing-adjoint-map-right-adjoint-is-normed-map = + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + unit-right-adjoint-is-normed-map : + id ⇒▵ map-right-adjoint-is-normed-map ∘ q + unit-right-adjoint-is-normed-map = + unit-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + counit-right-adjoint-is-normed-map : + q ∘ map-right-adjoint-is-normed-map ⇒▵ id + counit-right-adjoint-is-normed-map = + counit-right-adjoint-is-transposing-biadjoint + is-transposing-biadjoint-is-normed-map + + norm-map-is-normed-map : + map-left-adjoint-is-normed-map ⇒▵ map-right-adjoint-is-normed-map + norm-map-is-normed-map = pr2 H +``` + +### The wrong way counit associated to a normed map + +Given a transposing biadjoint `q! ⊣ q* ⊣ q∗`, then having a norm map +`Nm_q : q! ⇒ q∗` is equivalent to having a _wrong way counit_ map + +```text + ν : q*q! ⇒ id. +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} + (H : is-transposing-biadjoint q) + where + + equiv-wrong-way-counit-norm-map-is-transposing-biadjoint : + ( q ∘ map-left-adjoint-is-transposing-biadjoint H ⇒▵ id) ≃ + ( map-left-adjoint-is-transposing-biadjoint H ⇒▵ + map-right-adjoint-is-transposing-biadjoint H) + equiv-wrong-way-counit-norm-map-is-transposing-biadjoint = + equiv-Π-equiv-family + ( λ x → + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint H + ( map-left-adjoint-is-transposing-biadjoint H x) + ( x)) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-normed-map q) + where + + wrong-way-counit-is-normed-map : q ∘ map-left-adjoint-is-normed-map H ⇒▵ id + wrong-way-counit-is-normed-map x = + map-inv-equiv + ( is-transposing-adjoint-map-right-adjoint-is-normed-map H + ( map-left-adjoint-is-normed-map H x) + ( x)) + ( norm-map-is-normed-map H x) + + wrong-way-counit-is-normed-map' : + (x y : A) → q x →₂ q y → {! q (pr1 (pr1 (pr1 H)) (q x)) !} →₂ {! q y !} + wrong-way-counit-is-normed-map' x y f = + map-inv-equiv + ( is-transposing-adjoint-map-right-adjoint-is-normed-map H + ( map-left-adjoint-is-normed-map H (q x)) + ( q y)) + ( action-simplicial-hom-simplicial-natural-transformation + (norm-map-is-normed-map H) f) +``` + +### The wrong way unit associated to a normed map + +Given a transposing biadjoint `q! ⊣ q* ⊣ q∗`, then having a norm map +`Nm_q : q! ⇒ q∗` is equivalent to having a _wrong way unit_ map + +```text + μ : id ⇒ q*q∗. +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} + (H : is-transposing-biadjoint q) + where + + equiv-wrong-way-unit-norm-map-is-transposing-biadjoint : + ( map-left-adjoint-is-transposing-biadjoint H ⇒▵ + map-right-adjoint-is-transposing-biadjoint H) ≃ + ( id ⇒▵ q ∘ map-right-adjoint-is-transposing-biadjoint H) + equiv-wrong-way-unigithub agt-norm-map-is-transposing-biadjoint = + equiv-Π-equiv-family + ( λ x → + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint H + ( x) + ( map-right-adjoint-is-transposing-biadjoint H x)) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-normed-map q) + where + + wrong-way-unit-is-normed-map : id ⇒▵ q ∘ map-right-adjoint-is-normed-map H + wrong-way-unit-is-normed-map x = + map-equiv + ( is-transposing-adjoint-map-left-adjoint-is-normed-map H + ( x) + ( map-right-adjoint-is-normed-map H x)) + ( norm-map-is-normed-map H x) +``` + +### The norm map's action on edges in the image of `q` + +Given a directed edge `f : q* x → q* y` there is a square + +```text + Nm-q(f) + q!q* x ----> q∗q* y + ∧ | + | | + | ∨ + x ----------> y +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-normed-map q) + where + + action-hom-image-norm-map-is-normed-map : + {x y : A} → + q x →₂ q y → + map-left-adjoint-is-normed-map H (q x) →₂ map-right-adjoint-is-normed-map H (q y) + action-hom-image-norm-map-is-normed-map f = + action-simplicial-hom-simplicial-natural-transformation (norm-map-is-normed-map H) f +``` + +### The cardinality? of a normed map + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q : A → B} (H : is-normed-map q) + where + + cardinality?-is-normed-map' : + (x : B) → + ((q ∘ map-right-adjoint-is-normed-map H) x →₂ x) × + (x →₂ (q ∘ map-right-adjoint-is-normed-map H) x) + cardinality?-is-normed-map' x = + ( counit-right-adjoint-is-normed-map H x , wrong-way-unit-is-normed-map H x) +``` + +## Properties + +### Composition of normed maps + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {p : A → B} {q : B → C} + where + + is-normed-map-comp : + is-normed-map q → is-normed-map p → is-normed-map (q ∘ p) + is-normed-map-comp (Q , Nm-q) (P , Nm-p) = + ( is-transposing-biadjoint-comp Q P) , + ( horizontal-comp-simplicial-natural-transformation Nm-p Nm-q) +``` + +## Examples + +### The identity normed map + +```agda +is-normed-map-id : {l : Level} {A : UU l} → is-normed-map (id {A = A}) +pr1 is-normed-map-id = is-transposing-biadjoint-id +pr2 is-normed-map-id = id-simplicial-hom +``` + +### Equivalences are normed maps + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} (H : is-equiv f) + where + + is-normed-map-is-equiv : is-normed-map f + pr1 is-normed-map-is-equiv = is-transposing-biadjoint-is-equiv H + pr2 is-normed-map-is-equiv = id-simplicial-hom ∘ map-inv-is-equiv H +``` + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/representing-biinvertible-simplicial-arrow.lagda.md b/src/simplicial-type-theory/representing-biinvertible-simplicial-arrow.lagda.md new file mode 100644 index 0000000000..e7d93fbbc4 --- /dev/null +++ b/src/simplicial-type-theory/representing-biinvertible-simplicial-arrow.lagda.md @@ -0,0 +1,168 @@ +# The representing biinvertible simplicial arrow + +```agda +module simplicial-type-theory.representing-biinvertible-simplicial-arrow where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.contractible-types +open import foundation.coproduct-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition + +open import reflection.erasing-equality + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-cubes +open import simplicial-type-theory.standard-simplices + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.dependent-cocones-under-spans +open import synthetic-homotopy-theory.dependent-universal-property-pushouts +open import synthetic-homotopy-theory.induction-principle-pushouts +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +open import synthetic-homotopy-theory.recursion-principle-pushouts +open import synthetic-homotopy-theory.universal-property-pushouts + +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +```text + s ===== s + ∧ / ∧ + l \ L / R \ r + \ ∨ \ + t ===== t +``` + +## Postulates + +```agda +postulate + representing-biinvertible-simplicial-arrow : UU + + source-representing-biinvertible-simplicial-arrow : + representing-biinvertible-simplicial-arrow + + target-representing-biinvertible-simplicial-arrow : + representing-biinvertible-simplicial-arrow + + arrow-representing-biinvertible-simplicial-arrow : + 𝟚 → representing-biinvertible-simplicial-arrow + + arrow-left-inv-representing-biinvertible-simplicial-arrow : + 𝟚 → representing-biinvertible-simplicial-arrow + + arrow-right-inv-representing-biinvertible-simplicial-arrow : + 𝟚 → representing-biinvertible-simplicial-arrow + + left-2-cell-representing-biinvertible-simplicial-arrow : + Δ 2 → representing-biinvertible-simplicial-arrow + + compute-0-left-2-cell-representing-biinvertible-simplicial-arrow : + ( left-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((t , 0₂) , min-leq-𝟚)) = + arrow-representing-biinvertible-simplicial-arrow + + compute-1-left-2-cell-representing-biinvertible-simplicial-arrow : + ( left-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((t , t) , refl-leq-𝟚)) = + id-simplicial-arrow target-representing-biinvertible-simplicial-arrow + + compute-2-left-2-cell-representing-biinvertible-simplicial-arrow : + ( left-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((1₂ , t) , max-leq-𝟚)) = + arrow-left-inv-representing-biinvertible-simplicial-arrow + + right-2-cell-representing-biinvertible-simplicial-arrow : + Δ 2 → representing-biinvertible-simplicial-arrow + + compute-1-right-2-cell-representing-biinvertible-simplicial-arrow : + ( right-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((t , t) , refl-leq-𝟚)) = + id-simplicial-arrow source-representing-biinvertible-simplicial-arrow + + compute-0-right-2-cell-representing-biinvertible-simplicial-arrow : + ( right-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((t , 0₂) , min-leq-𝟚)) = + arrow-representing-biinvertible-simplicial-arrow + + compute-2-right-2-cell-representing-biinvertible-simplicial-arrow : + ( right-2-cell-representing-biinvertible-simplicial-arrow) ∘ + ( λ t → ((1₂ , t) , max-leq-𝟚)) = + arrow-representing-biinvertible-simplicial-arrow +``` + +## Definitions + +```agda +source-arrow-representing-biinvertible-simplicial-arrow : + arrow-representing-biinvertible-simplicial-arrow 0₂ = + {! id-simplicial-arrow\ntarget-representing-biinvertible-simplicial-arrow 0₂ !} + -- source-representing-biinvertible-simplicial-arrow +source-arrow-representing-biinvertible-simplicial-arrow = + htpy-eq + ( inv compute-0-left-2-cell-representing-biinvertible-simplicial-arrow) + 0₂ ∙ + {! !} ∙ + htpy-eq compute-1-left-2-cell-representing-biinvertible-simplicial-arrow 0₂ + +target-arrow-representing-biinvertible-simplicial-arrow : + arrow-representing-biinvertible-simplicial-arrow 1₂ = + target-representing-biinvertible-simplicial-arrow +target-arrow-representing-biinvertible-simplicial-arrow = {! !} + +source-arrow-left-inv-representing-biinvertible-simplicial-arrow : + arrow-left-inv-representing-biinvertible-simplicial-arrow 0₂ = + target-representing-biinvertible-simplicial-arrow +source-arrow-left-inv-representing-biinvertible-simplicial-arrow = {! !} + +target-arrow-left-inv-representing-biinvertible-simplicial-arrow : + arrow-left-inv-representing-biinvertible-simplicial-arrow 1₂ = + source-representing-biinvertible-simplicial-arrow +target-arrow-left-inv-representing-biinvertible-simplicial-arrow = {! !} + +source-arrow-right-inv-representing-biinvertible-simplicial-arrow : + arrow-right-inv-representing-biinvertible-simplicial-arrow 0₂ = + target-representing-biinvertible-simplicial-arrow +source-arrow-right-inv-representing-biinvertible-simplicial-arrow = {! !} + +target-arrow-right-inv-representing-biinvertible-simplicial-arrow : + arrow-right-inv-representing-biinvertible-simplicial-arrow 1₂ = + source-representing-biinvertible-simplicial-arrow +target-arrow-right-inv-representing-biinvertible-simplicial-arrow = {! !} +``` diff --git a/src/simplicial-type-theory/rewriting-directed-circle.lagda.md b/src/simplicial-type-theory/rewriting-directed-circle.lagda.md new file mode 100644 index 0000000000..02b74d1dc1 --- /dev/null +++ b/src/simplicial-type-theory/rewriting-directed-circle.lagda.md @@ -0,0 +1,66 @@ +# Rewriting rules for the directed circle + +```agda +{-# OPTIONS --rewriting #-} + +module simplicial-type-theory.rewriting-directed-circle where +``` + +
Imports + +```agda +open import foundation.homotopies +open import foundation.identity-types +open import foundation.universe-levels + +open import reflection.rewriting + +open import simplicial-type-theory.directed-circle +``` + +
+ +## Idea + +This module endows the +[directed circle](simplicial-type-theory.directed-circle.md) with strict +computation rules using Agda's [rewriting](reflection.rewriting.md) +functionality. This gives the strict equality + +```text + ind-directed-circle α (arrow-directed-circle t) ≐ + arrow-free-dependent-directed-loop free-loop-directed-circle α. +``` + +In addition, the pre-existing witness of this equality +`compute-arrow-ind-directed-circle` reduces to `refl`. This is achieved using +Agda's [equality erasure](reflection.erasing-equality.md) functionality. + +To enable this computation rule in your own formalizations, set the +`--rewriting` option and import this module. Keep in mind, however, that we +offer no way to selectively disable these rules, so if your module depends on +any other module that depends on this one, you will necessarily also have +rewriting for the directed circle enabled. + +By default, we abstain from using rewrite rules in agda-unimath. However, we +recognize their usefulness in, for instance, exploratory formalizations. Since +formalizations with and without rewrite rules can coexist, there is no harm in +providing these tools for those that see a need to use them. We do, however, +emphasize that formalizations without the use of rewrite rules are preferred +over those that do use them in the library, as the former can also be applied in +other formalizations that do not enable rewrite rules. + +## Rewrite rules + +```agda +{-# REWRITE compute-arrow-ind-directed-circle #-} +``` + +## See also + +- For some discussion on strict computation rules for higher inductive types, + see the introduction to Section 6.2 of {{#cite UF13}}. + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/simplicial-arrows.lagda.md b/src/simplicial-type-theory/simplicial-arrows.lagda.md new file mode 100644 index 0000000000..16e858a447 --- /dev/null +++ b/src/simplicial-type-theory/simplicial-arrows.lagda.md @@ -0,0 +1,89 @@ +# Simplicial arrows + +```agda +module simplicial-type-theory.simplicial-arrows where +``` + +
Imports + +```agda +open import foundation.function-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.negation +open import foundation.universe-levels + +open import simplicial-type-theory.directed-interval-type +``` + +
+ +## Idea + +A +{{#concept "simplicial arrow" Disambiguation="simplicial type theory" Agda=simplicial-arrow}} +in a type `A` is a map from the +[directed interval](simplicial-type-theory.directed-interval-type.md) to the +type, `𝟚 → A`. Given a simplicial arrow `α` in `A`, we call `α 0₂` the _source_, +and `α 1₂` the _target_ of the arrow. See +[directed edges](simplicial-type-theory.directed-edges.md) for simplicial arrows +with a specified source and target. + +## Definitions + +### Simplicial arrows in types dependent over the directed interval + +```agda +simplicial-arrow' : {l : Level} → (𝟚 → UU l) → UU l +simplicial-arrow' A = (t : 𝟚) → A t +``` + +### Simplicial arrows + +```agda +simplicial-arrow : {l : Level} → UU l → UU l +simplicial-arrow A = simplicial-arrow' (λ _ → A) +``` + +### The identity/constant simplicial arrows + +```agda +id-simplicial-arrow : {l : Level} {A : UU l} (x : A) → simplicial-arrow A +id-simplicial-arrow x _ = x +``` + +### The representing arrow of the directed interval + +```agda +representing-arrow-𝟚 : simplicial-arrow 𝟚 +representing-arrow-𝟚 = id +``` + +### Simplicial arrows arising from equalities + +```agda +module _ + {l : Level} {A : UU l} {x y : A} + where + + simplicial-arrow-eq : x = y → simplicial-arrow A + simplicial-arrow-eq refl = id-simplicial-arrow x + + compute-source-simplicial-arrow-eq : + (p : x = y) → simplicial-arrow-eq p 0₂ = x + compute-source-simplicial-arrow-eq refl = refl + + compute-target-simplicial-arrow-eq : + (p : x = y) → simplicial-arrow-eq p 1₂ = y + compute-target-simplicial-arrow-eq refl = refl +``` + +## Properties + +### The representing arrow of the directed interval is not constant + +```agda +is-not-constant-representing-arrow-𝟚 : + (t : 𝟚) → ¬ (representing-arrow-𝟚 ~ id-simplicial-arrow t) +is-not-constant-representing-arrow-𝟚 _ H = is-nontrivial-𝟚 (H 0₂ ∙ inv (H 1₂)) +``` diff --git a/src/simplicial-type-theory/simplicial-cones.lagda.md b/src/simplicial-type-theory/simplicial-cones.lagda.md new file mode 100644 index 0000000000..8718c87ad0 --- /dev/null +++ b/src/simplicial-type-theory/simplicial-cones.lagda.md @@ -0,0 +1,205 @@ +# The simplicial cone + +```agda +module simplicial-type-theory.simplicial-cones where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.dependent-cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Idea + +Given a type `X`, we define the {{#concept "simplicial cone type"}} as the +[pushout](synthetic-homotopy-theory.pushouts.md) + +```text + X --------> 1 + | | + (id , 0₂) | | + ∨ ⌜ ∨ + X × 𝟚 ---> cone₂ X +``` + +Intuitively, the simplicial cone of `X` can be understood as `X` with a point +`*` attached such that there is a +[directed edge](simplicial-type-theory.directed-edges.md) `* →₂ x` for every +`x : X`. + +## Definitions + +### The standard simplicial cone on a type + +```agda +simplicial-cone : {l : Level} → UU l → UU l +simplicial-cone X = + pushout (λ (x : X) → (x , 0₂)) (terminal-map X) + +module _ + {l : Level} {X : UU l} + where + + in-simplicial-cone' : X → 𝟚 → simplicial-cone X + in-simplicial-cone' x t = + inl-pushout (λ (x : X) → (x , 0₂)) (terminal-map X) (x , t) + + in-simplicial-cone : X → simplicial-cone X + in-simplicial-cone x = in-simplicial-cone' x 1₂ + + point-simplicial-cone : simplicial-cone X + point-simplicial-cone = + inr-pushout (λ (x : X) → (x , 0₂)) (terminal-map X) star + + glue-simplicial-cone : + (x : X) → + in-simplicial-cone' x 0₂ = point-simplicial-cone + glue-simplicial-cone = + glue-pushout (λ (x : X) → (x , 0₂)) (terminal-map X) + + hom-simplicial-cone : + (x : X) → point-simplicial-cone →₂ in-simplicial-cone x + hom-simplicial-cone x = + ( in-simplicial-cone' x , glue-simplicial-cone x , refl) + + cocone-simplicial-cone : + cocone (λ (x : X) → (x , 0₂)) (terminal-map X) (simplicial-cone X) + cocone-simplicial-cone = + cocone-pushout (λ (x : X) → (x , 0₂)) (terminal-map X) +``` + +### The dependent cogap map for the simplicial cone type + +```agda +module _ + {l : Level} (X : UU l) + where + + dependent-cogap-simplicial-cone' : + {l' : Level} {P : simplicial-cone X → UU l'} + (c : + dependent-cocone + ( λ x → x , 0₂) + ( terminal-map X) + ( cocone-simplicial-cone) + ( P)) + (x : simplicial-cone X) → P x + dependent-cogap-simplicial-cone' = + dependent-cogap (λ x → (x , 0₂)) (terminal-map X) + + dependent-cogap-simplicial-cone : + { l' : Level} {P : simplicial-cone X → UU l'} + ( f : P point-simplicial-cone) → + ( g : (x : X) (t : 𝟚) → P (in-simplicial-cone' x t)) → + ( p : + (x : X) → + dependent-identification P (glue-simplicial-cone x) (g x 0₂) f) + ( x : simplicial-cone X) → P x + dependent-cogap-simplicial-cone f g p = + dependent-cogap-simplicial-cone' + ( (λ (x , t) → g x t) , point f , p) +``` + +### The cogap map for the simplicial cone type + +```agda +module _ + {l1 l2 : Level} {X : UU l1} {Y : UU l2} + where + + cogap-simplicial-cone' : + cocone (λ x → (x , 0₂)) (terminal-map X) Y → + simplicial-cone X → Y + cogap-simplicial-cone' = + cogap (λ (x : X) → (x , 0₂)) (terminal-map X) + + cogap-simplicial-cone : + (g : X → 𝟚 → Y) (f : Y) (p : (x : X) → g x 0₂ = f) → + simplicial-cone X → Y + cogap-simplicial-cone g f p = + cogap-simplicial-cone' ((λ (x , t) → g x t) , point f , p) + + compute-point-cogap-simplicial-cone : + (g : X → 𝟚 → Y) (f : Y) (p : (x : X) → g x 0₂ = f) → + cogap-simplicial-cone g f p point-simplicial-cone = f + compute-point-cogap-simplicial-cone g f p = + compute-inr-cogap + ( λ x → (x , 0₂)) + ( terminal-map X) + ( (λ (x , t) → g x t) , point f , p) + ( star) + + compute-in-cogap-simplicial-cone' : + (g : X → 𝟚 → Y) (f : Y) (p : (x : X) → g x 0₂ = f) (x : X) (t : 𝟚) → + cogap-simplicial-cone g f p (in-simplicial-cone' x t) = + g x t + compute-in-cogap-simplicial-cone' g f p x t = + compute-inl-cogap + ( λ x → x , 0₂) + ( terminal-map X) + ( (λ (x , t) → g x t) , point f , p) + ( x , t) + + compute-in-cogap-simplicial-cone : + (g : X → 𝟚 → Y) (f : Y) (p : (x : X) → g x 0₂ = f) (x : X) → + cogap-simplicial-cone g f p (in-simplicial-cone x) = + g x 1₂ + compute-in-cogap-simplicial-cone g f p x = + compute-in-cogap-simplicial-cone' g f p x 1₂ +``` + +## Properties + +### The directed interval is equivalent to the simplicial cone over the unit type + +**Proof.** We have the pushout diagram + +```text + 1 ------> 1 + | | + | | + ∨ ⌜ ∨ + 𝟚 -----> C₂1, +``` + +and since the top horizontal map is an equivalence, so is its pushout. + +This remains to be formalized. diff --git a/src/simplicial-type-theory/simplicial-cubes.lagda.md b/src/simplicial-type-theory/simplicial-cubes.lagda.md new file mode 100644 index 0000000000..43f7a895be --- /dev/null +++ b/src/simplicial-type-theory/simplicial-cubes.lagda.md @@ -0,0 +1,155 @@ +# Simplicial cubes + +```agda +module simplicial-type-theory.simplicial-cubes where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.empty-types +open import foundation.function-types +open import foundation.identity-types +open import foundation.negated-equality +open import foundation.negation +open import foundation.propositions +open import foundation.sets +open import foundation.subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type + +open import synthetic-homotopy-theory.joins-of-types +``` + +
+ +## Idea + +Given a [natural number](elementary-number-theory.natural-numbers.md) 𝑛, the +{{#concept "standard simplicial 𝑛-cube" Disambiguation="simplicial type theory" Agda=simplicial-cube}} +consists of 𝑛-fold pairs of elements of the +[directed interval](simplicial-type-theory.directed-interval-type.md). The +standard simplicial 0-cube is defined to be the +[unit type](foundation.unit-type.md). + +## Definitions + +### The standard simplicial cubes + +```agda +simplicial-cube : ℕ → UU lzero +simplicial-cube 0 = unit +simplicial-cube 1 = 𝟚 +simplicial-cube (succ-ℕ (succ-ℕ n)) = 𝟚 × simplicial-cube (succ-ℕ n) +``` + +### The standard left-associated simplicial cubes + +```agda +left-associated-simplicial-cube : ℕ → UU lzero +left-associated-simplicial-cube 0 = unit +left-associated-simplicial-cube 1 = 𝟚 +left-associated-simplicial-cube (succ-ℕ (succ-ℕ n)) = + left-associated-simplicial-cube (succ-ℕ n) × 𝟚 +``` + +### The standard simplicial power cubes + +```agda +pow-simplicial-cube : ℕ → UU lzero +pow-simplicial-cube 0 = unit +pow-simplicial-cube 1 = 𝟚 +pow-simplicial-cube (succ-ℕ (succ-ℕ n)) = + pow-simplicial-cube (succ-ℕ n) × pow-simplicial-cube (succ-ℕ n) +``` + +### The boundary of the standard simplicial cube + +```agda +subtype-boundary-simplicial-cube : (n : ℕ) → subtype lzero (simplicial-cube n) +subtype-boundary-simplicial-cube 0 _ = + empty-Prop +subtype-boundary-simplicial-cube 1 x = + join-Prop (Id-𝟚-Prop x 0₂) (Id-𝟚-Prop x 1₂) +subtype-boundary-simplicial-cube (succ-ℕ (succ-ℕ n)) (x , u) = + join-Prop + ( subtype-boundary-simplicial-cube 1 x) + ( subtype-boundary-simplicial-cube (succ-ℕ n) u) + +boundary-simplicial-cube : ℕ → UU lzero +boundary-simplicial-cube = type-subtype ∘ subtype-boundary-simplicial-cube +``` + +### The predicate of being an initial element in the simplicial 𝑛-cube + +```agda +is-initial-element-simplicial-cube : (n : ℕ) → simplicial-cube n → UU lzero +is-initial-element-simplicial-cube 0 _ = unit +is-initial-element-simplicial-cube 1 x = (x = 0₂) +is-initial-element-simplicial-cube (succ-ℕ (succ-ℕ n)) (x , y) = + ( is-initial-element-simplicial-cube 1 x) × + ( is-initial-element-simplicial-cube (succ-ℕ n) y) + +is-prop-is-initial-element-simplicial-cube : + (n : ℕ) (x : simplicial-cube n) → + is-prop (is-initial-element-simplicial-cube n x) +is-prop-is-initial-element-simplicial-cube 0 _ = is-prop-unit +is-prop-is-initial-element-simplicial-cube 1 x = is-set-𝟚 x 0₂ +is-prop-is-initial-element-simplicial-cube (succ-ℕ (succ-ℕ n)) (x , y) = + is-prop-product + ( is-prop-is-initial-element-simplicial-cube 1 x) + ( is-prop-is-initial-element-simplicial-cube (succ-ℕ n) y) + +is-initial-element-simplicial-cube-Prop : + (n : ℕ) → simplicial-cube n → Prop lzero +is-initial-element-simplicial-cube-Prop n x = + ( is-initial-element-simplicial-cube n x , + is-prop-is-initial-element-simplicial-cube n x) +``` + +### The predicate of being a terminal element in the simplicial 𝑛-cube + +```agda +is-terminal-element-simplicial-cube : (n : ℕ) → simplicial-cube n → UU lzero +is-terminal-element-simplicial-cube 0 _ = unit +is-terminal-element-simplicial-cube 1 x = (x = 1₂) +is-terminal-element-simplicial-cube (succ-ℕ (succ-ℕ n)) (x , y) = + ( is-terminal-element-simplicial-cube 1 x) × + ( is-terminal-element-simplicial-cube (succ-ℕ n) y) + +is-prop-is-terminal-element-simplicial-cube : + (n : ℕ) (x : simplicial-cube n) → + is-prop (is-terminal-element-simplicial-cube n x) +is-prop-is-terminal-element-simplicial-cube 0 _ = is-prop-unit +is-prop-is-terminal-element-simplicial-cube 1 x = is-set-𝟚 x 1₂ +is-prop-is-terminal-element-simplicial-cube (succ-ℕ (succ-ℕ n)) (x , y) = + is-prop-product + ( is-prop-is-terminal-element-simplicial-cube 1 x) + ( is-prop-is-terminal-element-simplicial-cube (succ-ℕ n) y) + +is-terminal-element-simplicial-cube-Prop : + (n : ℕ) → simplicial-cube n → Prop lzero +is-terminal-element-simplicial-cube-Prop n x = + ( is-terminal-element-simplicial-cube n x , + is-prop-is-terminal-element-simplicial-cube n x) +``` + +## Properties + +### The simplicial 𝑛-cube is a set + +```agda +is-set-simplicial-cube : (n : ℕ) → is-set (simplicial-cube n) +is-set-simplicial-cube zero-ℕ = is-set-unit +is-set-simplicial-cube (succ-ℕ zero-ℕ) = is-set-𝟚 +is-set-simplicial-cube (succ-ℕ (succ-ℕ n)) = + is-set-product is-set-𝟚 (is-set-simplicial-cube (succ-ℕ n)) +``` diff --git a/src/simplicial-type-theory/simplicial-mapping-cylinders.lagda.md b/src/simplicial-type-theory/simplicial-mapping-cylinders.lagda.md new file mode 100644 index 0000000000..a1dcc7f8ae --- /dev/null +++ b/src/simplicial-type-theory/simplicial-mapping-cylinders.lagda.md @@ -0,0 +1,201 @@ +# Simplicial mapping cylinders + +```agda +module simplicial-type-theory.simplicial-mapping-cylinders where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.dependent-cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Idea + +Given a map `f : X → Y`, we define the +{{#concept "simplicial mapping cylinder"}} of `f` as the +[pushout](synthetic-homotopy-theory.pushouts.md) + +```text + f + X --------> Y + | | + (id , 1₂) | | + ∨ ⌜ ∨ + X × 𝟚 ----> cyl₂ f +``` + +Intuitively, the simplicial mapping cylinder of `f` can be understood as `X` +simplicially glued to `Y` along `f`. I.e., for every `x : X` there is a +[directed edge](simplicial-type-theory.directed-edges.md) `x →₂ f x`. + +## Definitions + +### The standard simplicial mapping cylinder of a function + +```agda +simplicial-mapping-cylinder : + {l1 l2 : Level} {X : UU l1} {Y : UU l2} → (X → Y) → UU (l1 ⊔ l2) +simplicial-mapping-cylinder {X = X} {Y} f = + pushout (λ (x : X) → (x , 1₂)) f + +module _ + {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) + where + + in-domain-interval-simplicial-mapping-cylinder : + X → 𝟚 → simplicial-mapping-cylinder f + in-domain-interval-simplicial-mapping-cylinder x t = + inl-pushout (λ (x : X) → (x , 1₂)) f (x , t) + + in-domain-simplicial-mapping-cylinder : X → simplicial-mapping-cylinder f + in-domain-simplicial-mapping-cylinder x = + in-domain-interval-simplicial-mapping-cylinder x 0₂ + + in-codomain-simplicial-mapping-cylinder : Y → simplicial-mapping-cylinder f + in-codomain-simplicial-mapping-cylinder = + inr-pushout (λ (x : X) → (x , 1₂)) f + + glue-simplicial-mapping-cylinder : + (x : X) → + in-domain-interval-simplicial-mapping-cylinder x 1₂ = + in-codomain-simplicial-mapping-cylinder (f x) + glue-simplicial-mapping-cylinder = + glue-pushout (λ (x : X) → (x , 1₂)) f + + hom-simplicial-mapping-cylinder : + (x : X) → + in-domain-simplicial-mapping-cylinder x →₂ + in-codomain-simplicial-mapping-cylinder (f x) + hom-simplicial-mapping-cylinder x = + ( in-domain-interval-simplicial-mapping-cylinder x , + refl , + glue-simplicial-mapping-cylinder x) + + cocone-simplicial-mapping-cylinder : + cocone (λ (x : X) → (x , 1₂)) f (simplicial-mapping-cylinder f) + cocone-simplicial-mapping-cylinder = + cocone-pushout (λ (x : X) → (x , 1₂)) f +``` + +### The dependent cogap map for the simplicial mapping cylinder + +```agda +module _ + {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) + where + + dependent-cogap-simplicial-mapping-cylinder' : + {l : Level} {P : simplicial-mapping-cylinder f → UU l} + (c : + dependent-cocone (λ x → (x , 1₂)) (f) + ( cocone-simplicial-mapping-cylinder f) + ( P)) + (x : simplicial-mapping-cylinder f) → P x + dependent-cogap-simplicial-mapping-cylinder' = + dependent-cogap (λ x → (x , 1₂)) f + + dependent-cogap-simplicial-mapping-cylinder : + { l : Level} {P : simplicial-mapping-cylinder f → UU l} + ( g : + (x : X) (t : 𝟚) → + P (in-domain-interval-simplicial-mapping-cylinder f x t)) → + ( h : (y : Y) → P (in-codomain-simplicial-mapping-cylinder f y)) → + ( p : + (x : X) → + dependent-identification P + ( glue-simplicial-mapping-cylinder f x) + ( g x 1₂) + ( h (f x))) + ( x : simplicial-mapping-cylinder f) → P x + dependent-cogap-simplicial-mapping-cylinder g h p = + dependent-cogap-simplicial-mapping-cylinder' + ( (λ (x , t) → g x t) , h , p) +``` + +### The cogap map for the simplicial mapping cylinder + +```agda +module _ + {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} {S : UU l3} + where + + cogap-simplicial-mapping-cylinder' : + cocone (λ x → (x , 1₂)) f S → + simplicial-mapping-cylinder f → S + cogap-simplicial-mapping-cylinder' = + cogap (λ (x : X) → (x , 1₂)) f + + cogap-simplicial-mapping-cylinder : + (g : X → 𝟚 → S) (h : Y → S) (p : (x : X) → g x 1₂ = h (f x)) → + simplicial-mapping-cylinder f → S + cogap-simplicial-mapping-cylinder g h p = + cogap-simplicial-mapping-cylinder' ((λ (x , t) → g x t) , h , p) + + compute-in-codomain-cogap-simplicial-mapping-cylinder : + (g : X → 𝟚 → S) (h : Y → S) (p : (x : X) → g x 1₂ = h (f x)) → + (y : Y) → + cogap-simplicial-mapping-cylinder g h p + ( in-codomain-simplicial-mapping-cylinder f y) = + ( h y) + compute-in-codomain-cogap-simplicial-mapping-cylinder g h p = + compute-inr-cogap (λ x → (x , 1₂)) f ((λ (x , t) → g x t) , h , p) + + compute-in-domain-interval-cogap-simplicial-mapping-cylinder : + (g : X → 𝟚 → S) (h : Y → S) (p : (x : X) → g x 1₂ = h (f x)) + (x : X) (t : 𝟚) → + cogap-simplicial-mapping-cylinder g h p + ( in-domain-interval-simplicial-mapping-cylinder f x t) = + ( g x t) + compute-in-domain-interval-cogap-simplicial-mapping-cylinder g h p x t = + compute-inl-cogap + ( λ x → (x , 1₂)) + ( f) + ( (λ (x , t) → g x t) , h , p) + ( x , t) + + compute-in-domain-cogap-simplicial-mapping-cylinder : + (g : X → 𝟚 → S) (h : Y → S) (p : (x : X) → g x 1₂ = h (f x)) (x : X) → + cogap-simplicial-mapping-cylinder g h p + ( in-domain-simplicial-mapping-cylinder f x) = + g x 0₂ + compute-in-domain-cogap-simplicial-mapping-cylinder g h p x = + compute-in-domain-interval-cogap-simplicial-mapping-cylinder g h p x 0₂ +``` diff --git a/src/simplicial-type-theory/simplicial-spines.lagda.md b/src/simplicial-type-theory/simplicial-spines.lagda.md new file mode 100644 index 0000000000..b1be05ee54 --- /dev/null +++ b/src/simplicial-type-theory/simplicial-spines.lagda.md @@ -0,0 +1,598 @@ +# The simplicial spines + +```agda +module simplicial-type-theory.simplicial-spines where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.contractible-types +open import foundation.coproduct-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition + +open import reflection.erasing-equality + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-cubes + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.dependent-cocones-under-spans +open import synthetic-homotopy-theory.dependent-universal-property-pushouts +open import synthetic-homotopy-theory.induction-principle-pushouts +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +open import synthetic-homotopy-theory.recursion-principle-pushouts +open import synthetic-homotopy-theory.universal-property-pushouts + +open import univalent-combinatorics.standard-finite-types +``` + +
+ +## Idea + +The {{#concept "𝑛-spine"}} is the classifying type of chains of directed edges +of length 𝑛. + +```text + 0 ---> 1 ----> ... ----> (n-1) ----> n +``` + +It has the universal property of the iterated +[pushout](synthetic-homotopy-theory.pushouts.md) + +```text + 0₂ + 1 ---------> 𝟚 + | | + target | | + ∨ ⌜ ∨ + spine n ----> spine (n + 1) +``` + +where + +```text + spine 0 := 1. +``` + +In +[`rewriting-simplicial-spines`](simplicial-type-theory.rewriting-simplicial-spines.md) +we equip the eliminator of the spines with strict computation rules on the point +constructors. + +## Postulates + +### The type of 𝑛-spines + +```agda +postulate + spine : ℕ → UU lzero + + star-spine-0 : spine 0 + + contraction-star-spine-0 : (x : spine 0) → star-spine-0 = x + + inl-spine : {n : ℕ} → spine n → spine (succ-ℕ n) + + in-arrow-spine : {n : ℕ} → 𝟚 → spine (succ-ℕ n) + +is-contr-spine-0 : is-contr (spine 0) +is-contr-spine-0 = (star-spine-0 , contraction-star-spine-0) + +initial-point-spine : (n : ℕ) → spine n +initial-point-spine zero-ℕ = star-spine-0 +initial-point-spine (succ-ℕ n) = inl-spine (initial-point-spine n) + +terminal-point-spine : (n : ℕ) → spine n +terminal-point-spine zero-ℕ = star-spine-0 +terminal-point-spine (succ-ℕ n) = in-arrow-spine 1₂ + +postulate + glue-spine : + {n : ℕ} → inl-spine (terminal-point-spine n) = in-arrow-spine {n} 0₂ +``` + +### The induction principle of the (𝑛+1)-spine + +We postulate that the (𝑛+1)-spine is the pushout + +```text + 0₂ + 1 ---------> 𝟚 + | | + target | | + ∨ ⌜ ∨ + spine n -----> spine (n + 1) +``` + +```agda +cocone-spine : + (n : ℕ) → + cocone (point (terminal-point-spine n)) (point 0₂) (spine (succ-ℕ n)) +cocone-spine n = (inl-spine , in-arrow-spine , point glue-spine) + +module _ + {l : Level} (n : ℕ) (P : spine (succ-ℕ n) → UU l) + (d : + dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P)) + where + + postulate + dependent-cogap-spine : (x : spine (succ-ℕ n)) → P x + + compute-inl-dependent-cogap-spine : + (x : spine n) → + dependent-cogap-spine (inl-spine x) = + horizontal-map-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( d) + ( x) + compute-inl-dependent-cogap-spine x = + primEraseEquality compute-inl-dependent-cogap-spine' + where + postulate + compute-inl-dependent-cogap-spine' : + dependent-cogap-spine (inl-spine x) = + horizontal-map-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( d) + ( x) + + compute-inr-dependent-cogap-spine : + (t : 𝟚) → + dependent-cogap-spine (in-arrow-spine t) = + vertical-map-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( d) + ( t) + compute-inr-dependent-cogap-spine t = + primEraseEquality compute-inr-dependent-cogap-spine' + where + postulate + compute-inr-dependent-cogap-spine' : + dependent-cogap-spine (in-arrow-spine t) = + vertical-map-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( d) + ( t) + + postulate + compute-glue-dependent-cogap-spine : + coherence-htpy-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cogap-spine)) + ( d) + ( compute-inl-dependent-cogap-spine) + ( compute-inr-dependent-cogap-spine) +``` + +## Definitions + +### The spines as subtypes of the cubes + +We can inductively define the 𝑛-spine as a subtype of the 𝑛-cube via t + +```text + · ┄┄┄┄┄┄> ∙ + ∧∧ ∧∧ + ⋰ ┆ / ┆ + · ┄┄┄┄┄┄> ∙ ┆ + y x ∧ · ┄┄ ∧ ┄> · + ∧ ∧ ┆ ∧ | ∧– + |/ ┆ ⋰ | ⋰ + └-> z ∙ ------> ∙ +``` + +```agda +subtype-spine : (n : ℕ) → subtype lzero (simplicial-cube n) +subtype-spine 0 _ = unit-Prop +subtype-spine 1 _ = unit-Prop +subtype-spine 2 (x , y) = join-Prop (Id-𝟚-Prop x 1₂) (Id-𝟚-Prop y 0₂) +subtype-spine (succ-ℕ (succ-ℕ (succ-ℕ n))) (x , u) = + join-Prop + ( is-terminal-element-simplicial-cube-Prop (succ-ℕ (succ-ℕ n)) u) + ( (Id-𝟚-Prop x 0₂) ∧ (subtype-spine (succ-ℕ (succ-ℕ n)) u)) +``` + +Let us work out what this definition unfolds to when `n` is `2`: + +```text + subtype-spine 2 (s , t) + ≐ is-terminal t ∨ ((s = 0₂) ∧ (subtype-spine 1 t)) + ≐ (t = 1₂) ∨ ((s = 0₂) ∧ unit) + ≃ (t = 1₂) ∨ (s = 0₂). +``` + +Observe again that the coordinates are read in order from right to left. + +```agda +spine' : ℕ → UU lzero +spine' n = type-subtype (subtype-spine n) + +is-set-spine' : (n : ℕ) → is-set (spine' n) +is-set-spine' n = + is-set-type-subtype (subtype-spine n) (is-set-simplicial-cube n) +``` + +### The point inclusions of the spines + +The 𝑛-spine has 𝑛+1 points that we enumerate + +```text + 0 ---> 1 ---> 2 ---> ⋯ ---> n +``` + +```agda +point-spine : (n : ℕ) → Fin (succ-ℕ n) → spine n +point-spine zero-ℕ _ = star-spine-0 +point-spine (succ-ℕ n) (inl x) = inl-spine (point-spine n x) +point-spine (succ-ℕ n) (inr x) = terminal-point-spine (succ-ℕ n) + +compute-inr-point-spine : + (n : ℕ) {x : unit} → point-spine n (inr x) = terminal-point-spine n +compute-inr-point-spine zero-ℕ = refl +compute-inr-point-spine (succ-ℕ n) = refl +``` + +### The arrow inclusions of the spine + +The 𝑛-spine has 𝑛 arrows. + +```agda +arrow-spine : (n : ℕ) → Fin n → 𝟚 → spine n +arrow-spine (succ-ℕ n) (inl x) = inl-spine ∘ arrow-spine n x +arrow-spine (succ-ℕ n) (inr x) = in-arrow-spine +``` + +### The hom inclusions of the spine + +```agda +source-hom-spine : (n : ℕ) (x : Fin n) → spine n +source-hom-spine n x = point-spine n (inl-Fin n x) + +target-hom-spine : (n : ℕ) (x : Fin n) → spine n +target-hom-spine n x = point-spine n (inr-Fin n x) + +inv-eq-source-arrow-spine : + (n : ℕ) (x : Fin n) → source-hom-spine n x = arrow-spine n x 0₂ +inv-eq-source-arrow-spine (succ-ℕ n) (inl x) = + ap inl-spine (inv-eq-source-arrow-spine n x) +inv-eq-source-arrow-spine (succ-ℕ n) (inr x) = + ap inl-spine (compute-inr-point-spine n) ∙ glue-spine + +eq-source-arrow-spine : + (n : ℕ) (x : Fin n) → arrow-spine n x 0₂ = source-hom-spine n x +eq-source-arrow-spine n x = inv (inv-eq-source-arrow-spine n x) + +eq-target-arrow-spine : + (n : ℕ) (x : Fin n) → arrow-spine n x 1₂ = target-hom-spine n x +eq-target-arrow-spine (succ-ℕ n) (inl x) = + ap inl-spine (eq-target-arrow-spine n x) +eq-target-arrow-spine (succ-ℕ n) (inr x) = refl + +hom-spine : (n : ℕ) (x : Fin n) → source-hom-spine n x →₂ target-hom-spine n x +hom-spine n x = + ( arrow-spine n x , eq-source-arrow-spine n x , eq-target-arrow-spine n x) +``` + +### The dependent universal property of the spines + +```agda +module _ + (n : ℕ) {l : Level} (P : spine (succ-ℕ n) → UU l) + where + + htpy-compute-dependent-cogap-spine : + ( c : + dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P)) → + htpy-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cogap-spine n P c)) + ( c) + htpy-compute-dependent-cogap-spine c = + ( compute-inl-dependent-cogap-spine n P c , + compute-inr-dependent-cogap-spine n P c , + compute-glue-dependent-cogap-spine n P c) + + is-section-dependent-cogap-spine : + is-section + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P)) + ( dependent-cogap-spine n P) + is-section-dependent-cogap-spine c = + eq-htpy-dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) + ( dependent-cogap-spine n P c)) + ( c) + ( htpy-compute-dependent-cogap-spine c) + +induction-principle-spine : + (n : ℕ) → + induction-principle-pushout + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) +induction-principle-spine n P = + ( dependent-cogap-spine n P , is-section-dependent-cogap-spine n P) + +is-retraction-dependent-cogap-spine : + (n : ℕ) {l : Level} (P : spine (succ-ℕ n) → UU l) → + is-retraction + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P)) + ( dependent-cogap-spine n P) +is-retraction-dependent-cogap-spine n = + is-retraction-ind-induction-principle-pushout + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( induction-principle-spine n) + +dependent-universal-property-spine : + (n : ℕ) → + dependent-universal-property-pushout + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) +dependent-universal-property-spine n P = + is-equiv-is-invertible + ( dependent-cogap-spine n P) + ( is-section-dependent-cogap-spine n P) + ( is-retraction-dependent-cogap-spine n P) + +equiv-dependent-universal-property-spine : + (n : ℕ) {l : Level} (P : spine (succ-ℕ n) → UU l) → + ( (x : spine (succ-ℕ n)) → P x) ≃ + ( dependent-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P)) +pr1 (equiv-dependent-universal-property-spine n P) = + dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( P) +pr2 (equiv-dependent-universal-property-spine n P) = + dependent-universal-property-spine n P +``` + +### The universal property of the spines + +```agda +module _ + (n : ℕ) {l : Level} {X : UU l} + where + + cogap-spine : + cocone (point (terminal-point-spine n)) (point 0₂) X → spine (succ-ℕ n) → X + cogap-spine = + dependent-cogap-spine n (λ _ → X) ∘ + dependent-cocone-constant-type-family-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + + is-section-cogap-spine : + is-section + ( cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) + ( cogap-spine) + is-section-cogap-spine = + ( ( triangle-dependent-cocone-map-constant-type-family' + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) ·r + ( cogap-spine)) ∙h + ( ( cocone-dependent-cocone-constant-type-family + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) ·l + ( is-section-dependent-cogap-spine n (λ _ → X)) ·r + ( dependent-cocone-constant-type-family-cocone + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n))) ∙h + ( is-retraction-cocone-dependent-cocone-constant-type-family + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) + + is-retraction-cogap-spine : + is-retraction + ( cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) + ( cogap-spine) + is-retraction-cogap-spine = + ( ( cogap-spine) ·l + ( triangle-dependent-cocone-map-constant-type-family' + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n))) ∙h + ( ( dependent-cogap-spine n (λ _ → X)) ·l + ( is-section-cocone-dependent-cocone-constant-type-family + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n)) ·r + ( dependent-cocone-map + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) + ( λ _ → X))) ∙h + ( is-retraction-dependent-cogap-spine n (λ _ → X)) + +recursion-principle-spine : + (n : ℕ) → + recursion-principle-pushout + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) +recursion-principle-spine n = (cogap-spine n , is-section-cogap-spine n) + +universal-property-spine : + (n : ℕ) → + universal-property-pushout + ( point (terminal-point-spine n)) + ( point 0₂) + ( cocone-spine n) +universal-property-spine n Y = + is-equiv-is-invertible + ( cogap-spine n) + ( is-section-cogap-spine n) + ( is-retraction-cogap-spine n) + +equiv-universal-property-spine : + (n : ℕ) {l : Level} {X : UU l} → + (spine (succ-ℕ n) → X) ≃ cocone (point (terminal-point-spine n)) (point 0₂) X +equiv-universal-property-spine n {X = X} = + ( cocone-map (point (terminal-point-spine n)) (point 0₂) (cocone-spine n) , + universal-property-spine n X) +``` + +### Auxillary definitions for the 𝑛-spine as a subtype of the 𝑛-cube + +> This remains to be formalized. + +```text +inl-spine' : (n : ℕ) → spine' n → spine' (succ-ℕ n) +inl-spine' zero-ℕ _ = (0₂ , star) +inl-spine' (succ-ℕ zero-ℕ) (t , _) = ((0₂ , t) , inr-join (refl , star)) +inl-spine' (succ-ℕ (succ-ℕ n)) x = ((0₂ , {! !}) , inr-join (refl , {! !})) + +terminal-point-spine' : (n : ℕ) → spine' n +terminal-point-spine' zero-ℕ = star , star +terminal-point-spine' (succ-ℕ zero-ℕ) = 1₂ , star +terminal-point-spine' (succ-ℕ (succ-ℕ n)) = ({! !} , {! !}) + +cocone-spine' : + (n : ℕ) → + cocone (point (terminal-point-spine' n)) (point 0₂) (spine' (succ-ℕ n)) +cocone-spine' = {! !} + +-- map-spine-spine' : +-- (n : ℕ) → spine' n → spine n +-- map-spine-spine' 0 x = star-spine-0 +-- map-spine-spine' 1 (x , u) = in-arrow-spine x +-- map-spine-spine' (succ-ℕ (succ-ℕ n)) ((t , x) , u) = +-- cogap-join +-- ( spine (succ-ℕ (succ-ℕ n))) +-- ( ( λ _ → in-arrow-spine t) , +-- ( λ rs → +-- inl-spine (map-spine-spine' (succ-ℕ n) (x , pr2 rs))) , +-- ( λ (is-terminal-x , t=0 , s) → +-- ( ( ap in-arrow-spine t=0) ∙ +-- ( inv glue-spine) ∙ +-- ( ap inl-spine +-- ( compute-inr-map-spine-spine' n (x , s) is-terminal-x))))) +-- ( u) +-- where +-- compute-inr-map-spine-spine' : +-- (n : ℕ) (xs : spine' (succ-ℕ n)) → +-- is-terminal-element-simplicial-cube (succ-ℕ n) (pr1 xs) → +-- in-arrow-spine 1₂ = map-spine-spine' (succ-ℕ n) xs +-- compute-inr-map-spine-spine' zero-ℕ xs is-terminal-x = +-- ap in-arrow-spine (inv is-terminal-x) +-- compute-inr-map-spine-spine' (succ-ℕ n) xs is-terminal-x = +-- inv (compute-inl-cogap-join {! !} {! !}) ∙ {! !} + +-- -- where map-spine-spine' (succ-ℕ n) (x , s) +``` + +## Properties + +### The 1-spine is the directed interval + +```text + 1 ----------> 𝟚 + | | + | | + ∨ ⌜ ∨ + 1 ≃ spine 0 ----> spine 1 +``` + +> This remains to be formalized. diff --git a/src/simplicial-type-theory/simplicially-covariant-type-families.lagda.md b/src/simplicial-type-theory/simplicially-covariant-type-families.lagda.md new file mode 100644 index 0000000000..1c4c34884c --- /dev/null +++ b/src/simplicial-type-theory/simplicially-covariant-type-families.lagda.md @@ -0,0 +1,111 @@ +# Simplicially covariant type families + +```agda +module simplicial-type-theory.simplicially-covariant-type-families where +``` + +
Imports + +```agda +open import foundation.0-connected-types +open import foundation.action-on-identifications-functions +open import foundation.connected-types +open import foundation.dependent-pair-types +open import foundation.diagonal-maps-of-types +open import foundation.equivalences +open import foundation.equivalences-arrows +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.negation +open import foundation.propositions +open import foundation.sections +open import foundation.torsorial-type-families +open import foundation.transport-along-identifications +open import foundation.unit-type +open import foundation.universe-levels + +open import orthogonal-factorization-systems.local-families-of-types +open import orthogonal-factorization-systems.null-types + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.circle +``` + +
+ +## Idea + +A type family `B : A → 𝒰` is +{{#concept "simplicially covariant" Disambiguation="type family" Agda=is-simplicially-covariant}} +if one of the following equivalent conditions hold: + +1. For every directed edge `f : x →₂ y` in `A` and element `x'` over `x`, the + type of dependent directed edges over `f` based at `x'` is + [torsorial](foundation.torsorial-type-families.md). + +2. For every simplicial arrow `α` in `A`, and element `x'` over `α 0₂`, the type + of dependent directed edges over `α` based at `x'` is torsorial. + + ```text + (α : arrow A) (x' : B (α 0₂)) → + is-contr (Σ (y' : B (α 1₂)), (dependent-hom B α x' y')) + ``` + + Note that this is just a slight simplification of the previous condition. + +3. The type family is + [local](orthogonal-factorization-systems.local-type-families.md) at the left + end-point inclusion `0₂ : 1 ↪ 𝟚`. + +4. The following square is a [pullback](foundation-core.pullbacks.md) + + ```text + postcomp 𝟚 pr1 + (𝟚 → Σ (x : A), (B x)) ---------------> (𝟚 → A) + | | + | | + ev 0₂ | | ev 0₂ + | | + ∨ ∨ + Σ (x : A), (B x) -------------------> A + pr1 + ``` + + + +## Definitions + +### The predicate on type families of being simplicially covariant + +```agda +module _ + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) + where + + is-simplicially-covariant : UU (l1 ⊔ l2) + is-simplicially-covariant = is-local-family (point 0₂) B +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) + where + + is-simplicially-covariant' : UU (l1 ⊔ l2) + is-simplicially-covariant' = + (α : simplicial-arrow A) (x' : B (α 0₂)) → + is-torsorial + ( dependent-simplicial-hom B (simplicial-hom-simplicial-arrow α) x') +``` + +## References + +{{#bibliography}} {{#reference RS17}} diff --git a/src/simplicial-type-theory/simplicially-discrete-types.lagda.md b/src/simplicial-type-theory/simplicially-discrete-types.lagda.md new file mode 100644 index 0000000000..c7c521861f --- /dev/null +++ b/src/simplicial-type-theory/simplicially-discrete-types.lagda.md @@ -0,0 +1,321 @@ +# Simplicially discrete types + +```agda +module simplicial-type-theory.simplicially-discrete-types where +``` + +
Imports + +```agda +open import foundation.0-connected-types +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.connected-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.diagonal-maps-of-types +open import foundation.empty-types +open import foundation.equivalences +open import foundation.equivalences-arrows +open import foundation.function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.logical-equivalences +open import foundation.negation +open import foundation.propositions +open import foundation.retracts-of-types +open import foundation.sections +open import foundation.torsorial-type-families +open import foundation.transport-along-identifications +open import foundation.universe-levels + +open import orthogonal-factorization-systems.null-maps +open import orthogonal-factorization-systems.null-types + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.inequality-directed-interval-type + +open import synthetic-homotopy-theory.circle +``` + +
+ +## Idea + +A type `A` is +{{#concept "simplicially discrete" Disambiguation="type" Agda=is-simplicially-discrete}} +if the canonical map + +```text + simplicial-hom-eq : (x = y) → (x →₂ y) +``` + +is an [equivalence](foundation-core.equivalences.md) for all `x y : A`. A +simplicially discrete type bears only trivial simplicial structure in the sense +that its simplices act precisely as its identifications. In particular, +simplicially discrete types are Rezk complete and Segal. + +## Definitions + +### The predicate on types of being simplicially discrete + +```agda +module _ + {l : Level} (A : UU l) + where + + is-simplicially-discrete : UU l + is-simplicially-discrete = + (x y : A) → is-equiv (simplicial-hom-eq {x = x} {y}) + + is-prop-is-simplicially-discrete : is-prop is-simplicially-discrete + is-prop-is-simplicially-discrete = + is-prop-Π (λ x → is-prop-Π (λ y → is-property-is-equiv simplicial-hom-eq)) + + is-simplicially-discrete-Prop : Prop l + is-simplicially-discrete-Prop = + ( is-simplicially-discrete , is-prop-is-simplicially-discrete) +``` + +### The type of simplicially discrete types + +```agda +Simplicially-Discrete-Type : (l : Level) → UU (lsuc l) +Simplicially-Discrete-Type l = Σ (UU l) (is-simplicially-discrete) + +module _ + {l : Level} (A : Simplicially-Discrete-Type l) + where + + type-Simplicially-Discrete-Type : UU l + type-Simplicially-Discrete-Type = pr1 A + + is-simplicially-discrete-Simplicially-Discrete-Type : + is-simplicially-discrete type-Simplicially-Discrete-Type + is-simplicially-discrete-Simplicially-Discrete-Type = pr2 A +``` + +## Properties + +### To show a type is simplicially discrete it suffices to construct a section of `simplicial-hom-eq` + +```agda +module _ + {l : Level} (A : UU l) + where + + is-simplicially-discrete-section-simplicial-hom-eq : + ((x y : A) → section (simplicial-hom-eq {x = x} {y})) → + is-simplicially-discrete A + is-simplicially-discrete-section-simplicial-hom-eq s x = + fundamental-theorem-id-section x (λ y → simplicial-hom-eq) (s x) +``` + +### Being simplicially discrete is equivalent to being `𝟚`-null + +**Proof.** We have the [equivalence of maps](foundation.equivalences-arrows.md) + +```text + ~ + A -------> Σ (x y : A), (x = y) + | | + Δ | | Σ² simplicial-hom-eq + ∨ ∨ + (𝟚 → A) ----> Σ (x y : A), (x →₂ y), + ~ +``` + +which implies that the diagonal map is an equivalence if and only if the total +map of `simplicial-hom-eq` is, and the total map is an equivalence if and only +if the fiberwise map is. + +```agda +module _ + {l : Level} {A : UU l} + where + + equiv-tot-simplicial-hom-eq-diagonal-exponential-𝟚 : + equiv-arrow + ( diagonal-exponential A 𝟚) + ( tot (λ x → tot (λ y → simplicial-hom-eq {x = x} {y}))) + equiv-tot-simplicial-hom-eq-diagonal-exponential-𝟚 = + ( compute-total-Id , compute-total-simplicial-hom , refl-htpy) + + abstract + is-simplicially-discrete-is-𝟚-null : + is-null 𝟚 A → is-simplicially-discrete A + is-simplicially-discrete-is-𝟚-null H x = + is-fiberwise-equiv-is-equiv-tot + ( is-fiberwise-equiv-is-equiv-tot + ( is-equiv-target-is-equiv-source-equiv-arrow + ( diagonal-exponential A 𝟚) + ( tot (λ x → tot (λ y → simplicial-hom-eq {x = x} {y}))) + ( equiv-tot-simplicial-hom-eq-diagonal-exponential-𝟚) + ( H)) + ( x)) + + abstract + is-𝟚-null-is-simplicially-discrete : + is-simplicially-discrete A → is-null 𝟚 A + is-𝟚-null-is-simplicially-discrete H = + is-equiv-source-is-equiv-target-equiv-arrow + ( diagonal-exponential A 𝟚) + ( tot (λ x → tot (λ y → simplicial-hom-eq {x = x} {y}))) + ( equiv-tot-simplicial-hom-eq-diagonal-exponential-𝟚) + ( is-equiv-tot-is-fiberwise-equiv + ( λ x → is-equiv-tot-is-fiberwise-equiv (H x))) + + iff-is-𝟚-null-is-simplicially-discrete : + is-simplicially-discrete A ↔ is-null 𝟚 A + iff-is-𝟚-null-is-simplicially-discrete = + ( is-𝟚-null-is-simplicially-discrete , is-simplicially-discrete-is-𝟚-null) +``` + +### Simplicially discrete types are closed under retracts + +```agda +is-simplicially-discrete-retract : + {l1 l2 : Level} {A : UU l1} {B : UU l2} → + A retract-of B → is-simplicially-discrete B → is-simplicially-discrete A +is-simplicially-discrete-retract r H = + is-simplicially-discrete-is-𝟚-null + ( is-null-retract-base r (is-𝟚-null-is-simplicially-discrete H)) +``` + +### Simplicially discrete types are closed under equivalences + +```agda +is-simplicially-discrete-equiv : + {l1 l2 : Level} {A : UU l1} {B : UU l2} → + A ≃ B → is-simplicially-discrete B → is-simplicially-discrete A +is-simplicially-discrete-equiv e H = + is-simplicially-discrete-is-𝟚-null + ( is-null-equiv-base e (is-𝟚-null-is-simplicially-discrete H)) +``` + +### Simplicially discrete types are closed under dependent products + +```agda +is-simplicially-discrete-Π : + {l1 l2 : Level} {I : UU l1} {B : I → UU l2} → + ((i : I) → is-simplicially-discrete (B i)) → + is-simplicially-discrete ((i : I) → B i) +is-simplicially-discrete-Π H = + is-simplicially-discrete-is-𝟚-null + ( is-null-Π (λ i → is-𝟚-null-is-simplicially-discrete (H i))) +``` + +### Simplicially discrete types are closed under exponentiation + +```agda +is-simplicially-discrete-function-type : + {l1 l2 : Level} {I : UU l1} {B : UU l2} → + is-simplicially-discrete B → + is-simplicially-discrete (I → B) +is-simplicially-discrete-function-type H = is-simplicially-discrete-Π (λ _ → H) +``` + +### Simplicially discrete types are closed under cartesian products + +```agda +is-simplicially-discrete-product : + {l1 l2 : Level} {A : UU l1} {B : UU l2} → + is-simplicially-discrete A → + is-simplicially-discrete B → + is-simplicially-discrete (A × B) +is-simplicially-discrete-product is-disc-A is-disc-B = + is-simplicially-discrete-is-𝟚-null + ( is-null-product + ( is-𝟚-null-is-simplicially-discrete is-disc-A) + ( is-𝟚-null-is-simplicially-discrete is-disc-B)) +``` + +### Simplicially discrete types are closed under dependent sums + +```agda +is-simplicially-discrete-Σ : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} → + is-simplicially-discrete A → + ((x : A) → is-simplicially-discrete (B x)) → + is-simplicially-discrete (Σ A B) +is-simplicially-discrete-Σ is-disc-A is-disc-B = + is-simplicially-discrete-is-𝟚-null + ( is-null-Σ 𝟚 + ( is-𝟚-null-is-simplicially-discrete is-disc-A) + ( λ x → is-𝟚-null-is-simplicially-discrete (is-disc-B x))) +``` + +### A family over a simplicially discrete type is a family of simplicially discrete types if and only if the dependent sum is + +One direction was established above, the converse is recorded below. + +```agda +is-simplicially-discrete-family-is-simplicially-discrete-Σ : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} → + is-simplicially-discrete A → + is-simplicially-discrete (Σ A B) → + (x : A) → is-simplicially-discrete (B x) +is-simplicially-discrete-family-is-simplicially-discrete-Σ + is-disc-A is-disc-ΣAB x = + is-simplicially-discrete-is-𝟚-null + ( is-null-family-is-null-Σ 𝟚 + ( is-𝟚-null-is-simplicially-discrete is-disc-A) + ( is-𝟚-null-is-simplicially-discrete is-disc-ΣAB) + ( x)) +``` + +### Simplicially discrete types are Segal + +This remains to be formalized. The proof boils down to showing that `Λ²₁ ↪ Δ²` +is anodyne with respect to `𝟚 → 1`. + +### A type is simplicially discrete if and only if it is pregroupoidal and Rezk complete + +This is proposition 10.10 of {{#cite RS17}}. This remains to be formalized. + + + +## Examples + +### The directed interval is not simplicially discrete + +```agda +is-not-simplicially-discrete-𝟚 : ¬ (is-simplicially-discrete 𝟚) +is-not-simplicially-discrete-𝟚 H = + is-nontrivial-𝟚 (map-inv-is-equiv (H 0₂ 1₂) representing-hom-𝟚) +``` + +### Propositions are simplicially discrete + +```agda +is-simplicially-discrete-is-prop : + {l : Level} {P : UU l} → is-prop P → is-simplicially-discrete P +is-simplicially-discrete-is-prop = + is-simplicially-discrete-is-𝟚-null ∘ is-null-is-prop-is-inhabited' 0₂ +``` + +### Contractible types are simplicially discrete + +```agda +is-simplicially-discrete-is-contr : + {l : Level} {P : UU l} → is-contr P → is-simplicially-discrete P +is-simplicially-discrete-is-contr is-contr-P = + is-simplicially-discrete-is-prop (is-prop-is-contr is-contr-P) +``` + +### Empty types are simplicially discrete + +```agda +is-simplicially-discrete-is-empty : + {l : Level} {P : UU l} → is-empty P → is-simplicially-discrete P +is-simplicially-discrete-is-empty is-empty-P = + is-simplicially-discrete-is-prop (is-prop-is-empty is-empty-P) +``` + +## References + +{{#bibliography}} {{#reference RS17}} diff --git a/src/simplicial-type-theory/standard-simplices.lagda.md b/src/simplicial-type-theory/standard-simplices.lagda.md new file mode 100644 index 0000000000..52cb2cdd72 --- /dev/null +++ b/src/simplicial-type-theory/standard-simplices.lagda.md @@ -0,0 +1,85 @@ +# Standard simplices + +```agda +module simplicial-type-theory.standard-simplices where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.embeddings +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unions-subtypes +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows +open import simplicial-type-theory.simplicial-cubes + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Definitions + +### The standard simplices + +$$x₁ ≥ x₂ ≥ … ≥ xₙ₋₁ ≥ xₙ$$ (in the right-associated cube) + +```agda +subtype-standard-simplex : (n : ℕ) → subtype lzero (simplicial-cube n) +subtype-standard-simplex 0 _ = + unit-Prop +subtype-standard-simplex 1 _ = + unit-Prop +subtype-standard-simplex 2 (x , y) = + leq-𝟚-Prop y x +subtype-standard-simplex (succ-ℕ (succ-ℕ (succ-ℕ n))) (x , y , u) = + conjunction-Prop + ( subtype-standard-simplex 2 (x , y)) + ( subtype-standard-simplex (succ-ℕ (succ-ℕ n)) (y , u)) + +predicate-standard-simplex : (n : ℕ) → simplicial-cube n → UU +predicate-standard-simplex = is-in-subtype ∘ subtype-standard-simplex + +standard-simplex : ℕ → UU +standard-simplex = type-subtype ∘ subtype-standard-simplex + +Δ : ℕ → UU +Δ = standard-simplex +``` + +## Properties + +### The standard 𝑛-simplex is a retract of the directed 𝑛-cube + +This remains to be formalized. Lemma 4.2.2 {{#cite MR23b}} + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/standard-simplicial-joins.lagda.md b/src/simplicial-type-theory/standard-simplicial-joins.lagda.md new file mode 100644 index 0000000000..27e801cda4 --- /dev/null +++ b/src/simplicial-type-theory/standard-simplicial-joins.lagda.md @@ -0,0 +1,117 @@ +# Simplicial joins + +```agda +module simplicial-type-theory.simplicial-joins where +``` + +
Imports + +```agda +open import elementary-number-theory.natural-numbers + +open import foundation.action-on-identifications-functions +open import foundation.booleans +open import foundation.cartesian-product-types +open import foundation.conjunction +open import foundation.coproduct-types +open import foundation.dependent-pair-types +open import foundation.disjunction +open import foundation.empty-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +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.sets +open import foundation.subtypes +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.unit-type +open import foundation.universe-levels + +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.inequality-directed-interval-type +open import simplicial-type-theory.simplicial-arrows + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.joins-of-types +open import synthetic-homotopy-theory.pushouts +``` + +
+ +## Idea + +We define the {{#concept "simplicial join"}} of two types `A *▵ B` as the +colimit of the diagram + +```text + A × B ----> B + | ⋮ + id × 1₂ × id | ⋮ + ∨ ⋮ + A × B ---> A × 𝟚 × B ⋮ + | id × 0₂ × id ⋱ ⋮ + | ⋱ ⋮ + ∨ ∨ ∨ + A ⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯⋯> A *▵ B. +``` + +Equivalently, the lax join is the oriented pushout + +```text + A × B ------> B + | | + | ⇗ | + ∨ ⌜ ∨ + A ------> A *▵ B. +``` + +Intuitively, we can understand `A *▵ B` as the universal type equipped with +edges `a →₂ b` for every `a : A` and `b : B`. + +This construction satisfies the laws + +- $𝟚 ≃ 1 *▵ 1$ + +- $Δⁿ⁺¹ ≃ Δⁿ⁺¹ *▵ 1 ≃ 1 *▵ Δⁿ⁺¹$ + +- $Λ²₀ ≃ 1 *▵ bool$ + +- $Λ²₂ ≃ bool *▵ 1$ + +- $1 *▵ (-)$ is the simplicial cone + +- $ (-) \*▵ 1$ is the simplicial cocone + +## Postulates + +### The standard simplicial join + +```agda +module _ + {l1 l2 : Level} (A : UU l1) (B : UU l2) + where + + postulate + standard-simplicial-join : UU (l1 ⊔ l2) + + in-standard-simplicial-join : A → B → 𝟚 → standard-simplicial-join + + glue-source-standard-simplicial-join : + (a : A) (b b' : B) → + in-standard-simplicial-join a b 0₂ = in-standard-simplicial-join a b' 0₂ + + glue-target-standard-simplicial-join : + (a a' : A) (b : B) → + in-standard-simplicial-join a b 1₂ = in-standard-simplicial-join a' b 1₂ +``` + +> It remains to define and postulate the induction principle of the simplicial +> join. + +## See also + +- The simplicial pushout join diff --git a/src/simplicial-type-theory/transposing-adjunctions-between-types.lagda.md b/src/simplicial-type-theory/transposing-adjunctions-between-types.lagda.md new file mode 100644 index 0000000000..e48795f415 --- /dev/null +++ b/src/simplicial-type-theory/transposing-adjunctions-between-types.lagda.md @@ -0,0 +1,371 @@ +# Transposing adjunctions between types + +```agda +module simplicial-type-theory.transposing-adjunctions-between-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.commuting-triangles-of-identifications +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.functoriality-dependent-function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.postcomposition-functions +open import foundation.precomposition-functions +open import foundation.univalence +open import foundation.universal-property-equivalences +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition +open import foundation.whiskering-identifications-concatenation + +open import foundation-core.equivalences +open import foundation-core.homotopies + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-edges-dependent-pair-types +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.natural-transformations +``` + +
+ +## Idea + +Consider a pair of maps `L : A ↔ B : R`. We say that `L` and `R` form a +transposing adjunction given a binary family of equivalences + +```text + (x : A) (y : B) → hom B (L x) y ≃ hom A x (R y) +``` + +## Definitions + +### Transposing adjoint pairs + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-adjunction : (A → B) → (B → A) → UU (l1 ⊔ l2) + is-transposing-adjunction L R = (x : A) (y : B) → hom▵ (L x) y ≃ hom▵ x (R y) + + _⊣▵_ : (A → B) → (B → A) → UU (l1 ⊔ l2) + _⊣▵_ = is-transposing-adjunction +``` + +### Transposing left adjoints + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-left-adjoint : (A → B) → UU (l1 ⊔ l2) + is-transposing-left-adjoint L = Σ (B → A) (is-transposing-adjunction L) +``` + +### Transposing right adjoints + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-right-adjoint : (A → B) → UU (l1 ⊔ l2) + is-transposing-right-adjoint R = + Σ (B → A) (λ L → is-transposing-adjunction L R) +``` + +### Transposing adjunctions + +```agda +transposing-adjunction : {l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2) +transposing-adjunction A B = Σ (A → B) is-transposing-left-adjoint + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (H : transposing-adjunction A B) + where + + map-left-adjoint-transposing-adjunction : A → B + map-left-adjoint-transposing-adjunction = pr1 H + + is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction : + is-transposing-left-adjoint map-left-adjoint-transposing-adjunction + is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction = pr2 H + + map-right-adjoint-transposing-adjunction : B → A + map-right-adjoint-transposing-adjunction = + pr1 is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction + + is-transposing-adjunction-transposing-adjunction : + is-transposing-adjunction + map-left-adjoint-transposing-adjunction + map-right-adjoint-transposing-adjunction + is-transposing-adjunction-transposing-adjunction = + pr2 is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction + + is-transposing-right-adjoint-map-right-adjoint-transposing-adjunction : + is-transposing-right-adjoint map-right-adjoint-transposing-adjunction + is-transposing-right-adjoint-map-right-adjoint-transposing-adjunction = + ( map-left-adjoint-transposing-adjunction , + is-transposing-adjunction-transposing-adjunction) +``` + +### The unit and counit natural transformations associated to a transposing adjunction + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + unit-is-transposing-adjunction : + {L : A → B} {R : B → A} → + is-transposing-adjunction L R → (id ⇒▵ R ∘ L) + unit-is-transposing-adjunction {L} {R} H x = + map-equiv (H x (L x)) (id-simplicial-hom (L x)) + + counit-is-transposing-adjunction : + {L : A → B} {R : B → A} → + is-transposing-adjunction L R → (L ∘ R ⇒▵ id) + counit-is-transposing-adjunction {L} {R} H y = + map-inv-equiv (H (R y) y) (id-simplicial-hom (R y)) +``` + +## Properties + +### The identity function is a transposing adjunction + +```agda +module _ + {l : Level} {A : UU l} + where + + is-transposing-adjunction-id : is-transposing-adjunction (id {A = A}) id + is-transposing-adjunction-id x y = id-equiv + + is-transposing-adjoint-id : is-transposing-left-adjoint (id {A = A}) + is-transposing-adjoint-id = id , is-transposing-adjunction-id + + is-transposing-adjoint-id' : is-transposing-right-adjoint (id {A = A}) + is-transposing-adjoint-id' = id , is-transposing-adjunction-id + + id-transposing-adjunction : transposing-adjunction A A + id-transposing-adjunction = (id , id , is-transposing-adjunction-id) +``` + +### Equivalences are transposing adjunctions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} (u : is-equiv f) + where + + is-transposing-adjunction-is-equiv' : + is-transposing-adjunction (map-inv-is-equiv u) f + is-transposing-adjunction-is-equiv' y x = + equiv-eq (ap (λ p → hom▵ p (f x)) (is-section-map-inv-is-equiv u y)) ∘e + equiv-action-simplicial-hom (f , u) (map-inv-is-equiv u y) x + + is-transposing-adjunction-is-equiv : + is-transposing-adjunction f (map-inv-is-equiv u) + is-transposing-adjunction-is-equiv x y = + inv-equiv + ( equiv-eq (ap (hom▵ (f x)) (is-section-map-inv-is-equiv u y)) ∘e + equiv-action-simplicial-hom (f , u) x (map-inv-is-equiv u y)) + + is-transposing-left-adjoint-is-equiv : is-transposing-left-adjoint f + is-transposing-left-adjoint-is-equiv = + (map-inv-is-equiv u , is-transposing-adjunction-is-equiv) + + is-transposing-right-adjoint-is-equiv : is-transposing-right-adjoint f + is-transposing-right-adjoint-is-equiv = + (map-inv-is-equiv u , is-transposing-adjunction-is-equiv') +``` + +### Composition of transposing adjunctions + +Given a diagram of transposing adjunctions + +```text + R R' + <-----> ------- + A ⊤ B ⊤ C + ------> ------> + L L', +``` + +then we have a composite transposing adjunction `L' ∘ L ⊣ R ∘ R'`. + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-transposing-adjunction-comp : + {L : A → B} {L' : B → C} {R : B → A} {R' : C → B} → + is-transposing-adjunction L' R' → + is-transposing-adjunction L R → + is-transposing-adjunction (L' ∘ L) (R ∘ R') + is-transposing-adjunction-comp {L} {L'} {R} {R'} H H' x y = + H' x (R' y) ∘e H (L x) y + + is-transposing-left-adjoint-comp : + {L : A → B} {L' : B → C} → + is-transposing-left-adjoint L' → + is-transposing-left-adjoint L → + is-transposing-left-adjoint (L' ∘ L) + is-transposing-left-adjoint-comp (R' , H') (R , H) = + ( R ∘ R' , is-transposing-adjunction-comp H' H) + + is-transposing-right-adjoint-comp : + {R : B → A} {R' : C → B} → + is-transposing-right-adjoint R → + is-transposing-right-adjoint R' → + is-transposing-right-adjoint (R ∘ R') + is-transposing-right-adjoint-comp (L' , H') (L , H) = + ( L ∘ L' , is-transposing-adjunction-comp H H') + + comp-transposing-adjunction : + transposing-adjunction B C → + transposing-adjunction A B → + transposing-adjunction A C + comp-transposing-adjunction (L' , R' , H') (L , R , H) = + ( L' ∘ L , R ∘ R' , is-transposing-adjunction-comp H' H) +``` + +### Dependent products of transposing adjunctions + +```agda +module _ + {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} + where + + is-transposing-adjunction-Π : + {L : (x : I) → A x → B x} {R : (x : I) → B x → A x} → + ((x : I) → is-transposing-adjunction (L x) (R x)) → + is-transposing-adjunction (map-Π L) (map-Π R) + is-transposing-adjunction-Π H f g = + inv-equiv extensionality-simplicial-natural-transformation ∘e + equiv-Π-equiv-family (λ i → H i (f i) (g i)) ∘e + extensionality-simplicial-natural-transformation + + is-transposing-left-adjoint-Π : + {L : (x : I) → A x → B x} → + ((x : I) → is-transposing-left-adjoint (L x)) → + is-transposing-left-adjoint (map-Π L) + is-transposing-left-adjoint-Π H = + ( map-Π (pr1 ∘ H) , is-transposing-adjunction-Π (pr2 ∘ H)) + + is-transposing-right-adjoint-Π : + {R : (x : I) → B x → A x} → + ((x : I) → is-transposing-right-adjoint (R x)) → + is-transposing-right-adjoint (map-Π R) + is-transposing-right-adjoint-Π H = + ( map-Π (pr1 ∘ H) , is-transposing-adjunction-Π (pr2 ∘ H)) + + transposing-adjunction-Π : + ((x : I) → transposing-adjunction (A x) (B x)) → + transposing-adjunction ((x : I) → A x) ((x : I) → B x) + transposing-adjunction-Π H = + ( map-Π (map-left-adjoint-transposing-adjunction ∘ H) , + map-Π (map-right-adjoint-transposing-adjunction ∘ H) , + is-transposing-adjunction-Π + ( is-transposing-adjunction-transposing-adjunction ∘ H)) +``` + +This is proven for families of one-sided inverse adjunctions between Rezk types +in Appendix B of {{#cite BW23}}. + +### Postcomposition by transposing adjunctions + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-adjunction-postcomp : + {l : Level} (X : UU l) {L : A → B} {R : B → A} → + is-transposing-adjunction L R → + is-transposing-adjunction (postcomp X L) (postcomp X R) + is-transposing-adjunction-postcomp X H = is-transposing-adjunction-Π (λ _ → H) + + transposing-adjunction-postcomp : + {l : Level} (X : UU l) → + transposing-adjunction A B → + transposing-adjunction (X → A) (X → B) + transposing-adjunction-postcomp X H = transposing-adjunction-Π (λ _ → H) +``` + +### Base change of transposing adjunctions + +Given a transposing right adjoint `R : A → B`, `L ⊣ R` and a map `j : B' → B` +such that `R ∘ L ∘ j ~ j`, then the base change of `R` along `j`, `R'`, has a +left adjoint right inverse. + +**Proof.** Given a homotopy `H : R ∘ L ∘ j ~ j` we obtain the cone + +```text + L ∘ j + B' -----> A + ║ | + ║ H | R + ║ ∨ + B' -----> B + j +``` + +which, by the universal property of base change gives us a right inverse `L'` to +`R' : A' → B'`. Now... + +> This remains to be formalized + +### Retracts of transposing adjunctions + +Maps that are retracts of transposing adjoints are transposing adjoints +themselves. + +**Proof.** Given a retract of maps + +```text + i r + A' --------> A --------> A' + | | ∧ | + R'| R |⊢| L |R' + ∨ ∨ | ∨ + B' --------> B --------> B' + j k +``` + +then the functorial action of functions on directed edges gives us a binary +family of retracts + +```text + hom A' x y ---> hom A (i x) (i y) ---> hom A' x y +``` + +We of course get a canonical map `L' : B' → A'` defined by the above retract of +maps. Now, we have a chain of maps + +```text + hom A' (L' x) y + → hom A (i (L' x)) (i y) + = hom B (j x) (R (i y)) + = hom B (j x) (j (R' y)) + → hom B' (k (j x)) (k (j (R' y))) + = hom B' x (R' y) +``` + +and by our initial retract this total composite is an equivalence. Hence `L'` is +a transposing adjoint to `R'`. □ + +> This remains to be formalized. + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/transposing-biadjunctions-between-types.lagda.md b/src/simplicial-type-theory/transposing-biadjunctions-between-types.lagda.md new file mode 100644 index 0000000000..a7e2625cc4 --- /dev/null +++ b/src/simplicial-type-theory/transposing-biadjunctions-between-types.lagda.md @@ -0,0 +1,284 @@ +# Transposing biadjunctions between types + +```agda +module simplicial-type-theory.transposing-biadjunctions-between-types where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.commuting-triangles-of-identifications +open import foundation.dependent-pair-types +open import foundation.function-types +open import foundation.functoriality-dependent-function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.identity-types +open import foundation.postcomposition-functions +open import foundation.precomposition-functions +open import foundation.univalence +open import foundation.universal-property-equivalences +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition +open import foundation.whiskering-identifications-concatenation + +open import foundation-core.equivalences +open import foundation-core.homotopies + +open import simplicial-type-theory.dependent-directed-edges +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-edges-dependent-pair-types +open import simplicial-type-theory.fully-faithful-maps +open import simplicial-type-theory.natural-transformations +open import simplicial-type-theory.transposing-adjunctions-between-types +``` + +
+ +## Idea + +Given a map of types `q : 𝒞 → 𝒟`, we say `q` is a +{{#concept "transposing biadjoint" Disambiguation="map of types" Agda=is-transposing-biadjoint}} +if it has a transposing left and transposing right adjoint. + +```text + 𝒞 + ∧ | ∧ + | | | + q! | ⊣ q* ⊣ | q∗ + | | | + | ∨ | + 𝒟 +``` + +## Definitions + +### Transposing adjoint triples + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-adjoint-triple : (B → A) → (A → B) → (B → A) → UU (l1 ⊔ l2) + is-transposing-adjoint-triple q! q* q∗ = + (is-transposing-adjunction q! q*) × (is-transposing-adjunction q* q∗) +``` + +### Transposing biadjoints + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-biadjoint : (A → B) → UU (l1 ⊔ l2) + is-transposing-biadjoint q* = + (is-transposing-right-adjoint q*) × (is-transposing-left-adjoint q*) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q* : A → B} + (H : is-transposing-biadjoint q*) + where + + is-transposing-right-adjoint-is-transposing-biadjoint : + is-transposing-right-adjoint q* + is-transposing-right-adjoint-is-transposing-biadjoint = + pr1 H + + is-transposing-left-adjoint-is-transposing-biadjoint : + is-transposing-left-adjoint q* + is-transposing-left-adjoint-is-transposing-biadjoint = + pr2 H + + map-left-adjoint-is-transposing-biadjoint : B → A + map-left-adjoint-is-transposing-biadjoint = + pr1 is-transposing-right-adjoint-is-transposing-biadjoint + + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint : + is-transposing-adjunction map-left-adjoint-is-transposing-biadjoint q* + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint = + pr2 is-transposing-right-adjoint-is-transposing-biadjoint + + unit-left-adjoint-is-transposing-biadjoint : + id ⇒▵ q* ∘ map-left-adjoint-is-transposing-biadjoint + unit-left-adjoint-is-transposing-biadjoint = + unit-is-transposing-adjunction + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint + + counit-left-adjoint-is-transposing-biadjoint : + map-left-adjoint-is-transposing-biadjoint ∘ q* ⇒▵ id + counit-left-adjoint-is-transposing-biadjoint = + counit-is-transposing-adjunction + is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint + + map-right-adjoint-is-transposing-biadjoint : B → A + map-right-adjoint-is-transposing-biadjoint = + pr1 is-transposing-left-adjoint-is-transposing-biadjoint + + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint : + is-transposing-adjunction q* map-right-adjoint-is-transposing-biadjoint + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint = + pr2 is-transposing-left-adjoint-is-transposing-biadjoint + + unit-right-adjoint-is-transposing-biadjoint : + id ⇒▵ map-right-adjoint-is-transposing-biadjoint ∘ q* + unit-right-adjoint-is-transposing-biadjoint = + unit-is-transposing-adjunction + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint + + counit-right-adjoint-is-transposing-biadjoint : + q* ∘ map-right-adjoint-is-transposing-biadjoint ⇒▵ id + counit-right-adjoint-is-transposing-biadjoint = + counit-is-transposing-adjunction + is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint +``` + +### Transposing biadjunctions + +```text +transposing-biadjunction : {l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2) +transposing-biadjunction A B = Σ (A → B) is-transposing-left-adjoint + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (H : transposing-adjunction A B) + where + + map-left-adjoint-transposing-adjunction : A → B + map-left-adjoint-transposing-adjunction = pr1 H + + is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction : + is-transposing-left-adjoint map-left-adjoint-transposing-adjunction + is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction = pr2 H + + map-right-adjoint-transposing-adjunction : B → A + map-right-adjoint-transposing-adjunction = + pr1 is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction + + is-transposing-adjunction-transposing-adjunction : + is-transposing-adjunction + map-left-adjoint-transposing-adjunction + map-right-adjoint-transposing-adjunction + is-transposing-adjunction-transposing-adjunction = + pr2 is-transposing-left-adjoint-map-left-adjoint-transposing-adjunction + + is-transposing-right-adjoint-map-right-adjoint-transposing-adjunction : + is-transposing-right-adjoint map-right-adjoint-transposing-adjunction + is-transposing-right-adjoint-map-right-adjoint-transposing-adjunction = + ( map-left-adjoint-transposing-adjunction , + is-transposing-adjunction-transposing-adjunction) +``` + +## Properties + +### The identity function is a transposing biadjoint + +```agda +module _ + {l : Level} {A : UU l} + where + + is-transposing-biadjoint-id : is-transposing-biadjoint (id {A = A}) + is-transposing-biadjoint-id = + ( is-transposing-adjoint-id' , is-transposing-adjoint-id) +``` + +### Equivalences are transposing biadjoints + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} (H : is-equiv f) + where + + is-transposing-biadjoint-is-equiv : is-transposing-biadjoint f + is-transposing-biadjoint-is-equiv = + ( is-transposing-right-adjoint-is-equiv H , + is-transposing-left-adjoint-is-equiv H) +``` + +### Composition of transposing biadjoints + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-transposing-biadjoint-comp : + {p : A → B} {q : B → C} → + is-transposing-biadjoint q → + is-transposing-biadjoint p → + is-transposing-biadjoint (q ∘ p) + is-transposing-biadjoint-comp (Q! , Q∗) (P! , P∗) = + ( is-transposing-right-adjoint-comp Q! P! , + is-transposing-left-adjoint-comp Q∗ P∗) +``` + +### Dependent products of transposing biadjoints + +```agda +module _ + {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} + where + + is-transposing-biadjoint-Π : + {q : (i : I) → A i → B i} → + ((i : I) → is-transposing-biadjoint (q i)) → + is-transposing-biadjoint (map-Π q) + is-transposing-biadjoint-Π H = + ( is-transposing-right-adjoint-Π + ( is-transposing-right-adjoint-is-transposing-biadjoint ∘ H)) , + ( is-transposing-left-adjoint-Π + ( is-transposing-left-adjoint-is-transposing-biadjoint ∘ H)) +``` + +### Postcomposition by transposing biadjoints + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + is-transposing-biadjoint-postcomp : + {l : Level} (X : UU l) {q : A → B} → + is-transposing-biadjoint q → + is-transposing-biadjoint (postcomp X q) + is-transposing-biadjoint-postcomp X H = + is-transposing-biadjoint-Π (λ _ → H) +``` + +### The composite transposing adjunction associated to a transposing biadjunction + +Given a transposing biadjunction `q! ⊣ q* ⊣ q∗`, then we have a further +adjunction + +```text + q*q! ⊣ q*q∗. +``` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {q* : A → B} + (H : is-transposing-biadjoint q*) + where + + is-transposing-adjunction-middle-comp-is-transposing-biadjoint : + is-transposing-adjunction + ( q* ∘ map-left-adjoint-is-transposing-biadjoint H) + ( q* ∘ map-right-adjoint-is-transposing-biadjoint H) + is-transposing-adjunction-middle-comp-is-transposing-biadjoint x y = + ( is-transposing-adjoint-map-left-adjoint-is-transposing-biadjoint H + ( x) + ( map-right-adjoint-is-transposing-biadjoint H y)) ∘e + ( is-transposing-adjoint-map-right-adjoint-is-transposing-biadjoint H + ( map-left-adjoint-is-transposing-biadjoint H x) + ( y)) +``` + +### Retracts of transposing biadjoints are transposing biadjoints + +> This remains to be formalized. + +## References + +{{#bibliography}} diff --git a/src/simplicial-type-theory/universal-property-directed-circle.lagda.md b/src/simplicial-type-theory/universal-property-directed-circle.lagda.md new file mode 100644 index 0000000000..abbf158488 --- /dev/null +++ b/src/simplicial-type-theory/universal-property-directed-circle.lagda.md @@ -0,0 +1,331 @@ +# The universal property of the directed circle + +```agda +module simplicial-type-theory.universal-property-directed-circle where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.action-on-identifications-functions +open import foundation.constant-type-families +open import foundation.contractible-maps +open import foundation.contractible-types +open import foundation.dependent-identifications +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.fibers-of-maps +open import foundation.function-extensionality +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.strictly-right-unital-concatenation-identifications +open import foundation.transport-along-identifications +open import foundation.universe-levels + +open import simplicial-type-theory.action-on-directed-edges-dependent-functions +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.free-directed-loops +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Definitions + +### Evaluating an ordinary function at a free directed loop + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (Y : UU l2) + where + + ev-free-directed-loop : (X → Y) → free-directed-loop Y + ev-free-directed-loop f = + ( f ∘ arrow-free-directed-loop α , + ap f (compute-target-free-directed-loop α)) +``` + +### The recursion principle of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + recursion-principle-directed-circle : UUω + recursion-principle-directed-circle = + {l : Level} (Y : UU l) → section (ev-free-directed-loop α Y) +``` + +### The universal property of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + universal-property-directed-circle : UUω + universal-property-directed-circle = + {l : Level} (Y : UU l) → is-equiv (ev-free-directed-loop α Y) +``` + +### Evaluating a dependent function at a free directed loop + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (P : X → UU l2) + where + + ev-free-directed-loop-Π : ((x : X) → P x) → free-dependent-directed-loop α P + ev-free-directed-loop-Π f = + ( f ∘ arrow-free-directed-loop α , + apd f (compute-target-free-directed-loop α)) +``` + +### The induction principle of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + induction-principle-directed-circle : UUω + induction-principle-directed-circle = + {l : Level} (P : X → UU l) → section (ev-free-directed-loop-Π α P) + +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) + (H : induction-principle-directed-circle α) (P : X → UU l2) + (β : free-dependent-directed-loop α P) + where + + function-induction-principle-directed-circle : (x : X) → P x + function-induction-principle-directed-circle = pr1 (H P) β + + compute-induction-principle-directed-circle : + ev-free-directed-loop-Π α P function-induction-principle-directed-circle = β + compute-induction-principle-directed-circle = pr2 (H P) β +``` + +### The dependent universal property of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + dependent-universal-property-directed-circle : UUω + dependent-universal-property-directed-circle = + {l : Level} (P : X → UU l) → is-equiv (ev-free-directed-loop-Π α P) +``` + +## Properties + +### The induction principle of the directed circle implies the dependent universal property of the directed circle + +To prove this, we have to show that the section of `ev-free-directed-loop-Π` is +also a retraction. This construction is also by the induction principle of the +directed circle, but it requires (a minimal amount of) preparations. + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + free-dependent-directed-loop-htpy : + {l2 : Level} {P : X → UU l2} {f g : (x : X) → P x} → + ( htpy-free-dependent-directed-loop α P + ( ev-free-directed-loop-Π α P f) + ( ev-free-directed-loop-Π α P g)) → + ( free-dependent-directed-loop α (λ x → f x = g x)) + free-dependent-directed-loop-htpy {l2} {P} {f} {g} (p , q) = + ( p , + map-compute-dependent-identification-eq-value f g + ( compute-target-free-directed-loop α) + ( p 1₂) + ( p 0₂) + ( equational-reasoning + apd f (compute-target-free-directed-loop α) ∙ p 0₂ + = apd f (compute-target-free-directed-loop α) ∙ᵣ p 0₂ + by + eq-right-strict-concat-concat + ( apd f (compute-target-free-directed-loop α)) + ( p 0₂) + = concat-dependent-identification P + ( refl) + ( compute-target-free-directed-loop α) + ( p 1₂) + ( apd g (compute-target-free-directed-loop α)) + by q + = ap (tr P (compute-target-free-directed-loop α)) (p 1₂) ∙ + apd g (compute-target-free-directed-loop α) + by + compute-concat-dependent-identification-left-base-refl P + ( compute-target-free-directed-loop α) + ( p 1₂) + ( apd g (compute-target-free-directed-loop α)))) + + is-retraction-ind-directed-circle : + ( I : induction-principle-directed-circle α) + { l2 : Level} (P : X → UU l2) → + is-retraction + ( ev-free-directed-loop-Π α P) + ( function-induction-principle-directed-circle α I P) + is-retraction-ind-directed-circle I P f = + eq-htpy + ( function-induction-principle-directed-circle α I + ( eq-value + ( function-induction-principle-directed-circle α I P + ( ev-free-directed-loop-Π α P f)) + ( f)) + ( free-dependent-directed-loop-htpy + ( htpy-free-dependent-directed-loop-eq α P _ _ + ( compute-induction-principle-directed-circle α I P + ( ev-free-directed-loop-Π α P f))))) + + abstract + dependent-universal-property-induction-principle-directed-circle : + induction-principle-directed-circle α → + dependent-universal-property-directed-circle α + dependent-universal-property-induction-principle-directed-circle I P = + is-equiv-is-invertible + ( function-induction-principle-directed-circle α I P) + ( compute-induction-principle-directed-circle α I P) + ( is-retraction-ind-directed-circle I P) +``` + +### Uniqueness of the maps obtained from the universal property of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + abstract + uniqueness-universal-property-directed-circle : + universal-property-directed-circle α → + {l2 : Level} (Y : UU l2) (α' : free-directed-loop Y) → + is-contr + ( Σ ( X → Y) + ( λ f → htpy-free-directed-loop (ev-free-directed-loop α Y f) α')) + uniqueness-universal-property-directed-circle I Y α' = + is-contr-is-equiv' + ( fiber (ev-free-directed-loop α Y) α') + ( tot + ( λ f → htpy-eq-free-directed-loop (ev-free-directed-loop α Y f) α')) + ( is-equiv-tot-is-fiberwise-equiv + ( λ f → + is-equiv-htpy-eq-free-directed-loop + ( ev-free-directed-loop α Y f) + ( α'))) + ( is-contr-map-is-equiv (I Y) α') +``` + +### Uniqueness of the dependent functions obtained from the dependent universal property of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + uniqueness-dependent-universal-property-directed-circle : + dependent-universal-property-directed-circle α → + {l2 : Level} {P : X → UU l2} (k : free-dependent-directed-loop α P) → + is-contr + ( Σ ( (x : X) → P x) + ( λ h → + htpy-free-dependent-directed-loop α P + ( ev-free-directed-loop-Π α P h) + ( k))) + uniqueness-dependent-universal-property-directed-circle I {P = P} k = + is-contr-is-equiv' + ( fiber (ev-free-directed-loop-Π α P) k) + ( tot + ( λ h → + htpy-free-dependent-directed-loop-eq α P + ( ev-free-directed-loop-Π α P h) + ( k))) + ( is-equiv-tot-is-fiberwise-equiv + ( λ h → + is-equiv-htpy-free-dependent-directed-loop-eq α P + ( ev-free-directed-loop-Π α P h) + ( k))) + ( is-contr-map-is-equiv (I P) k) +``` + +### The dependent universal property of the directed circle implies the universal property of the directed circle + +```agda +module _ + {l1 l2 : Level} {X : UU l1} (α : free-directed-loop X) (Y : UU l2) + where + + triangle-comparison-free-directed-loop : + ev-free-directed-loop-Π α (λ _ → Y) ~ + map-compute-free-dependent-directed-loop-constant-type-family α Y ∘ + ev-free-directed-loop α Y + triangle-comparison-free-directed-loop f = + eq-htpy-free-dependent-directed-loop α + ( λ _ → Y) + ( ev-free-directed-loop-Π α (λ _ → Y) f) + ( map-compute-free-dependent-directed-loop-constant-type-family α Y + ( ev-free-directed-loop α Y f)) + ( ( refl-htpy) , + ( apd-constant-type-family f (compute-target-free-directed-loop α))) + +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + abstract + universal-property-dependent-universal-property-directed-circle : + dependent-universal-property-directed-circle α → + universal-property-directed-circle α + universal-property-dependent-universal-property-directed-circle I Y = + is-equiv-top-map-triangle + ( ev-free-directed-loop-Π α (λ _ → Y)) + ( map-compute-free-dependent-directed-loop-constant-type-family α Y) + ( ev-free-directed-loop α Y) + ( triangle-comparison-free-directed-loop α Y) + ( is-equiv-map-equiv + ( compute-free-dependent-directed-loop-constant-type-family α Y)) + ( I (λ _ → Y)) +``` + +### The induction principle of the directed circle implies the universal property of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + abstract + universal-property-induction-principle-directed-circle : + induction-principle-directed-circle α → + universal-property-directed-circle α + universal-property-induction-principle-directed-circle I = + universal-property-dependent-universal-property-directed-circle α + ( dependent-universal-property-induction-principle-directed-circle α I) +``` + +### The induction principle of the directed circle implies ther recursion principle of the directed circle + +```agda +module _ + {l1 : Level} {X : UU l1} (α : free-directed-loop X) + where + + recursion-principle-induction-principle-directed-circle : + induction-principle-directed-circle α → + recursion-principle-directed-circle α + recursion-principle-induction-principle-directed-circle I Y = + section-is-equiv + ( universal-property-induction-principle-directed-circle α I Y) +``` diff --git a/src/simplicial-type-theory/whiskering-directed-edges.lagda.md b/src/simplicial-type-theory/whiskering-directed-edges.lagda.md new file mode 100644 index 0000000000..7f911be8e3 --- /dev/null +++ b/src/simplicial-type-theory/whiskering-directed-edges.lagda.md @@ -0,0 +1,256 @@ +# Whiskering directed edges + +```agda +module simplicial-type-theory.whiskering-directed-edges where +``` + +
Imports + +```agda +open import foundation.cartesian-product-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-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.retractions +open import foundation.retracts-of-types +open import foundation.sections +open import foundation.strictly-right-unital-concatenation-identifications +open import foundation.structure-identity-principle +open import foundation.torsorial-type-families +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.action-on-directed-edges-functions +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a [directed edge](simplicial-type-theory.directed-edges.md) `α : x →₂ y` +in `A` and an [identification](foundation-core.identity-types.md) `p : y = z`, +we may {#concept "whisker" Disambiguation="directed edge by identification"} `α` +by `p` to obtain a directed edge `x →₂ z`. + +## Definitions + +### Whiskering at the target of a directed edge + +```agda +module _ + {l : Level} {A : UU l} + where + + whisker-target-simplicial-hom : + {x y z : A} → y = z → x →₂ y → x →₂ z + whisker-target-simplicial-hom p α = + ( simplicial-arrow-simplicial-hom α , + eq-source-simplicial-hom α , + eq-target-simplicial-hom α ∙ᵣ p) +``` + +### Whiskering at the source of a directed edge + +```agda +module _ + {l : Level} {A : UU l} + where + + whisker-source-simplicial-hom : + {x y z : A} → x = z → x →₂ y → z →₂ y + whisker-source-simplicial-hom p α = + ( simplicial-arrow-simplicial-hom α , + eq-source-simplicial-hom α ∙ᵣ p , + eq-target-simplicial-hom α) +``` + +### Whiskering both end points of a directed edge + +```agda +module _ + {l : Level} {A : UU l} + where + + double-whisker-simplicial-hom : + {x y x' y' : A} → x = x' → y = y' → x →₂ y → x' →₂ y' + double-whisker-simplicial-hom p q α = + ( simplicial-arrow-simplicial-hom α , + eq-source-simplicial-hom α ∙ᵣ p , + eq-target-simplicial-hom α ∙ᵣ q) +``` + +## Properties + +### Whiskering directed edges is an equivalence + +```agda +module _ + {l : Level} {A : UU l} + where + + is-equiv-double-whisker-simplicial-hom : + {x y x' y' : A} (p : x = x') (q : y = y') → + is-equiv (double-whisker-simplicial-hom p q) + is-equiv-double-whisker-simplicial-hom refl refl = is-equiv-id + + is-equiv-whisker-target-simplicial-hom : + {x y z : A} (p : y = z) → + is-equiv (whisker-target-simplicial-hom {x = x} p) + is-equiv-whisker-target-simplicial-hom p = + is-equiv-double-whisker-simplicial-hom refl p + + is-equiv-whisker-source-simplicial-hom : + {x y z : A} (p : x = z) → + is-equiv (whisker-source-simplicial-hom {y = y} p) + is-equiv-whisker-source-simplicial-hom p = + is-equiv-double-whisker-simplicial-hom p refl +``` + +### Naturality of homotopies with respect to directed edges + +Given two maps `f g : A → B` and a homotopy `H : f ~ g`, then for every directed +edge `p : x →₂ y` in `A`, we have a commuting square + +```text + ap▵ f p + f x --------> f y + ║ ║ + H x ║ ║ H y + ║ ║ + g x --------> g y. + ap▵ g p +``` + +```agda +nat-htpy▵ : + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} (H : f ~ g) + {x y : A} (α : x →₂ y) → + action-simplicial-hom-function g α = + double-whisker-simplicial-hom (H x) (H y) (action-simplicial-hom-function f α) +nat-htpy▵ {A = A} {f = f} {g} H {x} {y} α = + ind-htpy f + ( λ g H → + action-simplicial-hom-function g α = + double-whisker-simplicial-hom (H x) (H y) + ( action-simplicial-hom-function f α)) + ( refl) + ( H) +``` + +Note that this proof only relies on +[function extensionality](foundation.function-extensionality.md) to get +extensionality for +[simplicial arrows](simplicial-type-theory.simplicial-arrows.md). + +### For any map `i : A → B`, a retraction of `i` induces a retraction of the action on identifications of `i` + +**Proof:** Suppose that `H : r ∘ i ~ id` and `p : i x = i y` is an +identification in `B`. Then we get the directed edge + +```text + H x ap▵ r p H y + x ===== r (i x) --------> r (i y) ===== y +``` + +This defines a map `s : (i x →₂ i y) → (x →₂ y)`. To see that `s ∘ ap▵ i ~ id`, +i.e., that the whiskering + +```text + H x ap▵ r (ap▵ i p) H y + x ===== r (i x) ----------------> r (i y) ===== y +``` + +is identical to `p : x = y` for all `p : x = y`, we proceed by identification +elimination. Then it suffices to show that `(H x)⁻¹ ∙ (H x)` is identical to +`refl`, which is indeed the case by the left inverse law of concatenation of +identifications. + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (i : A → B) + (r : B → A) (H : r ∘ i ~ id) + where + + is-hom-injective-has-retraction : {x y : A} → (i x →₂ i y) → (x →₂ y) + is-hom-injective-has-retraction {x} {y} p = + double-whisker-simplicial-hom + ( H x) + ( H y) + ( action-simplicial-hom-function r p) + + is-retraction-is-hom-injective-has-retraction' : + {x y : A} (α : x →₂ y) → + htpy-simplicial-hom + ( is-hom-injective-has-retraction (action-simplicial-hom-function i α)) + ( α) + pr1 (is-retraction-is-hom-injective-has-retraction' (α , p , q)) = + H ·r α + pr1 (pr2 (is-retraction-is-hom-injective-has-retraction' (α , refl , q))) = + left-unit-right-strict-concat ∙ inv right-unit + pr2 (pr2 (is-retraction-is-hom-injective-has-retraction' (α , p , refl))) = + left-unit-right-strict-concat ∙ inv right-unit + + is-retraction-is-hom-injective-has-retraction : + {x y : A} → + is-retraction + ( action-simplicial-hom-function i {x} {y}) + ( is-hom-injective-has-retraction) + is-retraction-is-hom-injective-has-retraction α = + eq-htpy-simplicial-hom + ( is-hom-injective-has-retraction (action-simplicial-hom-function i α)) + ( α) + ( is-retraction-is-hom-injective-has-retraction' α) + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (i : A → B) (R : retraction i) + where + + is-hom-injective-retraction : + {x y : A} → i x →₂ i y → x →₂ y + is-hom-injective-retraction = + is-hom-injective-has-retraction i + ( map-retraction i R) + ( is-retraction-map-retraction i R) + + is-retraction-is-hom-injective-retraction : + {x y : A} → + is-hom-injective-retraction ∘ action-simplicial-hom-function i {x} {y} ~ id + is-retraction-is-hom-injective-retraction = + is-retraction-is-hom-injective-has-retraction i + ( map-retraction i R) + ( is-retraction-map-retraction i R) + + retraction-ap▵ : + {x y : A} → retraction (action-simplicial-hom-function i {x} {y}) + pr1 retraction-ap▵ = is-hom-injective-retraction + pr2 retraction-ap▵ = is-retraction-is-hom-injective-retraction +``` + +### If `A` is a retract of `B` with inclusion `i : A → B`, then `x →₂ y` is a retract of `i x →₂ i y` for any two elements `x y : A` + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (R : A retract-of B) (x y : A) + where + + retract-simplicial-hom : + (x →₂ y) retract-of (inclusion-retract R x →₂ inclusion-retract R y) + pr1 retract-simplicial-hom = + action-simplicial-hom-function (inclusion-retract R) + pr2 retract-simplicial-hom = + retraction-ap▵ (inclusion-retract R) (retraction-retract R) +``` diff --git a/src/simplicial-type-theory/whiskering-simplicial-arrows-functions.lagda.md b/src/simplicial-type-theory/whiskering-simplicial-arrows-functions.lagda.md new file mode 100644 index 0000000000..5068e15de5 --- /dev/null +++ b/src/simplicial-type-theory/whiskering-simplicial-arrows-functions.lagda.md @@ -0,0 +1,136 @@ +# Whiskering operations on simplicial arrows of functions + +```agda +module simplicial-type-theory.whiskering-simplicial-arrows-functions where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.horizontal-composition-simplicial-arrows-functions +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a simplicial arrow `α` of functions `A → B` we may whisker it on the left +by a function `f : B → C` to obtain a simplicial arrow of functions `A → C`, or +we may whisker it on the right by a function `g : C → A` to obtain a simplicial +arrow of functions `C → B`. + +## Definitions + +### Left whiskering simplicial arrows of functions by functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + left-whisker-comp-simplicial-arrow : + (B → C) → simplicial-arrow (A → B) → simplicial-arrow (A → C) + left-whisker-comp-simplicial-arrow f = + horizontal-comp-simplicial-arrow (id-simplicial-arrow f) +``` + +### Right whiskering simplicial arrows of functions by functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + right-whisker-comp-simplicial-arrow : + simplicial-arrow (B → C) → (A → B) → simplicial-arrow (A → C) + right-whisker-comp-simplicial-arrow β g = + horizontal-comp-simplicial-arrow β (id-simplicial-arrow g) +``` + +## Properties + +### Unit laws of whiskering of simplicial arrows + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + left-unit-law-left-whisker-comp-simplicial-arrow : + (α : simplicial-arrow (A → B)) → + left-whisker-comp-simplicial-arrow id α = α + left-unit-law-left-whisker-comp-simplicial-arrow = + left-unit-law-horizontal-comp-simplicial-arrow + + right-unit-law-right-whisker-comp-simplicial-arrow : + (α : simplicial-arrow (A → B)) → + right-whisker-comp-simplicial-arrow α id = α + right-unit-law-right-whisker-comp-simplicial-arrow = + right-unit-law-horizontal-comp-simplicial-arrow +``` + +### Absorption laws of whiskering of simplicial arrows + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + right-absorption-law-left-whisker-comp-simplicial-arrow : + (g : B → C) (f : A → B) → + left-whisker-comp-simplicial-arrow g (id-simplicial-arrow f) = + id-simplicial-arrow (g ∘ f) + right-absorption-law-left-whisker-comp-simplicial-arrow g f = refl + + left-absorption-law-right-whisker-comp-simplicial-arrow : + (g : B → C) (f : A → B) → + right-whisker-comp-simplicial-arrow (id-simplicial-arrow g) f = + id-simplicial-arrow (g ∘ f) + left-absorption-law-right-whisker-comp-simplicial-arrow g f = refl +``` + +### Whiskering of simplicial arrows between functions preserves function composition + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + where + + preserves-comp-left-whisker-comp-simplicial-arrow : + (h : C → D) (g : B → C) (α : simplicial-arrow (A → B)) → + left-whisker-comp-simplicial-arrow (h ∘ g) α = + left-whisker-comp-simplicial-arrow h + ( left-whisker-comp-simplicial-arrow g α) + preserves-comp-left-whisker-comp-simplicial-arrow h g α = refl + + preserves-comp-right-whisker-comp-simplicial-arrow : + (α : simplicial-arrow (C → D)) (g : B → C) (f : A → B) → + right-whisker-comp-simplicial-arrow α (g ∘ f) = + right-whisker-comp-simplicial-arrow + ( right-whisker-comp-simplicial-arrow α g) + ( f) + preserves-comp-right-whisker-comp-simplicial-arrow h g α = refl +``` diff --git a/src/simplicial-type-theory/whiskering-simplicial-edges-functions.lagda.md b/src/simplicial-type-theory/whiskering-simplicial-edges-functions.lagda.md new file mode 100644 index 0000000000..373d9be9ad --- /dev/null +++ b/src/simplicial-type-theory/whiskering-simplicial-edges-functions.lagda.md @@ -0,0 +1,137 @@ +# Whiskering operations on directed edges between functions + +```agda +module simplicial-type-theory.whiskering-simplicial-edges-functions where +``` + +
Imports + +```agda +open import asimplicial-type-theory.horizontal-composition-directed-edges-functions + +open import foundation.action-on-identifications-functions +open import foundation.cartesian-product-types +open import foundation.dependent-pair-types +open import foundation.equality-cartesian-product-types +open import foundation.equality-dependent-pair-types +open import foundation.equivalences +open import foundation.function-extensionality +open import foundation.function-types +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-dependent-pair-types +open import foundation.homotopies +open import foundation.identity-types +open import foundation.retractions +open import foundation.sections +open import foundation.type-arithmetic-dependent-function-types +open import foundation.type-theoretic-principle-of-choice +open import foundation.universe-levels + +open import orthogonal-factorization-systems.extensions-of-maps + +open import simplicial-type-theory.directed-edges +open import simplicial-type-theory.directed-interval-type +open import simplicial-type-theory.horizontal-composition-simplicial-arrows-functions +open import simplicial-type-theory.simplicial-arrows +``` + +
+ +## Idea + +Given a directed edge `α : f →₂ f'` of functions `A → B` we may whisker it on +the left by a function `h : B → C` to obtain a directed edge of functions +`hα : hf →₂ hf'`, or we may whisker it on the right by a function `g : C → A` to +obtain a directed edge of functions `αg : fg →₂ f'g`. + +## Definitions + +### Left whiskering directed edges between functions by functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + left-whisker-comp-simplicial-hom : + (h : B → C) {f f' : A → B} → f →₂ f' → h ∘ f →₂ h ∘ f' + left-whisker-comp-simplicial-hom h = + horizontal-comp-simplicial-hom (id-simplicial-hom h) +``` + +### Right whiskering directed edges between functions by functions + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + right-whisker-comp-simplicial-hom : + {f f' : A → B} → f →₂ f' → (g : C → A) → f ∘ g →₂ f' ∘ g + right-whisker-comp-simplicial-hom α g = + horizontal-comp-simplicial-hom α (id-simplicial-hom g) +``` + +## Properties + +### Unit laws of whiskering of directed edges + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where + + left-unit-law-left-whisker-comp-simplicial-hom : + {f f' : A → B} (α : f →₂ f') → + left-whisker-comp-simplicial-hom id α = α + left-unit-law-left-whisker-comp-simplicial-hom = + left-unit-law-horizontal-comp-simplicial-hom + + right-unit-law-right-whisker-comp-simplicial-hom : + {f f' : A → B} (α : f →₂ f') → + right-whisker-comp-simplicial-hom α id = α + right-unit-law-right-whisker-comp-simplicial-hom = + right-unit-law-horizontal-comp-simplicial-hom +``` + +### Absorption laws of whiskering of directed edges + +These laws hold strictly. + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} + where + + right-absorption-law-left-whisker-comp-simplicial-hom : + (g : B → C) (f : A → B) → + left-whisker-comp-simplicial-hom g (id-simplicial-hom f) = + id-simplicial-hom (g ∘ f) + right-absorption-law-left-whisker-comp-simplicial-hom g f = refl + + left-absorption-law-right-whisker-comp-simplicial-hom : + (g : B → C) (f : A → B) → + right-whisker-comp-simplicial-hom (id-simplicial-hom g) f = + id-simplicial-hom (g ∘ f) + left-absorption-law-right-whisker-comp-simplicial-hom g f = refl +``` + +### Whiskering of directed edges between functions preserves function composition + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4} + where + + preserves-comp-left-whisker-comp-simplicial-hom : + (h : C → D) (g : B → C) {f f' : A → B} (α : f →₂ f') → + left-whisker-comp-simplicial-hom (h ∘ g) α = + left-whisker-comp-simplicial-hom h (left-whisker-comp-simplicial-hom g α) + preserves-comp-left-whisker-comp-simplicial-hom h g (α , refl , refl) = refl + + preserves-comp-right-whisker-comp-simplicial-hom : + {h h' : C → D} (α : h →₂ h') (g : B → C) (f : A → B) → + right-whisker-comp-simplicial-hom α (g ∘ f) = + right-whisker-comp-simplicial-hom (right-whisker-comp-simplicial-hom α g) f + preserves-comp-right-whisker-comp-simplicial-hom (α , refl , refl) g f = refl +``` diff --git a/src/synthetic-homotopy-theory/circle.lagda.md b/src/synthetic-homotopy-theory/circle.lagda.md index 4b11a4c0c1..3028b382fa 100644 --- a/src/synthetic-homotopy-theory/circle.lagda.md +++ b/src/synthetic-homotopy-theory/circle.lagda.md @@ -19,6 +19,7 @@ open import foundation.equivalences open import foundation.homotopies open import foundation.identity-types open import foundation.mere-equality +open import foundation.negated-equality open import foundation.propositional-truncations open import foundation.propositions open import foundation.retractions @@ -36,6 +37,7 @@ open import synthetic-homotopy-theory.free-loops open import synthetic-homotopy-theory.spheres open import synthetic-homotopy-theory.suspension-structures open import synthetic-homotopy-theory.suspensions-of-types +open import synthetic-homotopy-theory.universal-cover-circle open import synthetic-homotopy-theory.universal-property-circle open import univalent-combinatorics.standard-finite-types @@ -170,6 +172,16 @@ module _ pr2 (pr2 apply-universal-property-𝕊¹) ``` +### The loop of the circle is nontrivial + +```agda +is-nontrivial-loop-𝕊¹ : loop-𝕊¹ ≠ refl +is-nontrivial-loop-𝕊¹ = + is-nontrivial-loop-dependent-universal-property-circle + ( free-loop-𝕊¹) + ( dependent-universal-property-𝕊¹) +``` + ### The circle is 0-connected ```agda diff --git a/src/synthetic-homotopy-theory/dependent-pushout-products.lagda.md b/src/synthetic-homotopy-theory/dependent-pushout-products.lagda.md index 265ce6a2c9..4dea3dfe0f 100644 --- a/src/synthetic-homotopy-theory/dependent-pushout-products.lagda.md +++ b/src/synthetic-homotopy-theory/dependent-pushout-products.lagda.md @@ -23,11 +23,12 @@ open import synthetic-homotopy-theory.universal-property-pushouts ## Idea -The **dependent pushout-product** is a generalization of the +The _dependent pushout-product_ is a generalization of the [pushout-product](synthetic-homotopy-theory.pushout-products.md). Consider a map -`f : A → B` and a family of maps `g : (x : X) → B x → Y x`. The **dependent -pushout-product** is the [cogap map](synthetic-homotopy-theory.pushouts.md) of -the [commuting square](foundation-core.commuting-squares-of-maps.md) +`f : A → B` and a family of maps `g : (x : X) → B x → Y x`. The +{{#concept "dependent +pushout-product" Disambiguation="of maps of types" Agda=dependent-pushout-product}} +is the [cogap map](synthetic-homotopy-theory.pushouts.md) of the [commuting square](foundation-core.commuting-squares-of-maps.md) ```text Σ f id diff --git a/src/synthetic-homotopy-theory/descent-circle.lagda.md b/src/synthetic-homotopy-theory/descent-circle.lagda.md index 8adc5c7904..edb7ac84a1 100644 --- a/src/synthetic-homotopy-theory/descent-circle.lagda.md +++ b/src/synthetic-homotopy-theory/descent-circle.lagda.md @@ -51,8 +51,7 @@ way made precise in further sections of this file. The pair `(X, e)` is called **descent data** for the circle. ```agda -descent-data-circle : - ( l1 : Level) → UU (lsuc l1) +descent-data-circle : (l : Level) → UU (lsuc l) descent-data-circle = Type-With-Automorphism module _ diff --git a/src/synthetic-homotopy-theory/free-loops.lagda.md b/src/synthetic-homotopy-theory/free-loops.lagda.md index 9abf11b695..2ca408cc36 100644 --- a/src/synthetic-homotopy-theory/free-loops.lagda.md +++ b/src/synthetic-homotopy-theory/free-loops.lagda.md @@ -174,14 +174,14 @@ module _ {l1 l2 : Level} {X : UU l1} (α : free-loop X) (Y : UU l2) where - compute-free-dependent-loop-const : + compute-free-dependent-loop-constant-type-family : free-loop Y ≃ free-dependent-loop α (λ x → Y) - compute-free-dependent-loop-const = + compute-free-dependent-loop-constant-type-family = equiv-tot ( λ y → equiv-concat (tr-constant-type-family (loop-free-loop α) y) y) - map-compute-free-dependent-loop-const : + map-compute-free-dependent-loop-constant-type-family : free-loop Y → free-dependent-loop α (λ x → Y) - map-compute-free-dependent-loop-const = - map-equiv compute-free-dependent-loop-const + map-compute-free-dependent-loop-constant-type-family = + map-equiv compute-free-dependent-loop-constant-type-family ``` diff --git a/src/synthetic-homotopy-theory/joins-of-types.lagda.md b/src/synthetic-homotopy-theory/joins-of-types.lagda.md index 8b6ea38172..ddfb156c03 100644 --- a/src/synthetic-homotopy-theory/joins-of-types.lagda.md +++ b/src/synthetic-homotopy-theory/joins-of-types.lagda.md @@ -434,6 +434,18 @@ module _ ( up-join) ``` +### The propositional recursor for joins of types + +```agda +rec-join-Prop : + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (R : Prop l3) → + (A → type-Prop R) → (B → type-Prop R) → A * B → type-Prop R +rec-join-Prop R f g = + cogap-join + ( type-Prop R) + ( f , g , λ (t , s) → eq-is-prop' (is-prop-type-Prop R) (f t) (g s)) +``` + ## See also - [Joins of maps](synthetic-homotopy-theory.joins-of-maps.md) diff --git a/src/synthetic-homotopy-theory/pullback-property-pushouts.lagda.md b/src/synthetic-homotopy-theory/pullback-property-pushouts.lagda.md index 888f86f5ed..a447e3b223 100644 --- a/src/synthetic-homotopy-theory/pullback-property-pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/pullback-property-pushouts.lagda.md @@ -76,3 +76,8 @@ pullback-property-pushout f g c = ( precomp g Y) ( cone-pullback-property-pushout f g c Y) ``` + +## See also + +- For the dual property for [pullbacks](foundation-core.pullbacks.md), see + [postcomposition of pullbacks](foundation.postcomposition-pullbacks.md). diff --git a/src/synthetic-homotopy-theory/pushout-products.lagda.md b/src/synthetic-homotopy-theory/pushout-products.lagda.md index 5756497b68..a889d692c7 100644 --- a/src/synthetic-homotopy-theory/pushout-products.lagda.md +++ b/src/synthetic-homotopy-theory/pushout-products.lagda.md @@ -80,10 +80,11 @@ from the [fibers](foundation-core.fibers-of-maps.md) of `f □ g` to the There is an "adjoint relation" between the pushout-product and the [pullback-hom](orthogonal-factorization-systems.pullback-hom.md): For any three -maps `f`, `g`, and `h` we have a [homotopy](foundation-core.homotopies.md) +maps `f`, `g`, and `h` we have an +[equivalence of maps](foundation.equivalences-arrows.md) ```text - ⟨f □ g , h⟩ ~ ⟨f , ⟨g , h⟩⟩. + ⟨f □ g , h⟩ ≃ ⟨f , ⟨g , h⟩⟩. ``` ## Definitions @@ -159,6 +160,24 @@ module _ pr2 (pr2 (pr2 (center uniqueness-pushout-product))) ``` +## Properties + +### The adjoint relation between pushout-products and pullback-homs + +For any three maps `f`, `g`, and `h` we have an adjoint relation + +```text + hom-arrow (f □ g) h ≃ hom-arrow f ⟨g , h⟩ +``` + +that extends to an equivalence of maps + +```text + ⟨f □ g , h⟩ ≃ ⟨f , ⟨g , h⟩⟩. +``` + +This remains to be formalized. + ## See also - [The dependent pushout-product](synthetic-homotopy-theory.dependent-pushout-products.md) diff --git a/src/synthetic-homotopy-theory/universal-cover-circle.lagda.md b/src/synthetic-homotopy-theory/universal-cover-circle.lagda.md index 3ef47aaaae..885f22771f 100644 --- a/src/synthetic-homotopy-theory/universal-cover-circle.lagda.md +++ b/src/synthetic-homotopy-theory/universal-cover-circle.lagda.md @@ -29,6 +29,7 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.identity-types open import foundation.injective-maps +open import foundation.negated-equality open import foundation.negation open import foundation.precomposition-dependent-functions open import foundation.raising-universe-levels @@ -112,7 +113,7 @@ abstract ( f x p₀))) ``` -### The universal cover +### The universal cover of the circle ```agda module _ @@ -618,7 +619,7 @@ module _ where is-nontrivial-loop-dependent-universal-property-circle : - ¬ (loop-free-loop l = refl) + loop-free-loop l ≠ refl is-nontrivial-loop-dependent-universal-property-circle p = is-nonzero-one-ℤ ( is-injective-equiv diff --git a/src/synthetic-homotopy-theory/universal-property-circle.lagda.md b/src/synthetic-homotopy-theory/universal-property-circle.lagda.md index 5c155bafd4..9698cb554f 100644 --- a/src/synthetic-homotopy-theory/universal-property-circle.lagda.md +++ b/src/synthetic-homotopy-theory/universal-property-circle.lagda.md @@ -208,12 +208,13 @@ module _ where triangle-comparison-free-loop : - ( map-compute-free-dependent-loop-const α Y ∘ (ev-free-loop α Y)) ~ - ( ev-free-loop-Π α (λ x → Y)) + map-compute-free-dependent-loop-constant-type-family α Y ∘ + ev-free-loop α Y ~ + ev-free-loop-Π α (λ _ → Y) triangle-comparison-free-loop f = eq-Eq-free-dependent-loop α ( λ x → Y) - ( map-compute-free-dependent-loop-const α Y + ( map-compute-free-dependent-loop-constant-type-family α Y ( ev-free-loop α Y f)) ( ev-free-loop-Π α (λ x → Y) f) ( refl , @@ -230,10 +231,11 @@ module _ universal-property-dependent-universal-property-circle dup-circle Y = is-equiv-top-map-triangle ( ev-free-loop-Π α (λ x → Y)) - ( map-compute-free-dependent-loop-const α Y) + ( map-compute-free-dependent-loop-constant-type-family α Y) ( ev-free-loop α Y) ( inv-htpy (triangle-comparison-free-loop α Y)) - ( is-equiv-map-equiv (compute-free-dependent-loop-const α Y)) + ( is-equiv-map-equiv + ( compute-free-dependent-loop-constant-type-family α Y)) ( dup-circle (λ x → Y)) ```