From 9f7ceb0fd843864de22c33440717880cec54ca14 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 19:05:30 +0200 Subject: [PATCH 01/69] some informal proofs for truncation equivalences --- src/foundation-core/embeddings.lagda.md | 2 +- src/foundation-core/equivalences.lagda.md | 6 ++ src/foundation/connected-maps.lagda.md | 34 ++++++++++ .../truncation-equivalences.lagda.md | 64 +++++++++++++++++-- .../types-local-at-maps.lagda.md | 35 ++++++---- .../infinite-cyclic-types.lagda.md | 5 +- 6 files changed, 123 insertions(+), 23 deletions(-) diff --git a/src/foundation-core/embeddings.lagda.md b/src/foundation-core/embeddings.lagda.md index 9c21226fc1..6b7789043f 100644 --- a/src/foundation-core/embeddings.lagda.md +++ b/src/foundation-core/embeddings.lagda.md @@ -113,7 +113,7 @@ module _ where is-emb-id : is-emb (id {A = A}) - is-emb-id x y = is-equiv-htpy id ap-id is-equiv-id + is-emb-id x y = is-equiv-htpy-id ap-id id-emb : A ↪ A pr1 id-emb = id diff --git a/src/foundation-core/equivalences.lagda.md b/src/foundation-core/equivalences.lagda.md index c1f1c724ad..405a46c3f5 100644 --- a/src/foundation-core/equivalences.lagda.md +++ b/src/foundation-core/equivalences.lagda.md @@ -500,6 +500,12 @@ module _ htpy-map-inv-is-invertible H ( is-invertible-is-equiv F) ( is-invertible-is-equiv G) + +is-equiv-htpy-id : {l : Level} {A : UU l} {f : A → A} → f ~ id → is-equiv f +is-equiv-htpy-id H = is-equiv-htpy id H is-equiv-id + +is-equiv-htpy-id' : {l : Level} {A : UU l} {f : A → A} → id ~ f → is-equiv f +is-equiv-htpy-id' H = is-equiv-htpy' id H is-equiv-id ``` ### Any retraction of an equivalence is an equivalence diff --git a/src/foundation/connected-maps.lagda.md b/src/foundation/connected-maps.lagda.md index 414edbd171..1e4ecfeebe 100644 --- a/src/foundation/connected-maps.lagda.md +++ b/src/foundation/connected-maps.lagda.md @@ -204,6 +204,24 @@ module _ is-neg-two-connected-map b = is-neg-two-connected (fiber f b) ``` +### Connected maps are closed under homotopies + +```agda +module _ + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {f g : A → B} + where + + is-connected-map-htpy : + (f ~ g) → is-connected-map k g → is-connected-map k f + is-connected-map-htpy H G x = + is-connected-equiv' (equiv-fiber-htpy H x) (G x) + + is-connected-map-htpy' : + (f ~ g) → is-connected-map k f → is-connected-map k g + is-connected-map-htpy' H F x = + is-connected-equiv (equiv-fiber-htpy H x) (F x) +``` + ### Equivalences are `k`-connected for any `k` ```agda @@ -227,6 +245,22 @@ module _ pr2 (connected-map-equiv e) = is-connected-map-equiv e ``` +### The identity map is `k`-connected for every `k` + +```agda +is-connected-map-id : + {l : Level} {k : 𝕋} {A : UU l} → is-connected-map k (id' A) +is-connected-map-id = is-connected-map-equiv id-equiv + +is-connected-map-htpy-id : + {l : Level} {k : 𝕋} {A : UU l} {f : A → A} → f ~ id → is-connected-map k f +is-connected-map-htpy-id H = is-connected-map-htpy _ H is-connected-map-id + +is-connected-map-htpy-id' : + {l : Level} {k : 𝕋} {A : UU l} {f : A → A} → id ~ f → is-connected-map k f +is-connected-map-htpy-id' H = is-connected-map-htpy' _ H is-connected-map-id +``` + ### A `(k+1)`-connected map is `k`-connected ```agda diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 5462a0bf32..0d80bac5cf 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -213,6 +213,22 @@ module _ ( e)) ``` +### Sections of `k`-equivalences are `k`-equivalences + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-truncation-equivalence-map-section : + (k : 𝕋) (s : section f) → + is-truncation-equivalence k f → + is-truncation-equivalence k (map-section f s) + is-truncation-equivalence-map-section k (s , h) = + is-truncation-equivalence-right-factor f s + ( is-truncation-equivalence-is-equiv (f ∘ s) (is-equiv-htpy-id h)) +``` + ### Composing `k`-equivalences with equivalences ```agda @@ -452,6 +468,20 @@ module _ This is an instance of Proposition 2.31 in {{#cite CORS20}}. +**Proof.** Assume $f$ is $k$-connected. We want to show that the fibers of $g$ +are $k+1$-connected, so let $c$ be an element of the codomain of $g$. The fibers +of the composite $g ∘ f$ compute as + +$$ + \operatorname{fiber}_{g\circ f}(c) ≃ + \sum_{(b , p) : \operatorname{fiber}_{g}(c)}{\operatorname{fiber}_{f}(b)}. +$$ + +By the previous lemma, since $\operatorname{fiber}_{g\circ f}(c)$ is +$k+1$-connected, $\operatorname{fiber}_{g}(c)$ is $k+1$-connected if the first +projection map of this type is $k$-connected, and its fibers compute to the +fibers of $f$ so we are done. ∎ + ```agda module _ {l1 l2 l3 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} {C : UU l3} @@ -480,11 +510,36 @@ module _ ### A `k`-equivalence with a section is `k`-connected +**Proof.** If $k ≐ -2$ notice that every map is $-2$-connected. So let +$k ≐ n + 1$ for some truncation level $n$ and let $f$ be our $k$-equivalence +with a section. By assumption, we have a commuting triangle of maps + +```text + A + ∧ \ + s / \ f + / ∨ + B ======== B. +``` + +By the previous lemma, since the identity map is $k$-connected, it thus suffices +to show that $s$ is $n$-connected. But by the cancellation property of +$n+1$-equivalences $s$ is a $n+1$-equivalence and $n+1$-equivalences are in +particular $n$-connected. ∎ + ```agda module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) where + is-connected-map-section-is-truncation-equivalence-succ : + (k : 𝕋) (s : section f) → + is-truncation-equivalence (succ-𝕋 k) f → + is-connected-map k (map-section f s) + is-connected-map-section-is-truncation-equivalence-succ k (s , h) e = + is-connected-map-is-succ-truncation-equivalence s + ( is-truncation-equivalence-map-section (succ-𝕋 k) (s , h) e) + is-connected-map-is-truncation-equivalence-section : (k : 𝕋) → section f → is-truncation-equivalence k f → is-connected-map k f @@ -492,13 +547,8 @@ module _ is-neg-two-connected-map f is-connected-map-is-truncation-equivalence-section (succ-𝕋 k) (s , h) e = is-connected-map-right-factor-is-succ-connected-map-right-factor f s - ( is-connected-map-is-equiv (is-equiv-htpy id h is-equiv-id)) - ( is-connected-map-is-succ-truncation-equivalence s - ( is-truncation-equivalence-right-factor f s - ( is-truncation-equivalence-is-equiv - ( f ∘ s) - ( is-equiv-htpy id h is-equiv-id)) - ( e))) + ( is-connected-map-htpy-id h) + ( is-connected-map-section-is-truncation-equivalence-succ k (s , h) e) ``` ## References diff --git a/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md b/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md index 1ebbc9e3de..076c7343f3 100644 --- a/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md @@ -14,6 +14,7 @@ open import foundation.contractible-maps open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences +open import foundation.embeddings open import foundation.empty-types open import foundation.equivalences open import foundation.families-of-equivalences @@ -28,6 +29,7 @@ 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-maps open import foundation.propositions open import foundation.retractions open import foundation.retracts-of-maps @@ -316,27 +318,36 @@ module _ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) where - section-is-local-domains' : section (precomp f X) → is-local f Y → section f - pr1 (section-is-local-domains' section-precomp-X is-local-Y) = + section-is-local-domains'' : + section (precomp f X) → is-emb (precomp f Y) → section f + pr1 (section-is-local-domains'' section-precomp-X _) = pr1 section-precomp-X id - pr2 (section-is-local-domains' section-precomp-X is-local-Y) = + pr2 (section-is-local-domains'' section-precomp-X is-emb-precomp-Y) = htpy-eq ( ap ( pr1) - { ( f ∘ pr1 (section-is-local-domains' section-precomp-X is-local-Y)) , + { ( f ∘ pr1 section-precomp-X id) , ( ap (postcomp X f) (pr2 section-precomp-X id))} { id , refl} - ( eq-is-contr (is-contr-map-is-equiv is-local-Y f))) - - is-equiv-is-local-domains' : section (precomp f X) → is-local f Y → is-equiv f - pr1 (is-equiv-is-local-domains' section-precomp-X is-local-Y) = - section-is-local-domains' section-precomp-X is-local-Y - pr2 (is-equiv-is-local-domains' section-precomp-X is-local-Y) = + ( eq-is-prop (is-prop-map-is-emb is-emb-precomp-Y f))) + + section-is-local-domains' : + section (precomp f X) → is-local f Y → section f + section-is-local-domains' section-precomp-X is-local-Y = + section-is-local-domains'' section-precomp-X (is-emb-is-equiv is-local-Y) + + is-equiv-is-local-domains' : + section (precomp f X) → is-emb (precomp f Y) → is-equiv f + pr1 (is-equiv-is-local-domains' section-precomp-X is-emb-precomp-Y) = + section-is-local-domains'' section-precomp-X is-emb-precomp-Y + pr2 (is-equiv-is-local-domains' section-precomp-X is-emb-precomp-Y) = retraction-map-section-precomp f section-precomp-X is-equiv-is-local-domains : is-local f X → is-local f Y → is-equiv f - is-equiv-is-local-domains is-local-X = - is-equiv-is-local-domains' (pr1 is-local-X) + is-equiv-is-local-domains is-local-X is-local-Y = + is-equiv-is-local-domains' + ( section-is-equiv is-local-X) + ( is-emb-is-equiv is-local-Y) ``` ### If `f` has a retraction and the codomain of `f` is `f`-local, then `f` is an equivalence diff --git a/src/synthetic-homotopy-theory/infinite-cyclic-types.lagda.md b/src/synthetic-homotopy-theory/infinite-cyclic-types.lagda.md index bab47be956..24d9497897 100644 --- a/src/synthetic-homotopy-theory/infinite-cyclic-types.lagda.md +++ b/src/synthetic-homotopy-theory/infinite-cyclic-types.lagda.md @@ -145,7 +145,7 @@ module _ ( λ f → is-proof-irrelevant-is-prop ( is-property-is-equiv (pr1 f)) - ( is-equiv-htpy id + ( is-equiv-htpy-id ( htpy-eq ( ap ( pr1) @@ -153,8 +153,7 @@ module _ { y = pair id (pair refl refl-htpy)} ( eq-is-contr ( is-initial-ℤ-Pointed-Type-With-Aut - ℤ-Pointed-Type-With-Aut)))) - ( is-equiv-id)))) ∘e + ℤ-Pointed-Type-With-Aut))))))) ∘e ( ( equiv-right-swap-Σ) ∘e ( ( associative-Σ ( ℤ ≃ ℤ) From c97413c63caf9fa3e869adc4446527b8b668d657 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 19:16:27 +0200 Subject: [PATCH 02/69] edits --- .../truncation-equivalences.lagda.md | 42 +++++++++++++++---- 1 file changed, 34 insertions(+), 8 deletions(-) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 0d80bac5cf..8de3435314 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -123,7 +123,7 @@ is-truncation-equivalence-is-equiv-precomp k {A} {B} f H = ( H _ X)) ``` -### An equivalence is a `k`-equivalence for all `k` +### Equivalences are `k`-equivalences for all `k` ```agda module _ @@ -135,6 +135,32 @@ module _ is-truncation-equivalence-is-equiv e = is-equiv-map-equiv-trunc k (f , e) ``` +### The identity map is a `k`-equivalence for all `k` + +```agda +is-truncation-equivalence-id : + {l : Level} {k : 𝕋} {A : UU l} → is-truncation-equivalence k (id' A) +is-truncation-equivalence-id = is-truncation-equivalence-is-equiv id is-equiv-id +``` + +### The `k`-equivalences are closed under homotopies + +```agda +module _ + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {f g : A → B} + where + + is-truncation-equivalence-htpy : + f ~ g → is-truncation-equivalence k g → is-truncation-equivalence k f + is-truncation-equivalence-htpy H = + is-equiv-htpy (map-trunc k g) (htpy-trunc H) + + is-truncation-equivalence-htpy' : + f ~ g → is-truncation-equivalence k f → is-truncation-equivalence k g + is-truncation-equivalence-htpy' H = + is-equiv-htpy' (map-trunc k f) (htpy-trunc H) +``` + ### Every `k`-connected map is a `k`-equivalence ```agda @@ -164,7 +190,7 @@ module _ is-truncation-equivalence-comp g f ef eg = is-equiv-htpy ( map-trunc k g ∘ map-trunc k f) - ( preserves-comp-map-trunc k g f) + ( preserves-comp-map-trunc k g f) ( is-equiv-comp (map-trunc k g) (map-trunc k f) ef eg) truncation-equivalence-comp : @@ -191,14 +217,14 @@ module _ is-truncation-equivalence-left-factor : is-truncation-equivalence k f → is-truncation-equivalence k g - is-truncation-equivalence-left-factor ef = + is-truncation-equivalence-left-factor = is-equiv-left-factor ( map-trunc k g) ( map-trunc k f) - ( is-equiv-htpy + ( is-equiv-htpy' ( map-trunc k (g ∘ f)) - ( inv-htpy (preserves-comp-map-trunc k g f)) e) - ( ef) + ( preserves-comp-map-trunc k g f) + ( e)) is-truncation-equivalence-right-factor : is-truncation-equivalence k g → is-truncation-equivalence k f @@ -207,9 +233,9 @@ module _ ( map-trunc k g) ( map-trunc k f) ( eg) - ( is-equiv-htpy + ( is-equiv-htpy' ( map-trunc k (g ∘ f)) - ( inv-htpy (preserves-comp-map-trunc k g f)) + ( preserves-comp-map-trunc k g f) ( e)) ``` From 1fff72c64dbd54fb7f908ff71b1d0f72b7542b26 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 19:46:47 +0200 Subject: [PATCH 03/69] edit --- src/foundation-core/fibers-of-maps.lagda.md | 14 +++++++++++++- src/foundation/connected-maps.lagda.md | 2 +- 2 files changed, 14 insertions(+), 2 deletions(-) diff --git a/src/foundation-core/fibers-of-maps.lagda.md b/src/foundation-core/fibers-of-maps.lagda.md index 178e7d1397..4fa1a3f71d 100644 --- a/src/foundation-core/fibers-of-maps.lagda.md +++ b/src/foundation-core/fibers-of-maps.lagda.md @@ -455,7 +455,19 @@ module _ is-retraction-map-inv-equiv-fiber-htpy equiv-fiber-htpy : fiber f y ≃ fiber g y - equiv-fiber-htpy = map-equiv-fiber-htpy , is-equiv-map-equiv-fiber-htpy + equiv-fiber-htpy = + ( map-equiv-fiber-htpy , is-equiv-map-equiv-fiber-htpy) + + is-equiv-map-inv-equiv-fiber-htpy : is-equiv map-inv-equiv-fiber-htpy + is-equiv-map-inv-equiv-fiber-htpy = + is-equiv-is-invertible + map-equiv-fiber-htpy + is-retraction-map-inv-equiv-fiber-htpy + is-section-map-inv-equiv-fiber-htpy + + inv-equiv-fiber-htpy : fiber g y ≃ fiber f y + inv-equiv-fiber-htpy = + ( map-inv-equiv-fiber-htpy , is-equiv-map-inv-equiv-fiber-htpy) ``` We repeat the construction for `fiber'`. diff --git a/src/foundation/connected-maps.lagda.md b/src/foundation/connected-maps.lagda.md index 1e4ecfeebe..391ae1580f 100644 --- a/src/foundation/connected-maps.lagda.md +++ b/src/foundation/connected-maps.lagda.md @@ -214,7 +214,7 @@ module _ is-connected-map-htpy : (f ~ g) → is-connected-map k g → is-connected-map k f is-connected-map-htpy H G x = - is-connected-equiv' (equiv-fiber-htpy H x) (G x) + is-connected-equiv (inv-equiv-fiber-htpy H x) (G x) is-connected-map-htpy' : (f ~ g) → is-connected-map k f → is-connected-map k g From e799da9cf9118b7d1c731c5bc8026e4b39a3208e Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 20:08:46 +0200 Subject: [PATCH 04/69] edits --- .../truncation-equivalences.lagda.md | 21 +++++++++++++--- src/foundation/truncations.lagda.md | 6 +++-- .../universal-property-truncation.lagda.md | 25 +++++++++---------- 3 files changed, 33 insertions(+), 19 deletions(-) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 8de3435314..706accc4cb 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -67,6 +67,14 @@ module _ is-truncation-equivalence-truncation-equivalence : is-truncation-equivalence k map-truncation-equivalence is-truncation-equivalence-truncation-equivalence = pr2 f + + map-trunc-truncation-equivalence : type-trunc k A → type-trunc k B + map-trunc-truncation-equivalence = map-trunc k map-truncation-equivalence + + equiv-trunc-truncation-equivalence : type-trunc k A ≃ type-trunc k B + equiv-trunc-truncation-equivalence = + ( map-trunc-truncation-equivalence , + is-truncation-equivalence-truncation-equivalence) ``` ## Properties @@ -336,7 +344,7 @@ module _ ( is-trunc-type-Truncated-Type X))))))) ``` -### There is an `k`-equivalence between the fiber of a map and the fiber of its `(k+1)`-truncation +### There is a `k`-equivalence between the fiber of a map and the fiber of its `(k+1)`-truncation This is an instance of Corollary 2.29 in {{#cite CORS20}}. @@ -396,7 +404,8 @@ module _ truncation-equivalence k ( fiber f b) ( fiber (map-trunc (succ-𝕋 k) f) (unit-trunc b)) - pr1 truncation-equivalence-fiber-map-trunc-fiber = fiber-map-trunc-fiber + pr1 truncation-equivalence-fiber-map-trunc-fiber = + fiber-map-trunc-fiber pr2 truncation-equivalence-fiber-map-trunc-fiber = is-truncation-equivalence-fiber-map-trunc-fiber ``` @@ -443,6 +452,10 @@ module _ This follows part of the proof of Proposition 2.31 in {{#cite CORS20}}. +**Proof.** Let $f : A → B$ be a $k$-connected map on a domain that is +$k+1$-connected. We must show that the truncation of the codomain is +contractible. + ```agda module _ {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} (f : A → B) @@ -484,8 +497,8 @@ module _ ( equiv-unit-trunc ( fiber (map-trunc (succ-𝕋 k) f) (unit-trunc b) , is-trunc-fiber-map-trunc-is-succ-connected cA b))) ∘e - ( map-trunc k (fiber-map-trunc-fiber f b) , - is-truncation-equivalence-fiber-map-trunc-fiber f b)) + ( equiv-trunc-truncation-equivalence k + ( truncation-equivalence-fiber-map-trunc-fiber f b))) ( cf b))))) ( cA) ``` diff --git a/src/foundation/truncations.lagda.md b/src/foundation/truncations.lagda.md index 9523e78045..4435a43718 100644 --- a/src/foundation/truncations.lagda.md +++ b/src/foundation/truncations.lagda.md @@ -27,6 +27,8 @@ open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies open import foundation-core.propositions +open import foundation-core.retractions +open import foundation-core.sections open import foundation-core.torsorial-type-families open import foundation-core.truncation-levels open import foundation-core.universal-property-truncation @@ -317,11 +319,11 @@ module _ map-inv-unit-trunc = map-universal-property-trunc A id is-retraction-map-inv-unit-trunc : - ( map-inv-unit-trunc ∘ unit-trunc) ~ id + is-retraction unit-trunc map-inv-unit-trunc is-retraction-map-inv-unit-trunc = triangle-universal-property-trunc A id is-section-map-inv-unit-trunc : - ( unit-trunc ∘ map-inv-unit-trunc) ~ id + is-section unit-trunc map-inv-unit-trunc is-section-map-inv-unit-trunc = htpy-eq ( pr1 diff --git a/src/foundation/universal-property-truncation.lagda.md b/src/foundation/universal-property-truncation.lagda.md index 3d261465a4..6817f9315a 100644 --- a/src/foundation/universal-property-truncation.lagda.md +++ b/src/foundation/universal-property-truncation.lagda.md @@ -60,14 +60,14 @@ module _ ( Σ (type-Truncated-Type C) (λ z → g x = z)) ( equiv-tot ( λ z → - ( ( equiv-ev-refl' x) ∘e - ( equiv-Π-equiv-family - ( λ x' → - equiv-is-truncation - ( Id-Truncated-Type B (f x') (f x)) - ( ap f) - ( K x' x) - ( Id-Truncated-Type C (g x') z)))) ∘e + ( equiv-ev-refl' x) ∘e + ( equiv-Π-equiv-family + ( λ x' → + equiv-is-truncation + ( Id-Truncated-Type B (f x') (f x)) + ( ap f) + ( K x' x) + ( Id-Truncated-Type C (g x') z))) ∘e ( equiv-ev-pair))) ( is-torsorial-Id (g x))) @@ -82,11 +82,10 @@ module _ ( λ z → (t : fiber f y) → (g (pr1 t) = z))) ( ( equiv-tot ( λ h → - ( ( ( inv-equiv (equiv-funext)) ∘e - ( equiv-Π-equiv-family - ( λ x → - equiv-inv (g x) (h (f x)) ∘e equiv-ev-refl (f x)))) ∘e - ( equiv-swap-Π)) ∘e + ( equiv-eq-htpy) ∘e + ( equiv-Π-equiv-family + ( λ x → equiv-inv (g x) (h (f x)) ∘e equiv-ev-refl (f x))) ∘e + ( equiv-swap-Π) ∘e ( equiv-Π-equiv-family (λ x → equiv-ev-pair)))) ∘e ( distributive-Π-Σ)) ( is-contr-Π From 92788a97be58682048075c2b1c9e968be3f361bf Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 22:04:45 +0200 Subject: [PATCH 05/69] refactor, simplify, and informalize proof of `is-connected-map-is-connected-map-trunc-succ` --- src/foundation-core/truncated-maps.lagda.md | 20 +++ src/foundation/connected-maps.lagda.md | 20 ++- src/foundation/connected-types.lagda.md | 26 +++ .../truncation-equivalences.lagda.md | 160 +++++++++++------- src/foundation/truncations.lagda.md | 49 +++--- .../universal-property-truncation.lagda.md | 8 +- 6 files changed, 200 insertions(+), 83 deletions(-) diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 6c68755acc..8ff7e2d10b 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -10,10 +10,12 @@ module foundation-core.truncated-maps where open import foundation.action-on-identifications-functions open import foundation.dependent-pair-types open import foundation.equality-fibers-of-maps +open import foundation.type-arithmetic-dependent-pair-types open import foundation.universe-levels open import foundation-core.commuting-squares-of-maps open import foundation-core.contractible-maps +open import foundation-core.contractible-types open import foundation-core.equivalences open import foundation-core.fibers-of-maps open import foundation-core.function-types @@ -298,3 +300,21 @@ module _ ( is-trunc-map-comp k i g M ( is-trunc-map-is-equiv k K)) ``` + +### If the domain is contractible and the codomain is `k+1`-truncated then the map is `k`-truncated + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-trunc-map-is-trunc-succ-codomain-is-contr-domain : + is-contr A → + is-trunc (succ-𝕋 k) B → + is-trunc-map k f + is-trunc-map-is-trunc-succ-codomain-is-contr-domain c tB u = + is-trunc-equiv k + ( f (center c) = u) + ( left-unit-law-Σ-is-contr c (center c)) + ( tB (f (center c)) u) +``` diff --git a/src/foundation/connected-maps.lagda.md b/src/foundation/connected-maps.lagda.md index 391ae1580f..4d785084d6 100644 --- a/src/foundation/connected-maps.lagda.md +++ b/src/foundation/connected-maps.lagda.md @@ -334,6 +334,24 @@ module _ is-connected-equiv (inv-compute-fiber-tot f (x , y)) (H (x , y)) ``` +### A map is an equivalence if it is `k`-connected and `k`-truncated + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-contr-map-is-connected-map-is-trunc-map : + is-trunc-map k f → is-connected-map k f → is-contr-map f + is-contr-map-is-connected-map-is-trunc-map H K x = + is-contr-is-connected-is-trunc (H x) (K x) + + is-equiv-is-connected-map-is-trunc-map : + is-trunc-map k f → is-connected-map k f → is-equiv f + is-equiv-is-connected-map-is-trunc-map H K = + is-equiv-is-contr-map (is-contr-map-is-connected-map-is-trunc-map H K) +``` + ### Dependent universal property for connected maps ```agda @@ -453,7 +471,7 @@ module _ contraction-is-connected-map-dependent-universal-property-connected-map b ``` -### The map `unit-trunc {k}` is `k`-connected +### The unit map of the `k`-truncation is `k`-connected ```agda module _ diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index 06cf448eef..5ad572931a 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -27,6 +27,7 @@ open import foundation-core.functoriality-dependent-pair-types open import foundation-core.identity-types open import foundation-core.precomposition-functions open import foundation-core.retracts-of-types +open import foundation-core.truncated-maps open import foundation-core.truncated-types open import foundation-core.truncation-levels ``` @@ -172,6 +173,18 @@ module _ ( λ B → is-equiv-diagonal-exponential-is-contr H (type-Truncated-Type B)) ``` +### A type that is `k`-connected and `k`-truncated is contractible + +```agda +module _ + {l1 : Level} {k : 𝕋} {A : UU l1} + where + + is-contr-is-connected-is-trunc : is-trunc k A → is-connected k A → is-contr A + is-contr-is-connected-is-trunc H = + is-contr-equiv (type-trunc k A) (equiv-unit-trunc (A , H)) +``` + ### A type that is `(k+1)`-connected is `k`-connected ```agda @@ -271,3 +284,16 @@ module _ ( λ where refl → refl) ( center (K a x))))) ``` + +### If the domain of a map `f` is `k+1`-connected, then the `k+1`-truncation of `f` is `k`-truncated + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} (f : A → B) + where + + is-trunc-map-trunc-succ-is-succ-connected-domain : + is-connected (succ-𝕋 k) A → is-trunc-map k (map-trunc (succ-𝕋 k) f) + is-trunc-map-trunc-succ-is-succ-connected-domain c = + is-trunc-map-is-trunc-succ-codomain-is-contr-domain c is-trunc-type-trunc +``` diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 706accc4cb..76f8e12b38 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -32,6 +32,7 @@ open import foundation-core.homotopies open import foundation-core.precomposition-functions open import foundation-core.sections open import foundation-core.transport-along-identifications +open import foundation-core.truncated-maps open import foundation-core.truncated-types open import foundation-core.truncation-levels ``` @@ -377,28 +378,29 @@ module _ ( map-effectiveness-trunc k (f a) b) ∘ ( unit-trunc))) - is-truncation-equivalence-fiber-map-trunc-fiber : - is-truncation-equivalence k fiber-map-trunc-fiber - is-truncation-equivalence-fiber-map-trunc-fiber = - is-truncation-equivalence-comp - ( map-Σ-map-base-unit-trunc - ( λ t → map-trunc (succ-𝕋 k) f t = unit-trunc b)) - ( tot - ( λ a → - ( concat (naturality-unit-trunc (succ-𝕋 k) f a) (unit-trunc b)) ∘ - ( map-effectiveness-trunc k (f a) b) ∘ - ( unit-trunc))) - ( is-truncation-equivalence-is-truncation-equivalence-equiv - ( equiv-tot + abstract + is-truncation-equivalence-fiber-map-trunc-fiber : + is-truncation-equivalence k fiber-map-trunc-fiber + is-truncation-equivalence-fiber-map-trunc-fiber = + is-truncation-equivalence-comp + ( map-Σ-map-base-unit-trunc + ( λ t → map-trunc (succ-𝕋 k) f t = unit-trunc b)) + ( tot ( λ a → - ( equiv-concat - ( naturality-unit-trunc (succ-𝕋 k) f a) - ( unit-trunc b)) ∘e - ( effectiveness-trunc k (f a) b))) - ( λ (a , p) → a , unit-trunc p) - ( is-equiv-map-equiv (equiv-trunc-Σ k))) - ( is-truncation-equivalence-map-Σ-map-base-unit-trunc - ( λ t → map-trunc (succ-𝕋 k) f t = unit-trunc b)) + ( concat (naturality-unit-trunc (succ-𝕋 k) f a) (unit-trunc b)) ∘ + ( map-effectiveness-trunc k (f a) b) ∘ + ( unit-trunc))) + ( is-truncation-equivalence-is-truncation-equivalence-equiv + ( equiv-tot + ( λ a → + ( equiv-concat + ( naturality-unit-trunc (succ-𝕋 k) f a) + ( unit-trunc b)) ∘e + ( effectiveness-trunc k (f a) b))) + ( λ (a , p) → a , unit-trunc p) + ( is-equiv-map-equiv (equiv-trunc-Σ k))) + ( is-truncation-equivalence-map-Σ-map-base-unit-trunc + ( λ t → map-trunc (succ-𝕋 k) f t = unit-trunc b)) truncation-equivalence-fiber-map-trunc-fiber : truncation-equivalence k @@ -408,6 +410,13 @@ module _ fiber-map-trunc-fiber pr2 truncation-equivalence-fiber-map-trunc-fiber = is-truncation-equivalence-fiber-map-trunc-fiber + + equiv-trunc-fiber-map-trunc-fiber : + type-trunc k (fiber f b) ≃ + type-trunc k (fiber (map-trunc (succ-𝕋 k) f) (unit-trunc b)) + equiv-trunc-fiber-map-trunc-fiber = + equiv-trunc-truncation-equivalence k + ( truncation-equivalence-fiber-map-trunc-fiber) ``` ### Being `k`-connected is invariant under `k`-equivalences @@ -423,12 +432,25 @@ module _ is-connected-is-truncation-equivalence-is-connected f e = is-contr-equiv (type-trunc k B) (map-trunc k f , e) + is-connected-is-truncation-equivalence-is-connected' : + (f : A → B) → is-truncation-equivalence k f → + is-connected k A → is-connected k B + is-connected-is-truncation-equivalence-is-connected' f e = + is-contr-equiv' (type-trunc k A) (map-trunc k f , e) + is-connected-truncation-equivalence-is-connected : truncation-equivalence k A B → is-connected k B → is-connected k A is-connected-truncation-equivalence-is-connected f = is-connected-is-truncation-equivalence-is-connected ( map-truncation-equivalence k f) ( is-truncation-equivalence-truncation-equivalence k f) + + is-connected-truncation-equivalence-is-connected' : + truncation-equivalence k A B → is-connected k A → is-connected k B + is-connected-truncation-equivalence-is-connected' f = + is-connected-is-truncation-equivalence-is-connected' + ( map-truncation-equivalence k f) + ( is-truncation-equivalence-truncation-equivalence k f) ``` ### Every `(k+1)`-equivalence is `k`-connected @@ -445,7 +467,38 @@ module _ is-connected-map-is-succ-truncation-equivalence e b = is-connected-truncation-equivalence-is-connected ( truncation-equivalence-fiber-map-trunc-fiber f b) - ( is-connected-is-contr k (is-contr-map-is-equiv e (unit-trunc b))) + ( is-connected-map-is-equiv e (unit-trunc b)) +``` + +### A map is `k`-connected if and only if its `k+1`-truncation is + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-connected-map-trunc-succ-is-succ-connected-domain : + is-connected-map k f → + is-connected-map k (map-trunc (succ-𝕋 k) f) + is-connected-map-trunc-succ-is-succ-connected-domain cf t = + apply-universal-property-trunc-Prop + ( is-surjective-unit-trunc-succ t) + ( is-connected-Prop k (fiber (map-trunc (succ-𝕋 k) f) t)) + ( λ (b , p) → + tr + ( λ s → is-connected k (fiber (map-trunc (succ-𝕋 k) f) s)) + ( p) + ( is-connected-truncation-equivalence-is-connected' + ( truncation-equivalence-fiber-map-trunc-fiber f b) + ( cf b))) + + is-connected-map-is-connected-map-trunc-succ : + is-connected-map k (map-trunc (succ-𝕋 k) f) → + is-connected-map k f + is-connected-map-is-connected-map-trunc-succ cf' b = + is-connected-truncation-equivalence-is-connected + ( truncation-equivalence-fiber-map-trunc-fiber f b) + ( cf' (unit-trunc b)) ``` ### The codomain of a `k`-connected map is `(k+1)`-connected if its domain is `(k+1)`-connected @@ -453,53 +506,36 @@ module _ This follows part of the proof of Proposition 2.31 in {{#cite CORS20}}. **Proof.** Let $f : A → B$ be a $k$-connected map on a domain that is -$k+1$-connected. We must show that the truncation of the codomain is -contractible. +$k+1$-connected. To show that $B$ is $k+1$-connected it is enough to show that +$f$ is a $k+1$-equivalence, in other words, that $║f║ₖ₊₁$ is an equivalence. By +previous computations we know that $║f║ₖ₊₁$ is $k$-truncated since the domain is +$k+1$-connected, and that $║f║ₖ₊₁$ is $k$-connected since $f$ is $k$-connected, +so we are done. ∎ ```agda module _ {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} (f : A → B) where - is-trunc-fiber-map-trunc-is-succ-connected : - is-connected (succ-𝕋 k) A → - (b : B) → - is-trunc k (fiber (map-trunc (succ-𝕋 k) f) (unit-trunc b)) - is-trunc-fiber-map-trunc-is-succ-connected c b = - is-trunc-equiv k - ( map-trunc (succ-𝕋 k) f (center c) = unit-trunc b) - ( left-unit-law-Σ-is-contr c (center c)) - ( is-trunc-type-trunc (map-trunc (succ-𝕋 k) f (center c)) (unit-trunc b)) - - is-succ-connected-is-connected-map-is-succ-connected : + is-truncation-equivalence-succ-is-succ-connected-domain-is-connected-map : + is-connected-map k f → is-connected (succ-𝕋 k) A → + is-truncation-equivalence (succ-𝕋 k) f + is-truncation-equivalence-succ-is-succ-connected-domain-is-connected-map + cf cA = + is-equiv-is-connected-map-is-trunc-map + ( is-trunc-map-trunc-succ-is-succ-connected-domain f cA) + ( is-connected-map-trunc-succ-is-succ-connected-domain cf) + + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map : is-connected-map k f → + is-connected (succ-𝕋 k) A → is-connected (succ-𝕋 k) B - is-succ-connected-is-connected-map-is-succ-connected cA cf = - is-contr-is-equiv' - ( type-trunc (succ-𝕋 k) A) - ( map-trunc (succ-𝕋 k) f) - ( is-equiv-is-contr-map - ( λ t → - apply-universal-property-trunc-Prop - ( is-surjective-is-truncation - ( trunc (succ-𝕋 k) B) - ( is-truncation-trunc) - ( t)) - ( is-contr-Prop (fiber (map-trunc (succ-𝕋 k) f) t)) - ( λ (b , p) → - tr - ( λ s → is-contr (fiber (map-trunc (succ-𝕋 k) f) s)) - ( p) - ( is-contr-equiv' - ( type-trunc k (fiber f b)) - ( ( inv-equiv - ( equiv-unit-trunc - ( fiber (map-trunc (succ-𝕋 k) f) (unit-trunc b) , - is-trunc-fiber-map-trunc-is-succ-connected cA b))) ∘e - ( equiv-trunc-truncation-equivalence k - ( truncation-equivalence-fiber-map-trunc-fiber f b))) - ( cf b))))) + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map cf cA = + is-connected-is-truncation-equivalence-is-connected' f + ( is-truncation-equivalence-succ-is-succ-connected-domain-is-connected-map + ( cf) + ( cA)) ( cA) ``` @@ -538,13 +574,13 @@ module _ is-connected-map-right-factor-is-succ-connected-map-right-factor : is-connected-map k f → is-connected-map (succ-𝕋 k) g is-connected-map-right-factor-is-succ-connected-map-right-factor cf c = - is-succ-connected-is-connected-map-is-succ-connected + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map ( pr1) - ( is-connected-equiv' (compute-fiber-comp g f c) (cgf c)) ( λ p → is-connected-equiv ( equiv-fiber-pr1 (fiber f ∘ pr1) p) ( cf (pr1 p))) + ( is-connected-equiv' (compute-fiber-comp g f c) (cgf c)) ``` ### A `k`-equivalence with a section is `k`-connected diff --git a/src/foundation/truncations.lagda.md b/src/foundation/truncations.lagda.md index 4435a43718..b660749851 100644 --- a/src/foundation/truncations.lagda.md +++ b/src/foundation/truncations.lagda.md @@ -322,23 +322,24 @@ module _ is-retraction unit-trunc map-inv-unit-trunc is-retraction-map-inv-unit-trunc = triangle-universal-property-trunc A id - is-section-map-inv-unit-trunc : - is-section unit-trunc map-inv-unit-trunc - is-section-map-inv-unit-trunc = - htpy-eq - ( pr1 - ( pair-eq-Σ - ( eq-is-prop' - ( is-trunc-succ-is-trunc - ( neg-two-𝕋) - ( universal-property-trunc - ( k) - ( type-Truncated-Type A) - ( trunc k (type-Truncated-Type A)) - ( unit-trunc))) - ( unit-trunc ∘ map-inv-unit-trunc , - unit-trunc ·l is-retraction-map-inv-unit-trunc) - ( id , refl-htpy)))) + abstract + is-section-map-inv-unit-trunc : + is-section unit-trunc map-inv-unit-trunc + is-section-map-inv-unit-trunc = + htpy-eq + ( pr1 + ( pair-eq-Σ + ( eq-is-prop' + ( is-trunc-succ-is-trunc + ( neg-two-𝕋) + ( universal-property-trunc + ( k) + ( type-Truncated-Type A) + ( trunc k (type-Truncated-Type A)) + ( unit-trunc))) + ( unit-trunc ∘ map-inv-unit-trunc , + unit-trunc ·l is-retraction-map-inv-unit-trunc) + ( id , refl-htpy)))) is-equiv-unit-trunc : is-equiv unit-trunc is-equiv-unit-trunc = @@ -349,8 +350,18 @@ module _ equiv-unit-trunc : type-Truncated-Type A ≃ type-trunc k (type-Truncated-Type A) - pr1 equiv-unit-trunc = unit-trunc - pr2 equiv-unit-trunc = is-equiv-unit-trunc + equiv-unit-trunc = (unit-trunc , is-equiv-unit-trunc) + + is-equiv-map-inv-unit-trunc : is-equiv map-inv-unit-trunc + is-equiv-map-inv-unit-trunc = + is-equiv-is-invertible + unit-trunc + is-retraction-map-inv-unit-trunc + is-section-map-inv-unit-trunc + + inv-equiv-unit-trunc : + type-trunc k (type-Truncated-Type A) ≃ type-Truncated-Type A + inv-equiv-unit-trunc = (map-inv-unit-trunc , is-equiv-map-inv-unit-trunc) ``` ### A contractible type is equivalent to its `k`-truncation diff --git a/src/foundation/universal-property-truncation.lagda.md b/src/foundation/universal-property-truncation.lagda.md index 6817f9315a..8b3b3cf9c3 100644 --- a/src/foundation/universal-property-truncation.lagda.md +++ b/src/foundation/universal-property-truncation.lagda.md @@ -16,6 +16,7 @@ open import foundation.function-extensionality open import foundation.identity-types open import foundation.propositional-truncations open import foundation.surjective-maps +open import foundation.truncations open import foundation.type-arithmetic-dependent-function-types open import foundation.universal-property-dependent-pair-types open import foundation.universal-property-identity-types @@ -102,5 +103,10 @@ module _ map-inv-is-equiv ( dependent-universal-property-truncation-is-truncation B f H ( λ y → truncated-type-trunc-Prop k (fiber f y))) - ( λ x → unit-trunc-Prop (pair x refl)) + ( λ x → unit-trunc-Prop (x , refl)) + +is-surjective-unit-trunc-succ : + {l : Level} {k : 𝕋} {A : UU l} → is-surjective (unit-trunc {k = succ-𝕋 k} {A}) +is-surjective-unit-trunc-succ {k = k} {A} = + is-surjective-is-truncation (trunc (succ-𝕋 k) A) is-truncation-trunc ``` From 7b2dcb031a8c878c80481f52d5dd5066ec2e396c Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 22:06:45 +0200 Subject: [PATCH 06/69] Retractions of `k`-equivalences are `k`-equivalences --- src/foundation/truncation-equivalences.lagda.md | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 76f8e12b38..d1e1530260 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -30,6 +30,7 @@ open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies open import foundation-core.precomposition-functions +open import foundation-core.retractions open import foundation-core.sections open import foundation-core.transport-along-identifications open import foundation-core.truncated-maps @@ -264,6 +265,22 @@ module _ ( is-truncation-equivalence-is-equiv (f ∘ s) (is-equiv-htpy-id h)) ``` +### Retractions of `k`-equivalences are `k`-equivalences + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-truncation-equivalence-map-retraction : + (k : 𝕋) (r : retraction f) → + is-truncation-equivalence k f → + is-truncation-equivalence k (map-retraction f r) + is-truncation-equivalence-map-retraction k (r , h) = + is-truncation-equivalence-left-factor r f + ( is-truncation-equivalence-is-equiv (r ∘ f) (is-equiv-htpy-id h)) +``` + ### Composing `k`-equivalences with equivalences ```agda From 004e4590794bb40c915c26cc5a1857eedd685229 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 23 Sep 2025 22:08:55 +0200 Subject: [PATCH 07/69] simplify informal proof slightly --- src/foundation/truncation-equivalences.lagda.md | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index d1e1530260..19ab739513 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -522,10 +522,10 @@ module _ This follows part of the proof of Proposition 2.31 in {{#cite CORS20}}. -**Proof.** Let $f : A → B$ be a $k$-connected map on a domain that is -$k+1$-connected. To show that $B$ is $k+1$-connected it is enough to show that -$f$ is a $k+1$-equivalence, in other words, that $║f║ₖ₊₁$ is an equivalence. By -previous computations we know that $║f║ₖ₊₁$ is $k$-truncated since the domain is +**Proof.** Let $f : A → B$ be a $k$-connected map on a $k+1$-connected domain. +To show that the codomain is $k+1$-connected it is enough to show that $f$ is a +$k+1$-equivalence, in other words, that $║f║ₖ₊₁$ is an equivalence. By previous +computations we know that $║f║ₖ₊₁$ is $k$-truncated since the domain is $k+1$-connected, and that $║f║ₖ₊₁$ is $k$-connected since $f$ is $k$-connected, so we are done. ∎ From 5b8442e68edcfead6f1f3342e6740c9a11c921c0 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 15:00:46 +0200 Subject: [PATCH 08/69] edits --- src/foundation/connected-types.lagda.md | 2 +- src/foundation/truncation-equivalences.lagda.md | 13 +++++++------ 2 files changed, 8 insertions(+), 7 deletions(-) diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index 5ad572931a..7455522a52 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -285,7 +285,7 @@ module _ ( center (K a x))))) ``` -### If the domain of a map `f` is `k+1`-connected, then the `k+1`-truncation of `f` is `k`-truncated +### If the domain of `f` is `k+1`-connected, then the `k+1`-truncation of `f` is `k`-truncated ```agda module _ diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 19ab739513..0f5d647ef5 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -616,7 +616,7 @@ with a section. By assumption, we have a commuting triangle of maps By the previous lemma, since the identity map is $k$-connected, it thus suffices to show that $s$ is $n$-connected. But by the cancellation property of -$n+1$-equivalences $s$ is a $n+1$-equivalence and $n+1$-equivalences are in +$n+1$-equivalences $s$ is an $n+1$-equivalence and $n+1$-equivalences are in particular $n$-connected. ∎ ```agda @@ -630,7 +630,7 @@ module _ is-connected-map k (map-section f s) is-connected-map-section-is-truncation-equivalence-succ k (s , h) e = is-connected-map-is-succ-truncation-equivalence s - ( is-truncation-equivalence-map-section (succ-𝕋 k) (s , h) e) + ( is-truncation-equivalence-map-section (succ-𝕋 k) (s , h) e) is-connected-map-is-truncation-equivalence-section : (k : 𝕋) → @@ -646,9 +646,10 @@ module _ ## References - The notion of `k`-equivalence is a special case of the notion of - `L`-equivalence, where `L` is a reflective subuniverse. They were studied in - the paper {{#cite CORS20}}. -- The class of `k`-equivalences is left orthogonal to the class of `k`-étale - maps. This was shown in {{#cite CR21}}. + `L`-equivalence, where `L` is a reflective subuniverse. These were studied in + {{#cite CORS20}}. +- The class of `k`-equivalences is + [left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to the + class of `k`-étale maps. This was shown in {{#cite CR21}}. {{#bibliography}} From ec05b4396042818fef06c829ac9dc335177c74fb Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 15:09:42 +0200 Subject: [PATCH 09/69] edits --- src/foundation/connected-types.lagda.md | 6 ++---- src/foundation/truncation-equivalences.lagda.md | 2 +- src/univalent-combinatorics/subcounting.lagda.md | 6 +++--- src/univalent-combinatorics/subfinite-types.lagda.md | 11 ++++++----- 4 files changed, 12 insertions(+), 13 deletions(-) diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index 7455522a52..637cc49594 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -153,9 +153,7 @@ module _ where is-connected-retract-of : - A retract-of B → - is-connected k B → - is-connected k A + 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) ``` @@ -220,7 +218,7 @@ is-connected-Σ k H K = ≃ ║A║ₖ by the right unit law of Σ ``` -and so, in particular, if the total space is `k`-connected so is the base. □ +and so, in particular, if the total space is `k`-connected so is the base. ∎ ```agda is-connected-base : diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 0f5d647ef5..9b82767fa9 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -572,7 +572,7 @@ $$ By the previous lemma, since $\operatorname{fiber}_{g\circ f}(c)$ is $k+1$-connected, $\operatorname{fiber}_{g}(c)$ is $k+1$-connected if the first projection map of this type is $k$-connected, and its fibers compute to the -fibers of $f$ so we are done. ∎ +fibers of $f$. ∎ ```agda module _ diff --git a/src/univalent-combinatorics/subcounting.lagda.md b/src/univalent-combinatorics/subcounting.lagda.md index b21a38dc6b..99dddd80ad 100644 --- a/src/univalent-combinatorics/subcounting.lagda.md +++ b/src/univalent-combinatorics/subcounting.lagda.md @@ -149,7 +149,7 @@ abstract propositional truncations, we have an embedding `║X║₋₁ ↪ ║Fin k║₋₁`. By induction, if `k ≐ 0` then `║Fin k║₋₁ ≃ 0 ≃ Fin 0` and so `║X║₋₁ ↪ Fin 0` is a subcounting. Otherwise, if `k ≐ j + 1`, then `║Fin k║₋₁ ≃ 1 ≃ Fin 1` and again -`║X║₋₁ ↪ Fin 1` is a subcounting. +`║X║₋₁ ↪ Fin 1` is a subcounting. ∎ ```agda module _ @@ -173,8 +173,8 @@ We reproduce a proof given by [Gro-Tsen](https://mathoverflow.net/users/17064/gro-tsen) in this MathOverflow answer: . -**Proof.** Let $X$ be a type with a subcounting $ι : X ↪ Fin n$, and let -$f : X ↪ X$ be an arbitrary self-embedding. It suffices to prove $f$ is +**Proof.** Let $X$ be a type with a subcounting $ι : X ↪ \operatorname{Fin}n$, +and let $f : X ↪ X$ be an arbitrary self-embedding. It suffices to prove $f$ is surjective, so assume given an $x : X$ where we want to show there exists $z : X$ such that $f(z) = x$. The mapping $i ↦ fⁱ(x)$ defines an $ℕ$-indexed sequence of elements of $X$. Since the diff --git a/src/univalent-combinatorics/subfinite-types.lagda.md b/src/univalent-combinatorics/subfinite-types.lagda.md index 2b3ebefc41..5e7fc9fba5 100644 --- a/src/univalent-combinatorics/subfinite-types.lagda.md +++ b/src/univalent-combinatorics/subfinite-types.lagda.md @@ -117,11 +117,12 @@ We reproduce a proof given by [Gro-Tsen](https://mathoverflow.net/users/17064/gro-tsen) in this MathOverflow answer: . -**Proof.** Let $X$ be a subfinite type witnessed by $ι : X ↪ Fin n$, and let -$f : X ↪ X$ be an arbitrary self-embedding. It suffices to prove $f$ is -surjective, so assume given an $x : X$ where we want to show there exists -$z : X$ such that $f(z) = x$. The mapping $i ↦ fⁱ(x)$ defines an $ℕ$-indexed -sequence of elements of $X$. Since the +**Proof.** Let $X$ be a subfinite type witnessed by +$ι : X ↪ \operatorname{Fin}n$, and let $f : X ↪ X$ be an arbitrary +self-embedding. It suffices to prove $f$ is surjective, so assume given an +$x : X$ where we want to show there exists $z : X$ such that $f(z) = x$. The +mapping $i ↦ fⁱ(x)$ defines an $ℕ$-indexed sequence of elements of $X$. Since +the [standard pigeonhole principle](univalent-combinatorics.pigeonhole-principle.md) applies to $\operatorname{Fin}n$ there has to be $i < j$ in $\operatorname{Fin}n$ such that $ι(fⁱ(x)) = ι(fʲ(x))$. Since $ι$ is an embedding From 26a8447a369ede18551036b5ae166a6fbbe6b57c Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 16:12:33 +0200 Subject: [PATCH 10/69] more edits --- .../truncation-equivalences.lagda.md | 30 +++++++++++++++---- 1 file changed, 25 insertions(+), 5 deletions(-) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 9b82767fa9..d887de639e 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -560,7 +560,11 @@ module _ This is an instance of Proposition 2.31 in {{#cite CORS20}}. -**Proof.** Assume $f$ is $k$-connected. We want to show that the fibers of $g$ +**Proof.** If $g$ is $(k+1)$-connected then by the cancellation property of +$(k+1)$-equivalences, $f$ is a $k+1$-equivalence, and so in particular +$k$-connected. + +Conversely, assume $f$ is $k$-connected. We want to show that the fibers of $g$ are $k+1$-connected, so let $c$ be an element of the codomain of $g$. The fibers of the composite $g ∘ f$ compute as @@ -580,13 +584,19 @@ module _ (g : B → C) (f : A → B) (cgf : is-connected-map (succ-𝕋 k) (g ∘ f)) where + is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor : + is-connected-map (succ-𝕋 k) g → is-truncation-equivalence (succ-𝕋 k) f + is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor cg = + is-truncation-equivalence-right-factor g f + ( is-truncation-equivalence-is-connected-map (g ∘ f) cgf) + ( is-truncation-equivalence-is-connected-map g cg) + is-connected-map-right-factor-is-succ-connected-map-left-factor : is-connected-map (succ-𝕋 k) g → is-connected-map k f is-connected-map-right-factor-is-succ-connected-map-left-factor cg = is-connected-map-is-succ-truncation-equivalence f - ( is-truncation-equivalence-right-factor g f - ( is-truncation-equivalence-is-connected-map (g ∘ f) cgf) - ( is-truncation-equivalence-is-connected-map g cg)) + ( is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor + ( cg)) is-connected-map-right-factor-is-succ-connected-map-right-factor : is-connected-map k f → is-connected-map (succ-𝕋 k) g @@ -600,11 +610,21 @@ module _ ( is-connected-equiv' (compute-fiber-comp g f c) (cgf c)) ``` +As a corollary, if $g ∘ f$ is $(k + 1)$-connected for some $g$, and $f$ is +$k$-connected, then $f$ is a $k+1$-equivalence. + +```agda + is-succ-truncation-equiv-is-succ-connected-comp : + is-connected-map k f → is-truncation-equivalence (succ-𝕋 k) f + is-succ-truncation-equiv-is-succ-connected-comp cf = is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor + ( is-connected-map-right-factor-is-succ-connected-map-right-factor cf) +``` + ### A `k`-equivalence with a section is `k`-connected **Proof.** If $k ≐ -2$ notice that every map is $-2$-connected. So let $k ≐ n + 1$ for some truncation level $n$ and let $f$ be our $k$-equivalence -with a section. By assumption, we have a commuting triangle of maps +with a section $s$. By assumption, we have a commuting triangle of maps ```text A From 08301c98d96a6248d168aef9311437728d2c625c Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 20:47:59 +0200 Subject: [PATCH 11/69] edits `connected-maps` --- src/foundation/connected-maps.lagda.md | 91 +++++++++++++++----------- 1 file changed, 51 insertions(+), 40 deletions(-) diff --git a/src/foundation/connected-maps.lagda.md b/src/foundation/connected-maps.lagda.md index 4d785084d6..7843126958 100644 --- a/src/foundation/connected-maps.lagda.md +++ b/src/foundation/connected-maps.lagda.md @@ -41,9 +41,12 @@ open import foundation-core.truncated-maps ## Idea -A map is said to be **`k`-connected** if its -[fibers](foundation-core.fibers-of-maps.md) are -[`k`-connected types](foundation.connected-types.md). +A map is said to be +{{#concept "`k`-connected" Disambiguation="map of types" Agda=is-connected-map Agda=connected-map}} +if its [fibers](foundation-core.fibers-of-maps.md) are +`k`-[connected types](foundation.connected-types.md). In other words, if their +`k`-[truncations](foundation.truncations.md) are +[contractible](foundation-core.contractible-types.md). ## Definitions @@ -85,43 +88,6 @@ module _ emb-inclusion-connected-map : connected-map k A B ↪ (A → B) emb-inclusion-connected-map = emb-subtype (is-connected-map-Prop k) - - htpy-connected-map : (f g : connected-map k A B) → UU (l1 ⊔ l2) - htpy-connected-map f g = (map-connected-map f) ~ (map-connected-map g) - - refl-htpy-connected-map : (f : connected-map k A B) → htpy-connected-map f f - refl-htpy-connected-map f = refl-htpy - - is-torsorial-htpy-connected-map : - (f : connected-map k A B) → is-torsorial (htpy-connected-map f) - is-torsorial-htpy-connected-map f = - is-torsorial-Eq-subtype - ( is-torsorial-htpy (map-connected-map f)) - ( is-prop-is-connected-map k) - ( map-connected-map f) - ( refl-htpy-connected-map f) - ( is-connected-map-connected-map f) - - htpy-eq-connected-map : - (f g : connected-map k A B) → f = g → htpy-connected-map f g - htpy-eq-connected-map f .f refl = refl-htpy-connected-map f - - is-equiv-htpy-eq-connected-map : - (f g : connected-map k A B) → is-equiv (htpy-eq-connected-map f g) - is-equiv-htpy-eq-connected-map f = - fundamental-theorem-id - ( is-torsorial-htpy-connected-map f) - ( htpy-eq-connected-map f) - - extensionality-connected-map : - (f g : connected-map k A B) → (f = g) ≃ htpy-connected-map f g - pr1 (extensionality-connected-map f g) = htpy-eq-connected-map f g - pr2 (extensionality-connected-map f g) = is-equiv-htpy-eq-connected-map f g - - eq-htpy-connected-map : - (f g : connected-map k A B) → htpy-connected-map f g → (f = g) - eq-htpy-connected-map f g = - map-inv-equiv (extensionality-connected-map f g) ``` ### The type of connected maps into a type @@ -193,6 +159,51 @@ module _ ## Properties +### Characterizing equality of `k`-connected maps + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : UU l2} + where + + htpy-connected-map : (f g : connected-map k A B) → UU (l1 ⊔ l2) + htpy-connected-map f g = (map-connected-map f) ~ (map-connected-map g) + + refl-htpy-connected-map : (f : connected-map k A B) → htpy-connected-map f f + refl-htpy-connected-map f = refl-htpy + + is-torsorial-htpy-connected-map : + (f : connected-map k A B) → is-torsorial (htpy-connected-map f) + is-torsorial-htpy-connected-map f = + is-torsorial-Eq-subtype + ( is-torsorial-htpy (map-connected-map f)) + ( is-prop-is-connected-map k) + ( map-connected-map f) + ( refl-htpy-connected-map f) + ( is-connected-map-connected-map f) + + htpy-eq-connected-map : + (f g : connected-map k A B) → f = g → htpy-connected-map f g + htpy-eq-connected-map f .f refl = refl-htpy-connected-map f + + is-equiv-htpy-eq-connected-map : + (f g : connected-map k A B) → is-equiv (htpy-eq-connected-map f g) + is-equiv-htpy-eq-connected-map f = + fundamental-theorem-id + ( is-torsorial-htpy-connected-map f) + ( htpy-eq-connected-map f) + + extensionality-connected-map : + (f g : connected-map k A B) → (f = g) ≃ htpy-connected-map f g + pr1 (extensionality-connected-map f g) = htpy-eq-connected-map f g + pr2 (extensionality-connected-map f g) = is-equiv-htpy-eq-connected-map f g + + eq-htpy-connected-map : + (f g : connected-map k A B) → htpy-connected-map f g → (f = g) + eq-htpy-connected-map f g = + map-inv-equiv (extensionality-connected-map f g) +``` + ### All maps are `(-2)`-connected ```agda From 85b1305afd2e316a321afd30e29ccc7d29ac692d Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 22:41:01 +0200 Subject: [PATCH 12/69] edits --- src/foundation-core/contractible-types.lagda.md | 16 +++++++--------- ...nctoriality-dependent-function-types.lagda.md | 13 +++++++++---- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/src/foundation-core/contractible-types.lagda.md b/src/foundation-core/contractible-types.lagda.md index d86bdea803..0cdc98945d 100644 --- a/src/foundation-core/contractible-types.lagda.md +++ b/src/foundation-core/contractible-types.lagda.md @@ -27,8 +27,8 @@ open import foundation-core.transport-along-identifications ## Idea -Contractible types are types that have, up to identification, exactly one -element. +{{#concept "Contractible types" Agda=is-contr}} are types that have, up to +[identification](foundation-core.identity-types.md), exactly one element. ## Definition @@ -93,8 +93,7 @@ module _ abstract is-contr-equiv : A ≃ B → is-contr B → is-contr A - is-contr-equiv (e , is-equiv-e) is-contr-B = - is-contr-is-equiv e is-equiv-e is-contr-B + is-contr-equiv (e , is-equiv-e) = is-contr-is-equiv e is-equiv-e module _ {l1 l2 : Level} (A : UU l1) {B : UU l2} @@ -103,16 +102,14 @@ module _ abstract is-contr-is-equiv' : (f : A → B) → is-equiv f → is-contr A → is-contr B - is-contr-is-equiv' f is-equiv-f is-contr-A = + is-contr-is-equiv' f is-equiv-f = is-contr-is-equiv A ( map-inv-is-equiv is-equiv-f) ( is-equiv-map-inv-is-equiv is-equiv-f) - ( is-contr-A) abstract is-contr-equiv' : (e : A ≃ B) → is-contr A → is-contr B - is-contr-equiv' (pair e is-equiv-e) is-contr-A = - is-contr-is-equiv' e is-equiv-e is-contr-A + is-contr-equiv' (e , is-equiv-e) = is-contr-is-equiv' e is-equiv-e module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} @@ -128,7 +125,8 @@ module _ ( contraction is-contr-A) equiv-is-contr : is-contr A → is-contr B → A ≃ B - pr1 (equiv-is-contr is-contr-A is-contr-B) a = center is-contr-B + pr1 (equiv-is-contr is-contr-A is-contr-B) a = + center is-contr-B pr2 (equiv-is-contr is-contr-A is-contr-B) = is-equiv-is-contr _ is-contr-A is-contr-B ``` diff --git a/src/foundation-core/functoriality-dependent-function-types.lagda.md b/src/foundation-core/functoriality-dependent-function-types.lagda.md index 1ed18e0b4c..1308e74fe2 100644 --- a/src/foundation-core/functoriality-dependent-function-types.lagda.md +++ b/src/foundation-core/functoriality-dependent-function-types.lagda.md @@ -96,15 +96,20 @@ module _ {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} where + abstract + is-contr-map-Π-is-fiberwise-contr-map : + {f : (i : I) → A i → B i} → + ((i : I) → is-contr-map (f i)) → is-contr-map (map-Π f) + is-contr-map-Π-is-fiberwise-contr-map H g = + is-contr-equiv' _ (compute-fiber-map-Π _ g) (is-contr-Π (λ i → H i (g i))) + abstract is-equiv-map-Π-is-fiberwise-equiv : {f : (i : I) → A i → B i} → is-fiberwise-equiv f → is-equiv (map-Π f) is-equiv-map-Π-is-fiberwise-equiv is-equiv-f = is-equiv-is-contr-map - ( λ g → - is-contr-equiv' _ - ( compute-fiber-map-Π _ g) - ( is-contr-Π (λ i → is-contr-map-is-equiv (is-equiv-f i) (g i)))) + ( is-contr-map-Π-is-fiberwise-contr-map + ( is-contr-map-is-equiv ∘ is-equiv-f)) equiv-Π-equiv-family : (e : (i : I) → A i ≃ B i) → ((i : I) → A i) ≃ ((i : I) → B i) From fb3887fbedd8323b4aacf7ad41055a65a17bbc0c Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 22:44:12 +0200 Subject: [PATCH 13/69] pre-commit --- src/foundation/truncation-equivalences.lagda.md | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index d887de639e..7fc1667d37 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -586,7 +586,8 @@ module _ is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor : is-connected-map (succ-𝕋 k) g → is-truncation-equivalence (succ-𝕋 k) f - is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor cg = + is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor + cg = is-truncation-equivalence-right-factor g f ( is-truncation-equivalence-is-connected-map (g ∘ f) cgf) ( is-truncation-equivalence-is-connected-map g cg) @@ -616,7 +617,8 @@ $k$-connected, then $f$ is a $k+1$-equivalence. ```agda is-succ-truncation-equiv-is-succ-connected-comp : is-connected-map k f → is-truncation-equivalence (succ-𝕋 k) f - is-succ-truncation-equiv-is-succ-connected-comp cf = is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor + is-succ-truncation-equiv-is-succ-connected-comp cf = + is-succ-truncation-equivalence-right-factor-is-succ-connected-map-left-factor ( is-connected-map-right-factor-is-succ-connected-map-right-factor cf) ``` From 7f2f7cd125548dbd33dd21afc47e549bc5805533 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 22:50:43 +0200 Subject: [PATCH 14/69] `subuniverse-connected-maps` --- src/orthogonal-factorization-systems.lagda.md | 1 + .../subuniverse-connected-maps.lagda.md | 381 ++++++++++++++++++ 2 files changed, 382 insertions(+) create mode 100644 src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index e2a8bf7286..8b57f3f7ca 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -67,6 +67,7 @@ open import orthogonal-factorization-systems.sigma-closed-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-subuniverses public open import orthogonal-factorization-systems.stable-orthogonal-factorization-systems public +open import orthogonal-factorization-systems.subuniverse-connected-maps public open import orthogonal-factorization-systems.types-colocal-at-maps public open import orthogonal-factorization-systems.types-local-at-maps public open import orthogonal-factorization-systems.types-separated-at-maps public diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md new file mode 100644 index 0000000000..b83f7d9dde --- /dev/null +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -0,0 +1,381 @@ +# Connected maps with respect to a subuniverse + +```agda +module orthogonal-factorization-systems.subuniverse-connected-maps where +``` + +
Imports + +```agda +open import foundation.connected-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.dependent-universal-property-equivalences +open import foundation.equivalences +open import foundation.equivalences-arrows +open import foundation.function-extensionality +open import foundation.functoriality-dependent-function-types +open import foundation.fundamental-theorem-of-identity-types +open import foundation.homotopy-induction +open import foundation.iterated-successors-truncation-levels +open import foundation.precomposition-dependent-functions +open import foundation.precomposition-type-families +open import foundation.propositional-truncations +open import foundation.structure-identity-principle +open import foundation.subtype-identity-principle +open import foundation.subuniverses +open import foundation.surjective-maps +open import foundation.truncated-types +open import foundation.truncation-levels +open import foundation.truncations +open import foundation.univalence +open import foundation.universal-property-dependent-pair-types +open import foundation.universal-property-family-of-fibers-of-maps +open import foundation.universe-levels + +open import foundation-core.contractible-maps +open import foundation-core.embeddings +open import foundation-core.fibers-of-maps +open import foundation-core.function-types +open import foundation-core.functoriality-dependent-pair-types +open import foundation-core.homotopies +open import foundation-core.identity-types +open import foundation-core.propositions +open import foundation-core.subtypes +open import foundation-core.torsorial-type-families +open import foundation-core.truncated-maps +``` + +
+ +## Idea + +Given a [subuniverse](foundation.subuniverses.md) `K`, A map `f : A → B` is said +to be +{{#concept "`K`-connected" Disambiguation="map of types" Agda=is-subuniverse-connected-map}} +if it satisfies the +{{#concept "universal property" Disambiguation="subuniverse connected map of types"}} +of `K`-connected maps: + +For every `K`-valued family `U` over `B`, the +[dependent precomposition map](foundation-core.precomposition-dependent-functions.md) + +```text + - ∘ f : ((b : B) → U b) → ((a : A) → U (f a)) +``` + +is an equivalence. + +## Definitions + +### The predicate on maps of being `K`-connected + +```agda +is-subuniverse-connected-map-Prop : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → + (A → B) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) +is-subuniverse-connected-map-Prop K {B = B} f = + Π-Prop (B → type-subuniverse K) λ U → is-equiv-Prop (precomp-Π f (pr1 ∘ U)) + +is-subuniverse-connected-map : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → + (A → B) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) +is-subuniverse-connected-map K f = + type-Prop (is-subuniverse-connected-map-Prop K f) + +is-prop-is-subuniverse-connected-map : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) → + is-prop (is-subuniverse-connected-map K f) +is-prop-is-subuniverse-connected-map K f = + is-prop-type-Prop (is-subuniverse-connected-map-Prop K f) +``` + +### The type of `K`-connected maps between two types + +```agda +subuniverse-connected-map : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) → + UU l1 → UU l2 → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) +subuniverse-connected-map K A B = + type-subtype (is-subuniverse-connected-map-Prop K {A} {B}) + +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} + where + + map-subuniverse-connected-map : subuniverse-connected-map K A B → A → B + map-subuniverse-connected-map = + inclusion-subtype (is-subuniverse-connected-map-Prop K) + + is-subuniverse-connected-map-subuniverse-connected-map : + (f : subuniverse-connected-map K A B) → + is-subuniverse-connected-map K (map-subuniverse-connected-map f) + is-subuniverse-connected-map-subuniverse-connected-map = + is-in-subtype-inclusion-subtype (is-subuniverse-connected-map-Prop K) + + emb-inclusion-subuniverse-connected-map : + subuniverse-connected-map K A B ↪ (A → B) + emb-inclusion-subuniverse-connected-map = + emb-subtype (is-subuniverse-connected-map-Prop K) +``` + +### The type of `K`-connected maps into a type + +```agda +Subuniverse-Connected-Map : + {l1 l3 l4 : Level} (l2 : Level) (K : subuniverse l3 l4) (A : UU l1) → + UU (l1 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l2) +Subuniverse-Connected-Map l2 K A = Σ (UU l2) (subuniverse-connected-map K A) + +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} (f : Subuniverse-Connected-Map l2 K A) + where + + type-Subuniverse-Connected-Map : UU l2 + type-Subuniverse-Connected-Map = pr1 f + + subuniverse-connected-map-Subuniverse-Connected-Map : + subuniverse-connected-map K A type-Subuniverse-Connected-Map + subuniverse-connected-map-Subuniverse-Connected-Map = pr2 f + + map-Subuniverse-Connected-Map : A → type-Subuniverse-Connected-Map + map-Subuniverse-Connected-Map = + map-subuniverse-connected-map K + ( subuniverse-connected-map-Subuniverse-Connected-Map) + + is-subuniverse-connected-map-Subuniverse-Connected-Map : + is-subuniverse-connected-map K map-Subuniverse-Connected-Map + is-subuniverse-connected-map-Subuniverse-Connected-Map = + is-subuniverse-connected-map-subuniverse-connected-map K + ( subuniverse-connected-map-Subuniverse-Connected-Map) +``` + +## Properties + +### Characterizing equality of `K`-connected maps + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} + where + + htpy-subuniverse-connected-map : + (f g : subuniverse-connected-map K A B) → UU (l1 ⊔ l2) + htpy-subuniverse-connected-map f g = + map-subuniverse-connected-map K f ~ map-subuniverse-connected-map K g + + refl-htpy-subuniverse-connected-map : + (f : subuniverse-connected-map K A B) → htpy-subuniverse-connected-map f f + refl-htpy-subuniverse-connected-map f = refl-htpy + + is-torsorial-htpy-subuniverse-connected-map : + (f : subuniverse-connected-map K A B) → + is-torsorial (htpy-subuniverse-connected-map f) + is-torsorial-htpy-subuniverse-connected-map f = + is-torsorial-Eq-subtype + ( is-torsorial-htpy (map-subuniverse-connected-map K f)) + ( is-prop-is-subuniverse-connected-map K) + ( map-subuniverse-connected-map K f) + ( refl-htpy-subuniverse-connected-map f) + ( is-subuniverse-connected-map-subuniverse-connected-map K f) + + htpy-eq-connected-map : + (f g : subuniverse-connected-map K A B) → + f = g → htpy-subuniverse-connected-map f g + htpy-eq-connected-map f .f refl = refl-htpy-subuniverse-connected-map f + + is-equiv-htpy-eq-connected-map : + (f g : subuniverse-connected-map K A B) → + is-equiv (htpy-eq-connected-map f g) + is-equiv-htpy-eq-connected-map f = + fundamental-theorem-id + ( is-torsorial-htpy-subuniverse-connected-map f) + ( htpy-eq-connected-map f) + + extensionality-connected-map : + (f g : subuniverse-connected-map K A B) → + (f = g) ≃ htpy-subuniverse-connected-map f g + pr1 (extensionality-connected-map f g) = htpy-eq-connected-map f g + pr2 (extensionality-connected-map f g) = is-equiv-htpy-eq-connected-map f g + + eq-htpy-subuniverse-connected-map : + (f g : subuniverse-connected-map K A B) → + htpy-subuniverse-connected-map f g → (f = g) + eq-htpy-subuniverse-connected-map f g = + map-inv-equiv (extensionality-connected-map f g) +``` + +### All maps are `Contr`-connected + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where + + is-Contr-connected-map : is-subuniverse-connected-map (is-contr-Prop {l3}) f + is-Contr-connected-map U = + is-equiv-is-contr + ( precomp-Π f (pr1 ∘ U)) + ( is-contr-Π (pr2 ∘ U)) + ( is-contr-Π (pr2 ∘ U ∘ f)) +``` + +### Equivalences are `K`-connected for any `K` + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} + where + + is-subuniverse-connected-map-is-equiv : + {f : A → B} → is-equiv f → is-subuniverse-connected-map K f + is-subuniverse-connected-map-is-equiv H U = + is-equiv-precomp-Π-is-equiv H (pr1 ∘ U) + + is-subuniverse-connected-map-equiv : + (e : A ≃ B) → is-subuniverse-connected-map K (map-equiv e) + is-subuniverse-connected-map-equiv e = + is-subuniverse-connected-map-is-equiv (is-equiv-map-equiv e) + + subuniverse-connected-map-equiv : + (A ≃ B) → subuniverse-connected-map K A B + subuniverse-connected-map-equiv e = + ( map-equiv e , is-subuniverse-connected-map-equiv e) +``` + +### The composition of two `K`-connected maps is `K`-connected + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-subuniverse-connected-map-comp : + {g : B → C} {f : A → B} → + is-subuniverse-connected-map K g → + is-subuniverse-connected-map K f → + is-subuniverse-connected-map K (g ∘ f) + is-subuniverse-connected-map-comp {g} {f} G F U = + is-equiv-comp + ( precomp-Π f (pr1 ∘ U ∘ g)) + ( precomp-Π g (pr1 ∘ U)) + ( G U) + ( F (U ∘ g)) + + comp-subuniverse-connected-map : + subuniverse-connected-map K B C → + subuniverse-connected-map K A B → + subuniverse-connected-map K A C + pr1 (comp-subuniverse-connected-map g f) = + map-subuniverse-connected-map K g ∘ map-subuniverse-connected-map K f + pr2 (comp-subuniverse-connected-map g f) = + is-subuniverse-connected-map-comp + ( is-subuniverse-connected-map-subuniverse-connected-map K g) + ( is-subuniverse-connected-map-subuniverse-connected-map K f) +``` + +### Right cancellation of `K`-connected maps + +```agda +is-subuniverse-connected-map-left-factor : + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {C : UU l3} + {g : B → C} {h : A → B} → + is-subuniverse-connected-map K h → + is-subuniverse-connected-map K (g ∘ h) → + is-subuniverse-connected-map K g +is-subuniverse-connected-map-left-factor K {g = g} {h} H GH U = + is-equiv-right-factor + ( precomp-Π h (pr1 ∘ U ∘ g)) + ( precomp-Π g (pr1 ∘ U)) + ( H (U ∘ g)) + ( GH U) +``` + +### The total map induced by a family of maps is `K`-connected if and only if all maps in the family are `K`-connected + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : A → UU l2} {C : A → UU l3} + (f : (x : A) → B x → C x) + where + + is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map : + ((x : A) → is-subuniverse-connected-map K (f x)) → + is-subuniverse-connected-map K (tot f) + is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map H U = + is-equiv-source-is-equiv-target-equiv-arrow + ( precomp-Π (tot f) (pr1 ∘ U)) + ( map-Π (λ i → precomp-Π (f i) (pr1 ∘ U ∘ (i ,_)))) + ( equiv-ev-pair , equiv-ev-pair , refl-htpy) + ( is-equiv-map-Π-is-fiberwise-equiv (λ i → H i (U ∘ (i ,_)))) + + -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot : + -- is-subuniverse-connected-map K (tot f) → + -- (x : A) → is-subuniverse-connected-map K (f x) + -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot H = {! !} + -- -- is-subuniverse-connected-equiv (inv-compute-fiber-tot f (x , y)) (H (x , y)) +``` + + + +### Characterization of the identity type of `Subuniverse-Connected-Map l2 K A` + +```agda +equiv-Subuniverse-Connected-Map : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} → + (f g : Subuniverse-Connected-Map l2 K A) → UU (l1 ⊔ l2) +equiv-Subuniverse-Connected-Map K f g = + Σ ( type-Subuniverse-Connected-Map K f ≃ type-Subuniverse-Connected-Map K g) + ( λ e → + map-equiv e ∘ map-Subuniverse-Connected-Map K f ~ + map-Subuniverse-Connected-Map K g) + +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} + (f : Subuniverse-Connected-Map l2 K A) + where + + id-equiv-Subuniverse-Connected-Map : equiv-Subuniverse-Connected-Map K f f + id-equiv-Subuniverse-Connected-Map = (id-equiv , refl-htpy) + + is-torsorial-equiv-Subuniverse-Connected-Map : + is-torsorial (equiv-Subuniverse-Connected-Map K f) + is-torsorial-equiv-Subuniverse-Connected-Map = + is-torsorial-Eq-structure + ( is-torsorial-equiv (type-Subuniverse-Connected-Map K f)) + ( type-Subuniverse-Connected-Map K f , id-equiv) + ( is-torsorial-htpy-subuniverse-connected-map K + ( subuniverse-connected-map-Subuniverse-Connected-Map K f)) + + equiv-eq-Subuniverse-Connected-Map : + (g : Subuniverse-Connected-Map l2 K A) → + f = g → equiv-Subuniverse-Connected-Map K f g + equiv-eq-Subuniverse-Connected-Map .f refl = + id-equiv-Subuniverse-Connected-Map + + is-equiv-equiv-eq-Subuniverse-Connected-Map : + (g : Subuniverse-Connected-Map l2 K A) → + is-equiv (equiv-eq-Subuniverse-Connected-Map g) + is-equiv-equiv-eq-Subuniverse-Connected-Map = + fundamental-theorem-id + ( is-torsorial-equiv-Subuniverse-Connected-Map) + ( equiv-eq-Subuniverse-Connected-Map) + + extensionality-Subuniverse-Connected-Map : + (g : Subuniverse-Connected-Map l2 K A) → + (f = g) ≃ equiv-Subuniverse-Connected-Map K f g + extensionality-Subuniverse-Connected-Map g = + ( equiv-eq-Subuniverse-Connected-Map g , + is-equiv-equiv-eq-Subuniverse-Connected-Map g) + + eq-equiv-Subuniverse-Connected-Map : + (g : Subuniverse-Connected-Map l2 K A) → + equiv-Subuniverse-Connected-Map K f g → f = g + eq-equiv-Subuniverse-Connected-Map g = + map-inv-equiv (extensionality-Subuniverse-Connected-Map g) +``` From c0f65ce2d631d3aa62f925d74f0d98d068d5151e Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 25 Sep 2025 23:08:56 +0200 Subject: [PATCH 15/69] `K`-connected maps are closed under cobase change --- .../subuniverse-connected-maps.lagda.md | 24 +++++++++++++++++++ 1 file changed, 24 insertions(+) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index b83f7d9dde..d64f89b801 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -44,6 +44,10 @@ open import foundation-core.propositions open import foundation-core.subtypes open import foundation-core.torsorial-type-families open import foundation-core.truncated-maps + +open import synthetic-homotopy-theory.cocones-under-spans +open import synthetic-homotopy-theory.dependent-pullback-property-pushouts +open import synthetic-homotopy-theory.pushouts ``` @@ -379,3 +383,23 @@ module _ eq-equiv-Subuniverse-Connected-Map g = map-inv-equiv (extensionality-Subuniverse-Connected-Map g) ``` + +### `K`-connected maps are closed under cobase change + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} (K : subuniverse l5 l6) + {S : UU l1} {A : UU l2} {B : UU l3} + (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) + where + + is-subuniverse-connected-map-cobase-change : + dependent-pullback-property-pushout f g c → + is-subuniverse-connected-map K g → + is-subuniverse-connected-map K (horizontal-map-cocone f g c) + is-subuniverse-connected-map-cobase-change H G U = + is-equiv-vertical-map-is-pullback _ _ + ( cone-dependent-pullback-property-pushout f g c (pr1 ∘ U)) + ( G (U ∘ vertical-map-cocone f g c)) + ( H (pr1 ∘ U)) +``` From f00102e59b0437dd8cb3c98768f6188e202d62e1 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 26 Sep 2025 10:40:58 +0200 Subject: [PATCH 16/69] `K`-connected maps are closed under retracts of maps informal proof --- .../subuniverse-connected-maps.lagda.md | 34 +++++++++++++++++++ 1 file changed, 34 insertions(+) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index d64f89b801..0a97ab267f 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -403,3 +403,37 @@ module _ ( G (U ∘ vertical-map-cocone f g c)) ( H (pr1 ∘ U)) ``` + +### `K`-connected maps are closed under retracts of maps + +**Proof.** Given a retract of maps + +```text + i r + A' ------> A ------> A' + | | | + f'| I f R | f' + ∨ ∨ ∨ + B' ------> B ------> B' + i' r' +``` + +with higher coherence `α`, and a `K`-valued family `U` over `B'` there is an +induced retract of dependent precomposition maps + +```text + Π(A'),(U∘f') <--- Π(A'),(U∘r'∘i'∘f) <--- Π(A),(U∘r'∘f) <--- Π(A'),(U∘f') + ∧ ∧ ∧ + | α* □ Π(I),(U∘r') | Π(R),U | + Π(f'),U | Π(f),(U∘r') | Π(f'),U + | | | + Π(B'),U <--------- Π(B'),(U∘r'∘i') <----- Π(B),(U∘r') <--- Π(B'),(U), +``` + +and since equivalences are closed under retracts of maps, if `f` is +`K`-connected then so is `f'`. ∎ + +Note that, since equivalences are already closed under noncoherent retracts of +maps, we are not obligated to produce the higher coherence of this retract. + +> This remains to be formalized. From 8943ae1cc64c2351a266241e7ef3bca75654b64d Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Sun, 28 Sep 2025 20:43:57 +0200 Subject: [PATCH 17/69] edits --- .../pushouts.lagda.md | 26 ++++++++++++++++++- 1 file changed, 25 insertions(+), 1 deletion(-) diff --git a/src/synthetic-homotopy-theory/pushouts.lagda.md b/src/synthetic-homotopy-theory/pushouts.lagda.md index 5242c08e65..2222635327 100644 --- a/src/synthetic-homotopy-theory/pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/pushouts.lagda.md @@ -34,6 +34,7 @@ open import synthetic-homotopy-theory.dependent-universal-property-pushouts open import synthetic-homotopy-theory.flattening-lemma-pushouts open import synthetic-homotopy-theory.induction-principle-pushouts open import synthetic-homotopy-theory.pullback-property-pushouts +open import synthetic-homotopy-theory.dependent-pullback-property-pushouts open import synthetic-homotopy-theory.universal-property-pushouts ``` @@ -460,7 +461,7 @@ module _ ( up-pushout f g) ``` -### Pushout cocones satisfy the pullback property of the pushout +### Cocones satisfy the pullback property of the pushout if and only if they are pushouts ```agda module _ @@ -483,6 +484,29 @@ module _ ( universal-property-pushout-pullback-property-pushout f g c pb) ``` +### Cocones satisfy the dependent pullback property of the pushout if and only if they are pushouts + +```agda +module _ + {l1 l2 l3 l4 : Level} {S : UU l1} {A : UU l2} {B : UU l3} + (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) + where + + abstract + dependent-pullback-property-pushout-is-pushout : + is-pushout f g c → dependent-pullback-property-pushout f g c + dependent-pullback-property-pushout-is-pushout po = + dependent-pullback-property-pullback-property-pushout f g c + ( pullback-property-pushout-is-pushout f g c po) + + abstract + is-pushout-dependent-pullback-property-pushout : + dependent-pullback-property-pushout f g c → is-pushout f g c + is-pushout-dependent-pullback-property-pushout pb = + is-pushout-pullback-property-pushout f g c + ( pullback-property-dependent-pullback-property-pushout f g c pb) +``` + ### Fibers of the cogap map We characterize the [fibers](foundation-core.fibers-of-maps.md) of the cogap map From 3ed7b9654cfb855f35dcd5fae166b605802ed9f3 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Sun, 28 Sep 2025 20:44:10 +0200 Subject: [PATCH 18/69] Coproducts of `K`-connected maps --- .../subuniverse-connected-maps.lagda.md | 101 ++++++++++++------ 1 file changed, 71 insertions(+), 30 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 0a97ab267f..ab30d00169 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -12,8 +12,13 @@ open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences open import foundation.equivalences +open import foundation.coproduct-types +open import orthogonal-factorization-systems.orthogonal-maps open import foundation.equivalences-arrows open import foundation.function-extensionality +open import foundation.functoriality-cartesian-product-types +open import foundation.functoriality-coproduct-types +open import foundation.universal-property-coproduct-types open import foundation.functoriality-dependent-function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopy-induction @@ -70,6 +75,10 @@ For every `K`-valued family `U` over `B`, the is an equivalence. +Equivalently, a `K`-connected map `f : A → B` is a map that is +[left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to maps +`h : C → B` whose fibers are in `K`. + ## Definitions ### The predicate on maps of being `K`-connected @@ -299,34 +308,6 @@ is-subuniverse-connected-map-left-factor K {g = g} {h} H GH U = ( GH U) ``` -### The total map induced by a family of maps is `K`-connected if and only if all maps in the family are `K`-connected - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : A → UU l2} {C : A → UU l3} - (f : (x : A) → B x → C x) - where - - is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map : - ((x : A) → is-subuniverse-connected-map K (f x)) → - is-subuniverse-connected-map K (tot f) - is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map H U = - is-equiv-source-is-equiv-target-equiv-arrow - ( precomp-Π (tot f) (pr1 ∘ U)) - ( map-Π (λ i → precomp-Π (f i) (pr1 ∘ U ∘ (i ,_)))) - ( equiv-ev-pair , equiv-ev-pair , refl-htpy) - ( is-equiv-map-Π-is-fiberwise-equiv (λ i → H i (U ∘ (i ,_)))) - - -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot : - -- is-subuniverse-connected-map K (tot f) → - -- (x : A) → is-subuniverse-connected-map K (f x) - -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot H = {! !} - -- -- is-subuniverse-connected-equiv (inv-compute-fiber-tot f (x , y)) (H (x , y)) -``` - - - ### Characterization of the identity type of `Subuniverse-Connected-Map l2 K A` ```agda @@ -393,15 +374,23 @@ module _ (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) where - is-subuniverse-connected-map-cobase-change : + is-subuniverse-connected-map-cobase-change' : dependent-pullback-property-pushout f g c → is-subuniverse-connected-map K g → is-subuniverse-connected-map K (horizontal-map-cocone f g c) - is-subuniverse-connected-map-cobase-change H G U = + is-subuniverse-connected-map-cobase-change' H G U = is-equiv-vertical-map-is-pullback _ _ ( cone-dependent-pullback-property-pushout f g c (pr1 ∘ U)) ( G (U ∘ vertical-map-cocone f g c)) ( H (pr1 ∘ U)) + + is-subuniverse-connected-map-cobase-change : + is-pushout f g c → + is-subuniverse-connected-map K g → + is-subuniverse-connected-map K (horizontal-map-cocone f g c) + is-subuniverse-connected-map-cobase-change H = + is-subuniverse-connected-map-cobase-change' + ( dependent-pullback-property-pushout-is-pushout f g c H) ``` ### `K`-connected maps are closed under retracts of maps @@ -437,3 +426,55 @@ Note that, since equivalences are already closed under noncoherent retracts of maps, we are not obligated to produce the higher coherence of this retract. > This remains to be formalized. + +### The total map induced by a family of maps is `K`-connected if and only if all maps in the family are `K`-connected + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : A → UU l2} {C : A → UU l3} + (f : (x : A) → B x → C x) + where + + is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map : + ((x : A) → is-subuniverse-connected-map K (f x)) → + is-subuniverse-connected-map K (tot f) + is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map H U = + is-equiv-source-is-equiv-target-equiv-arrow + ( precomp-Π (tot f) (pr1 ∘ U)) + ( map-Π (λ i → precomp-Π (f i) (pr1 ∘ U ∘ (i ,_)))) + ( equiv-ev-pair , equiv-ev-pair , refl-htpy) + ( is-equiv-map-Π-is-fiberwise-equiv (λ i → H i (U ∘ (i ,_)))) + + -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot : + -- is-subuniverse-connected-map K (tot f) → + -- (x : A) → is-subuniverse-connected-map K (f x) + -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot H = {! !} + -- -- is-subuniverse-connected-equiv (inv-compute-fiber-tot f (x , y)) (H (x , y)) +``` + +### Coproducts of `K`-connected maps + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} (K : subuniverse l5 l6) + {A : UU l1} {B : UU l2} {A' : UU l3} {B' : UU l4} + {f : A → B} {f' : A' → B'} + where + + is-subuniverse-connected-map-coproduct : + is-subuniverse-connected-map K f → + is-subuniverse-connected-map K f' → + is-subuniverse-connected-map K (map-coproduct f f') + is-subuniverse-connected-map-coproduct F F' U = + is-equiv-source-is-equiv-target-equiv-arrow + ( precomp-Π (map-coproduct f f') (pr1 ∘ U)) + ( map-product + ( precomp-Π f (pr1 ∘ U ∘ inl)) + ( precomp-Π f' (pr1 ∘ U ∘ inr))) + ( equiv-dependent-universal-property-coproduct (pr1 ∘ U) , + equiv-dependent-universal-property-coproduct + ( pr1 ∘ U ∘ map-coproduct f f') , + refl-htpy) + ( is-equiv-map-product _ _ (F (U ∘ inl)) (F' (U ∘ inr))) +``` From 35d57cbda77aaef10aa5e4ed5379af2df0a7f6af Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Mon, 29 Sep 2025 12:28:09 +0200 Subject: [PATCH 19/69] edits --- .../subuniverse-connected-maps.lagda.md | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index ab30d00169..07ee93aad2 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -407,8 +407,8 @@ module _ i' r' ``` -with higher coherence `α`, and a `K`-valued family `U` over `B'` there is an -induced retract of dependent precomposition maps +with higher coherence `α`, and a `K`-valued family `U` over `B'` there is by +functoriality an induced retract of dependent precomposition maps ```text Π(A'),(U∘f') <--- Π(A'),(U∘r'∘i'∘f) <--- Π(A),(U∘r'∘f) <--- Π(A'),(U∘f') @@ -472,7 +472,8 @@ module _ ( map-product ( precomp-Π f (pr1 ∘ U ∘ inl)) ( precomp-Π f' (pr1 ∘ U ∘ inr))) - ( equiv-dependent-universal-property-coproduct (pr1 ∘ U) , + ( equiv-dependent-universal-property-coproduct + ( pr1 ∘ U) , equiv-dependent-universal-property-coproduct ( pr1 ∘ U ∘ map-coproduct f f') , refl-htpy) From 65cf586e327b8a0d80f1f4699ee0432f7155f7a4 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 00:29:48 +0200 Subject: [PATCH 20/69] lots of edits --- .../commuting-squares-of-maps.lagda.md | 8 ++ .../dependent-identifications.lagda.md | 10 ++ ...recomposition-dependent-functions.lagda.md | 8 +- .../transport-along-identifications.lagda.md | 17 +++- ...entifications-dependent-functions.lagda.md | 5 + src/foundation/dependent-homotopies.lagda.md | 11 ++- src/foundation/morphisms-arrows.lagda.md | 98 +++++++++++++++++-- ...recomposition-dependent-functions.lagda.md | 38 +++++++ .../transport-along-homotopies.lagda.md | 40 +++++++- .../transport-along-identifications.lagda.md | 24 +++-- .../truncation-equivalences.lagda.md | 51 +++++++--- .../pushouts.lagda.md | 2 +- 12 files changed, 270 insertions(+), 42 deletions(-) diff --git a/src/foundation-core/commuting-squares-of-maps.lagda.md b/src/foundation-core/commuting-squares-of-maps.lagda.md index 6cce4fb386..f1d668956a 100644 --- a/src/foundation-core/commuting-squares-of-maps.lagda.md +++ b/src/foundation-core/commuting-squares-of-maps.lagda.md @@ -82,6 +82,14 @@ module _ pasting-horizontal-coherence-square-maps sq-left sq-right = (bottom-right ·l sq-left) ∙h (sq-right ·r top-left) + pasting-horizontal-coherence-square-maps' : + coherence-square-maps' top-left left mid bottom-left → + coherence-square-maps' top-right mid right bottom-right → + coherence-square-maps' + (top-right ∘ top-left) left right (bottom-right ∘ bottom-left) + pasting-horizontal-coherence-square-maps' sq-left sq-right = + (sq-right ·r top-left) ∙h (bottom-right ·l sq-left) + pasting-horizontal-up-to-htpy-coherence-square-maps : {top : A → C} (H : coherence-triangle-maps top top-right top-left) {bottom : X → Z} diff --git a/src/foundation-core/dependent-identifications.lagda.md b/src/foundation-core/dependent-identifications.lagda.md index 55178798aa..434dd59b25 100644 --- a/src/foundation-core/dependent-identifications.lagda.md +++ b/src/foundation-core/dependent-identifications.lagda.md @@ -47,6 +47,16 @@ refl-dependent-identification : {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x : A} {y : B x} → dependent-identification B refl y y refl-dependent-identification B = refl + +dependent-identification' : + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x x' : A} (p : x = x') → + B x → B x' → UU l2 +dependent-identification' B p u v = (u = inv-tr B p v) + +refl-dependent-identification' : + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x : A} {y : B x} → + dependent-identification' B refl y y +refl-dependent-identification' B = refl ``` ### Iterated dependent identifications diff --git a/src/foundation-core/precomposition-dependent-functions.lagda.md b/src/foundation-core/precomposition-dependent-functions.lagda.md index 009dd7f292..57d0f7b860 100644 --- a/src/foundation-core/precomposition-dependent-functions.lagda.md +++ b/src/foundation-core/precomposition-dependent-functions.lagda.md @@ -50,7 +50,11 @@ module _ {f g : A → B} (H : f ~ g) (C : B → UU l3) (h : (y : B) → C y) where - htpy-precomp-Π : + htpy-htpy-precomp-Π : dependent-homotopy (λ _ → C) H (precomp-Π f C h) (precomp-Π g C h) - htpy-precomp-Π x = apd h (H x) + htpy-htpy-precomp-Π x = apd h (H x) + + htpy-htpy-precomp-Π' : + dependent-homotopy' (λ _ → C) H (precomp-Π f C h) (precomp-Π g C h) + htpy-htpy-precomp-Π' x = apd' h (H x) ``` diff --git a/src/foundation-core/transport-along-identifications.lagda.md b/src/foundation-core/transport-along-identifications.lagda.md index 49b511e7e0..0eff7e01d7 100644 --- a/src/foundation-core/transport-along-identifications.lagda.md +++ b/src/foundation-core/transport-along-identifications.lagda.md @@ -32,9 +32,15 @@ recorded in ### Transport ```agda -tr : - {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x y : A} (p : x = y) → B x → B y -tr B refl b = b +module _ + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x y : A} + where + + tr : x = y → B x → B y + tr refl b = b + + inv-tr : y = x → B x → B y + inv-tr p = tr (inv p) ``` ## Properties @@ -91,6 +97,11 @@ module _ {x y z : A} (p : x = y) (q : y = z) (b : B x) → tr B (p ∙ q) b = tr B q (tr B p b) tr-concat refl q b = refl + + comp-tr : + {x y z : A} (p : x = y) (q : y = z) (b : B x) → + tr B q (tr B p b) = tr B (p ∙ q) b + comp-tr p q b = inv (tr-concat p q b) ``` ### Transposing transport along the inverse of an identification diff --git a/src/foundation/action-on-identifications-dependent-functions.lagda.md b/src/foundation/action-on-identifications-dependent-functions.lagda.md index eecad428e0..d991baa53b 100644 --- a/src/foundation/action-on-identifications-dependent-functions.lagda.md +++ b/src/foundation/action-on-identifications-dependent-functions.lagda.md @@ -41,6 +41,11 @@ apd : {l1 l2 : Level} {A : UU l1} {B : A → UU l2} (f : (x : A) → B x) {x y : A} (p : x = y) → dependent-identification B p (f x) (f y) apd f refl = refl + +apd' : + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} (f : (x : A) → B x) {x y : A} + (p : x = y) → dependent-identification' B p (f x) (f y) +apd' f refl = refl ``` ## See also diff --git a/src/foundation/dependent-homotopies.lagda.md b/src/foundation/dependent-homotopies.lagda.md index eaac3938b5..f78c58c715 100644 --- a/src/foundation/dependent-homotopies.lagda.md +++ b/src/foundation/dependent-homotopies.lagda.md @@ -45,6 +45,11 @@ module _ ((x : A) → C x (f x)) → ((x : A) → C x (g x)) → UU (l1 ⊔ l3) dependent-homotopy f' g' = (x : A) → dependent-identification (C x) (H x) (f' x) (g' x) + + dependent-homotopy' : + ((x : A) → C x (f x)) → ((x : A) → C x (g x)) → UU (l1 ⊔ l3) + dependent-homotopy' f' g' = + (x : A) → dependent-identification' (C x) (H x) (f' x) (g' x) ``` ### The reflexive dependent homotopy @@ -57,7 +62,11 @@ module _ refl-dependent-homotopy : {f' : (x : A) → C x (f x)} → dependent-homotopy C refl-htpy f' f' - refl-dependent-homotopy = refl-htpy + refl-dependent-homotopy x = refl-dependent-identification (C x) + + refl-dependent-homotopy' : + {f' : (x : A) → C x (f x)} → dependent-homotopy' C refl-htpy f' f' + refl-dependent-homotopy' x = refl-dependent-identification' (C x) ``` ### Iterated dependent homotopies diff --git a/src/foundation/morphisms-arrows.lagda.md b/src/foundation/morphisms-arrows.lagda.md index 7abfa99887..0c53c53eb7 100644 --- a/src/foundation/morphisms-arrows.lagda.md +++ b/src/foundation/morphisms-arrows.lagda.md @@ -10,16 +10,17 @@ module foundation.morphisms-arrows where open import foundation.cones-over-cospan-diagrams open import foundation.dependent-pair-types open import foundation.function-extensionality +open import foundation.homotopies open import foundation.postcomposition-functions open import foundation.precomposition-functions open import foundation.universe-levels open import foundation.whiskering-homotopies-composition open import foundation-core.commuting-squares-of-maps +open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.functoriality-dependent-function-types open import foundation-core.functoriality-dependent-pair-types -open import foundation-core.homotopies open import foundation-core.identity-types ``` @@ -74,6 +75,44 @@ module _ coh-hom-arrow = pr2 ∘ pr2 ``` +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + (f : A → B) (g : X → Y) + where + + coherence-hom-arrow' : (A → X) → (B → Y) → UU (l1 ⊔ l4) + coherence-hom-arrow' i = coherence-square-maps' i f g + + hom-arrow' : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + hom-arrow' = Σ (A → X) (λ i → Σ (B → Y) (coherence-hom-arrow' i)) + + map-domain-hom-arrow' : hom-arrow' → A → X + map-domain-hom-arrow' = pr1 + + map-codomain-hom-arrow' : hom-arrow' → B → Y + map-codomain-hom-arrow' = pr1 ∘ pr2 + + coh-hom-arrow' : + (h : hom-arrow') → + coherence-hom-arrow' (map-domain-hom-arrow' h) (map-codomain-hom-arrow' h) + coh-hom-arrow' = pr2 ∘ pr2 + + equiv-hom-arrow-hom-arrow' : hom-arrow' ≃ hom-arrow f g + equiv-hom-arrow-hom-arrow' = + equiv-tot (λ i → equiv-tot (λ j → equiv-inv-htpy (g ∘ i) (j ∘ f))) + + equiv-hom-arrow'-hom-arrow : hom-arrow f g ≃ hom-arrow' + equiv-hom-arrow'-hom-arrow = + equiv-tot (λ i → equiv-tot (λ j → equiv-inv-htpy (j ∘ f) (g ∘ i))) + + hom-arrow-hom-arrow' : hom-arrow' → hom-arrow f g + hom-arrow-hom-arrow' = map-equiv equiv-hom-arrow-hom-arrow' + + hom-arrow'-hom-arrow : hom-arrow f g → hom-arrow' + hom-arrow'-hom-arrow = map-equiv equiv-hom-arrow'-hom-arrow +``` + ## Operations ### The identity morphism of arrows @@ -94,11 +133,14 @@ where the homotopy `id ∘ f ~ f ∘ id` is the reflexivity homotopy. ```agda module _ - {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} + {l1 l2 : Level} {A : UU l1} {B : UU l2} where - id-hom-arrow : hom-arrow f f - id-hom-arrow = (id , id , refl-htpy) + id-hom-arrow' : (f : A → B) → hom-arrow f f + id-hom-arrow' f = (id , id , refl-htpy) + + id-hom-arrow : {f : A → B} → hom-arrow f f + id-hom-arrow {f} = id-hom-arrow' f ``` ### Composition of morphisms of arrows @@ -165,12 +207,48 @@ module _ ( coh-hom-arrow g h b) comp-hom-arrow : hom-arrow f h - pr1 comp-hom-arrow = - map-domain-comp-hom-arrow - pr1 (pr2 comp-hom-arrow) = - map-codomain-comp-hom-arrow - pr2 (pr2 comp-hom-arrow) = - coh-comp-hom-arrow + comp-hom-arrow = + ( map-domain-comp-hom-arrow , + map-codomain-comp-hom-arrow , + coh-comp-hom-arrow) +``` + +```agda +module _ + {l1 l2 l3 l4 l5 l6 : Level} + {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} {U : UU l5} {V : UU l6} + (f : A → B) (g : X → Y) (h : U → V) (b : hom-arrow' g h) (a : hom-arrow' f g) + where + + map-domain-comp-hom-arrow' : A → U + map-domain-comp-hom-arrow' = + map-domain-hom-arrow' g h b ∘ map-domain-hom-arrow' f g a + + map-codomain-comp-hom-arrow' : B → V + map-codomain-comp-hom-arrow' = + map-codomain-hom-arrow' g h b ∘ map-codomain-hom-arrow' f g a + + coh-comp-hom-arrow' : + coherence-hom-arrow' f h + ( map-domain-comp-hom-arrow') + ( map-codomain-comp-hom-arrow') + coh-comp-hom-arrow' = + pasting-horizontal-coherence-square-maps' + ( map-domain-hom-arrow' f g a) + ( map-domain-hom-arrow' g h b) + ( f) + ( g) + ( h) + ( map-codomain-hom-arrow' f g a) + ( map-codomain-hom-arrow' g h b) + ( coh-hom-arrow' f g a) + ( coh-hom-arrow' g h b) + + comp-hom-arrow' : hom-arrow' f h + comp-hom-arrow' = + ( map-domain-comp-hom-arrow' , + map-codomain-comp-hom-arrow' , + coh-comp-hom-arrow') ``` ### Transposing morphisms of arrows diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index fef97c901d..6d64834d25 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -10,9 +10,13 @@ open import foundation-core.precomposition-dependent-functions public ```agda open import foundation.action-on-identifications-functions +open import foundation.dependent-homotopies +open import foundation.dependent-identifications open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences open import foundation.function-extensionality +open import foundation.transport-along-homotopies +open import foundation.transport-along-identifications open import foundation.universe-levels open import foundation-core.commuting-squares-of-maps @@ -30,6 +34,40 @@ open import foundation-core.type-theoretic-principle-of-choice ## Properties +### The action of dependent precomposition on homotopies + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} + {f g : A → B} (H : f ~ g) (C : B → UU l3) (h : (y : B) → C y) + where + + eq-htpy-precomp-Π : (λ x → tr C (H x) (precomp-Π f C h x)) = precomp-Π g C h + eq-htpy-precomp-Π = eq-htpy (htpy-htpy-precomp-Π H C h) + + htpy-precomp-Π : + dependent-identification + ( λ v → (a : A) → C (v a)) + ( eq-htpy H) + ( precomp-Π f C h) + ( precomp-Π g C h) + htpy-precomp-Π = + compute-tr-htpy (λ _ → C) H (precomp-Π f C h) ∙ eq-htpy-precomp-Π + + eq-htpy-precomp-Π' : + precomp-Π f C h = (λ x → inv-tr C (H x) (precomp-Π g C h x)) + eq-htpy-precomp-Π' = eq-htpy (htpy-htpy-precomp-Π' H C h) + + htpy-precomp-Π' : + dependent-identification' + ( λ v → (a : A) → C (v a)) + ( eq-htpy H) + ( precomp-Π f C h) + ( precomp-Π g C h) + htpy-precomp-Π' = + eq-htpy-precomp-Π' ∙ inv (compute-inv-tr-htpy (λ _ → C) H (precomp-Π g C h)) +``` + ### Equivalences induce an equivalence from the type of homotopies between two dependent functions to the type of homotopies between their precomposites ```agda diff --git a/src/foundation/transport-along-homotopies.lagda.md b/src/foundation/transport-along-homotopies.lagda.md index 577c47c512..890a38d330 100644 --- a/src/foundation/transport-along-homotopies.lagda.md +++ b/src/foundation/transport-along-homotopies.lagda.md @@ -11,11 +11,11 @@ open import foundation.action-on-identifications-functions open import foundation.function-extensionality open import foundation.homotopy-induction open import foundation.transport-along-higher-identifications +open import foundation.transport-along-identifications open import foundation.universe-levels open import foundation-core.homotopies open import foundation-core.identity-types -open import foundation-core.transport-along-identifications ``` @@ -41,6 +41,10 @@ module _ (f ~ g) → ((x : A) → C x (f x)) → ((x : A) → C x (g x)) tr-htpy H f' x = tr (C x) (H x) (f' x) + inv-tr-htpy : + (f ~ g) → ((x : A) → C x (g x)) → ((x : A) → C x (f x)) + inv-tr-htpy H f' x = inv-tr (C x) (H x) (f' x) + tr-htpy² : {H K : f ~ g} (L : H ~ K) (f' : (x : A) → C x (f x)) → tr-htpy H f' ~ tr-htpy K f' @@ -82,3 +86,37 @@ module _ ( λ _ → statement-compute-tr-htpy) ( base-case-compute-tr-htpy) ``` + +### Inverse transport along a homotopy `H` is homotopic to inverse transport along the identification `eq-htpy H` + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} (C : (x : A) → B x → UU l3) + where + + statement-compute-inv-tr-htpy : + {f g : (x : A) → B x} (H : f ~ g) → UU (l1 ⊔ l3) + statement-compute-inv-tr-htpy H = + inv-tr (λ u → (x : A) → C x (u x)) (eq-htpy H) ~ inv-tr-htpy C H + + base-case-compute-inv-tr-htpy : + {f : (x : A) → B x} → statement-compute-inv-tr-htpy (refl-htpy' f) + base-case-compute-inv-tr-htpy = + htpy-eq (ap (inv-tr (λ u → (x : A) → C x (u x))) (eq-htpy-refl-htpy _)) + + abstract + compute-inv-tr-htpy : + {f g : (x : A) → B x} (H : f ~ g) → statement-compute-inv-tr-htpy H + compute-inv-tr-htpy {f} = + ind-htpy f + ( λ _ → statement-compute-inv-tr-htpy) + ( base-case-compute-inv-tr-htpy) + + compute-inv-tr-htpy-refl-htpy : + {f : (x : A) → B x} → + compute-inv-tr-htpy (refl-htpy' f) = base-case-compute-inv-tr-htpy + compute-inv-tr-htpy-refl-htpy {f} = + compute-ind-htpy f + ( λ _ → statement-compute-inv-tr-htpy) + ( base-case-compute-inv-tr-htpy) +``` diff --git a/src/foundation/transport-along-identifications.lagda.md b/src/foundation/transport-along-identifications.lagda.md index d137762977..09370b0d33 100644 --- a/src/foundation/transport-along-identifications.lagda.md +++ b/src/foundation/transport-along-identifications.lagda.md @@ -17,6 +17,8 @@ open import foundation-core.equivalences open import foundation-core.function-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.retractions +open import foundation-core.sections ``` @@ -41,25 +43,31 @@ module _ {l1 l2 : Level} {A : UU l1} (B : A → UU l2) {x y : A} where - inv-tr : x = y → B y → B x - inv-tr p = tr B (inv p) - - is-retraction-inv-tr : (p : x = y) → inv-tr p ∘ tr B p ~ id + is-retraction-inv-tr : (p : x = y) → is-retraction (tr B p) (inv-tr B p) is-retraction-inv-tr refl b = refl - is-section-inv-tr : (p : x = y) → tr B p ∘ inv-tr p ~ id + is-section-inv-tr : (p : x = y) → is-section (tr B p) (inv-tr B p) is-section-inv-tr refl b = refl is-equiv-tr : (p : x = y) → is-equiv (tr B p) is-equiv-tr p = is-equiv-is-invertible - ( inv-tr p) + ( inv-tr B p) ( is-section-inv-tr p) ( is-retraction-inv-tr p) + is-equiv-inv-tr : (p : x = y) → is-equiv (inv-tr B p) + is-equiv-inv-tr p = + is-equiv-is-invertible + ( tr B p) + ( is-retraction-inv-tr p) + ( is-section-inv-tr p) + equiv-tr : x = y → B x ≃ B y - pr1 (equiv-tr p) = tr B p - pr2 (equiv-tr p) = is-equiv-tr p + equiv-tr p = (tr B p , is-equiv-tr p) + + equiv-inv-tr : x = y → B y ≃ B x + equiv-inv-tr p = (inv-tr B p , is-equiv-inv-tr p) ``` ### Transporting along `refl` is the identity equivalence diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index 7fc1667d37..f7dd57292f 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -29,6 +29,7 @@ open import foundation-core.fibers-of-maps open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies +open import foundation-core.precomposition-dependent-functions open import foundation-core.precomposition-functions open import foundation-core.retractions open import foundation-core.sections @@ -312,27 +313,44 @@ module _ (g : B → C) (f : A ≃ B) → is-truncation-equivalence k g → is-truncation-equivalence k (g ∘ map-equiv f) - is-truncation-equivalence-equiv-is-truncation-equivalence g f eg = - is-truncation-equivalence-is-equiv-is-truncation-equivalence g - ( map-equiv f) - ( eg) - ( is-equiv-map-equiv f) + is-truncation-equivalence-equiv-is-truncation-equivalence g (f , ef) eg = + is-truncation-equivalence-is-equiv-is-truncation-equivalence g f eg ef is-truncation-equivalence-is-truncation-equivalence-equiv : (g : B ≃ C) (f : A → B) → is-truncation-equivalence k f → is-truncation-equivalence k (map-equiv g ∘ f) - is-truncation-equivalence-is-truncation-equivalence-equiv g f ef = - is-truncation-equivalence-is-truncation-equivalence-is-equiv - ( map-equiv g) - ( f) - ( is-equiv-map-equiv g) - ( ef) + is-truncation-equivalence-is-truncation-equivalence-equiv (g , eg) f ef = + is-truncation-equivalence-is-truncation-equivalence-is-equiv g f eg ef ``` ### The map on dependent pair types induced by the unit of the `(k+1)`-truncation is a `k`-equivalence -This is an instance of Lemma 2.27 in {{#cite CORS20}} listed below. +This is an instance of Lemma 2.27 in {{#cite CORS20}}. + +```agda +module _ + {l1 l2 : Level} {k : 𝕋} + {X : UU l1} (P : type-trunc k X → UU l2) + where + + map-Σ-map-base-unit-trunc' : + Σ X (P ∘ unit-trunc) → Σ (type-trunc k X) P + map-Σ-map-base-unit-trunc' = map-Σ-map-base unit-trunc P + + is-truncation-equivalence-map-Σ-map-base-unit-trunc' : + is-truncation-equivalence k map-Σ-map-base-unit-trunc' + is-truncation-equivalence-map-Σ-map-base-unit-trunc' = + is-truncation-equivalence-is-equiv-precomp k + ( map-Σ-map-base-unit-trunc') + ( λ l (Y , TY) → + is-equiv-equiv + ( equiv-ev-pair) + ( equiv-ev-pair) + ( refl-htpy) + ( dependent-universal-property-trunc + ( λ t → ((P t → Y) , is-trunc-function-type k TY)))) +``` ```agda module _ @@ -349,17 +367,18 @@ module _ is-truncation-equivalence-map-Σ-map-base-unit-trunc = is-truncation-equivalence-is-equiv-precomp k ( map-Σ-map-base-unit-trunc) - ( λ l X → + ( λ l (Y , TY) → is-equiv-equiv + {f = precomp (λ x → unit-trunc (pr1 x) , pr2 x) Y} + {g = precomp-Π unit-trunc (λ |x| → (b : P |x|) → Y)} ( equiv-ev-pair) ( equiv-ev-pair) ( refl-htpy) ( dependent-universal-property-trunc ( λ t → - ( ( P t → type-Truncated-Type X) , + ( ( P t → Y) , ( is-trunc-succ-is-trunc k - ( is-trunc-function-type k - ( is-trunc-type-Truncated-Type X))))))) + ( is-trunc-function-type k TY)))))) ``` ### There is a `k`-equivalence between the fiber of a map and the fiber of its `(k+1)`-truncation diff --git a/src/synthetic-homotopy-theory/pushouts.lagda.md b/src/synthetic-homotopy-theory/pushouts.lagda.md index 2222635327..025c11f075 100644 --- a/src/synthetic-homotopy-theory/pushouts.lagda.md +++ b/src/synthetic-homotopy-theory/pushouts.lagda.md @@ -30,11 +30,11 @@ open import reflection.erasing-equality open import synthetic-homotopy-theory.cocones-under-spans open import synthetic-homotopy-theory.dependent-cocones-under-spans +open import synthetic-homotopy-theory.dependent-pullback-property-pushouts open import synthetic-homotopy-theory.dependent-universal-property-pushouts open import synthetic-homotopy-theory.flattening-lemma-pushouts open import synthetic-homotopy-theory.induction-principle-pushouts open import synthetic-homotopy-theory.pullback-property-pushouts -open import synthetic-homotopy-theory.dependent-pullback-property-pushouts open import synthetic-homotopy-theory.universal-property-pushouts ``` From a36ebc13c5c27b5240fec0243ed5e93b7f9fc3e8 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 00:30:03 +0200 Subject: [PATCH 21/69] `is-equiv-postcomp-extension` --- .../extensions-maps.lagda.md | 31 +++++++++++++++++++ 1 file changed, 31 insertions(+) diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index ccdc296bd0..a2a5d62ec6 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -393,6 +393,31 @@ module _ extension-along-self = id , is-extension-along-self ``` +### Postcomposition of extensions by an equivalence is an equivalence + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-equiv-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-equiv g → + is-equiv (postcomp-extension f i g) + is-equiv-postcomp-extension f i g G = + is-equiv-map-Σ + ( is-extension f (g ∘ i)) + ( is-equiv-postcomp-is-equiv g G B) + ( λ j → + is-equiv-map-Π-is-fiberwise-equiv + ( λ x → is-emb-is-equiv G (i x) (j (f x)))) + + equiv-postcomp-extension : + (f : A → B) (i : A → X) (g : X ≃ Y) → + extension f i ≃ extension f (map-equiv g ∘ i) + equiv-postcomp-extension f i (g , G) = + ( postcomp-extension f i g , is-equiv-postcomp-extension f i g G) +``` + ### Postcomposition of extensions by an embedding is an embedding ```agda @@ -411,6 +436,12 @@ module _ is-emb-is-equiv ( is-equiv-map-Π-is-fiberwise-equiv ( λ x → H (i x) (j (f x))))) + + emb-postcomp-extension : + (f : A → B) (i : A → X) (g : X ↪ Y) → + extension f i ↪ extension f (map-emb g ∘ i) + emb-postcomp-extension f i (g , G) = + postcomp-extension f i g , is-emb-postcomp-extension f i g G ``` ## See also From f1945a4e75bc40ca97c422e498850ee235a9ec3b Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 00:33:07 +0200 Subject: [PATCH 22/69] map-separated types --- src/orthogonal-factorization-systems.lagda.md | 2 + ...parated-types-global-subuniverses.lagda.md | 89 +++++++++++++++++++ .../map-separated-types-subuniverses.lagda.md | 55 ++++++++++++ 3 files changed, 146 insertions(+) create mode 100644 src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md create mode 100644 src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 8b57f3f7ca..97d4fb470d 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -45,6 +45,8 @@ open import orthogonal-factorization-systems.localizations-at-global-subuniverse open import orthogonal-factorization-systems.localizations-at-maps public open import orthogonal-factorization-systems.localizations-at-subuniverses public open import orthogonal-factorization-systems.locally-small-modal-operators public +open import orthogonal-factorization-systems.map-separated-types-global-subuniverses public +open import orthogonal-factorization-systems.map-separated-types-subuniverses public open import orthogonal-factorization-systems.maps-local-at-maps public open import orthogonal-factorization-systems.mere-lifting-properties public open import orthogonal-factorization-systems.modal-induction public diff --git a/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md b/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md new file mode 100644 index 0000000000..c6dc6ba0e5 --- /dev/null +++ b/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md @@ -0,0 +1,89 @@ +# Map-separated types at global subuniverses + +```agda +module orthogonal-factorization-systems.map-separated-types-global-subuniverses where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.equivalences +open import foundation.global-subuniverses +open import foundation.universe-levels + +open import foundation-core.function-types +open import foundation-core.identity-types +open import foundation-core.propositions + +open import orthogonal-factorization-systems.extensions-maps +``` + +
+ +## Idea + +Consider a [global subuniverse](foundation.global-subuniverses.md) `K` and a map +`δ : X → Y`. A type `A` is said to be +{{#concept "`K`-separated relative to `δ`" Agda=is-map-separated}} if the type +of extensions of any map `f : X → A` along `δ` is in `K`. + +As a special case, a type `A` is _`K`-separated_ if it is `K`-separated relative +to the terminal projection at a two-element type. + +## Definitions + +### The predicate of being separated + +```agda +module _ + {α : Level → Level} {l1 l2 l3 : Level} (K : global-subuniverse α) + {X : UU l1} {Y : UU l2} (δ : X → Y) + where + + is-map-separated : (A : UU l3) → UU (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) + is-map-separated A = + (f : X → A) → is-in-global-subuniverse K (extension δ f) + + is-prop-is-map-separated : + (A : UU l3) → is-prop (is-map-separated A) + is-prop-is-map-separated A = + is-prop-Π (λ f → is-prop-is-in-global-subuniverse K (extension δ f)) + + is-map-separated-Prop : + (A : UU l3) → Prop (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) + is-map-separated-Prop A = + ( is-map-separated A , is-prop-is-map-separated A) +``` + +### The global subuniverse of `K`-separated types relative to `δ` + +```agda +module _ + {α : Level → Level} {l1 l2 : Level} (K : global-subuniverse α) + {X : UU l1} {Y : UU l2} (δ : X → Y) + where + + is-closed-under-equiv-map-separated-global-subuniverse : + {l4 l5 : Level} → + is-closed-under-equiv-subuniverses + ( λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) + ( λ l3 → is-map-separated-Prop {l3 = l3} K δ) + ( l4) + ( l5) + is-closed-under-equiv-map-separated-global-subuniverse A B e H f = + is-closed-under-equiv-global-subuniverse K + ( extension δ (map-inv-equiv e ∘ f)) + ( extension δ f) + ( inv-equiv (equiv-postcomp-extension δ f (inv-equiv e))) + ( H (map-inv-equiv e ∘ f)) + + map-separated-global-subuniverse : + global-subuniverse (λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) + map-separated-global-subuniverse = + λ where + .subuniverse-global-subuniverse l3 → + is-map-separated-Prop K δ + .is-closed-under-equiv-global-subuniverse → + is-closed-under-equiv-map-separated-global-subuniverse +``` diff --git a/src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md b/src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md new file mode 100644 index 0000000000..95109c020e --- /dev/null +++ b/src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md @@ -0,0 +1,55 @@ +# Map-separated types at subuniverses + +```agda +module orthogonal-factorization-systems.map-separated-types-subuniverses where +``` + +
Imports + +```agda +open import foundation.dependent-pair-types +open import foundation.subuniverses +open import foundation.universe-levels + +open import foundation-core.identity-types +open import foundation-core.propositions + +open import orthogonal-factorization-systems.extensions-maps +``` + +
+ +## Idea + +Consider a [subuniverse](foundation.subuniverses.md) `K` and a map `δ : X → Y`. +A type `A` is said to be +{{#concept "`K`-separated relative to `δ`" Agda=is-map-separated'}} if the type +of extensions of any map `f : X → A` along `δ` is in `K`. + +## Definitions + +### The predicate of being map-separated + +```agda +module _ + {l1 l2 l3 : Level} (K : subuniverse l1 l2) + {X : UU l1} {Y : UU l1} (δ : X → Y) + where + + is-map-separated' : (A : UU l1) → UU (l1 ⊔ l2) + is-map-separated' A = (f : X → A) → is-in-subuniverse K (extension δ f) + + is-prop-is-map-separated' : + (A : UU l1) → is-prop (is-map-separated' A) + is-prop-is-map-separated' A = + is-prop-Π (λ f → is-prop-is-in-subuniverse K (extension δ f)) + + is-map-separated'-Prop : + (A : UU l1) → Prop (l1 ⊔ l2) + is-map-separated'-Prop A = + ( is-map-separated' A , is-prop-is-map-separated' A) +``` + +## See also + +- [Map-separated types at global subuniverses](orthogonal-factorization-systems.map-separated-types-global-subuniverses.md) From 2140a78ab4398f4fe4c360fd78c6f19b4a94e393 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 00:34:26 +0200 Subject: [PATCH 23/69] `is-subuniverse-localization-map` --- .../localizations-at-subuniverses.lagda.md | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md index c12ced2306..14cc6f8f64 100644 --- a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md @@ -12,6 +12,7 @@ open import foundation.dependent-pair-types open import foundation.subuniverses open import foundation.universe-levels +open import orthogonal-factorization-systems.subuniverse-equivalences open import orthogonal-factorization-systems.types-local-at-maps ``` @@ -33,6 +34,16 @@ is an [equivalence](foundation-core.equivalences.md). ## Definition +### The predicate on a map of being a localization at a subuniverse + +```agda +is-subuniverse-localization-map : + {l1 l2 lP : Level} (P : subuniverse l1 lP) {A : UU l2} {PA : UU l1} + (η : A → PA) → UU (lsuc l1 ⊔ l2 ⊔ lP) +is-subuniverse-localization-map P {A} {PA} η = + is-in-subuniverse P PA × is-subuniverse-equiv P η +``` + ### The predicate of being a localization at a subuniverse ```agda From f946d4c4d842c9c239fadee44ff5cd884542cbe7 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 00:39:23 +0200 Subject: [PATCH 24/69] fix indentation --- src/foundation-core/commuting-squares-of-maps.lagda.md | 2 +- .../extensions-maps.lagda.md | 8 ++++---- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/foundation-core/commuting-squares-of-maps.lagda.md b/src/foundation-core/commuting-squares-of-maps.lagda.md index f1d668956a..8fe0c3410d 100644 --- a/src/foundation-core/commuting-squares-of-maps.lagda.md +++ b/src/foundation-core/commuting-squares-of-maps.lagda.md @@ -88,7 +88,7 @@ module _ coherence-square-maps' (top-right ∘ top-left) left right (bottom-right ∘ bottom-left) pasting-horizontal-coherence-square-maps' sq-left sq-right = - (sq-right ·r top-left) ∙h (bottom-right ·l sq-left) + (sq-right ·r top-left) ∙h (bottom-right ·l sq-left) pasting-horizontal-up-to-htpy-coherence-square-maps : {top : A → C} (H : coherence-triangle-maps top top-right top-left) diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index a2a5d62ec6..d729c62f38 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -412,8 +412,8 @@ module _ ( λ x → is-emb-is-equiv G (i x) (j (f x)))) equiv-postcomp-extension : - (f : A → B) (i : A → X) (g : X ≃ Y) → - extension f i ≃ extension f (map-equiv g ∘ i) + (f : A → B) (i : A → X) (g : X ≃ Y) → + extension f i ≃ extension f (map-equiv g ∘ i) equiv-postcomp-extension f i (g , G) = ( postcomp-extension f i g , is-equiv-postcomp-extension f i g G) ``` @@ -438,8 +438,8 @@ module _ ( λ x → H (i x) (j (f x))))) emb-postcomp-extension : - (f : A → B) (i : A → X) (g : X ↪ Y) → - extension f i ↪ extension f (map-emb g ∘ i) + (f : A → B) (i : A → X) (g : X ↪ Y) → + extension f i ↪ extension f (map-emb g ∘ i) emb-postcomp-extension f i (g , G) = postcomp-extension f i g , is-emb-postcomp-extension f i g G ``` From b4fc42d6dccee5ca8f105a496dff42be08b84c0a Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 11:33:22 +0200 Subject: [PATCH 25/69] fix a silly name --- ...unctoriality-dependent-pair-types.lagda.md | 6 ++-- src/foundation-core/subtypes.lagda.md | 2 +- src/foundation-core/truncated-maps.lagda.md | 4 +-- src/foundation/0-maps.lagda.md | 12 ++++---- src/foundation/decidable-embeddings.lagda.md | 28 +++++++++---------- src/foundation/decidable-subtypes.lagda.md | 2 +- src/foundation/embeddings.lagda.md | 4 +-- src/foundation/faithful-maps.lagda.md | 4 +-- ...oriality-dependent-function-types.lagda.md | 16 +++++------ ...unctoriality-dependent-pair-types.lagda.md | 22 +++++++-------- .../functoriality-function-types.lagda.md | 4 +-- src/logic/de-morgan-embeddings.lagda.md | 28 +++++++++---------- src/logic/de-morgan-subtypes.lagda.md | 2 +- ...double-negation-stable-embeddings.lagda.md | 28 +++++++++---------- .../double-negation-stable-subtypes.lagda.md | 2 +- .../anodyne-maps.lagda.md | 8 +++--- src/trees/functoriality-w-types.lagda.md | 10 +++---- .../binomial-types.lagda.md | 4 +-- 18 files changed, 93 insertions(+), 93 deletions(-) diff --git a/src/foundation-core/functoriality-dependent-pair-types.lagda.md b/src/foundation-core/functoriality-dependent-pair-types.lagda.md index f3d941d326..5334f6737b 100644 --- a/src/foundation-core/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation-core/functoriality-dependent-pair-types.lagda.md @@ -389,9 +389,9 @@ module _ where abstract - is-contr-map-map-Σ-map-base : + is-contr-map-Σ-map-base : is-contr-map f → is-contr-map (map-Σ-map-base f C) - is-contr-map-map-Σ-map-base is-contr-f (pair y z) = + is-contr-map-Σ-map-base is-contr-f (pair y z) = is-contr-equiv' ( fiber f y) ( compute-fiber-map-Σ-map-base f C (pair y z)) @@ -409,7 +409,7 @@ module _ is-equiv-map-Σ-map-base : is-equiv f → is-equiv (map-Σ-map-base f C) is-equiv-map-Σ-map-base is-equiv-f = is-equiv-is-contr-map - ( is-contr-map-map-Σ-map-base f C (is-contr-map-is-equiv is-equiv-f)) + ( is-contr-map-Σ-map-base f C (is-contr-map-is-equiv is-equiv-f)) equiv-Σ-equiv-base : {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3) (e : A ≃ B) → diff --git a/src/foundation-core/subtypes.lagda.md b/src/foundation-core/subtypes.lagda.md index 5310ef728e..84e8af493e 100644 --- a/src/foundation-core/subtypes.lagda.md +++ b/src/foundation-core/subtypes.lagda.md @@ -226,7 +226,7 @@ module _ pr1 (trunc-map-into-subtype f p) x = (map-trunc-map f x , p x) pr2 (trunc-map-into-subtype f p) = is-trunc-map-into-subtype - ( is-trunc-map-map-trunc-map f) + ( is-trunc-trunc-map f) ( p) ``` diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 8ff7e2d10b..bd67b2a0c0 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -53,9 +53,9 @@ module _ map-trunc-map = pr1 abstract - is-trunc-map-map-trunc-map : + is-trunc-trunc-map : (f : trunc-map k A B) → is-trunc-map k (map-trunc-map f) - is-trunc-map-map-trunc-map = pr2 + is-trunc-trunc-map = pr2 ``` ## Properties diff --git a/src/foundation/0-maps.lagda.md b/src/foundation/0-maps.lagda.md index 365e87053d..4890707a44 100644 --- a/src/foundation/0-maps.lagda.md +++ b/src/foundation/0-maps.lagda.md @@ -43,9 +43,9 @@ module _ map-0-map : {A : UU l1} {B : UU l2} → 0-map A B → A → B map-0-map = pr1 - is-0-map-map-0-map : + is-0-map-0-map : {A : UU l1} {B : UU l2} (f : 0-map A B) → is-0-map (map-0-map f) - is-0-map-map-0-map = pr2 + is-0-map-0-map = pr2 ``` ## Properties @@ -140,8 +140,8 @@ module _ where abstract - is-0-map-map-Σ-map-base : is-0-map f → is-0-map (map-Σ-map-base f C) - is-0-map-map-Σ-map-base = is-trunc-map-map-Σ-map-base zero-𝕋 C + is-0-map-Σ-map-base : is-0-map f → is-0-map (map-Σ-map-base f C) + is-0-map-Σ-map-base = is-trunc-map-Σ-map-base zero-𝕋 C ``` ### The functorial action of `Σ` preserves `0`-maps @@ -152,7 +152,7 @@ module _ (D : B → UU l4) {f : A → B} {g : (x : A) → C x → D (f x)} where - is-0-map-map-Σ : + is-0-map-Σ : is-0-map f → ((x : A) → is-0-map (g x)) → is-0-map (map-Σ D f g) - is-0-map-map-Σ = is-trunc-map-map-Σ zero-𝕋 D + is-0-map-Σ = is-trunc-map-Σ zero-𝕋 D ``` diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md index 051f6be3f1..256341db51 100644 --- a/src/foundation/decidable-embeddings.lagda.md +++ b/src/foundation/decidable-embeddings.lagda.md @@ -127,21 +127,21 @@ module _ map-decidable-emb : X → Y map-decidable-emb = pr1 e - is-decidable-emb-map-decidable-emb : + is-decidable-emb-decidable-emb : is-decidable-emb map-decidable-emb - is-decidable-emb-map-decidable-emb = pr2 e + is-decidable-emb-decidable-emb = pr2 e - is-emb-map-decidable-emb : is-emb map-decidable-emb - is-emb-map-decidable-emb = - is-emb-is-decidable-emb is-decidable-emb-map-decidable-emb + is-emb-decidable-emb : is-emb map-decidable-emb + is-emb-decidable-emb = + is-emb-is-decidable-emb is-decidable-emb-decidable-emb - is-decidable-map-map-decidable-emb : + is-decidable-decidable-emb : is-decidable-map map-decidable-emb - is-decidable-map-map-decidable-emb = - is-decidable-map-is-decidable-emb is-decidable-emb-map-decidable-emb + is-decidable-decidable-emb = + is-decidable-map-is-decidable-emb is-decidable-emb-decidable-emb emb-decidable-emb : X ↪ Y - emb-decidable-emb = map-decidable-emb , is-emb-map-decidable-emb + emb-decidable-emb = map-decidable-emb , is-emb-decidable-emb ``` ## Properties @@ -386,7 +386,7 @@ abstract ( is-prop-is-decidable-emb) ( map-decidable-emb f) ( refl-htpy) - ( is-decidable-emb-map-decidable-emb f) + ( is-decidable-emb-decidable-emb f) abstract is-equiv-htpy-eq-decidable-emb : @@ -464,7 +464,7 @@ module _ decidable-emb-tot : ((x : A) → B x ↪ᵈ C x) → Σ A B ↪ᵈ Σ A C decidable-emb-tot f = ( tot (λ x → map-decidable-emb (f x)) , - is-decidable-emb-tot (λ x → is-decidable-emb-map-decidable-emb (f x))) + is-decidable-emb-tot (λ x → is-decidable-emb-decidable-emb (f x))) ``` ### The map on total spaces induced by a decidable embedding on the base is a decidable embedding @@ -484,7 +484,7 @@ module _ (f : A ↪ᵈ B) → Σ A (C ∘ map-decidable-emb f) ↪ᵈ Σ B C decidable-emb-map-Σ-map-base f = ( map-Σ-map-base (map-decidable-emb f) C , - is-decidable-emb-map-Σ-map-base ((is-decidable-emb-map-decidable-emb f))) + is-decidable-emb-map-Σ-map-base ((is-decidable-emb-decidable-emb f))) ``` ### The functoriality of dependent pair types preserves decidable embeddings @@ -512,8 +512,8 @@ module _ decidable-emb-Σ f g = ( ( map-Σ D (map-decidable-emb f) (λ x → map-decidable-emb (g x))) , ( is-decidable-emb-map-Σ - ( is-decidable-emb-map-decidable-emb f) - ( λ x → is-decidable-emb-map-decidable-emb (g x)))) + ( is-decidable-emb-decidable-emb f) + ( λ x → is-decidable-emb-decidable-emb (g x)))) ``` ### Products of decidable embeddings are decidable embeddings diff --git a/src/foundation/decidable-subtypes.lagda.md b/src/foundation/decidable-subtypes.lagda.md index b2c8b2f18f..8cedbad085 100644 --- a/src/foundation/decidable-subtypes.lagda.md +++ b/src/foundation/decidable-subtypes.lagda.md @@ -168,7 +168,7 @@ module _ fiber (map-decidable-emb f) y pr2 (decidable-subtype-decidable-emb y) = is-decidable-prop-map-is-decidable-emb - ( is-decidable-emb-map-decidable-emb f) + ( is-decidable-emb-decidable-emb f) ( y) compute-type-decidable-subtype-decidable-emb : diff --git a/src/foundation/embeddings.lagda.md b/src/foundation/embeddings.lagda.md index 9d0c224af3..9001bb666b 100644 --- a/src/foundation/embeddings.lagda.md +++ b/src/foundation/embeddings.lagda.md @@ -223,7 +223,7 @@ module _ is-emb-map-Σ-map-base : {f : A → B} (C : B → UU l3) → is-emb f → is-emb (map-Σ-map-base f C) is-emb-map-Σ-map-base C H = - is-emb-is-prop-map (is-prop-map-map-Σ-map-base C (is-prop-map-is-emb H)) + is-emb-is-prop-map (is-prop-map-Σ-map-base C (is-prop-map-is-emb H)) emb-Σ-emb-base : (f : A ↪ B) (C : B → UU l3) → Σ A (λ a → C (map-emb f a)) ↪ Σ B C @@ -240,7 +240,7 @@ module _ is-emb f → ((x : A) → is-emb (g x)) → is-emb (map-Σ D f g) is-emb-map-Σ D H K = is-emb-is-prop-map - ( is-prop-map-map-Σ D + ( is-prop-map-Σ D ( is-prop-map-is-emb H) ( λ x → is-prop-map-is-emb (K x))) diff --git a/src/foundation/faithful-maps.lagda.md b/src/foundation/faithful-maps.lagda.md index 0e523f08b3..cec08adb9c 100644 --- a/src/foundation/faithful-maps.lagda.md +++ b/src/foundation/faithful-maps.lagda.md @@ -237,7 +237,7 @@ module _ is-faithful f → is-faithful (map-Σ-map-base f C) is-faithful-map-Σ-map-base H = is-faithful-is-0-map - ( is-0-map-map-Σ-map-base C (is-0-map-is-faithful H)) + ( is-0-map-Σ-map-base C (is-0-map-is-faithful H)) faithful-map-Σ-faithful-map-base : (f : faithful-map A B) (C : B → UU l3) → @@ -256,7 +256,7 @@ module _ is-faithful f → ((x : A) → is-faithful (g x)) → is-faithful (map-Σ D f g) is-faithful-map-Σ H K = is-faithful-is-0-map - ( is-0-map-map-Σ D + ( is-0-map-Σ D ( is-0-map-is-faithful H) ( λ x → is-0-map-is-faithful (K x))) ``` diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md index 9fc8e6c98c..73d044b7bd 100644 --- a/src/foundation/functoriality-dependent-function-types.lagda.md +++ b/src/foundation/functoriality-dependent-function-types.lagda.md @@ -141,10 +141,10 @@ module _ where abstract - is-trunc-map-map-Π : + is-trunc-map-Π : (k : 𝕋) (f : (i : I) → A i → B i) → ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π f) - is-trunc-map-map-Π k f H h = + is-trunc-map-Π k f H h = is-trunc-equiv' k ( (i : I) → fiber (f i) (h i)) ( compute-fiber-map-Π f h) @@ -155,7 +155,7 @@ module _ {f : (i : I) → A i → B i} → ((i : I) → is-emb (f i)) → is-emb (map-Π f) is-emb-map-Π {f} H = is-emb-is-prop-map - ( is-trunc-map-map-Π neg-one-𝕋 f (λ i → is-prop-map-is-emb (H i))) + ( is-trunc-map-Π neg-one-𝕋 f (λ i → is-prop-map-is-emb (H i))) emb-Π : ((i : I) → A i ↪ B i) → ((i : I) → A i) ↪ ((i : I) → B i) emb-Π f = (map-Π (map-emb ∘ f) , is-emb-map-Π (is-emb-map-emb ∘ f)) @@ -164,22 +164,22 @@ module _ ### A family of truncated maps over any map induces a truncated map on dependent function types ```agda -is-trunc-map-map-Π-is-trunc-map' : +is-trunc-map-Π-is-trunc-map' : (k : 𝕋) {l1 l2 l3 l4 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π' α f) -is-trunc-map-map-Π-is-trunc-map' k {J = J} α f H h = +is-trunc-map-Π-is-trunc-map' k {J = J} α f H h = is-trunc-equiv' k ( (j : J) → fiber (f (α j)) (h j)) ( compute-fiber-map-Π' α f h) ( is-trunc-Π k (λ j → H (α j) (h j))) -is-trunc-map-is-trunc-map-map-Π' : +is-trunc-map-is-trunc-map-Π' : (k : 𝕋) {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} (f : (i : I) → A i → B i) → ({l : Level} {J : UU l} (α : J → I) → is-trunc-map k (map-Π' α f)) → (i : I) → is-trunc-map k (f i) -is-trunc-map-is-trunc-map-map-Π' k {A = A} {B} f H i b = +is-trunc-map-is-trunc-map-Π' k {A = A} {B} f H i b = is-trunc-equiv' k ( fiber (map-Π (λ _ → f i)) (point b)) ( equiv-Σ @@ -197,7 +197,7 @@ is-emb-map-Π-is-emb' : ((i : I) → is-emb (f i)) → is-emb (map-Π' α f) is-emb-map-Π-is-emb' α f H = is-emb-is-prop-map - ( is-trunc-map-map-Π-is-trunc-map' neg-one-𝕋 α f + ( is-trunc-map-Π-is-trunc-map' neg-one-𝕋 α f ( λ i → is-prop-map-is-emb (H i))) ``` diff --git a/src/foundation/functoriality-dependent-pair-types.lagda.md b/src/foundation/functoriality-dependent-pair-types.lagda.md index beffdd38f2..b3dfd0b0fc 100644 --- a/src/foundation/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation/functoriality-dependent-pair-types.lagda.md @@ -181,37 +181,37 @@ module _ where abstract - is-trunc-map-map-Σ-map-base : + is-trunc-map-Σ-map-base : (k : 𝕋) {f : A → B} (C : B → UU l3) → is-trunc-map k f → is-trunc-map k (map-Σ-map-base f C) - is-trunc-map-map-Σ-map-base k {f} C H y = + is-trunc-map-Σ-map-base k {f} C H y = is-trunc-equiv' k ( fiber f (pr1 y)) ( compute-fiber-map-Σ-map-base f C y) ( H (pr1 y)) abstract - is-prop-map-map-Σ-map-base : + is-prop-map-Σ-map-base : {f : A → B} (C : B → UU l3) → is-prop-map f → is-prop-map (map-Σ-map-base f C) - is-prop-map-map-Σ-map-base C = is-trunc-map-map-Σ-map-base neg-one-𝕋 C + is-prop-map-Σ-map-base C = is-trunc-map-Σ-map-base neg-one-𝕋 C module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} where abstract - is-trunc-map-map-Σ : + is-trunc-map-Σ : (k : 𝕋) (D : B → UU l4) {f : A → B} {g : (x : A) → C x → D (f x)} → is-trunc-map k f → ((x : A) → is-trunc-map k (g x)) → is-trunc-map k (map-Σ D f g) - is-trunc-map-map-Σ k D {f} {g} H K = + is-trunc-map-Σ k D {f} {g} H K = is-trunc-map-left-map-triangle k ( map-Σ D f g) ( map-Σ-map-base f D) ( tot g) ( triangle-map-Σ D f g) - ( is-trunc-map-map-Σ-map-base k D H) + ( is-trunc-map-Σ-map-base k D H) ( is-trunc-map-tot k K) module _ @@ -219,16 +219,16 @@ module _ where abstract - is-contr-map-map-Σ : + is-contr-map-Σ : is-contr-map f → ((x : A) → is-contr-map (g x)) → is-contr-map (map-Σ D f g) - is-contr-map-map-Σ = is-trunc-map-map-Σ neg-two-𝕋 D + is-contr-map-Σ = is-trunc-map-Σ neg-two-𝕋 D abstract - is-prop-map-map-Σ : + is-prop-map-Σ : is-prop-map f → ((x : A) → is-prop-map (g x)) → is-prop-map (map-Σ D f g) - is-prop-map-map-Σ = is-trunc-map-map-Σ neg-one-𝕋 D + is-prop-map-Σ = is-trunc-map-Σ neg-one-𝕋 D ``` ### Commuting squares of maps on total spaces diff --git a/src/foundation/functoriality-function-types.lagda.md b/src/foundation/functoriality-function-types.lagda.md index a86808a6d2..f8d8ab6ca1 100644 --- a/src/foundation/functoriality-function-types.lagda.md +++ b/src/foundation/functoriality-function-types.lagda.md @@ -71,7 +71,7 @@ module _ is-trunc-map k f → {l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f) is-trunc-map-postcomp-is-trunc-map is-trunc-f A = - is-trunc-map-map-Π-is-trunc-map' k + is-trunc-map-Π-is-trunc-map' k ( terminal-map A) ( point f) ( point is-trunc-f) @@ -80,7 +80,7 @@ module _ ({l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f)) → is-trunc-map k f is-trunc-map-is-trunc-map-postcomp is-trunc-postcomp-f = - is-trunc-map-is-trunc-map-map-Π' k + is-trunc-map-is-trunc-map-Π' k ( point f) ( λ {l} {J} α → is-trunc-postcomp-f {l} J) ( star) diff --git a/src/logic/de-morgan-embeddings.lagda.md b/src/logic/de-morgan-embeddings.lagda.md index 3391dea43b..b590deac87 100644 --- a/src/logic/de-morgan-embeddings.lagda.md +++ b/src/logic/de-morgan-embeddings.lagda.md @@ -143,19 +143,19 @@ module _ map-de-morgan-emb : X → Y map-de-morgan-emb = pr1 e - is-de-morgan-emb-map-de-morgan-emb : is-de-morgan-emb map-de-morgan-emb - is-de-morgan-emb-map-de-morgan-emb = pr2 e + is-de-morgan-emb-de-morgan-emb : is-de-morgan-emb map-de-morgan-emb + is-de-morgan-emb-de-morgan-emb = pr2 e - is-emb-map-de-morgan-emb : is-emb map-de-morgan-emb - is-emb-map-de-morgan-emb = - is-emb-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb + is-emb-de-morgan-emb : is-emb map-de-morgan-emb + is-emb-de-morgan-emb = + is-emb-is-de-morgan-emb is-de-morgan-emb-de-morgan-emb - is-de-morgan-map-map-de-morgan-emb : is-de-morgan-map map-de-morgan-emb - is-de-morgan-map-map-de-morgan-emb = - is-de-morgan-map-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb + is-de-morgan-de-morgan-emb : is-de-morgan-map map-de-morgan-emb + is-de-morgan-de-morgan-emb = + is-de-morgan-map-is-de-morgan-emb is-de-morgan-emb-de-morgan-emb emb-de-morgan-emb : X ↪ Y - emb-de-morgan-emb = map-de-morgan-emb , is-emb-map-de-morgan-emb + emb-de-morgan-emb = map-de-morgan-emb , is-emb-de-morgan-emb ``` ## Properties @@ -425,7 +425,7 @@ abstract ( is-prop-is-de-morgan-emb) ( map-de-morgan-emb f) ( refl-htpy) - ( is-de-morgan-emb-map-de-morgan-emb f) + ( is-de-morgan-emb-de-morgan-emb f) abstract is-equiv-htpy-eq-de-morgan-emb : @@ -478,7 +478,7 @@ module _ de-morgan-emb-tot f = ( tot (map-de-morgan-emb ∘ f) , is-de-morgan-emb-tot - ( is-de-morgan-emb-map-de-morgan-emb ∘ f)) + ( is-de-morgan-emb-de-morgan-emb ∘ f)) ``` ### The map on total spaces induced by a De Morgan embedding on the base is a De Morgan embedding @@ -502,7 +502,7 @@ module _ de-morgan-emb-map-Σ-map-base f = ( map-Σ-map-base (map-de-morgan-emb f) C , is-de-morgan-emb-map-Σ-map-base - ( is-de-morgan-emb-map-de-morgan-emb f)) + ( is-de-morgan-emb-de-morgan-emb f)) ``` ### The functoriality of dependent pair types on De Morgan embeddings @@ -530,8 +530,8 @@ module _ de-morgan-emb-Σ f g = ( ( map-Σ D (map-decidable-emb f) (map-de-morgan-emb ∘ g)) , ( is-de-morgan-emb-map-Σ - ( is-decidable-emb-map-decidable-emb f) - ( is-de-morgan-emb-map-de-morgan-emb ∘ g))) + ( is-decidable-emb-decidable-emb f) + ( is-de-morgan-emb-de-morgan-emb ∘ g))) ``` ### Products of De Morgan embeddings are De Morgan embeddings diff --git a/src/logic/de-morgan-subtypes.lagda.md b/src/logic/de-morgan-subtypes.lagda.md index 8981009f85..0578cc921e 100644 --- a/src/logic/de-morgan-subtypes.lagda.md +++ b/src/logic/de-morgan-subtypes.lagda.md @@ -161,7 +161,7 @@ module _ fiber (map-de-morgan-emb f) y pr2 (de-morgan-subtype-de-morgan-emb y) = is-de-morgan-prop-map-is-de-morgan-emb - ( is-de-morgan-emb-map-de-morgan-emb f) + ( is-de-morgan-emb-de-morgan-emb f) ( y) compute-type-de-morgan-type-de-morgan-emb : diff --git a/src/logic/double-negation-stable-embeddings.lagda.md b/src/logic/double-negation-stable-embeddings.lagda.md index f6f753758a..ff510de7fa 100644 --- a/src/logic/double-negation-stable-embeddings.lagda.md +++ b/src/logic/double-negation-stable-embeddings.lagda.md @@ -145,25 +145,25 @@ module _ map-double-negation-stable-emb : X → Y map-double-negation-stable-emb = pr1 e - is-double-negation-stable-emb-map-double-negation-stable-emb : + is-double-negation-stable-emb-double-negation-stable-emb : is-double-negation-stable-emb map-double-negation-stable-emb - is-double-negation-stable-emb-map-double-negation-stable-emb = pr2 e + is-double-negation-stable-emb-double-negation-stable-emb = pr2 e - is-emb-map-double-negation-stable-emb : + is-emb-double-negation-stable-emb : is-emb map-double-negation-stable-emb - is-emb-map-double-negation-stable-emb = + is-emb-double-negation-stable-emb = is-emb-is-double-negation-stable-emb - is-double-negation-stable-emb-map-double-negation-stable-emb + is-double-negation-stable-emb-double-negation-stable-emb - is-double-negation-eliminating-map-map-double-negation-stable-emb : + is-double-negation-eliminating-double-negation-stable-emb : is-double-negation-eliminating-map map-double-negation-stable-emb - is-double-negation-eliminating-map-map-double-negation-stable-emb = + is-double-negation-eliminating-double-negation-stable-emb = is-double-negation-eliminating-map-is-double-negation-stable-emb - is-double-negation-stable-emb-map-double-negation-stable-emb + is-double-negation-stable-emb-double-negation-stable-emb emb-double-negation-stable-emb : X ↪ Y emb-double-negation-stable-emb = - map-double-negation-stable-emb , is-emb-map-double-negation-stable-emb + map-double-negation-stable-emb , is-emb-double-negation-stable-emb ``` ## Properties @@ -450,7 +450,7 @@ abstract ( is-prop-is-double-negation-stable-emb) ( map-double-negation-stable-emb f) ( refl-htpy) - ( is-double-negation-stable-emb-map-double-negation-stable-emb f) + ( is-double-negation-stable-emb-double-negation-stable-emb f) abstract is-equiv-htpy-eq-double-negation-stable-emb : @@ -536,7 +536,7 @@ module _ double-negation-stable-emb-tot f = ( tot (map-double-negation-stable-emb ∘ f) , is-double-negation-stable-emb-tot - ( is-double-negation-stable-emb-map-double-negation-stable-emb ∘ f)) + ( is-double-negation-stable-emb-double-negation-stable-emb ∘ f)) ``` ### The map on total spaces induced by a double negation stable embedding on the base is a double negation stable embedding @@ -560,7 +560,7 @@ module _ double-negation-stable-emb-map-Σ-map-base f = ( map-Σ-map-base (map-double-negation-stable-emb f) C , is-double-negation-stable-emb-map-Σ-map-base - ( is-double-negation-stable-emb-map-double-negation-stable-emb f)) + ( is-double-negation-stable-emb-double-negation-stable-emb f)) ``` ### The functoriality of dependent pair types preserves double negation stable embeddings @@ -590,8 +590,8 @@ module _ ( map-double-negation-stable-emb f) ( map-double-negation-stable-emb ∘ g)) , ( is-double-negation-stable-emb-map-Σ - ( is-double-negation-stable-emb-map-double-negation-stable-emb f) - ( is-double-negation-stable-emb-map-double-negation-stable-emb ∘ g))) + ( is-double-negation-stable-emb-double-negation-stable-emb f) + ( is-double-negation-stable-emb-double-negation-stable-emb ∘ g))) ``` ### Products of double negation stable embeddings are double negation stable embeddings diff --git a/src/logic/double-negation-stable-subtypes.lagda.md b/src/logic/double-negation-stable-subtypes.lagda.md index 6614510a1b..807d3c342b 100644 --- a/src/logic/double-negation-stable-subtypes.lagda.md +++ b/src/logic/double-negation-stable-subtypes.lagda.md @@ -165,7 +165,7 @@ module _ fiber (map-double-negation-stable-emb f) y pr2 (double-negation-stable-subtype-double-negation-stable-emb y) = is-double-negation-stable-prop-map-is-double-negation-stable-emb - ( is-double-negation-stable-emb-map-double-negation-stable-emb f) + ( is-double-negation-stable-emb-double-negation-stable-emb f) ( y) compute-type-double-negation-stable-type-double-negation-stable-emb : diff --git a/src/orthogonal-factorization-systems/anodyne-maps.lagda.md b/src/orthogonal-factorization-systems/anodyne-maps.lagda.md index 8f86a1c09d..ba5406b0bc 100644 --- a/src/orthogonal-factorization-systems/anodyne-maps.lagda.md +++ b/src/orthogonal-factorization-systems/anodyne-maps.lagda.md @@ -136,9 +136,9 @@ module _ (f : A → B) (j : C → D) (i : E → F) where - is-anodyne-map-map-product : + is-anodyne-map-product : is-anodyne-map f j → is-anodyne-map f i → is-anodyne-map f (map-product j i) - is-anodyne-map-map-product J I g H = + is-anodyne-map-product J I g H = is-orthogonal-left-product j i g (J g H) (I g H) ``` @@ -151,11 +151,11 @@ module _ (f : A → B) (j : C → D) (i : E → F) where - is-anodyne-map-map-coproduct : + is-anodyne-map-coproduct : is-anodyne-map f j → is-anodyne-map f i → is-anodyne-map f (map-coproduct j i) - is-anodyne-map-map-coproduct J I g H = + is-anodyne-map-coproduct J I g H = is-orthogonal-left-coproduct j i g (J g H) (I g H) ``` diff --git a/src/trees/functoriality-w-types.lagda.md b/src/trees/functoriality-w-types.lagda.md index 6a32e1a93e..e156b5b1b4 100644 --- a/src/trees/functoriality-w-types.lagda.md +++ b/src/trees/functoriality-w-types.lagda.md @@ -125,18 +125,18 @@ abstract ### For any family of equivalences `e` over `f`, if `f` is truncated then `map-𝕎 f e` is truncated ```agda -is-trunc-map-map-𝕎 : +is-trunc-map-𝕎 : {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) (f : A → C) (e : (x : A) → B x ≃ D (f x)) → is-trunc-map k f → is-trunc-map k (map-𝕎 D f e) -is-trunc-map-map-𝕎 k D f e H (tree-𝕎 c γ) = +is-trunc-map-𝕎 k D f e H (tree-𝕎 c γ) = is-trunc-equiv k ( fiber-map-𝕎 D f e (tree-𝕎 c γ)) ( equiv-fiber-map-𝕎 D f e (tree-𝕎 c γ)) ( is-trunc-Σ ( H c) - ( λ t → is-trunc-Π k (λ d → is-trunc-map-map-𝕎 k D f e H (γ d)))) + ( λ t → is-trunc-Π k (λ d → is-trunc-map-𝕎 k D f e H (γ d)))) ``` ### For any family of equivalences `e` over `f`, if `f` is an equivalence then `map-𝕎 f e` is an equivalence @@ -148,7 +148,7 @@ is-equiv-map-𝕎 : is-equiv f → is-equiv (map-𝕎 D f e) is-equiv-map-𝕎 D f e H = is-equiv-is-contr-map - ( is-trunc-map-map-𝕎 neg-two-𝕋 D f e (is-contr-map-is-equiv H)) + ( is-trunc-map-𝕎 neg-two-𝕋 D f e (is-contr-map-is-equiv H)) equiv-𝕎 : {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) @@ -169,7 +169,7 @@ is-emb-map-𝕎 : is-emb f → is-emb (map-𝕎 D f e) is-emb-map-𝕎 D f e H = is-emb-is-prop-map - (is-trunc-map-map-𝕎 neg-one-𝕋 D f e (is-prop-map-is-emb H)) + (is-trunc-map-𝕎 neg-one-𝕋 D f e (is-prop-map-is-emb H)) emb-𝕎 : {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) diff --git a/src/univalent-combinatorics/binomial-types.lagda.md b/src/univalent-combinatorics/binomial-types.lagda.md index f815dbd237..29c35ce17f 100644 --- a/src/univalent-combinatorics/binomial-types.lagda.md +++ b/src/univalent-combinatorics/binomial-types.lagda.md @@ -97,7 +97,7 @@ module _ is-emb-map-emb-binomial-type-Level : is-emb map-decidable-emb-binomial-type-Level is-emb-map-emb-binomial-type-Level = - is-emb-map-decidable-emb decidable-emb-binomial-type-Level + is-emb-decidable-emb decidable-emb-binomial-type-Level ``` ### The standard binomial types @@ -127,7 +127,7 @@ module _ abstract is-emb-map-emb-binomial-type : is-emb map-decidable-emb-binomial-type is-emb-map-emb-binomial-type = - is-emb-map-decidable-emb decidable-emb-binomial-type + is-emb-decidable-emb decidable-emb-binomial-type ``` ### The type of decidable subtypes of `A` such that the total space is merely equivalent to a given finite type From d13feda7f5772ed5b01930385c103df111f96ac0 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 11:49:02 +0200 Subject: [PATCH 26/69] fix another silly name --- src/foundation-core/truncated-maps.lagda.md | 10 +++++----- src/foundation/diagonals-of-maps.lagda.md | 2 +- src/foundation/faithful-maps.lagda.md | 4 ++-- src/foundation/mere-path-cosplit-maps.lagda.md | 2 +- src/foundation/path-cosplit-maps.lagda.md | 2 +- .../precomposition-dependent-functions.lagda.md | 2 +- 6 files changed, 11 insertions(+), 11 deletions(-) diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index bd67b2a0c0..7142a95bd0 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -107,18 +107,18 @@ module _ where abstract - is-trunc-map-is-trunc-map-ap : + is-trunc-map-succ-is-trunc-map-ap : ((x y : A) → is-trunc-map k (ap f {x} {y})) → is-trunc-map (succ-𝕋 k) f - is-trunc-map-is-trunc-map-ap is-trunc-map-ap-f b (pair x p) (pair x' p') = + is-trunc-map-succ-is-trunc-map-ap is-trunc-map-ap-f b (pair x p) (pair x' p') = is-trunc-equiv k ( fiber (ap f) (p ∙ (inv p'))) ( equiv-fiber-ap-eq-fiber f (pair x p) (pair x' p')) ( is-trunc-map-ap-f x x' (p ∙ (inv p'))) abstract - is-trunc-map-ap-is-trunc-map : + is-trunc-map-ap-is-trunc-map-succ : is-trunc-map (succ-𝕋 k) f → (x y : A) → is-trunc-map k (ap f {x} {y}) - is-trunc-map-ap-is-trunc-map is-trunc-map-f x y p = + is-trunc-map-ap-is-trunc-map-succ is-trunc-map-f x y p = is-trunc-is-equiv' k ( pair x p = pair y refl) ( eq-fiber-fiber-ap f x y p) @@ -141,7 +141,7 @@ is-trunc-is-trunc-map-into-is-trunc ( k) ( ap f) ( is-trunc-B (f a) (f a')) - ( is-trunc-map-ap-is-trunc-map k f is-trunc-map-f a a') + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-map-f a a') ``` ### A family of types is a family of `k`-truncated types if and only of the projection map is `k`-truncated diff --git a/src/foundation/diagonals-of-maps.lagda.md b/src/foundation/diagonals-of-maps.lagda.md index a965e62bc4..8454997692 100644 --- a/src/foundation/diagonals-of-maps.lagda.md +++ b/src/foundation/diagonals-of-maps.lagda.md @@ -88,7 +88,7 @@ abstract is-trunc-is-equiv k (fiber (ap f) p) ( fiber-ap-fiber-diagonal-map f (triple x y p)) ( is-equiv-fiber-ap-fiber-diagonal-map f (triple x y p)) - ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y p) + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y p) abstract is-trunc-map-is-trunc-diagonal-map : diff --git a/src/foundation/faithful-maps.lagda.md b/src/foundation/faithful-maps.lagda.md index cec08adb9c..281041beb5 100644 --- a/src/foundation/faithful-maps.lagda.md +++ b/src/foundation/faithful-maps.lagda.md @@ -107,12 +107,12 @@ module _ is-0-map-is-faithful : is-faithful f → is-0-map f is-0-map-is-faithful H = - is-trunc-map-is-trunc-map-ap neg-one-𝕋 f + is-trunc-map-succ-is-trunc-map-ap neg-one-𝕋 f ( λ x y → is-prop-map-is-emb (H x y)) is-faithful-is-0-map : is-0-map f → is-faithful f is-faithful-is-0-map H x y = - is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map neg-one-𝕋 f H x y) + is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map-succ neg-one-𝕋 f H x y) ``` ## Properties diff --git a/src/foundation/mere-path-cosplit-maps.lagda.md b/src/foundation/mere-path-cosplit-maps.lagda.md index d33ddf9534..4042949892 100644 --- a/src/foundation/mere-path-cosplit-maps.lagda.md +++ b/src/foundation/mere-path-cosplit-maps.lagda.md @@ -105,7 +105,7 @@ is-mere-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = unit-trunc-Prop (retraction-is-contr-map is-trunc-f) is-mere-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = is-mere-path-cosplit-is-trunc k - ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is merely `k+1`-path-cosplit diff --git a/src/foundation/path-cosplit-maps.lagda.md b/src/foundation/path-cosplit-maps.lagda.md index 854352b42a..9ee6882c86 100644 --- a/src/foundation/path-cosplit-maps.lagda.md +++ b/src/foundation/path-cosplit-maps.lagda.md @@ -170,7 +170,7 @@ is-path-cosplit-is-trunc : is-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = retraction-is-contr-map is-trunc-f is-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = - is-path-cosplit-is-trunc k (is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) + is-path-cosplit-is-trunc k (is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is `k+1`-path-cosplit diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index 6d64834d25..1d43737bc8 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -165,7 +165,7 @@ is-trunc-map-succ-precomp-Π : ((g h : (b : B) → C b) → is-trunc-map k (precomp-Π f (eq-value g h))) → is-trunc-map (succ-𝕋 k) (precomp-Π f C) is-trunc-map-succ-precomp-Π {k = k} {f = f} {C = C} H = - is-trunc-map-is-trunc-map-ap k (precomp-Π f C) + is-trunc-map-succ-is-trunc-map-ap k (precomp-Π f C) ( λ g h → is-trunc-map-top-is-trunc-map-bottom-is-equiv k ( ap (precomp-Π f C)) From 68dae007fcba006c40d786ea1cd82bf4e773be63 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 11:51:20 +0200 Subject: [PATCH 27/69] `is-trunc-map-postcomp-extension` --- src/foundation-core/truncated-maps.lagda.md | 14 ++++++++++ .../extensions-maps.lagda.md | 28 +++++++++++++++++++ 2 files changed, 42 insertions(+) diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 7142a95bd0..dbf9e3c370 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -126,6 +126,20 @@ module _ ( is-trunc-map-f (f y) (pair x p) (pair y refl)) ``` +### The action on identifications of a `k`-truncated map is also `k`-truncated + +```agda +module _ + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (f : A → B) + where + + abstract + is-trunc-map-ap-is-trunc-map : + is-trunc-map k f → (x y : A) → is-trunc-map k (ap f {x} {y}) + is-trunc-map-ap-is-trunc-map H = + is-trunc-map-ap-is-trunc-map-succ k f (is-trunc-map-succ-is-trunc-map k H) +``` + ### The domain of any `k`-truncated map into a `k+1`-truncated type is `k+1`-truncated ```agda diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index d729c62f38..04ccf940aa 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -16,6 +16,7 @@ open import foundation.equivalences open import foundation.function-types open import foundation.functoriality-dependent-function-types open import foundation.functoriality-dependent-pair-types +open import foundation.functoriality-function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopies open import foundation.homotopy-induction @@ -26,6 +27,7 @@ open import foundation.propositions open import foundation.sets open import foundation.structure-identity-principle open import foundation.transport-along-identifications +open import foundation.truncated-maps open import foundation.truncated-types open import foundation.truncation-levels open import foundation.type-arithmetic-dependent-pair-types @@ -444,6 +446,32 @@ module _ postcomp-extension f i g , is-emb-postcomp-extension f i g G ``` +### Postcomposition of extensions by a `k`-truncated map is `k`-truncated + +```agda +module _ + {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-trunc-map-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-trunc-map k g → + is-trunc-map k (postcomp-extension f i g) + is-trunc-map-postcomp-extension f i g G = + is-trunc-map-Σ k + ( is-extension f (g ∘ i)) + ( is-trunc-map-postcomp-is-trunc-map k g G B) + ( λ j → + is-trunc-map-Π k + ( λ a → ap g) + ( λ a → is-trunc-map-ap-is-trunc-map k g G (i a) (j (f a)))) + + trunc-map-postcomp-extension : + (f : A → B) (i : A → X) (g : trunc-map k X Y) → + trunc-map k (extension f i) (extension f (map-trunc-map g ∘ i)) + trunc-map-postcomp-extension f i (g , G) = + ( postcomp-extension f i g , is-trunc-map-postcomp-extension f i g G) +``` + ## See also - [`orthogonal-factorization-systems.lifts-maps`](orthogonal-factorization-systems.lifts-maps.md) From 242d30df672922645db55c26ec8c3d4fdb8ac4a6 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 12:11:48 +0200 Subject: [PATCH 28/69] `postcomposition-extensions-maps` --- src/orthogonal-factorization-systems.lagda.md | 1 + .../extensions-maps.lagda.md | 91 ----------- .../postcomposition-extensions-maps.lagda.md | 144 ++++++++++++++++++ 3 files changed, 145 insertions(+), 91 deletions(-) create mode 100644 src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 97d4fb470d..7d1b5ee681 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -58,6 +58,7 @@ open import orthogonal-factorization-systems.null-types public open import orthogonal-factorization-systems.open-modalities public open import orthogonal-factorization-systems.orthogonal-factorization-systems public open import orthogonal-factorization-systems.orthogonal-maps public +open import orthogonal-factorization-systems.postcomposition-extensions-maps public open import orthogonal-factorization-systems.precomposition-lifts-families-of-elements public open import orthogonal-factorization-systems.pullback-hom public open import orthogonal-factorization-systems.raise-modalities public diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index 04ccf940aa..ffdc6ae7f6 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -213,20 +213,6 @@ module _ is-extension-right-whisker F h = F ∘ h ``` -### Postcomposition of extensions - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → - extension f i → extension f (g ∘ i) - postcomp-extension f i g = - map-Σ (is-extension f (g ∘ i)) (postcomp B g) (λ j H → g ·l H) -``` - ## Properties ### Characterizing identifications of extensions of maps @@ -395,83 +381,6 @@ module _ extension-along-self = id , is-extension-along-self ``` -### Postcomposition of extensions by an equivalence is an equivalence - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-equiv-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-equiv g → - is-equiv (postcomp-extension f i g) - is-equiv-postcomp-extension f i g G = - is-equiv-map-Σ - ( is-extension f (g ∘ i)) - ( is-equiv-postcomp-is-equiv g G B) - ( λ j → - is-equiv-map-Π-is-fiberwise-equiv - ( λ x → is-emb-is-equiv G (i x) (j (f x)))) - - equiv-postcomp-extension : - (f : A → B) (i : A → X) (g : X ≃ Y) → - extension f i ≃ extension f (map-equiv g ∘ i) - equiv-postcomp-extension f i (g , G) = - ( postcomp-extension f i g , is-equiv-postcomp-extension f i g G) -``` - -### Postcomposition of extensions by an embedding is an embedding - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-emb-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-emb g → - is-emb (postcomp-extension f i g) - is-emb-postcomp-extension f i g H = - is-emb-map-Σ - ( is-extension f (g ∘ i)) - ( is-mono-is-emb g H B) - ( λ j → - is-emb-is-equiv - ( is-equiv-map-Π-is-fiberwise-equiv - ( λ x → H (i x) (j (f x))))) - - emb-postcomp-extension : - (f : A → B) (i : A → X) (g : X ↪ Y) → - extension f i ↪ extension f (map-emb g ∘ i) - emb-postcomp-extension f i (g , G) = - postcomp-extension f i g , is-emb-postcomp-extension f i g G -``` - -### Postcomposition of extensions by a `k`-truncated map is `k`-truncated - -```agda -module _ - {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-trunc-map-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-trunc-map k g → - is-trunc-map k (postcomp-extension f i g) - is-trunc-map-postcomp-extension f i g G = - is-trunc-map-Σ k - ( is-extension f (g ∘ i)) - ( is-trunc-map-postcomp-is-trunc-map k g G B) - ( λ j → - is-trunc-map-Π k - ( λ a → ap g) - ( λ a → is-trunc-map-ap-is-trunc-map k g G (i a) (j (f a)))) - - trunc-map-postcomp-extension : - (f : A → B) (i : A → X) (g : trunc-map k X Y) → - trunc-map k (extension f i) (extension f (map-trunc-map g ∘ i)) - trunc-map-postcomp-extension f i (g , G) = - ( postcomp-extension f i g , is-trunc-map-postcomp-extension f i g G) -``` - ## See also - [`orthogonal-factorization-systems.lifts-maps`](orthogonal-factorization-systems.lifts-maps.md) diff --git a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md new file mode 100644 index 0000000000..410be2dea8 --- /dev/null +++ b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md @@ -0,0 +1,144 @@ +# Postcomposition of extensions of maps + +```agda +module orthogonal-factorization-systems.postcomposition-extensions-maps where +``` + +
Imports + +```agda +open import foundation.action-on-identifications-dependent-functions +open import foundation.action-on-identifications-functions +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.embeddings +open import foundation.equivalences +open import foundation.function-types +open import foundation.functoriality-dependent-function-types +open import foundation.functoriality-dependent-pair-types +open import foundation.functoriality-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.monomorphisms +open import foundation.postcomposition-functions +open import foundation.propositions +open import foundation.sets +open import foundation.structure-identity-principle +open import foundation.transport-along-identifications +open import foundation.truncated-maps +open import foundation.truncated-types +open import foundation.truncation-levels +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universe-levels +open import foundation.whiskering-homotopies-composition + +open import foundation-core.torsorial-type-families + +open import orthogonal-factorization-systems.extensions-maps +``` + +
+ +## Idea + +Given a map `i : A → B` and a map `f : A → X`, then we may +{{#concept "postcompose" Disambiguation="extension of map by map" Agda=postcomp-extension}} +any [extension](orthogonal-factorization-systems.extensions-maps.md) +`α : extension i f` by `g` to obtain an extension `g ∘ α : extension i (g ∘ f)`. + +## Definition + +### Postcomposition of extensions + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + postcomp-extension : + (i : A → B) (f : A → X) (g : X → Y) → + extension i f → extension i (g ∘ f) + postcomp-extension i f g = + map-Σ (is-extension i (g ∘ f)) (postcomp B g) (λ j H → g ·l H) +``` + +## Properties + +### Postcomposition of extensions by an equivalence is an equivalence + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-equiv-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-equiv g → + is-equiv (postcomp-extension f i g) + is-equiv-postcomp-extension f i g G = + is-equiv-map-Σ + ( is-extension f (g ∘ i)) + ( is-equiv-postcomp-is-equiv g G B) + ( λ j → + is-equiv-map-Π-is-fiberwise-equiv + ( λ x → is-emb-is-equiv G (i x) (j (f x)))) + + equiv-postcomp-extension : + (f : A → B) (i : A → X) (g : X ≃ Y) → + extension f i ≃ extension f (map-equiv g ∘ i) + equiv-postcomp-extension f i (g , G) = + ( postcomp-extension f i g , is-equiv-postcomp-extension f i g G) +``` + +### Postcomposition of extensions by an embedding is an embedding + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-emb-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-emb g → + is-emb (postcomp-extension f i g) + is-emb-postcomp-extension f i g H = + is-emb-map-Σ + ( is-extension f (g ∘ i)) + ( is-mono-is-emb g H B) + ( λ j → + is-emb-is-equiv + ( is-equiv-map-Π-is-fiberwise-equiv + ( λ x → H (i x) (j (f x))))) + + emb-postcomp-extension : + (f : A → B) (i : A → X) (g : X ↪ Y) → + extension f i ↪ extension f (map-emb g ∘ i) + emb-postcomp-extension f i (g , G) = + postcomp-extension f i g , is-emb-postcomp-extension f i g G +``` + +### Postcomposition of extensions by a `k`-truncated map is `k`-truncated + +```agda +module _ + {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-trunc-map-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-trunc-map k g → + is-trunc-map k (postcomp-extension f i g) + is-trunc-map-postcomp-extension f i g G = + is-trunc-map-Σ k + ( is-extension f (g ∘ i)) + ( is-trunc-map-postcomp-is-trunc-map k g G B) + ( λ j → + is-trunc-map-Π k + ( λ a → ap g) + ( λ a → is-trunc-map-ap-is-trunc-map k g G (i a) (j (f a)))) + + trunc-map-postcomp-extension : + (f : A → B) (i : A → X) (g : trunc-map k X Y) → + trunc-map k (extension f i) (extension f (map-trunc-map g ∘ i)) + trunc-map-postcomp-extension f i (g , G) = + ( postcomp-extension f i g , is-trunc-map-postcomp-extension f i g G) +``` From c2e67a92d449b4840aaf3b17e69d687e92c2bc6d Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 12:22:25 +0200 Subject: [PATCH 29/69] edits `extensions-maps` --- .../extensions-maps.lagda.md | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index ffdc6ae7f6..888385a62f 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -311,7 +311,7 @@ module _ ((x : A) → is-trunc (succ-𝕋 k) (P (i x))) → (g : (x : B) → P x) → is-trunc k (is-extension i f g) is-trunc-is-extension-dependent-type f is-trunc-P g = - is-trunc-Π k λ x → is-trunc-P x (f x) (g (i x)) + is-trunc-Π k (λ x → is-trunc-P x (f x) (g (i x))) is-trunc-extension-dependent-type : {P : B → UU l3} (f : (x : A) → P (i x)) → @@ -341,7 +341,7 @@ module _ ((x : A) → is-prop (P (i x))) → (g : (x : B) → P x) → is-contr (is-extension i f g) is-contr-is-extension f is-prop-P g = - is-contr-Π λ x → is-prop-P x (f x) (g (i x)) + is-contr-Π (λ x → is-prop-P x (f x) (g (i x))) is-prop-is-extension : {P : B → UU l3} (f : (x : A) → P (i x)) → @@ -364,7 +364,7 @@ module _ is-extension-self = refl-htpy extension-self : extension-dependent-type id P f - extension-self = f , is-extension-self + extension-self = (f , is-extension-self) ``` ### The identity is an extension of every map along themselves @@ -378,7 +378,7 @@ module _ is-extension-along-self = refl-htpy extension-along-self : extension f f - extension-along-self = id , is-extension-along-self + extension-along-self = (id , is-extension-along-self) ``` ## See also From 7e76f72a30fb482a87ea51f3684dfe4e275ed761 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 13:40:50 +0200 Subject: [PATCH 30/69] fix import --- src/foundation/surjective-maps.lagda.md | 1 + .../map-separated-types-global-subuniverses.lagda.md | 1 + 2 files changed, 2 insertions(+) diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md index 2c2aff28b8..1e7120f2dd 100644 --- a/src/foundation/surjective-maps.lagda.md +++ b/src/foundation/surjective-maps.lagda.md @@ -52,6 +52,7 @@ open import foundation-core.truncated-maps open import foundation-core.truncation-levels open import orthogonal-factorization-systems.extensions-maps +open import orthogonal-factorization-systems.postcomposition-extensions-maps ``` diff --git a/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md b/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md index c6dc6ba0e5..7771b585dd 100644 --- a/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md @@ -11,6 +11,7 @@ open import foundation.dependent-pair-types open import foundation.equivalences open import foundation.global-subuniverses open import foundation.universe-levels +open import orthogonal-factorization-systems.postcomposition-extensions-maps open import foundation-core.function-types open import foundation-core.identity-types From f8a4c1f0ad2d59fa9782dbe3623c35e547d957bb Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Tue, 30 Sep 2025 14:57:41 +0200 Subject: [PATCH 31/69] rename files --- src/orthogonal-factorization-systems.lagda.md | 4 ++-- ...relative-separations-types-global-subuniverses.lagda.md} | 6 +++--- ....md => relative-separations-types-subuniverses.lagda.md} | 4 ++-- 3 files changed, 7 insertions(+), 7 deletions(-) rename src/orthogonal-factorization-systems/{map-separated-types-global-subuniverses.lagda.md => relative-separations-types-global-subuniverses.lagda.md} (94%) rename src/orthogonal-factorization-systems/{map-separated-types-subuniverses.lagda.md => relative-separations-types-subuniverses.lagda.md} (91%) diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 7d1b5ee681..d5d3dc288d 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -45,8 +45,6 @@ open import orthogonal-factorization-systems.localizations-at-global-subuniverse open import orthogonal-factorization-systems.localizations-at-maps public open import orthogonal-factorization-systems.localizations-at-subuniverses public open import orthogonal-factorization-systems.locally-small-modal-operators public -open import orthogonal-factorization-systems.map-separated-types-global-subuniverses public -open import orthogonal-factorization-systems.map-separated-types-subuniverses public open import orthogonal-factorization-systems.maps-local-at-maps public open import orthogonal-factorization-systems.mere-lifting-properties public open import orthogonal-factorization-systems.modal-induction public @@ -66,6 +64,8 @@ open import orthogonal-factorization-systems.reflective-global-subuniverses publ open import orthogonal-factorization-systems.reflective-modalities public open import orthogonal-factorization-systems.reflective-subuniverses public open import orthogonal-factorization-systems.regular-cd-structures public +open import orthogonal-factorization-systems.relative-separations-types-global-subuniverses public +open import orthogonal-factorization-systems.relative-separations-types-subuniverses public open import orthogonal-factorization-systems.sigma-closed-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-subuniverses public diff --git a/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md similarity index 94% rename from src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md rename to src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md index 7771b585dd..43a467525a 100644 --- a/src/orthogonal-factorization-systems/map-separated-types-global-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md @@ -1,7 +1,7 @@ -# Map-separated types at global subuniverses +# Relative separations of types at global subuniverses ```agda -module orthogonal-factorization-systems.map-separated-types-global-subuniverses where +module orthogonal-factorization-systems.relative-separations-types-global-subuniverses where ```
Imports @@ -11,13 +11,13 @@ open import foundation.dependent-pair-types open import foundation.equivalences open import foundation.global-subuniverses open import foundation.universe-levels -open import orthogonal-factorization-systems.postcomposition-extensions-maps open import foundation-core.function-types open import foundation-core.identity-types open import foundation-core.propositions open import orthogonal-factorization-systems.extensions-maps +open import orthogonal-factorization-systems.postcomposition-extensions-maps ```
diff --git a/src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md similarity index 91% rename from src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md rename to src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md index 95109c020e..de39d97c9f 100644 --- a/src/orthogonal-factorization-systems/map-separated-types-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md @@ -1,7 +1,7 @@ -# Map-separated types at subuniverses +# Relative separations of types at subuniverses ```agda -module orthogonal-factorization-systems.map-separated-types-subuniverses where +module orthogonal-factorization-systems.relative-separations-types-subuniverses where ```
Imports From b3d9701093158a2aa568a866c5b9a212603db261 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Wed, 1 Oct 2025 16:13:51 +0200 Subject: [PATCH 32/69] fix links in `foundation.projective-types` --- src/foundation/projective-types.lagda.md | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/foundation/projective-types.lagda.md b/src/foundation/projective-types.lagda.md index 07ee95a086..7e0578ef97 100644 --- a/src/foundation/projective-types.lagda.md +++ b/src/foundation/projective-types.lagda.md @@ -26,17 +26,17 @@ open import foundation-core.truncated-types A type `X` is said to be {{#concept "set-projective" Agda=is-set-projective}} if for every [surjective map](foundation.surjective-maps.md) `f : A ↠ B` onto a -[set](foundation-core.sets.md) `B` the [postcomposition -function[(foundation-core.postcomposition-functions.md) +[set](foundation-core.sets.md) `B` the +[postcomposition function](foundation-core.postcomposition-functions.md) ```text (X → A) → (X → B) ``` is surjective. This is [equivalent](foundation.logical-equivalences.md) to the -condition that for every [equivalence -relation[(foundation-core.equivalence-relations.md) `R` on a type `A` the -natural map +condition that for every +[equivalence relation](foundation-core.equivalence-relations.md) `R` on a type +`A` the natural map ```text (X → A)/~ → (X → A/R) From c31efdc7425ba92016fc28057932a3b08a622d72 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 2 Oct 2025 16:27:34 +0200 Subject: [PATCH 33/69] work separation, subuniverse connected maps and equivalences --- src/foundation-core/truncated-maps.lagda.md | 14 +- src/foundation/path-cosplit-maps.lagda.md | 3 +- ...recomposition-dependent-functions.lagda.md | 67 +++ .../precomposition-functions.lagda.md | 10 +- .../truncation-equivalences.lagda.md | 12 +- ...property-family-of-fibers-of-maps.lagda.md | 246 +++++++-- .../universal-property-unit-type.lagda.md | 62 ++- src/orthogonal-factorization-systems.lagda.md | 1 + .../localizations-at-subuniverses.lagda.md | 30 +- .../reflective-subuniverses.lagda.md | 2 +- ...rations-types-global-subuniverses.lagda.md | 32 +- ...ve-separations-types-subuniverses.lagda.md | 22 +- .../subuniverse-connected-maps.lagda.md | 267 +++++++++- .../subuniverse-equivalences.lagda.md | 495 ++++++++++++++++++ .../types-separated-at-maps.lagda.md | 8 +- 15 files changed, 1139 insertions(+), 132 deletions(-) create mode 100644 src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index dbf9e3c370..c4a44f53bd 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -109,21 +109,21 @@ module _ abstract is-trunc-map-succ-is-trunc-map-ap : ((x y : A) → is-trunc-map k (ap f {x} {y})) → is-trunc-map (succ-𝕋 k) f - is-trunc-map-succ-is-trunc-map-ap is-trunc-map-ap-f b (pair x p) (pair x' p') = + is-trunc-map-succ-is-trunc-map-ap is-trunc-map-ap-f b (x , p) (x' , p') = is-trunc-equiv k - ( fiber (ap f) (p ∙ (inv p'))) - ( equiv-fiber-ap-eq-fiber f (pair x p) (pair x' p')) - ( is-trunc-map-ap-f x x' (p ∙ (inv p'))) + ( fiber (ap f) (p ∙ inv p')) + ( equiv-fiber-ap-eq-fiber f (x , p) (x' , p')) + ( is-trunc-map-ap-f x x' (p ∙ inv p')) abstract is-trunc-map-ap-is-trunc-map-succ : is-trunc-map (succ-𝕋 k) f → (x y : A) → is-trunc-map k (ap f {x} {y}) is-trunc-map-ap-is-trunc-map-succ is-trunc-map-f x y p = is-trunc-is-equiv' k - ( pair x p = pair y refl) + ( (x , p) = (y , refl)) ( eq-fiber-fiber-ap f x y p) ( is-equiv-eq-fiber-fiber-ap f x y p) - ( is-trunc-map-f (f y) (pair x p) (pair y refl)) + ( is-trunc-map-f (f y) (x , p) (y , refl)) ``` ### The action on identifications of a `k`-truncated map is also `k`-truncated @@ -283,7 +283,7 @@ abstract ( map-compute-fiber-comp g h (g b)) ( is-equiv-map-compute-fiber-comp g h (g b)) ( is-trunc-map-htpy k (inv-htpy H) is-trunc-f (g b))) - ( pair b refl) + ( b , refl) ``` ### If a composite `g ∘ h` and its left factor `g` are truncated maps, then its right factor `h` is a truncated map diff --git a/src/foundation/path-cosplit-maps.lagda.md b/src/foundation/path-cosplit-maps.lagda.md index 9ee6882c86..671585677c 100644 --- a/src/foundation/path-cosplit-maps.lagda.md +++ b/src/foundation/path-cosplit-maps.lagda.md @@ -170,7 +170,8 @@ is-path-cosplit-is-trunc : is-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = retraction-is-contr-map is-trunc-f is-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = - is-path-cosplit-is-trunc k (is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) + is-path-cosplit-is-trunc k + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is `k+1`-path-cosplit diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index 1d43737bc8..1afc8d566f 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -15,12 +15,15 @@ open import foundation.dependent-identifications open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences open import foundation.function-extensionality +open import foundation.sections open import foundation.transport-along-homotopies open import foundation.transport-along-identifications open import foundation.universe-levels open import foundation-core.commuting-squares-of-maps +open import foundation-core.commuting-triangles-of-maps open import foundation-core.equivalences +open import foundation-core.fibers-of-maps open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies @@ -34,6 +37,70 @@ open import foundation-core.type-theoretic-principle-of-choice ## Properties +### Computations of the fibers of `precomp-Π` + +The fiber of `- ∘ f : ((b : B) → U b) → ((a : A) → U (f a))` at +`g ∘ f : (b : B) → U b` is equivalent to the type of maps `h : (b : B) → U b` +equipped with a homotopy witnessing that the square + +```text + f + A -----> B + | | + f | | h + ∨ ∨ + B ---> U ∘ f + g +``` + +commutes. + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (U : B → UU l3) + where + + compute-extension-fiber-precomp-Π' : + (g : (a : A) → U (f a)) → + fiber (precomp-Π f U) g ≃ + Σ ((b : B) → U b) (λ h → (a : A) → (h ∘ f) a = g a) + compute-extension-fiber-precomp-Π' g = + equiv-tot (λ h → equiv-funext) + + compute-extension-fiber-precomp-Π : + (g : (a : A) → U (f a)) → + fiber (precomp-Π f U) g ≃ Σ ((b : B) → U b) (λ h → g ~ h ∘ f) + compute-extension-fiber-precomp-Π g = + equiv-tot (λ h → equiv-funext) ∘e equiv-fiber (precomp-Π f U) g + + compute-fiber-precomp-Π : + (g : (b : B) → U b) → + fiber (precomp-Π f U) (g ∘ f) ≃ Σ ((b : B) → U b) (λ h → g ∘ f ~ h ∘ f) + compute-fiber-precomp-Π g = + compute-extension-fiber-precomp-Π (g ∘ f) + + compute-total-fiber-precomp-Π : + Σ ((b : B) → U b) (λ g → fiber (precomp-Π f U) (g ∘ f)) ≃ + Σ ((b : B) → U b) (λ u → Σ ((b : B) → U b) (λ v → u ∘ f ~ v ∘ f)) + compute-total-fiber-precomp-Π = equiv-tot compute-fiber-precomp-Π + + diagonal-into-fibers-precomp-Π : + ((b : B) → U b) → Σ ((b : B) → U b) (λ g → fiber (precomp-Π f U) (g ∘ f)) + diagonal-into-fibers-precomp-Π = map-section-family (λ g → (g , refl)) +``` + +- In + [`foundation.universal-property-family-of-fibers-of-maps`](foundation.universal-property-family-of-fibers-of-maps.md) + we compute the fiber family of dependent precomposition maps as a dependent + product + ```text + compute-fiber-Π-precomp-Π : + fiber (precomp-Π f U) g ≃ + ( (b : B) → + Σ (u : U b), + (((a , p) : fiber f b) → dependent-identification U p (g a) u)). + ``` + ### The action of dependent precomposition on homotopies ```agda diff --git a/src/foundation/precomposition-functions.lagda.md b/src/foundation/precomposition-functions.lagda.md index 9e48f70cfd..74deeba4c9 100644 --- a/src/foundation/precomposition-functions.lagda.md +++ b/src/foundation/precomposition-functions.lagda.md @@ -134,22 +134,22 @@ module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (X : UU l3) where - compute-coherence-triangle-fiber-precomp' : + compute-extension-fiber-precomp' : (g : A → X) → fiber (precomp f X) g ≃ Σ (B → X) (λ h → coherence-triangle-maps' g h f) - compute-coherence-triangle-fiber-precomp' g = equiv-tot (λ _ → equiv-funext) + compute-extension-fiber-precomp' g = equiv-tot (λ _ → equiv-funext) - compute-coherence-triangle-fiber-precomp : + compute-extension-fiber-precomp : (g : A → X) → fiber (precomp f X) g ≃ Σ (B → X) (λ h → coherence-triangle-maps g h f) - compute-coherence-triangle-fiber-precomp g = + compute-extension-fiber-precomp g = equiv-tot (λ _ → equiv-funext) ∘e equiv-fiber (precomp f X) g compute-fiber-precomp : (g : B → X) → fiber (precomp f X) (g ∘ f) ≃ Σ (B → X) (λ h → coherence-square-maps f f h g) - compute-fiber-precomp g = compute-coherence-triangle-fiber-precomp (g ∘ f) + compute-fiber-precomp g = compute-extension-fiber-precomp (g ∘ f) compute-total-fiber-precomp : Σ (B → X) (λ g → fiber (precomp f X) (g ∘ f)) ≃ diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md index f7dd57292f..da39546ed6 100644 --- a/src/foundation/truncation-equivalences.lagda.md +++ b/src/foundation/truncation-equivalences.lagda.md @@ -388,12 +388,12 @@ This is an instance of Corollary 2.29 in {{#cite CORS20}}. We consider the following composition of maps ```text - fiber f b = Σ A (λ a → f a = b) - → Σ A (λ a → ║f a = b║) - ≃ Σ A (λ a → |f a| = |b|) - ≃ Σ A (λ a → ║f║ |a| = |b|) - → Σ ║A║ (λ t → ║f║ t = |b|) - = fiber ║f║ |b| + fiber f b = Σ A (λ a → f a = b) + → Σ A (λ a → ║f a = b║) + ≃ Σ A (λ a → |f a| = |b|) + ≃ Σ A (λ a → ║f║ |a| = |b|) + → Σ ║A║ (λ t → ║f║ t = |b|) + = fiber ║f║ |b| ``` where the first and last maps are `k`-equivalences. diff --git a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md index a9d2691df4..fecf1fb87c 100644 --- a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md +++ b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md @@ -12,12 +12,15 @@ open import foundation.dependent-pair-types open import foundation.diagonal-maps-of-types open import foundation.families-of-equivalences open import foundation.function-extensionality +open import foundation.precomposition-dependent-functions open import foundation.subtype-identity-principle +open import foundation.type-theoretic-principle-of-choice open import foundation.universe-levels open import foundation-core.constant-maps open import foundation-core.contractible-maps open import foundation-core.contractible-types +open import foundation-core.dependent-identifications open import foundation-core.equivalences open import foundation-core.fibers-of-maps open import foundation-core.function-types @@ -25,7 +28,6 @@ open import foundation-core.functoriality-dependent-function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies open import foundation-core.identity-types -open import foundation-core.precomposition-dependent-functions open import foundation-core.retractions open import foundation-core.sections @@ -144,6 +146,10 @@ module _ lift-family-of-elements-fiber : lift-family-of-elements (fiber f) f pr1 (lift-family-of-elements-fiber a) = a pr2 (lift-family-of-elements-fiber a) = refl + + lift-family-of-elements-fiber' : lift-family-of-elements (fiber' f) f + pr1 (lift-family-of-elements-fiber' a) = a + pr2 (lift-family-of-elements-fiber' a) = refl ``` ## Properties @@ -152,58 +158,60 @@ module _ ```agda module _ - {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) + (C : (y : B) (z : fiber f y) → UU l3) where - module _ - {l3 : Level} (C : (y : B) (z : fiber f y) → UU l3) - where - - ev-lift-family-of-elements-fiber : - ((y : B) (z : fiber f y) → C y z) → ((x : A) → C (f x) (x , refl)) - ev-lift-family-of-elements-fiber = - ev-double-lift-family-of-elements (lift-family-of-elements-fiber f) + ev-lift-family-of-elements-fiber : + ((y : B) (z : fiber f y) → C y z) → ((x : A) → C (f x) (x , refl)) + ev-lift-family-of-elements-fiber = + ev-double-lift-family-of-elements (lift-family-of-elements-fiber f) - extend-lift-family-of-elements-fiber : - ((x : A) → C (f x) (x , refl)) → ((y : B) (z : fiber f y) → C y z) - extend-lift-family-of-elements-fiber h .(f x) (x , refl) = h x - - is-section-extend-lift-family-of-elements-fiber : - is-section - ( ev-lift-family-of-elements-fiber) - ( extend-lift-family-of-elements-fiber) - is-section-extend-lift-family-of-elements-fiber h = refl + extend-lift-family-of-elements-fiber : + ((x : A) → C (f x) (x , refl)) → ((y : B) (z : fiber f y) → C y z) + extend-lift-family-of-elements-fiber h .(f x) (x , refl) = h x - is-retraction-extend-lift-family-of-elements-fiber' : - (h : (y : B) (z : fiber f y) → C y z) (y : B) → - extend-lift-family-of-elements-fiber - ( ev-lift-family-of-elements-fiber h) - ( y) ~ - h y - is-retraction-extend-lift-family-of-elements-fiber' h .(f z) (z , refl) = - refl + is-section-extend-lift-family-of-elements-fiber : + is-section + ( ev-lift-family-of-elements-fiber) + ( extend-lift-family-of-elements-fiber) + is-section-extend-lift-family-of-elements-fiber h = refl + + htpy-is-retraction-extend-lift-family-of-elements-fiber : + (h : (y : B) (z : fiber f y) → C y z) (y : B) → + extend-lift-family-of-elements-fiber + ( ev-lift-family-of-elements-fiber h) + ( y) ~ + h y + htpy-is-retraction-extend-lift-family-of-elements-fiber h .(f z) (z , refl) = + refl + abstract is-retraction-extend-lift-family-of-elements-fiber : is-retraction ( ev-lift-family-of-elements-fiber) ( extend-lift-family-of-elements-fiber) is-retraction-extend-lift-family-of-elements-fiber h = - eq-htpy (eq-htpy ∘ is-retraction-extend-lift-family-of-elements-fiber' h) + eq-htpy + ( eq-htpy ∘ htpy-is-retraction-extend-lift-family-of-elements-fiber h) - is-equiv-extend-lift-family-of-elements-fiber : - is-equiv extend-lift-family-of-elements-fiber - is-equiv-extend-lift-family-of-elements-fiber = - is-equiv-is-invertible - ( ev-lift-family-of-elements-fiber) - ( is-retraction-extend-lift-family-of-elements-fiber) - ( is-section-extend-lift-family-of-elements-fiber) + is-equiv-extend-lift-family-of-elements-fiber : + is-equiv extend-lift-family-of-elements-fiber + is-equiv-extend-lift-family-of-elements-fiber = + is-equiv-is-invertible + ( ev-lift-family-of-elements-fiber) + ( is-retraction-extend-lift-family-of-elements-fiber) + ( is-section-extend-lift-family-of-elements-fiber) + + inv-equiv-dependent-universal-property-family-of-fibers : + ((x : A) → C (f x) (x , refl)) ≃ ((y : B) (z : fiber f y) → C y z) + inv-equiv-dependent-universal-property-family-of-fibers = + ( extend-lift-family-of-elements-fiber , + is-equiv-extend-lift-family-of-elements-fiber) - inv-equiv-dependent-universal-property-family-of-fibers : - ((x : A) → C (f x) (x , refl)) ≃ ((y : B) (z : fiber f y) → C y z) - pr1 inv-equiv-dependent-universal-property-family-of-fibers = - extend-lift-family-of-elements-fiber - pr2 inv-equiv-dependent-universal-property-family-of-fibers = - is-equiv-extend-lift-family-of-elements-fiber +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where dependent-universal-property-family-of-fibers-fiber : dependent-universal-property-family-of-fibers @@ -211,18 +219,95 @@ module _ ( lift-family-of-elements-fiber f) dependent-universal-property-family-of-fibers-fiber C = is-equiv-is-invertible - ( extend-lift-family-of-elements-fiber C) - ( is-section-extend-lift-family-of-elements-fiber C) - ( is-retraction-extend-lift-family-of-elements-fiber C) + ( extend-lift-family-of-elements-fiber f C) + ( is-section-extend-lift-family-of-elements-fiber f C) + ( is-retraction-extend-lift-family-of-elements-fiber f C) equiv-dependent-universal-property-family-of-fibers : {l3 : Level} (C : (y : B) (z : fiber f y) → UU l3) → ((y : B) (z : fiber f y) → C y z) ≃ ((x : A) → C (f x) (x , refl)) - pr1 (equiv-dependent-universal-property-family-of-fibers C) = - ev-lift-family-of-elements-fiber C - pr2 (equiv-dependent-universal-property-family-of-fibers C) = - dependent-universal-property-family-of-fibers-fiber C + equiv-dependent-universal-property-family-of-fibers C = + ( ev-lift-family-of-elements-fiber f C , + dependent-universal-property-family-of-fibers-fiber C) +``` + +### The variant family of fibers of a map satisfies the dependent universal property of the family of fibers of a map + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) + (C : (y : B) (z : fiber' f y) → UU l3) + where + + ev-lift-family-of-elements-fiber' : + ((y : B) (z : fiber' f y) → C y z) → ((x : A) → C (f x) (x , refl)) + ev-lift-family-of-elements-fiber' = + ev-double-lift-family-of-elements (lift-family-of-elements-fiber' f) + + extend-lift-family-of-elements-fiber' : + ((x : A) → C (f x) (x , refl)) → ((y : B) (z : fiber' f y) → C y z) + extend-lift-family-of-elements-fiber' h .(f x) (x , refl) = h x + + is-section-extend-lift-family-of-elements-fiber' : + is-section + ( ev-lift-family-of-elements-fiber') + ( extend-lift-family-of-elements-fiber') + is-section-extend-lift-family-of-elements-fiber' h = refl + + htpy-is-retraction-extend-lift-family-of-elements-fiber' : + (h : (y : B) (z : fiber' f y) → C y z) (y : B) → + extend-lift-family-of-elements-fiber' + ( ev-lift-family-of-elements-fiber' h) + ( y) ~ + h y + htpy-is-retraction-extend-lift-family-of-elements-fiber' h .(f z) (z , refl) = + refl + + abstract + is-retraction-extend-lift-family-of-elements-fiber' : + is-retraction + ( ev-lift-family-of-elements-fiber') + ( extend-lift-family-of-elements-fiber') + is-retraction-extend-lift-family-of-elements-fiber' h = + eq-htpy + ( eq-htpy ∘ htpy-is-retraction-extend-lift-family-of-elements-fiber' h) + + is-equiv-extend-lift-family-of-elements-fiber' : + is-equiv extend-lift-family-of-elements-fiber' + is-equiv-extend-lift-family-of-elements-fiber' = + is-equiv-is-invertible + ( ev-lift-family-of-elements-fiber') + ( is-retraction-extend-lift-family-of-elements-fiber') + ( is-section-extend-lift-family-of-elements-fiber') + + inv-equiv-dependent-universal-property-family-of-fibers' : + ((x : A) → C (f x) (x , refl)) ≃ ((y : B) (z : fiber' f y) → C y z) + inv-equiv-dependent-universal-property-family-of-fibers' = + ( extend-lift-family-of-elements-fiber' , + is-equiv-extend-lift-family-of-elements-fiber') + +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where + + dependent-universal-property-family-of-fibers-fiber' : + dependent-universal-property-family-of-fibers + ( fiber' f) + ( lift-family-of-elements-fiber' f) + dependent-universal-property-family-of-fibers-fiber' C = + is-equiv-is-invertible + ( extend-lift-family-of-elements-fiber' f C) + ( is-section-extend-lift-family-of-elements-fiber' f C) + ( is-retraction-extend-lift-family-of-elements-fiber' f C) + + equiv-dependent-universal-property-family-of-fibers' : + {l3 : Level} (C : (y : B) (z : fiber' f y) → UU l3) → + ((y : B) (z : fiber' f y) → C y z) ≃ + ((x : A) → C (f x) (x , refl)) + equiv-dependent-universal-property-family-of-fibers' C = + ( ev-lift-family-of-elements-fiber' f C , + dependent-universal-property-family-of-fibers-fiber' C) ``` ### The family of fibers of a map satisfies the universal property of the family of fibers of a map @@ -431,3 +516,68 @@ module _ ( is-equiv-map-Π-is-fiberwise-equiv H) ( universal-property-family-of-fibers-fiber f C) ``` + +### Computing the fibers of precomposition dependent functions as dependent products + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (U : B → UU l3) + where + + family-fiber-Π-precomp-Π : ((a : A) → U (f a)) → B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp-Π g b = + Σ (U b) (λ u → ((a , p) : fiber f b) → dependent-identification U p (g a) u) + + fiber-Π-precomp-Π : (g : (a : A) → U (f a)) → UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp-Π g = (b : B) → family-fiber-Π-precomp-Π g b + + compute-fiber-Π-precomp-Π : + (g : (a : A) → U (f a)) → fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π g + compute-fiber-Π-precomp-Π g = + equivalence-reasoning + fiber (precomp-Π f U) g + ≃ Σ ((b : B) → U b) (λ h → (a : A) → g a = (h ∘ f) a) + by compute-extension-fiber-precomp-Π f U g + ≃ Σ ( (b : B) → U b) + ( λ h → (b : B) ((a , p) : fiber f b) → + dependent-identification U p (g a) (h b)) + by + equiv-tot + ( λ h → + inv-equiv-dependent-universal-property-family-of-fibers f + ( λ y (a , p) → dependent-identification U p (g a) (h y))) + ≃ ( (b : B) → + Σ ( U b) + ( λ u → + ((a , p) : fiber f b) → dependent-identification U p (g a) u)) + by inv-distributive-Π-Σ + + family-fiber-Π-precomp-Π' : ((a : A) → U (f a)) → B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp-Π' g b = + Σ ( U b) + ( λ u → ((a , p) : fiber' f b) → dependent-identification U p u (g a)) + + fiber-Π-precomp-Π' : (g : (a : A) → U (f a)) → UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp-Π' g = (b : B) → family-fiber-Π-precomp-Π' g b + + compute-fiber-Π-precomp-Π' : + (g : (a : A) → U (f a)) → fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π' g + compute-fiber-Π-precomp-Π' g = + equivalence-reasoning + fiber (precomp-Π f U) g + ≃ Σ ((b : B) → U b) (λ h → (a : A) → (h ∘ f) a = g a) + by compute-extension-fiber-precomp-Π' f U g + ≃ Σ ( (b : B) → U b) + ( λ h → (b : B) ((a , p) : fiber' f b) → + dependent-identification U p (h b) (g a)) + by + equiv-tot + ( λ h → + inv-equiv-dependent-universal-property-family-of-fibers' f + ( λ y (a , p) → dependent-identification U p (h y) (g a))) + ≃ ( (b : B) → + Σ ( U b) + ( λ u → + ((a , p) : fiber' f b) → dependent-identification U p u (g a))) + by inv-distributive-Π-Σ +``` diff --git a/src/foundation/universal-property-unit-type.lagda.md b/src/foundation/universal-property-unit-type.lagda.md index ca15e0b392..bf7b958180 100644 --- a/src/foundation/universal-property-unit-type.lagda.md +++ b/src/foundation/universal-property-unit-type.lagda.md @@ -18,7 +18,10 @@ open import foundation-core.constant-maps open import foundation-core.contractible-types open import foundation-core.equivalences open import foundation-core.homotopies +open import foundation-core.identity-types open import foundation-core.precomposition-functions +open import foundation-core.retractions +open import foundation-core.sections ```
@@ -41,7 +44,7 @@ ev-star P f = f star ev-star' : {l : Level} (Y : UU l) → (unit → Y) → Y -ev-star' Y = ev-star (λ t → Y) +ev-star' Y = ev-star (λ _ → Y) abstract dependent-universal-property-unit : @@ -55,20 +58,35 @@ pr1 (equiv-dependent-universal-property-unit P) = ev-star P pr2 (equiv-dependent-universal-property-unit P) = dependent-universal-property-unit P -abstract - universal-property-unit : - {l : Level} (Y : UU l) → is-equiv (ev-star' Y) - universal-property-unit Y = dependent-universal-property-unit (λ t → Y) +module _ + {l : Level} (Y : UU l) + where + + is-section-const-unit : is-section (ev-star' Y) (const unit) + is-section-const-unit = refl-htpy + + is-retraction-const-unit : is-retraction (ev-star' Y) (const unit) + is-retraction-const-unit = refl-htpy + + universal-property-unit : is-equiv (ev-star' Y) + universal-property-unit = + is-equiv-is-invertible + ( const unit) + ( is-section-const-unit) + ( is-retraction-const-unit) + + equiv-universal-property-unit : (unit → Y) ≃ Y + equiv-universal-property-unit = (ev-star' Y , universal-property-unit) -equiv-universal-property-unit : - {l : Level} (Y : UU l) → (unit → Y) ≃ Y -pr1 (equiv-universal-property-unit Y) = ev-star' Y -pr2 (equiv-universal-property-unit Y) = universal-property-unit Y + is-equiv-const-unit : is-equiv (const unit) + is-equiv-const-unit = + is-equiv-is-invertible + ( ev-star' Y) + ( is-retraction-const-unit) + ( is-section-const-unit) -inv-equiv-universal-property-unit : - {l : Level} (Y : UU l) → Y ≃ (unit → Y) -inv-equiv-universal-property-unit Y = - inv-equiv (equiv-universal-property-unit Y) + inv-equiv-universal-property-unit : Y ≃ (unit → Y) + inv-equiv-universal-property-unit = (const unit , is-equiv-const-unit) abstract is-equiv-point-is-contr : @@ -123,3 +141,21 @@ abstract ( universal-property-unit-is-equiv-point x is-equiv-point Y) ( refl-htpy) ``` + +### Precomposition with the terminal map is equivalent to the diagonal exponential + +```agda +is-equiv-precomp-terminal-map-is-equiv-diagonal-exponential : + {l1 l2 : Level} {X : UU l1} {U : UU l2} → + is-equiv (diagonal-exponential U X) → + is-equiv (precomp (terminal-map X) U) +is-equiv-precomp-terminal-map-is-equiv-diagonal-exponential H = + is-equiv-left-factor _ _ H (is-equiv-const-unit _) + +is-equiv-diagonal-exponential-is-equiv-precomp-terminal-map : + {l1 l2 : Level} {X : UU l1} {U : UU l2} → + is-equiv (precomp (terminal-map X) U) → + is-equiv (diagonal-exponential U X) +is-equiv-diagonal-exponential-is-equiv-precomp-terminal-map H = + is-equiv-comp _ _ (is-equiv-const-unit _) H +``` diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index d5d3dc288d..21205c2716 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -71,6 +71,7 @@ open import orthogonal-factorization-systems.sigma-closed-reflective-modalities open import orthogonal-factorization-systems.sigma-closed-reflective-subuniverses public open import orthogonal-factorization-systems.stable-orthogonal-factorization-systems public open import orthogonal-factorization-systems.subuniverse-connected-maps public +open import orthogonal-factorization-systems.subuniverse-equivalences public open import orthogonal-factorization-systems.types-colocal-at-maps public open import orthogonal-factorization-systems.types-local-at-maps public open import orthogonal-factorization-systems.types-separated-at-maps public diff --git a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md index 14cc6f8f64..97f0b6d507 100644 --- a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md @@ -21,16 +21,17 @@ open import orthogonal-factorization-systems.types-local-at-maps ## Idea Let `P` be a [subuniverse](foundation.subuniverses.md). Given a type `X`, its -**localization** at `P`, or **`P`-localization**, is a type `Y` in `P` and a map -`η : X → Y` such that every type in `P` is -`η`[-local](orthogonal-factorization-systems.types-local-at-maps.md). I.e., for -every `Z` in `P`, the [precomposition map](foundation-core.function-types.md) +{{#concept "localization" Disambiguation="of a type at a subuniverse" Agda=subuniverse-localization}} +at `P`, or **`P`-localization**, is a type `Y` in `P` and a `P`-equivalence +`η : X → Y`, i.e., a map such that for every `Z` in `P` the +[precomposition map](foundation-core.function-types.md) ```text - ∘ η : (Y → Z) → (X → Z) ``` -is an [equivalence](foundation-core.equivalences.md). +is an [equivalence](foundation-core.equivalences.md). In other words, every type +in `P` is `η`[-local](orthogonal-factorization-systems.types-local-at-maps.md). ## Definition @@ -51,9 +52,7 @@ is-subuniverse-localization : {l1 l2 lP : Level} (P : subuniverse l1 lP) → UU l2 → UU l1 → UU (lsuc l1 ⊔ l2 ⊔ lP) is-subuniverse-localization {l1} {l2} P X Y = - ( is-in-subuniverse P Y) × - ( Σ ( X → Y) - ( λ η → (Z : UU l1) → is-in-subuniverse P Z → is-local η Z)) + (is-in-subuniverse P Y) × (subuniverse-equiv P X Y) ``` ```agda @@ -69,8 +68,7 @@ module _ unit-is-subuniverse-localization = pr1 (pr2 is-localization-Y) is-local-at-unit-is-in-subuniverse-is-subuniverse-localization : - (Z : UU l1) → is-in-subuniverse P Z → - is-local unit-is-subuniverse-localization Z + (Z : type-subuniverse P) → is-local unit-is-subuniverse-localization (pr1 Z) is-local-at-unit-is-in-subuniverse-is-subuniverse-localization = pr2 (pr2 is-localization-Y) ``` @@ -102,15 +100,19 @@ module _ is-in-subuniverse-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) + type-subuniverse-subuniverse-localization : type-subuniverse P + type-subuniverse-subuniverse-localization = + ( type-subuniverse-localization , + is-in-subuniverse-subuniverse-localization) + unit-subuniverse-localization : X → type-subuniverse-localization unit-subuniverse-localization = unit-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) - is-local-at-unit-is-in-subuniverse-subuniverse-localization : - (Z : UU l1) → - is-in-subuniverse P Z → is-local unit-subuniverse-localization Z - is-local-at-unit-is-in-subuniverse-subuniverse-localization = + is-subuniverse-equiv-unit-subuniverse-localization : + is-subuniverse-equiv P unit-subuniverse-localization + is-subuniverse-equiv-unit-subuniverse-localization = is-local-at-unit-is-in-subuniverse-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) ``` diff --git a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md index 41a30148b0..1cd22d5acc 100644 --- a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md @@ -186,7 +186,7 @@ module _ is-in-subuniverse-subuniverse-localization 𝒫 (L A) pr2 (pr2 (pr2 is-reflective-has-all-localizations-subuniverse)) A B is-in-subuniverse-A = - is-local-at-unit-is-in-subuniverse-subuniverse-localization + is-subuniverse-equiv-unit-subuniverse-localization 𝒫 (L B) A is-in-subuniverse-A ``` diff --git a/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md index 43a467525a..b68f408e34 100644 --- a/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md @@ -26,7 +26,7 @@ open import orthogonal-factorization-systems.postcomposition-extensions-maps Consider a [global subuniverse](foundation.global-subuniverses.md) `K` and a map `δ : X → Y`. A type `A` is said to be -{{#concept "`K`-separated relative to `δ`" Agda=is-map-separated}} if the type +{{#concept "`K`-separated relative to `δ`" Agda=is-rel-separated}} if the type of extensions of any map `f : X → A` along `δ` is in `K`. As a special case, a type `A` is _`K`-separated_ if it is `K`-separated relative @@ -42,19 +42,19 @@ module _ {X : UU l1} {Y : UU l2} (δ : X → Y) where - is-map-separated : (A : UU l3) → UU (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - is-map-separated A = + is-rel-separated : (A : UU l3) → UU (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) + is-rel-separated A = (f : X → A) → is-in-global-subuniverse K (extension δ f) - is-prop-is-map-separated : - (A : UU l3) → is-prop (is-map-separated A) - is-prop-is-map-separated A = + is-prop-is-rel-separated : + (A : UU l3) → is-prop (is-rel-separated A) + is-prop-is-rel-separated A = is-prop-Π (λ f → is-prop-is-in-global-subuniverse K (extension δ f)) - is-map-separated-Prop : + is-rel-separated-Prop : (A : UU l3) → Prop (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - is-map-separated-Prop A = - ( is-map-separated A , is-prop-is-map-separated A) + is-rel-separated-Prop A = + ( is-rel-separated A , is-prop-is-rel-separated A) ``` ### The global subuniverse of `K`-separated types relative to `δ` @@ -65,26 +65,26 @@ module _ {X : UU l1} {Y : UU l2} (δ : X → Y) where - is-closed-under-equiv-map-separated-global-subuniverse : + is-closed-under-equiv-rel-separated-global-subuniverse : {l4 l5 : Level} → is-closed-under-equiv-subuniverses ( λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - ( λ l3 → is-map-separated-Prop {l3 = l3} K δ) + ( λ l3 → is-rel-separated-Prop {l3 = l3} K δ) ( l4) ( l5) - is-closed-under-equiv-map-separated-global-subuniverse A B e H f = + is-closed-under-equiv-rel-separated-global-subuniverse A B e H f = is-closed-under-equiv-global-subuniverse K ( extension δ (map-inv-equiv e ∘ f)) ( extension δ f) ( inv-equiv (equiv-postcomp-extension δ f (inv-equiv e))) ( H (map-inv-equiv e ∘ f)) - map-separated-global-subuniverse : + rel-separated-global-subuniverse : global-subuniverse (λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - map-separated-global-subuniverse = + rel-separated-global-subuniverse = λ where .subuniverse-global-subuniverse l3 → - is-map-separated-Prop K δ + is-rel-separated-Prop K δ .is-closed-under-equiv-global-subuniverse → - is-closed-under-equiv-map-separated-global-subuniverse + is-closed-under-equiv-rel-separated-global-subuniverse ``` diff --git a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md index de39d97c9f..43c124f49e 100644 --- a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md @@ -23,12 +23,12 @@ open import orthogonal-factorization-systems.extensions-maps Consider a [subuniverse](foundation.subuniverses.md) `K` and a map `δ : X → Y`. A type `A` is said to be -{{#concept "`K`-separated relative to `δ`" Agda=is-map-separated'}} if the type +{{#concept "`K`-separated relative to `δ`" Agda=is-rel-separated'}} if the type of extensions of any map `f : X → A` along `δ` is in `K`. ## Definitions -### The predicate of being map-separated +### The predicate of being rel-separated ```agda module _ @@ -36,20 +36,20 @@ module _ {X : UU l1} {Y : UU l1} (δ : X → Y) where - is-map-separated' : (A : UU l1) → UU (l1 ⊔ l2) - is-map-separated' A = (f : X → A) → is-in-subuniverse K (extension δ f) + is-rel-separated' : (A : UU l1) → UU (l1 ⊔ l2) + is-rel-separated' A = (f : X → A) → is-in-subuniverse K (extension δ f) - is-prop-is-map-separated' : - (A : UU l1) → is-prop (is-map-separated' A) - is-prop-is-map-separated' A = + is-prop-is-rel-separated' : + (A : UU l1) → is-prop (is-rel-separated' A) + is-prop-is-rel-separated' A = is-prop-Π (λ f → is-prop-is-in-subuniverse K (extension δ f)) - is-map-separated'-Prop : + is-rel-separated'-Prop : (A : UU l1) → Prop (l1 ⊔ l2) - is-map-separated'-Prop A = - ( is-map-separated' A , is-prop-is-map-separated' A) + is-rel-separated'-Prop A = + ( is-rel-separated' A , is-prop-is-rel-separated' A) ``` ## See also -- [Map-separated types at global subuniverses](orthogonal-factorization-systems.map-separated-types-global-subuniverses.md) +- [Map-separated types at global subuniverses](orthogonal-factorization-systems.rel-separated-types-global-subuniverses.md) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 07ee93aad2..daec6996df 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -8,24 +8,30 @@ module orthogonal-factorization-systems.subuniverse-connected-maps where ```agda open import foundation.connected-types +open import foundation.constant-maps open import foundation.contractible-types +open import foundation.coproduct-types +open import foundation.dependent-identifications open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences +open import foundation.diagonal-maps-of-types open import foundation.equivalences -open import foundation.coproduct-types -open import orthogonal-factorization-systems.orthogonal-maps open import foundation.equivalences-arrows open import foundation.function-extensionality open import foundation.functoriality-cartesian-product-types open import foundation.functoriality-coproduct-types -open import foundation.universal-property-coproduct-types open import foundation.functoriality-dependent-function-types open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopy-induction +open import foundation.inhabited-types open import foundation.iterated-successors-truncation-levels +open import foundation.postcomposition-functions open import foundation.precomposition-dependent-functions +open import foundation.precomposition-functions open import foundation.precomposition-type-families open import foundation.propositional-truncations +open import foundation.sections +open import foundation.split-surjective-maps open import foundation.structure-identity-principle open import foundation.subtype-identity-principle open import foundation.subuniverses @@ -33,9 +39,12 @@ open import foundation.surjective-maps open import foundation.truncated-types open import foundation.truncation-levels open import foundation.truncations +open import foundation.unit-type open import foundation.univalence +open import foundation.universal-property-coproduct-types open import foundation.universal-property-dependent-pair-types open import foundation.universal-property-family-of-fibers-of-maps +open import foundation.universal-property-unit-type open import foundation.universe-levels open import foundation-core.contractible-maps @@ -45,11 +54,18 @@ open import foundation-core.function-types open import foundation-core.functoriality-dependent-pair-types open import foundation-core.homotopies open import foundation-core.identity-types +open import foundation-core.injective-maps +open import foundation-core.postcomposition-dependent-functions open import foundation-core.propositions open import foundation-core.subtypes open import foundation-core.torsorial-type-families open import foundation-core.truncated-maps +open import orthogonal-factorization-systems.localizations-at-subuniverses +open import orthogonal-factorization-systems.orthogonal-maps +open import orthogonal-factorization-systems.relative-separations-types-subuniverses +open import orthogonal-factorization-systems.subuniverse-equivalences + open import synthetic-homotopy-theory.cocones-under-spans open import synthetic-homotopy-theory.dependent-pullback-property-pushouts open import synthetic-homotopy-theory.pushouts @@ -73,7 +89,7 @@ For every `K`-valued family `U` over `B`, the - ∘ f : ((b : B) → U b) → ((a : A) → U (f a)) ``` -is an equivalence. +is an [equivalence](foundation-core.equivalences.md). Equivalently, a `K`-connected map `f : A → B` is a map that is [left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to maps @@ -93,8 +109,8 @@ is-subuniverse-connected-map-Prop K {B = B} f = is-subuniverse-connected-map : {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → (A → B) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -is-subuniverse-connected-map K f = - type-Prop (is-subuniverse-connected-map-Prop K f) +is-subuniverse-connected-map K {B = B} f = + (U : B → type-subuniverse K) → is-equiv (precomp-Π f (pr1 ∘ U)) is-prop-is-subuniverse-connected-map : {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) @@ -167,6 +183,200 @@ module _ ## Properties +### Equivalent characterizations of `K`-connected maps + +#### Contractible fiber-localization condition + +Given a subuniverse `K` then a map `f` is `K`-connected if if the fibers are +`K`-equivalent to contractible types. + +**Proof.** We have an equivalence of arrows + +```text + ~ + ((b : B) → unit → U b) --------> ((b : B) → U b) + | | + | | precomp-Π f U + | | + ∨ ~ ∨ + ((b : B) → fiber f b → U b) ---> ((a : A) → U (f a)) +``` + +where the left-hand map is +`map-Π (λ b → precomp (terminal-map (fiber f b)) (U b))`, hence if each terminal +map is a `K`-equivalence then this is an equivalence as well. ∎ + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) + where + + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers : + ((b : B) → is-subuniverse-equiv K (terminal-map (fiber f b))) → + is-subuniverse-connected-map K f + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers H U = + is-equiv-target-is-equiv-source-equiv-arrow + ( map-Π (λ b → precomp (terminal-map (fiber f b)) (pr1 (U b)))) + ( precomp-Π f (pr1 ∘ U)) + ( ( equiv-Π-equiv-family + ( λ b → equiv-universal-property-unit (pr1 (U b)))) , + ( equiv-universal-property-family-of-fibers f (pr1 ∘ U)) , + ( refl-htpy)) + ( is-equiv-map-Π-is-fiberwise-equiv (λ b → H b (U b))) +``` + +#### Constant map condition on fibers + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) + where + + is-subuniverse-connected-map-is-equiv-diagonal-exponential-fibers : + ( (b : B) (U : type-subuniverse K) → + is-equiv (diagonal-exponential (pr1 U) (fiber f b))) → + is-subuniverse-connected-map K f + is-subuniverse-connected-map-is-equiv-diagonal-exponential-fibers H = + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f + ( λ b U → + is-equiv-precomp-terminal-map-is-equiv-diagonal-exponential (H b U)) +``` + +#### Section condition + +A map is `K`-connected if and only if its dependent precomposition maps admit +sections and the fibers have `K`-localizations. + +In fact, the following conditions suffice: + +There is a family `Kfib` together with a map `η : (b : B) → fiber f b → Kfib b` +such that the precomposition map of `η b` at `Kfib b` is injective, so a weak +kind of `Kfib b`-[epimorphism](foundation.epimorphisms.md), and the fiber of the +dependent precomposition map of `f` along `Kfib` over `λ a → η (f a) (a , refl)` +is inhabited. + +```agda +module _ + {l1 l2 l3 : Level} + {A : UU l1} {B : UU l2} (f : A → B) + where + + is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' : + (Kfib : B → UU l3) → + (η : (b : B) → fiber f b → Kfib b) → + ( (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → + ( fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl))) → + ((b : B) → is-contr (Kfib b)) + is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' + Kfib η is-htpy-injective-precomp-η-Kfib Fη b = + ( pr1 (Fη b) , + is-htpy-injective-precomp-η-Kfib b + ( λ where (a , refl) → pr2 (Fη b) (a , refl))) + + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' : + (Kfib : B → UU l3) → + (η : (b : B) → fiber f b → Kfib b) → + ( (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → + fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl)) → + ((b : B) → is-contr (Kfib b)) + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' + Kfib η is-htpy-injective-precomp-η-Kfib (s , H) b = + ( s b , + is-htpy-injective-precomp-η-Kfib b (λ where (a , refl) → htpy-eq H a)) + + is-contr-subuniverse-localization-fiber-has-section-precomp-Π : + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → + ((U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) → + ((b : B) → is-contr (pr1 (Kfib b))) + is-contr-subuniverse-localization-fiber-has-section-precomp-Π K Kfib s = + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' + ( type-subuniverse-localization K ∘ Kfib) + ( unit-subuniverse-localization K ∘ Kfib) + ( λ b H → + htpy-eq + ( is-injective-is-equiv + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) + ( type-subuniverse-subuniverse-localization K (Kfib b))) + ( eq-htpy H))) + ( is-split-surjective-section + ( precomp-Π f (type-subuniverse-localization K ∘ Kfib)) + ( s (type-subuniverse-subuniverse-localization K ∘ Kfib)) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) +``` + +#### Surjection condition + +A map is `K`-connected if and only if its dependent precomposition maps are +surjective and the fibers have `K`-localizations. + +In fact, it suffices that a certain family is inhabited. + +```agda +module _ + {l1 l2 l3 : Level} + {A : UU l1} {B : UU l2} (f : A → B) + where + + abstract + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' : + (Kfib : B → UU l3) → + (η : (b : B) → fiber f b → Kfib b) → + ( (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → + ( (b : B) → + is-inhabited + ( family-fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl)) b)) → + ((b : B) → is-contr (Kfib b)) + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' + Kfib η is-htpy-injective-precomp-η-Kfib Fη b = + rec-trunc-Prop + ( is-contr-Prop (Kfib b)) + ( λ (sb , Hb) → + ( sb , + is-htpy-injective-precomp-η-Kfib b + ( λ where (a , refl) → Hb (a , refl)))) + ( Fη b) + + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π : + (Kfib : (b : B) → UU l3) → + (η : (b : B) → fiber f b → Kfib b) → + ( (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → + is-inhabited (fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl))) → + ((b : B) → is-contr (Kfib b)) + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π + Kfib η is-htpy-injective-precomp-η-Kfib s b = + rec-trunc-Prop + ( is-contr-Prop (Kfib b)) + ( λ s → + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' + f Kfib η is-htpy-injective-precomp-η-Kfib s b) + ( s) + + is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π : + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → + ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → + ((b : B) → is-contr (pr1 (Kfib b))) + is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π K Kfib s = + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π + ( type-subuniverse-localization K ∘ Kfib) + ( unit-subuniverse-localization K ∘ Kfib) + ( λ b H → + htpy-eq + ( is-injective-is-equiv + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) + ( type-subuniverse-subuniverse-localization K (Kfib b))) + ( eq-htpy H))) + ( s ( type-subuniverse-subuniverse-localization K ∘ Kfib) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) +``` + ### Characterizing equality of `K`-connected maps ```agda @@ -258,6 +468,19 @@ module _ ( map-equiv e , is-subuniverse-connected-map-equiv e) ``` +### `K`-connected maps are `K`-equivalences + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) + where + + is-subuniverse-equiv-is-subuniverse-connected-map : + is-subuniverse-connected-map K f → is-subuniverse-equiv K f + is-subuniverse-equiv-is-subuniverse-connected-map F U = F (λ _ → U) +``` + ### The composition of two `K`-connected maps is `K`-connected ```agda @@ -479,3 +702,35 @@ module _ refl-htpy) ( is-equiv-map-product _ _ (F (U ∘ inl)) (F' (U ∘ inr))) ``` + +### The map on dependent pair types induced by a `K`-connected map is a `K`-equivalence + +This requires the added assumption that `K` is closed under exponentiation by +arbitrary types. + +This is a generalization of Lemma 2.27 in {{#cite CORS20}}, since the unit map +of the $ΣK$-localization is a $ΣK$-equivalence, and $ΣK$-equivalences are in +particular $K$-connected. + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : subuniverse-connected-map K A B) (P : B → UU l3) + where + + map-Σ-map-base-subuniverse-connected-map : + Σ A (P ∘ map-subuniverse-connected-map K f) → Σ B P + map-Σ-map-base-subuniverse-connected-map = + map-Σ-map-base (map-subuniverse-connected-map K f) P + + is-subuniverse-equiv-map-Σ-map-base-subuniverse-connected-map : + ((V : UU l3) (U : type-subuniverse K) → is-in-subuniverse K (V → pr1 U)) → + is-subuniverse-equiv K map-Σ-map-base-subuniverse-connected-map + is-subuniverse-equiv-map-Σ-map-base-subuniverse-connected-map H U = + is-equiv-source-is-equiv-target-equiv-arrow + ( precomp map-Σ-map-base-subuniverse-connected-map (pr1 U)) + ( precomp-Π (map-subuniverse-connected-map K f) (λ y → (b : P y) → pr1 U)) + ( equiv-ev-pair , equiv-ev-pair , refl-htpy) + ( is-subuniverse-connected-map-subuniverse-connected-map K f + ( λ y → ((P y → pr1 U) , H (P y) U))) +``` diff --git a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md new file mode 100644 index 0000000000..dcb47157bb --- /dev/null +++ b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md @@ -0,0 +1,495 @@ +# `K`-Equivalences + +```agda +module orthogonal-factorization-systems.subuniverse-equivalences where +``` + +
Imports + +```agda +open import foundation.commuting-squares-of-maps +open import foundation.connected-maps +open import foundation.connected-types +open import foundation.contractible-types +open import foundation.dependent-pair-types +open import foundation.functoriality-truncation +open import foundation.identity-types +open import foundation.precomposition-functions +open import foundation.precomposition-functions-into-subuniverses +open import foundation.propositional-truncations +open import foundation.subuniverses +open import foundation.truncations +open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universal-property-dependent-pair-types +open import foundation.universal-property-equivalences +open import foundation.universal-property-truncation +open import foundation.universe-levels + +open import foundation-core.contractible-maps +open import foundation-core.equivalences +open import foundation-core.fibers-of-maps +open import foundation-core.function-types +open import foundation-core.functoriality-dependent-pair-types +open import foundation-core.homotopies +open import foundation-core.retractions +open import foundation-core.sections +open import foundation-core.transport-along-identifications +open import foundation-core.truncated-maps +open import foundation-core.truncated-types +open import foundation-core.truncation-levels +``` + +
+ +## Idea + +Given a [subuniverse](foundation.subuniverses.md) `K`, A map `f : A → B` is said +to be a +{{#concept "`K`-equivalence" Disambiguation="map of types" Agda=is-subuniverse-equiv}} +if it satisfies the +{{#concept "universal property" Disambiguation="subuniverse connected map of types"}} +of `K`-equivalences: + +For every `K`-type `U`, the +[precomposition map](foundation-core.precomposition-functions.md) + +```text + - ∘ f : (B → U) → (A → U) +``` + +is an [equivalence](foundation-core.equivalences.md). + +## Definition + +```agda +is-subuniverse-equiv : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → + (A → B) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) +is-subuniverse-equiv K f = + (U : type-subuniverse K) → is-equiv (precomp f (pr1 U)) + +subuniverse-equiv : + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) → + UU l1 → UU l2 → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) +subuniverse-equiv K A B = Σ (A → B) (is-subuniverse-equiv K) + +module _ + {l1 l2 l3 l4 : Level} + (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} + (f : subuniverse-equiv K A B) + where + + map-subuniverse-equiv : A → B + map-subuniverse-equiv = pr1 f + + is-subuniverse-equiv-subuniverse-equiv : + is-subuniverse-equiv K map-subuniverse-equiv + is-subuniverse-equiv-subuniverse-equiv = pr2 f +``` + +## Properties + +### Equivalences are `K`-equivalences for all `K` + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) + where + + is-subuniverse-equiv-is-equiv : is-equiv f → is-subuniverse-equiv K f + is-subuniverse-equiv-is-equiv H U = is-equiv-precomp-is-equiv f H (pr1 U) +``` + +### The identity map is a `K`-equivalence for all `K` + +```agda +is-subuniverse-equiv-id : + {l1 l2 l3 : Level} (K : subuniverse l2 l3) {A : UU l1} → + is-subuniverse-equiv K (id' A) +is-subuniverse-equiv-id K = is-subuniverse-equiv-is-equiv K id is-equiv-id +``` + +### The `K`-equivalences are closed under homotopies + +```agda +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} {f g : A → B} + where + + is-subuniverse-equiv-htpy : + f ~ g → is-subuniverse-equiv K g → is-subuniverse-equiv K f + is-subuniverse-equiv-htpy H G U = + is-equiv-htpy (precomp g (pr1 U)) (htpy-precomp H (pr1 U)) (G U) + + is-subuniverse-equiv-htpy' : + f ~ g → is-subuniverse-equiv K f → is-subuniverse-equiv K g + is-subuniverse-equiv-htpy' H G U = + is-equiv-htpy' (precomp f (pr1 U)) (htpy-precomp H (pr1 U)) (G U) +``` + +### The `K`-equivalences are closed under composition + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-subuniverse-equiv-comp : + (g : B → C) (f : A → B) → + is-subuniverse-equiv K f → + is-subuniverse-equiv K g → + is-subuniverse-equiv K (g ∘ f) + is-subuniverse-equiv-comp g f F G U = + is-equiv-comp (precomp f (pr1 U)) (precomp g (pr1 U)) (G U) (F U) + + subuniverse-equiv-comp : + subuniverse-equiv K B C → + subuniverse-equiv K A B → + subuniverse-equiv K A C + pr1 (subuniverse-equiv-comp g f) = + map-subuniverse-equiv K g ∘ map-subuniverse-equiv K f + pr2 (subuniverse-equiv-comp g f) = + is-subuniverse-equiv-comp + ( map-subuniverse-equiv K g) + ( map-subuniverse-equiv K f) + ( is-subuniverse-equiv-subuniverse-equiv K f) + ( is-subuniverse-equiv-subuniverse-equiv K g) +``` + +### The class of `K`-equivalences has the 3-for-2 property + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {C : UU l3} + (g : B → C) (f : A → B) (GF : is-subuniverse-equiv K (g ∘ f)) + where + + is-subuniverse-equiv-left-factor : + is-subuniverse-equiv K f → is-subuniverse-equiv K g + is-subuniverse-equiv-left-factor F U = + is-equiv-right-factor (precomp f (pr1 U)) (precomp g (pr1 U)) (F U) (GF U) + + is-subuniverse-equiv-right-factor : + is-subuniverse-equiv K g → is-subuniverse-equiv K f + is-subuniverse-equiv-right-factor G U = + is-equiv-left-factor (precomp f (pr1 U)) (precomp g (pr1 U)) (GF U) (G U) +``` + +### Sections of `K`-equivalences are `K`-equivalences + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-subuniverse-equiv-map-section : + (K : subuniverse l3 l4) (s : section f) → + is-subuniverse-equiv K f → + is-subuniverse-equiv K (map-section f s) + is-subuniverse-equiv-map-section K (s , h) = + is-subuniverse-equiv-right-factor K f s + ( is-subuniverse-equiv-is-equiv K (f ∘ s) (is-equiv-htpy-id h)) +``` + +### Retractions of `K`-equivalences are `K`-equivalences + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {f : A → B} + where + + is-subuniverse-equiv-map-retraction : + (K : subuniverse l3 l4) (r : retraction f) → + is-subuniverse-equiv K f → + is-subuniverse-equiv K (map-retraction f r) + is-subuniverse-equiv-map-retraction K (r , h) = + is-subuniverse-equiv-left-factor K r f + ( is-subuniverse-equiv-is-equiv K (r ∘ f) (is-equiv-htpy-id h)) +``` + +### Composing `K`-equivalences with equivalences + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {C : UU l3} + where + + is-subuniverse-equiv-is-equiv-is-subuniverse-equiv : + (g : B → C) (f : A → B) → + is-subuniverse-equiv K g → + is-equiv f → + is-subuniverse-equiv K (g ∘ f) + is-subuniverse-equiv-is-equiv-is-subuniverse-equiv g f eg ef = + is-subuniverse-equiv-comp K g f + ( is-subuniverse-equiv-is-equiv K f ef) + ( eg) + + is-subuniverse-equiv-is-subuniverse-equiv-is-equiv : + (g : B → C) (f : A → B) → + is-equiv g → + is-subuniverse-equiv K f → + is-subuniverse-equiv K (g ∘ f) + is-subuniverse-equiv-is-subuniverse-equiv-is-equiv g f eg ef = + is-subuniverse-equiv-comp K g f + ( ef) + ( is-subuniverse-equiv-is-equiv K g eg) + + is-subuniverse-equiv-equiv-is-subuniverse-equiv : + (g : B → C) (f : A ≃ B) → + is-subuniverse-equiv K g → + is-subuniverse-equiv K (g ∘ map-equiv f) + is-subuniverse-equiv-equiv-is-subuniverse-equiv g (f , ef) eg = + is-subuniverse-equiv-is-equiv-is-subuniverse-equiv g f eg ef + + is-subuniverse-equiv-is-subuniverse-equiv-equiv : + (g : B ≃ C) (f : A → B) → + is-subuniverse-equiv K f → + is-subuniverse-equiv K (map-equiv g ∘ f) + is-subuniverse-equiv-is-subuniverse-equiv-equiv (g , eg) f ef = + is-subuniverse-equiv-is-subuniverse-equiv-is-equiv g f eg ef +``` + +### Being `K`-connected is invariant under `K`-equivalences + +```text +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} + where + + is-connected-is-subuniverse-equiv-is-connected : + (f : A → B) → is-subuniverse-equiv K f → + is-connected K B → is-connected K A + is-connected-is-subuniverse-equiv-is-connected f e = + is-contr-equiv (type-trunc K B) (map-trunc K f , e) + + is-connected-is-subuniverse-equiv-is-connected' : + (f : A → B) → is-subuniverse-equiv K f → + is-connected K A → is-connected K B + is-connected-is-subuniverse-equiv-is-connected' f e = + is-contr-equiv' (type-trunc K A) (map-trunc K f , e) + + is-connected-subuniverse-equiv-is-connected : + subuniverse-equiv K A B → is-connected K B → is-connected K A + is-connected-subuniverse-equiv-is-connected f = + is-connected-is-subuniverse-equiv-is-connected + ( map-subuniverse-equiv K f) + ( is-subuniverse-equiv-subuniverse-equiv K f) + + is-connected-subuniverse-equiv-is-connected' : + subuniverse-equiv K A B → is-connected K A → is-connected K B + is-connected-subuniverse-equiv-is-connected' f = + is-connected-is-subuniverse-equiv-is-connected' + ( map-subuniverse-equiv K f) + ( is-subuniverse-equiv-subuniverse-equiv K f) +``` + +### Every `ΣK`-equivalence is `K`-connected + +This is a generalization of Proposition 2.30 in {{#cite CORS20}}. + +```text +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} (f : A → B) + where + + is-connected-map-is-succ-subuniverse-equiv : + is-subuniverse-equiv (succ-𝕋 K) f → is-connected-map K f + is-connected-map-is-succ-subuniverse-equiv e b = + is-connected-subuniverse-equiv-is-connected + ( subuniverse-equiv-fiber-map-trunc-fiber f b) + ( is-connected-map-is-equiv e (unit-trunc b)) +``` + +### A map is `K`-connected if and only if its `ΣK`-localization is + +```text +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} {f : A → B} + where + + is-connected-map-trunc-succ-is-succ-connected-domain : + is-connected-map K f → + is-connected-map K (map-trunc (succ-𝕋 K) f) + is-connected-map-trunc-succ-is-succ-connected-domain cf t = + apply-universal-property-trunc-Prop + ( is-surjective-unit-trunc-succ t) + ( is-connected-Prop K (fiber (map-trunc (succ-𝕋 K) f) t)) + ( λ (b , p) → + tr + ( λ s → is-connected K (fiber (map-trunc (succ-𝕋 K) f) s)) + ( p) + ( is-connected-subuniverse-equiv-is-connected' + ( subuniverse-equiv-fiber-map-trunc-fiber f b) + ( cf b))) + + is-connected-map-is-connected-map-trunc-succ : + is-connected-map K (map-trunc (succ-𝕋 K) f) → + is-connected-map K f + is-connected-map-is-connected-map-trunc-succ cf' b = + is-connected-subuniverse-equiv-is-connected + ( subuniverse-equiv-fiber-map-trunc-fiber f b) + ( cf' (unit-trunc b)) +``` + +### The codomain of a `K`-connected map is `ΣK`-connected if its domain is `ΣK`-connected + +This follows part of the proof of Proposition 2.31 in {{#cite CORS20}}. + +**Proof.** Let $f : A → B$ be a $K$-connected map on a $K+1$-connected domain. +To show that the codomain is $K+1$-connected it is enough to show that $f$ is a +$K+1$-equivalence, in other words, that $║f║ₖ₊₁$ is an equivalence. By previous +computations we know that $║f║ₖ₊₁$ is $K$-truncated since the domain is +$K+1$-connected, and that $║f║ₖ₊₁$ is $K$-connected since $f$ is $K$-connected, +so we are done. ∎ + +```text +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} (f : A → B) + where + + is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map : + is-connected-map K f → + is-connected (succ-𝕋 K) A → + is-subuniverse-equiv (succ-𝕋 K) f + is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map + cf cA = + is-equiv-is-connected-map-is-trunc-map + ( is-trunc-map-trunc-succ-is-succ-connected-domain f cA) + ( is-connected-map-trunc-succ-is-succ-connected-domain cf) + + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map : + is-connected-map K f → + is-connected (succ-𝕋 K) A → + is-connected (succ-𝕋 K) B + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map cf cA = + is-connected-is-subuniverse-equiv-is-connected' f + ( is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map + ( cf) + ( cA)) + ( cA) +``` + +### If `g ∘ f` is `ΣK`-connected, then `f` is `K`-connected if and only if `g` is `ΣK`-connected + +This is an instance of Proposition 2.31 in {{#cite CORS20}}. + +**Proof.** If $g$ is $ΣK$-connected then by the cancellation property of +$ΣK$-equivalences, $f$ is a $K+1$-equivalence, and so in particular +$K$-connected. + +Conversely, assume $f$ is $K$-connected. We want to show that the fibers of $g$ +are $K+1$-connected, so let $c$ be an element of the codomain of $g$. The fibers +of the composite $g ∘ f$ compute as + +$$ + \operatorname{fiber}_{g\circ f}(c) ≃ + \sum_{(b , p) : \operatorname{fiber}_{g}(c)}{\operatorname{fiber}_{f}(b)}. +$$ + +By the previous lemma, since $\operatorname{fiber}_{g\circ f}(c)$ is +$K+1$-connected, $\operatorname{fiber}_{g}(c)$ is $K+1$-connected if the first +projection map of this type is $K$-connected, and its fibers compute to the +fibers of $f$. ∎ + +```text +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) {A : UU l1} {B : UU l2} {C : UU l3} + (g : B → C) (f : A → B) (cgf : is-connected-map (succ-𝕋 K) (g ∘ f)) + where + + is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor : + is-connected-map (succ-𝕋 K) g → is-subuniverse-equiv (succ-𝕋 K) f + is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor + cg = + is-subuniverse-equiv-right-factor g f + ( is-subuniverse-equiv-is-connected-map (g ∘ f) cgf) + ( is-subuniverse-equiv-is-connected-map g cg) + + is-connected-map-right-factor-is-succ-connected-map-left-factor : + is-connected-map (succ-𝕋 K) g → is-connected-map K f + is-connected-map-right-factor-is-succ-connected-map-left-factor cg = + is-connected-map-is-succ-subuniverse-equiv f + ( is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor + ( cg)) + + is-connected-map-right-factor-is-succ-connected-map-right-factor : + is-connected-map K f → is-connected-map (succ-𝕋 K) g + is-connected-map-right-factor-is-succ-connected-map-right-factor cf c = + is-succ-connected-codomain-is-succ-connected-domain-is-connected-map + ( pr1) + ( λ p → + is-connected-equiv + ( equiv-fiber-pr1 (fiber f ∘ pr1) p) + ( cf (pr1 p))) + ( is-connected-equiv' (compute-fiber-comp g f c) (cgf c)) +``` + +As a corollary, if $g ∘ f$ is $(K + 1)$-connected for some $g$, and $f$ is +$K$-connected, then $f$ is a $K+1$-equivalence. + +```text + is-succ-truncation-equiv-is-succ-connected-comp : + is-connected-map K f → is-subuniverse-equiv (succ-𝕋 K) f + is-succ-truncation-equiv-is-succ-connected-comp cf = + is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor + ( is-connected-map-right-factor-is-succ-connected-map-right-factor cf) +``` + +### A `K`-equivalence with a section is `K`-connected + +**Proof.** If $K ≐ -2$ notice that every map is $-2$-connected. So let +$K ≐ n + 1$ for some truncation level $n$ and let $f$ be our $K$-equivalence +with a section $s$. By assumption, we have a commuting triangle of maps + +```text + A + ∧ \ + s / \ f + / ∨ + B ======== B. +``` + +By the previous lemma, since the identity map is $K$-connected, it thus suffices +to show that $s$ is $n$-connected. But by the cancellation property of +$n+1$-equivalences $s$ is an $n+1$-equivalence and $n+1$-equivalences are in +particular $n$-connected. ∎ + +```text +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where + + is-connected-map-section-is-subuniverse-equiv-succ : + (K : subuniverse l3 l4) (s : section f) → + is-subuniverse-equiv (succ-𝕋 K) f → + is-connected-map K (map-section f s) + is-connected-map-section-is-subuniverse-equiv-succ K (s , h) e = + is-connected-map-is-succ-subuniverse-equiv s + ( is-subuniverse-equiv-map-section (succ-𝕋 K) (s , h) e) + + is-connected-map-is-subuniverse-equiv-section : + (K : subuniverse l3 l4) → + section f → is-subuniverse-equiv K f → is-connected-map K f + is-connected-map-is-subuniverse-equiv-section neg-two-𝕋 (s , h) e = + is-neg-two-connected-map f + is-connected-map-is-subuniverse-equiv-section (succ-𝕋 K) (s , h) e = + is-connected-map-right-factor-is-succ-connected-map-right-factor f s + ( is-connected-map-htpy-id h) + ( is-connected-map-section-is-subuniverse-equiv-succ K (s , h) e) +``` + +## References + +- The notion of `K`-equivalence is a special case of the notion of + `L`-equivalence, where `L` is a reflective subuniverse. These were studied in + {{#cite CORS20}}. +- The class of `K`-equivalences is + [left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to the + class of `K`-étale maps. This was shown in {{#cite CR21}}. + +{{#bibliography}} diff --git a/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md b/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md index a909c66dac..c5e4c76aa2 100644 --- a/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md @@ -28,11 +28,11 @@ module _ {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) where - is-map-separated-family : (X → UU l3) → UU (l1 ⊔ l2 ⊔ l3) - is-map-separated-family A = (x : X) (y z : A x) → is-local f (y = z) + is-rel-separated-family : (X → UU l3) → UU (l1 ⊔ l2 ⊔ l3) + is-rel-separated-family A = (x : X) (y z : A x) → is-local f (y = z) - is-map-separated : UU l3 → UU (l1 ⊔ l2 ⊔ l3) - is-map-separated A = is-map-separated-family (λ _ → A) + is-rel-separated : UU l3 → UU (l1 ⊔ l2 ⊔ l3) + is-rel-separated A = is-rel-separated-family (λ _ → A) ``` ## References From d5fc08b9604026f636e44c35048b3ba9caa31726 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 2 Oct 2025 21:08:32 +0200 Subject: [PATCH 34/69] more fiber computation --- ...property-family-of-fibers-of-maps.lagda.md | 145 ++++++++++++++++-- 1 file changed, 131 insertions(+), 14 deletions(-) diff --git a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md index fecf1fb87c..f70e51e4ec 100644 --- a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md +++ b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md @@ -13,8 +13,10 @@ open import foundation.diagonal-maps-of-types open import foundation.families-of-equivalences open import foundation.function-extensionality open import foundation.precomposition-dependent-functions +open import foundation.precomposition-functions open import foundation.subtype-identity-principle open import foundation.type-theoretic-principle-of-choice +open import foundation.universal-property-dependent-pair-types open import foundation.universe-levels open import foundation-core.constant-maps @@ -519,21 +521,32 @@ module _ ### Computing the fibers of precomposition dependent functions as dependent products +We give four equivalences for the fibers of precomposition dependent functions +as dependent products: + +```text + fiber (precomp-Π f U) g + ≃ (b : B) → Σ (u : U b), ((a , p) : fiber f b) → g a =ₚᵁ u + ≃ (b : B) → Σ (u : U b), ((a , p) : fiber' f b) → u =ₚᵁ g a + ≃ (b : B) → Σ (u : U b), (a : A) (p : f a = b) → g a =ₚᵁ u + ≃ (b : B) → Σ (u : U b), (a : A) (p : b = f a) → u =ₚᵁ g a +``` + ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (U : B → UU l3) + (g : (a : A) → U (f a)) where - family-fiber-Π-precomp-Π : ((a : A) → U (f a)) → B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp-Π g b = + family-fiber-Π-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp-Π b = Σ (U b) (λ u → ((a , p) : fiber f b) → dependent-identification U p (g a) u) - fiber-Π-precomp-Π : (g : (a : A) → U (f a)) → UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp-Π g = (b : B) → family-fiber-Π-precomp-Π g b + fiber-Π-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp-Π = (b : B) → family-fiber-Π-precomp-Π b - compute-fiber-Π-precomp-Π : - (g : (a : A) → U (f a)) → fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π g - compute-fiber-Π-precomp-Π g = + compute-fiber-Π-precomp-Π : fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π + compute-fiber-Π-precomp-Π = equivalence-reasoning fiber (precomp-Π f U) g ≃ Σ ((b : B) → U b) (λ h → (a : A) → g a = (h ∘ f) a) @@ -552,17 +565,30 @@ module _ ((a , p) : fiber f b) → dependent-identification U p (g a) u)) by inv-distributive-Π-Σ - family-fiber-Π-precomp-Π' : ((a : A) → U (f a)) → B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp-Π' g b = + family-fiber-Π-curry-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-curry-precomp-Π b = + Σ ( U b) + ( λ u → (a : A) (p : f a = b) → dependent-identification U p (g a) u) + + fiber-Π-curry-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-curry-precomp-Π = (b : B) → family-fiber-Π-curry-precomp-Π b + + compute-fiber-Π-curry-precomp-Π : + fiber (precomp-Π f U) g ≃ fiber-Π-curry-precomp-Π + compute-fiber-Π-curry-precomp-Π = + ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e + ( compute-fiber-Π-precomp-Π) + + family-fiber-Π-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp-Π' b = Σ ( U b) ( λ u → ((a , p) : fiber' f b) → dependent-identification U p u (g a)) - fiber-Π-precomp-Π' : (g : (a : A) → U (f a)) → UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp-Π' g = (b : B) → family-fiber-Π-precomp-Π' g b + fiber-Π-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp-Π' = (b : B) → family-fiber-Π-precomp-Π' b - compute-fiber-Π-precomp-Π' : - (g : (a : A) → U (f a)) → fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π' g - compute-fiber-Π-precomp-Π' g = + compute-fiber-Π-precomp-Π' : fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π' + compute-fiber-Π-precomp-Π' = equivalence-reasoning fiber (precomp-Π f U) g ≃ Σ ((b : B) → U b) (λ h → (a : A) → (h ∘ f) a = g a) @@ -580,4 +606,95 @@ module _ ( λ u → ((a , p) : fiber' f b) → dependent-identification U p u (g a))) by inv-distributive-Π-Σ + + family-fiber-Π-curry-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-curry-precomp-Π' b = + Σ (U b) (λ u → (a : A) (p : b = f a) → dependent-identification U p u (g a)) + + fiber-Π-curry-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-curry-precomp-Π' = (b : B) → family-fiber-Π-curry-precomp-Π' b + + compute-fiber-Π-curry-precomp-Π' : + fiber (precomp-Π f U) g ≃ fiber-Π-curry-precomp-Π' + compute-fiber-Π-curry-precomp-Π' = + ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e + ( compute-fiber-Π-precomp-Π') +``` + +### Computing the fibers of precomposition functions as dependent products + +We give four equivalences for the fibers of precomposition functions as +dependent products: + +```text + fiber (precomp f U) g + ≃ (b : B) → Σ (u : U), ((a , p) : fiber f b) → g a = u + ≃ (b : B) → Σ (u : U), ((a , p) : fiber' f b) → u = g a + ≃ (b : B) → Σ (u : U), (a : A) → f a = b → g a = u + ≃ (b : B) → Σ (u : U), (a : A) → b = f a → u = g a +``` + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (U : UU l3) + (g : A → U) + where + + family-fiber-Π-precomp : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp b = + Σ U (λ u → ((a , _) : fiber f b) → g a = u) + + fiber-Π-precomp : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp = (b : B) → family-fiber-Π-precomp b + + compute-fiber-Π-precomp : fiber (precomp f U) g ≃ fiber-Π-precomp + compute-fiber-Π-precomp = + equivalence-reasoning + fiber (precomp f U) g + ≃ Σ (B → U) (λ h → (a : A) → g a = (h ∘ f) a) + by compute-extension-fiber-precomp f U g + ≃ Σ ( B → U) + ( λ h → (b : B) ((a , _) : fiber f b) → g a = h b) + by + equiv-tot + ( λ h → + inv-equiv-dependent-universal-property-family-of-fibers f + ( λ y (a , _) → (g a = h y))) + ≃ ( (b : B) → Σ U (λ u → ((a , _) : fiber f b) → g a = u)) + by inv-distributive-Π-Σ + + compute-fiber-Π-curry-precomp : + fiber (precomp f U) g ≃ ((b : B) → Σ U (λ u → (a : A) → f a = b → g a = u)) + compute-fiber-Π-curry-precomp = + ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e + ( compute-fiber-Π-precomp) + + family-fiber-Π-precomp' : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-Π-precomp' b = + Σ U (λ u → ((a , _) : fiber' f b) → u = g a) + + fiber-Π-precomp' : UU (l1 ⊔ l2 ⊔ l3) + fiber-Π-precomp' = (b : B) → family-fiber-Π-precomp' b + + compute-fiber-Π-precomp' : fiber (precomp f U) g ≃ fiber-Π-precomp' + compute-fiber-Π-precomp' = + equivalence-reasoning + fiber (precomp f U) g + ≃ Σ (B → U) (λ h → (h ∘ f) ~ g) + by compute-extension-fiber-precomp' f U g + ≃ Σ ( B → U) + ( λ h → (b : B) ((a , _) : fiber' f b) → h b = g a) + by + equiv-tot + ( λ h → + inv-equiv-dependent-universal-property-family-of-fibers' f + ( λ y (a , _) → (h y = g a))) + ≃ ( (b : B) → Σ U (λ u → ((a , _) : fiber' f b) → u = g a)) + by inv-distributive-Π-Σ + + compute-fiber-Π-curry-precomp' : + fiber (precomp f U) g ≃ ((b : B) → Σ U (λ u → (a : A) → b = f a → u = g a)) + compute-fiber-Π-curry-precomp' = + ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e + ( compute-fiber-Π-precomp') ``` From a6a95fb33a4ce07f9c3ecd59d0cf61754180b2e9 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 00:13:35 +0200 Subject: [PATCH 35/69] =?UTF-8?q?`is-subuniverse-equiv-terminal-map-fibers?= =?UTF-8?q?-is-inhabited-family-fiber-=CE=A0-precomp-=CE=A0`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../subuniverse-connected-maps.lagda.md | 168 ++++++++++++++++-- .../subuniverse-equivalences.lagda.md | 36 +++- 2 files changed, 191 insertions(+), 13 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index daec6996df..0c75c2c7f5 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -259,8 +259,7 @@ is inhabited. ```agda module _ - {l1 l2 l3 : Level} - {A : UU l1} {B : UU l2} (f : A → B) + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) where is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' : @@ -307,6 +306,22 @@ module _ ( precomp-Π f (type-subuniverse-localization K ∘ Kfib)) ( s (type-subuniverse-subuniverse-localization K ∘ Kfib)) ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) + + is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π : + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → + ((U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) → + (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) + is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π K Kfib s b = + is-subuniverse-equiv-comp K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-is-equiv K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( is-equiv-terminal-map-is-contr + ( is-contr-subuniverse-localization-fiber-has-section-precomp-Π + K Kfib s b))) ``` #### Surjection condition @@ -314,12 +329,21 @@ module _ A map is `K`-connected if and only if its dependent precomposition maps are surjective and the fibers have `K`-localizations. -In fact, it suffices that a certain family is inhabited. +In fact, it suffices that the family + +```text + b ↦ + Σ ( u : K(fiber f b)), + ( ((a , p) : fiber f b) → + dependent-identification (b ↦ K(fiber f b)) p u (η (f a) (a , refl))) +``` + +is inhabited, which is a slightly weaker condition than inhabitedness of the +fiber of `precomp-Π f` over the map `a ↦ η (f a) (a , refl)`. ```agda module _ - {l1 l2 l3 : Level} - {A : UU l1} {B : UU l2} (f : A → B) + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) where abstract @@ -334,13 +358,35 @@ module _ ((b : B) → is-contr (Kfib b)) is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' Kfib η is-htpy-injective-precomp-η-Kfib Fη b = - rec-trunc-Prop - ( is-contr-Prop (Kfib b)) - ( λ (sb , Hb) → - ( sb , - is-htpy-injective-precomp-η-Kfib b - ( λ where (a , refl) → Hb (a , refl)))) - ( Fη b) + rec-trunc-Prop + ( is-contr-Prop (Kfib b)) + ( λ (sb , Hb) → + ( sb , + is-htpy-injective-precomp-η-Kfib b + ( λ where (a , refl) → Hb (a , refl)))) + ( Fη b) + + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π : + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → + ( (b : B) → + is-inhabited + ( family-fiber-Π-precomp-Π' f + ( type-subuniverse-localization K ∘ Kfib) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) + ( b))) → + ((b : B) → is-contr (type-subuniverse-localization K (Kfib b))) + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π + K Kfib = + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' + ( type-subuniverse-localization K ∘ Kfib) + ( unit-subuniverse-localization K ∘ Kfib) + ( λ b H → + htpy-eq + ( is-injective-is-equiv + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) + ( type-subuniverse-subuniverse-localization K (Kfib b))) + ( eq-htpy H))) is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π : (Kfib : (b : B) → UU l3) → @@ -375,6 +421,28 @@ module _ ( eq-htpy H))) ( s ( type-subuniverse-subuniverse-localization K ∘ Kfib) ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) + + is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π : + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → + ( (b : B) → + is-inhabited + ( family-fiber-Π-precomp-Π' f + ( type-subuniverse-localization K ∘ Kfib) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) + ( b))) → + (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) + is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π + K Kfib s b = + is-subuniverse-equiv-comp K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-is-equiv K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( is-equiv-terminal-map-is-contr + ( is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π + K Kfib s b))) ``` ### Characterizing equality of `K`-connected maps @@ -734,3 +802,79 @@ module _ ( is-subuniverse-connected-map-subuniverse-connected-map K f ( λ y → ((P y → pr1 U) , H (P y) U))) ``` + +### There is a `K`-equivalence between the fiber of a map and the fiber of its `ΣK`-localization + +This is generalization of Corollary 2.29 in {{#cite CORS20}}. + +Assume given arbitrary `f : A → B` and `Σ(δ)K`-equivalences `A → A'` and +`B → B'`. Then there is + +```text + fiber f b = Σ A (λ a → f a = b) + → Σ A (λ a → ║f a = b║) + ≃ Σ A (λ a → |f a| = |b|) + ≃ Σ A (λ a → ║f║ |a| = |b|) + → Σ ║A║ (λ t → ║f║ t = |b|) + = fiber ║f║ |b| +``` + +where the first and last maps are `K`-equivalences. + +```text +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) (b : B) + where + + fiber-map-trunc-fiber : + fiber f b → fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b) + fiber-map-trunc-fiber = + ( map-Σ-map-base-subuniverse-equiv + ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) ∘ + ( tot + ( λ a → + ( concat (naturality-unit-trunc (succ-𝕋 K) f a) (unit-trunc b)) ∘ + ( map-effectiveness-trunc K (f a) b) ∘ + ( unit-trunc))) + + abstract + is-subuniverse-equiv-fiber-map-trunc-fiber : + is-subuniverse-equiv K fiber-map-trunc-fiber + is-subuniverse-equiv-fiber-map-trunc-fiber = + is-subuniverse-equiv-comp + ( map-Σ-map-base-subuniverse-equiv + ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) + ( tot + ( λ a → + ( concat (naturality-unit-trunc (succ-𝕋 K) f a) (unit-trunc b)) ∘ + ( map-effectiveness-trunc K (f a) b) ∘ + ( unit-trunc))) + ( is-subuniverse-equiv-is-subuniverse-equiv-equiv + ( equiv-tot + ( λ a → + ( equiv-concat + ( naturality-unit-trunc (succ-𝕋 K) f a) + ( unit-trunc b)) ∘e + ( effectiveness-trunc K (f a) b))) + ( λ (a , p) → a , unit-trunc p) + ( is-equiv-map-equiv (equiv-trunc-Σ K))) + ( is-subuniverse-equiv-map-Σ-map-base-unit-trunc + ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) + + subuniverse-equiv-fiber-map-trunc-fiber : + subuniverse-equiv K + ( fiber f b) + ( fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b)) + pr1 subuniverse-equiv-fiber-map-trunc-fiber = + fiber-map-trunc-fiber + pr2 subuniverse-equiv-fiber-map-trunc-fiber = + is-subuniverse-equiv-fiber-map-trunc-fiber + + equiv-trunc-fiber-map-trunc-fiber : + type-trunc K (fiber f b) ≃ + type-trunc K (fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b)) + equiv-trunc-fiber-map-trunc-fiber = + equiv-trunc-subuniverse-equiv K + ( subuniverse-equiv-fiber-map-trunc-fiber) +``` diff --git a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md index dcb47157bb..ace3ce7f73 100644 --- a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md @@ -159,7 +159,7 @@ module _ ( is-subuniverse-equiv-subuniverse-equiv K g) ``` -### The class of `K`-equivalences has the 3-for-2 property +### The class of `K`-equivalences has the left and right cancellation property ```agda module _ @@ -179,6 +179,40 @@ module _ is-equiv-left-factor (precomp f (pr1 U)) (precomp g (pr1 U)) (GF U) (G U) ``` +### The class of `K`-equivalences satisfies the 3-for-2 property + +```agda +module _ + {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) + {A : UU l1} {B : UU l2} {X : UU l3} + (f : A → X) (g : B → X) (h : A → B) (p : f ~ g ∘ h) + where + + is-subuniverse-equiv-top-map-triangle : + is-subuniverse-equiv K g → + is-subuniverse-equiv K f → + is-subuniverse-equiv K h + is-subuniverse-equiv-top-map-triangle G F = + is-subuniverse-equiv-right-factor K g h + ( is-subuniverse-equiv-htpy' K p F) + ( G) + + is-subuniverse-equiv-right-map-triangle : + is-subuniverse-equiv K f → + is-subuniverse-equiv K h → + is-subuniverse-equiv K g + is-subuniverse-equiv-right-map-triangle F = + is-subuniverse-equiv-left-factor K g h (is-subuniverse-equiv-htpy' K p F) + + + is-subuniverse-equiv-left-map-triangle : + is-subuniverse-equiv K h → + is-subuniverse-equiv K g → + is-subuniverse-equiv K f + is-subuniverse-equiv-left-map-triangle H G = + is-subuniverse-equiv-htpy K p (is-subuniverse-equiv-comp K g h H G) +``` + ### Sections of `K`-equivalences are `K`-equivalences ```agda From 2e104ccb0209961b15c980fbbfd77522c24c0712 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 00:30:30 +0200 Subject: [PATCH 36/69] round off conditions subuniverse connected maps --- .../subuniverse-connected-maps.lagda.md | 132 +++++++++--------- 1 file changed, 67 insertions(+), 65 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 0c75c2c7f5..85a82b7fb6 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -249,51 +249,41 @@ module _ A map is `K`-connected if and only if its dependent precomposition maps admit sections and the fibers have `K`-localizations. -In fact, the following conditions suffice: - -There is a family `Kfib` together with a map `η : (b : B) → fiber f b → Kfib b` -such that the precomposition map of `η b` at `Kfib b` is injective, so a weak -kind of `Kfib b`-[epimorphism](foundation.epimorphisms.md), and the fiber of the -dependent precomposition map of `f` along `Kfib` over `λ a → η (f a) (a , refl)` -is inhabited. - ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) + (Kfib : B → UU l3) (η : (b : B) → fiber f b → Kfib b) + ( is-htpy-injective-precomp-η-Kfib : + (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) where is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' : - (Kfib : B → UU l3) → - (η : (b : B) → fiber f b → Kfib b) → - ( (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → ( fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl))) → ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' - Kfib η is-htpy-injective-precomp-η-Kfib Fη b = + is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' Fη b = ( pr1 (Fη b) , is-htpy-injective-precomp-η-Kfib b ( λ where (a , refl) → pr2 (Fη b) (a , refl))) is-contr-subuniverse-localization-fiber-has-section-precomp-Π' : - (Kfib : B → UU l3) → - (η : (b : B) → fiber f b → Kfib b) → - ( (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl)) → ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' - Kfib η is-htpy-injective-precomp-η-Kfib (s , H) b = + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' (s , H) b = ( s b , is-htpy-injective-precomp-η-Kfib b (λ where (a , refl) → htpy-eq H a)) +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} (f : A → B) + {l4 : Level} (K : subuniverse l3 l4) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) + (s : (U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) + where + is-contr-subuniverse-localization-fiber-has-section-precomp-Π : - {l4 : Level} (K : subuniverse l3 l4) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → - ((U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) → ((b : B) → is-contr (pr1 (Kfib b))) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π K Kfib s = - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' + is-contr-subuniverse-localization-fiber-has-section-precomp-Π = + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' f ( type-subuniverse-localization K ∘ Kfib) ( unit-subuniverse-localization K ∘ Kfib) ( λ b H → @@ -308,11 +298,8 @@ module _ ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π : - {l4 : Level} (K : subuniverse l3 l4) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → - ((U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) → (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) - is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π K Kfib s b = + is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π b = is-subuniverse-equiv-comp K ( terminal-map (type-subuniverse-localization K (Kfib b))) ( unit-subuniverse-localization K (Kfib b)) @@ -320,8 +307,13 @@ module _ ( is-subuniverse-equiv-is-equiv K ( terminal-map (type-subuniverse-localization K (Kfib b))) ( is-equiv-terminal-map-is-contr - ( is-contr-subuniverse-localization-fiber-has-section-precomp-Π - K Kfib s b))) + ( is-contr-subuniverse-localization-fiber-has-section-precomp-Π b))) + + is-subuniverse-connected-map-has-section-precomp-Π : + is-subuniverse-connected-map K f + is-subuniverse-connected-map-has-section-precomp-Π = + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f + ( is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π) ``` #### Surjection condition @@ -329,7 +321,7 @@ module _ A map is `K`-connected if and only if its dependent precomposition maps are surjective and the fibers have `K`-localizations. -In fact, it suffices that the family +In fact, it suffices that the fibers have `K`-localizations and the family ```text b ↦ @@ -344,20 +336,21 @@ fiber of `precomp-Π f` over the map `a ↦ η (f a) (a , refl)`. ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) + (Kfib : B → UU l3) + (η : (b : B) → fiber f b → Kfib b) + ( is-htpy-injective-precomp-η-Kfib : + (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) where abstract is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' : - (Kfib : B → UU l3) → - (η : (b : B) → fiber f b → Kfib b) → - ( (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → ( (b : B) → is-inhabited ( family-fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl)) b)) → ((b : B) → is-contr (Kfib b)) is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' - Kfib η is-htpy-injective-precomp-η-Kfib Fη b = + Fη b = rec-trunc-Prop ( is-contr-Prop (Kfib b)) ( λ (sb , Hb) → @@ -366,9 +359,25 @@ module _ ( λ where (a , refl) → Hb (a , refl)))) ( Fη b) + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π : + is-inhabited (fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl))) → + ((b : B) → is-contr (Kfib b)) + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π + s b = + rec-trunc-Prop + ( is-contr-Prop (Kfib b)) + ( λ s → + is-contr-subuniverse-localization-fiber-has-section-precomp-Π' + f Kfib η is-htpy-injective-precomp-η-Kfib s b) + ( s) + +module _ + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) + (Kfib : (b : B) → subuniverse-localization K (fiber f b)) + where + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π : - {l4 : Level} (K : subuniverse l3 l4) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → ( (b : B) → is-inhabited ( family-fiber-Π-precomp-Π' f @@ -376,9 +385,9 @@ module _ ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) ( b))) → ((b : B) → is-contr (type-subuniverse-localization K (Kfib b))) - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π - K Kfib = + is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π = is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' + ( f) ( type-subuniverse-localization K ∘ Kfib) ( unit-subuniverse-localization K ∘ Kfib) ( λ b H → @@ -388,29 +397,11 @@ module _ ( type-subuniverse-subuniverse-localization K (Kfib b))) ( eq-htpy H))) - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π : - (Kfib : (b : B) → UU l3) → - (η : (b : B) → fiber f b → Kfib b) → - ( (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) → - is-inhabited (fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl))) → - ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π - Kfib η is-htpy-injective-precomp-η-Kfib s b = - rec-trunc-Prop - ( is-contr-Prop (Kfib b)) - ( λ s → - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' - f Kfib η is-htpy-injective-precomp-η-Kfib s b) - ( s) - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π : - {l4 : Level} (K : subuniverse l3 l4) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → ((b : B) → is-contr (pr1 (Kfib b))) - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π K Kfib s = - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π + is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π s = + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π f ( type-subuniverse-localization K ∘ Kfib) ( unit-subuniverse-localization K ∘ Kfib) ( λ b H → @@ -423,8 +414,6 @@ module _ ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π : - {l4 : Level} (K : subuniverse l3 l4) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) → ( (b : B) → is-inhabited ( family-fiber-Π-precomp-Π' f @@ -433,7 +422,7 @@ module _ ( b))) → (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π - K Kfib s b = + s b = is-subuniverse-equiv-comp K ( terminal-map (type-subuniverse-localization K (Kfib b))) ( unit-subuniverse-localization K (Kfib b)) @@ -442,7 +431,20 @@ module _ ( terminal-map (type-subuniverse-localization K (Kfib b))) ( is-equiv-terminal-map-is-contr ( is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π - K Kfib s b))) + s b))) + + is-subuniverse-connected-map-is-inhabited-family-fiber-Π-precomp-Π : + ( (b : B) → + is-inhabited + ( family-fiber-Π-precomp-Π' f + ( type-subuniverse-localization K ∘ Kfib) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) + ( b))) → + is-subuniverse-connected-map K f + is-subuniverse-connected-map-is-inhabited-family-fiber-Π-precomp-Π Fη = + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f + ( is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π + ( Fη)) ``` ### Characterizing equality of `K`-connected maps From 225c8c06a91a0c9a6a791a4ef8df88c7a7d8a667 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 00:34:05 +0200 Subject: [PATCH 37/69] pre-commit --- .../subuniverse-connected-maps.lagda.md | 2 +- .../subuniverse-equivalences.lagda.md | 1 - 2 files changed, 1 insertion(+), 2 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 85a82b7fb6..92afbeaae1 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -419,7 +419,7 @@ module _ ( family-fiber-Π-precomp-Π' f ( type-subuniverse-localization K ∘ Kfib) ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) - ( b))) → + ( b))) → (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π s b = diff --git a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md index ace3ce7f73..3ac2ccf951 100644 --- a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md @@ -204,7 +204,6 @@ module _ is-subuniverse-equiv-right-map-triangle F = is-subuniverse-equiv-left-factor K g h (is-subuniverse-equiv-htpy' K p F) - is-subuniverse-equiv-left-map-triangle : is-subuniverse-equiv K h → is-subuniverse-equiv K g → From f6b197ee76be2076649666216a3e2725de2852e1 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 00:44:48 +0200 Subject: [PATCH 38/69] forgot one --- .../subuniverse-connected-maps.lagda.md | 65 +++++++++++++------ 1 file changed, 44 insertions(+), 21 deletions(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 92afbeaae1..ceb6f5d4b9 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -252,10 +252,10 @@ sections and the fibers have `K`-localizations. ```agda module _ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) - (Kfib : B → UU l3) (η : (b : B) → fiber f b → Kfib b) - ( is-htpy-injective-precomp-η-Kfib : - (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) + (Kfib : B → UU l3) (η : (b : B) → fiber f b → Kfib b) + ( is-htpy-injective-precomp-η-Kfib : + (b : B) {g h : Kfib b → Kfib b} → + precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) where is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' : @@ -274,8 +274,8 @@ module _ is-htpy-injective-precomp-η-Kfib b (λ where (a , refl) → htpy-eq H a)) module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} (f : A → B) - {l4 : Level} (K : subuniverse l3 l4) + {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) + {A : UU l1} {B : UU l2} (f : A → B) (Kfib : (b : B) → subuniverse-localization K (fiber f b)) (s : (U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) where @@ -397,21 +397,6 @@ module _ ( type-subuniverse-subuniverse-localization K (Kfib b))) ( eq-htpy H))) - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π : - ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → - ((b : B) → is-contr (pr1 (Kfib b))) - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π s = - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π f - ( type-subuniverse-localization K ∘ Kfib) - ( unit-subuniverse-localization K ∘ Kfib) - ( λ b H → - htpy-eq - ( is-injective-is-equiv - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) - ( type-subuniverse-subuniverse-localization K (Kfib b))) - ( eq-htpy H))) - ( s ( type-subuniverse-subuniverse-localization K ∘ Kfib) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π : ( (b : B) → @@ -445,6 +430,44 @@ module _ is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f ( is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π ( Fη)) + + is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π : + ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → + ((b : B) → is-contr (pr1 (Kfib b))) + is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π H = + is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π f + ( type-subuniverse-localization K ∘ Kfib) + ( unit-subuniverse-localization K ∘ Kfib) + ( λ b H → + htpy-eq + ( is-injective-is-equiv + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) + ( type-subuniverse-subuniverse-localization K (Kfib b))) + ( eq-htpy H))) + ( H ( type-subuniverse-subuniverse-localization K ∘ Kfib) + ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) + + + is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π : + ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → + (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) + is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π H b = + is-subuniverse-equiv-comp K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) + ( is-subuniverse-equiv-is-equiv K + ( terminal-map (type-subuniverse-localization K (Kfib b))) + ( is-equiv-terminal-map-is-contr + ( is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π + H b))) + + is-subuniverse-connected-map-is-surjective-precomp-Π : + ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → + is-subuniverse-connected-map K f + is-subuniverse-connected-map-is-surjective-precomp-Π H = + is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f + ( is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π H) ``` ### Characterizing equality of `K`-connected maps From 5827064eeed9d65a130ecd2795b6f73b1351263c Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 14:53:12 +0200 Subject: [PATCH 39/69] fix a link --- .../relative-separations-types-subuniverses.lagda.md | 2 +- .../subuniverse-connected-maps.lagda.md | 2 -- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md index 43c124f49e..56d9cee15d 100644 --- a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md @@ -52,4 +52,4 @@ module _ ## See also -- [Map-separated types at global subuniverses](orthogonal-factorization-systems.rel-separated-types-global-subuniverses.md) +- [Relative separations of types at global subuniverses](orthogonal-factorization-systems.relative-separations-types-global-subuniverses.md) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index ceb6f5d4b9..93906ac736 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -397,7 +397,6 @@ module _ ( type-subuniverse-subuniverse-localization K (Kfib b))) ( eq-htpy H))) - is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π : ( (b : B) → is-inhabited @@ -447,7 +446,6 @@ module _ ( H ( type-subuniverse-subuniverse-localization K ∘ Kfib) ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) - is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π : ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) From d79da3ef3649feab2858f36307abb18f980dc7d4 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 14:53:29 +0200 Subject: [PATCH 40/69] modulation! --- src/foundation-core/subtypes.lagda.md | 2 +- src/foundation-core/truncated-maps.lagda.md | 4 +- ...oriality-dependent-function-types.lagda.md | 220 +++++++++--------- .../reflective-subuniverses.lagda.md | 2 +- 4 files changed, 111 insertions(+), 117 deletions(-) diff --git a/src/foundation-core/subtypes.lagda.md b/src/foundation-core/subtypes.lagda.md index 84e8af493e..25af195532 100644 --- a/src/foundation-core/subtypes.lagda.md +++ b/src/foundation-core/subtypes.lagda.md @@ -226,7 +226,7 @@ module _ pr1 (trunc-map-into-subtype f p) x = (map-trunc-map f x , p x) pr2 (trunc-map-into-subtype f p) = is-trunc-map-into-subtype - ( is-trunc-trunc-map f) + ( is-trunc-map-trunc-map f) ( p) ``` diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index c4a44f53bd..47395ec194 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -53,9 +53,9 @@ module _ map-trunc-map = pr1 abstract - is-trunc-trunc-map : + is-trunc-map-trunc-map : (f : trunc-map k A B) → is-trunc-map k (map-trunc-map f) - is-trunc-trunc-map = pr2 + is-trunc-map-trunc-map = pr2 ``` ## Properties diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md index 73d044b7bd..0df22316f6 100644 --- a/src/foundation/functoriality-dependent-function-types.lagda.md +++ b/src/foundation/functoriality-dependent-function-types.lagda.md @@ -49,9 +49,9 @@ first argument, and covariantly in its second argument. ```agda module _ - { l1 l2 l3 l4 : Level} - { A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) - ( e : A' ≃ A) (f : (a' : A') → B' a' ≃ B (map-equiv e a')) + {l1 l2 l3 l4 : Level} + {A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) + (e : A' ≃ A) (f : (a' : A') → B' a' ≃ B (map-equiv e a')) where map-equiv-Π : ((a' : A') → B' a') → ((a : A) → B a) @@ -112,8 +112,8 @@ module _ α x refl = refl id-map-equiv-Π : - { l1 l2 : Level} {A : UU l1} (B : A → UU l2) → - ( map-equiv-Π B (id-equiv {A = A}) (λ a → id-equiv {A = B a})) ~ id + {l1 l2 : Level} {A : UU l1} (B : A → UU l2) → + (map-equiv-Π B (id-equiv {A = A}) (λ a → id-equiv {A = B a})) ~ id id-map-equiv-Π B h = eq-htpy (compute-map-equiv-Π B id-equiv (λ _ → id-equiv) h) ``` @@ -164,118 +164,112 @@ module _ ### A family of truncated maps over any map induces a truncated map on dependent function types ```agda -is-trunc-map-Π-is-trunc-map' : - (k : 𝕋) {l1 l2 l3 l4 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} - {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → - ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π' α f) -is-trunc-map-Π-is-trunc-map' k {J = J} α f H h = - is-trunc-equiv' k - ( (j : J) → fiber (f (α j)) (h j)) - ( compute-fiber-map-Π' α f h) - ( is-trunc-Π k (λ j → H (α j) (h j))) - -is-trunc-map-is-trunc-map-Π' : - (k : 𝕋) {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} - (f : (i : I) → A i → B i) → - ({l : Level} {J : UU l} (α : J → I) → is-trunc-map k (map-Π' α f)) → - (i : I) → is-trunc-map k (f i) -is-trunc-map-is-trunc-map-Π' k {A = A} {B} f H i b = - is-trunc-equiv' k - ( fiber (map-Π (λ _ → f i)) (point b)) - ( equiv-Σ - ( λ a → f i a = b) - ( equiv-universal-property-unit (A i)) - ( λ h → equiv-ap - ( equiv-universal-property-unit (B i)) - ( map-Π (λ _ → f i) h) - ( point b))) - ( H (λ _ → i) (point b)) - -is-emb-map-Π-is-emb' : - {l1 l2 l3 l4 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} → - {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → - ((i : I) → is-emb (f i)) → is-emb (map-Π' α f) -is-emb-map-Π-is-emb' α f H = - is-emb-is-prop-map - ( is-trunc-map-Π-is-trunc-map' neg-one-𝕋 α f - ( λ i → is-prop-map-is-emb (H i))) +module _ + {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} + where + + is-trunc-map-Π-is-trunc-map' : + (k : 𝕋) {l4 : Level} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → + ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π' α f) + is-trunc-map-Π-is-trunc-map' k {J = J} α f H h = + is-trunc-equiv' k + ( (j : J) → fiber (f (α j)) (h j)) + ( compute-fiber-map-Π' α f h) + ( is-trunc-Π k (λ j → H (α j) (h j))) + + is-trunc-map-is-trunc-map-Π' : + (k : 𝕋) (f : (i : I) → A i → B i) → + ({l : Level} {J : UU l} (α : J → I) → is-trunc-map k (map-Π' α f)) → + (i : I) → is-trunc-map k (f i) + is-trunc-map-is-trunc-map-Π' k f H i b = + is-trunc-equiv' k + ( fiber (map-Π (λ _ → f i)) (point b)) + ( equiv-Σ + ( λ a → f i a = b) + ( equiv-universal-property-unit (A i)) + ( λ h → equiv-ap + ( equiv-universal-property-unit (B i)) + ( map-Π (λ _ → f i) h) + ( point b))) + ( H (λ _ → i) (point b)) + + is-emb-map-Π-is-emb' : + {l4 : Level} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → + ((i : I) → is-emb (f i)) → is-emb (map-Π' α f) + is-emb-map-Π-is-emb' α f H = + is-emb-is-prop-map + ( is-trunc-map-Π-is-trunc-map' neg-one-𝕋 α f + ( λ i → is-prop-map-is-emb (H i))) ``` -### +### The action on homotopies of equivalences ```agda -HTPY-map-equiv-Π : - { l1 l2 l3 l4 : Level} - { A' : UU l1} (B' : A' → UU l2) {A : UU l3} (B : A → UU l4) - ( e e' : A' ≃ A) (H : htpy-equiv e e') → - UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) -HTPY-map-equiv-Π {A' = A'} B' {A} B e e' H = - ( f : (a' : A') → B' a' ≃ B (map-equiv e a')) → - ( f' : (a' : A') → B' a' ≃ B (map-equiv e' a')) → - ( K : (a' : A') → - ((tr B (H a')) ∘ (map-equiv (f a'))) ~ (map-equiv (f' a'))) → - ( map-equiv-Π B e f) ~ (map-equiv-Π B e' f') - -htpy-map-equiv-Π-refl-htpy : - { l1 l2 l3 l4 : Level} - { A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) - ( e : A' ≃ A) → - HTPY-map-equiv-Π B' B e e (refl-htpy-equiv e) -htpy-map-equiv-Π-refl-htpy {B' = B'} B e f f' K = - ( htpy-map-Π - ( λ a → - ( tr B (is-section-map-inv-is-equiv (is-equiv-map-equiv e) a)) ·l - ( K (map-inv-is-equiv (is-equiv-map-equiv e) a)))) ·r - ( precomp-Π (map-inv-is-equiv (is-equiv-map-equiv e)) B') - -abstract - htpy-map-equiv-Π : - { l1 l2 l3 l4 : Level} - { A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) - ( e e' : A' ≃ A) (H : htpy-equiv e e') → - HTPY-map-equiv-Π B' B e e' H - htpy-map-equiv-Π {B' = B'} B e e' H f f' K = - ind-htpy-equiv e - ( HTPY-map-equiv-Π B' B e) - ( htpy-map-equiv-Π-refl-htpy B e) - e' H f f' K - - compute-htpy-map-equiv-Π : - { l1 l2 l3 l4 : Level} - { A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) - ( e : A' ≃ A) → - ( htpy-map-equiv-Π {B' = B'} B e e (refl-htpy-equiv e)) = - ( ( htpy-map-equiv-Π-refl-htpy B e)) - compute-htpy-map-equiv-Π {B' = B'} B e = - compute-ind-htpy-equiv e - ( HTPY-map-equiv-Π B' B e) - ( htpy-map-equiv-Π-refl-htpy B e) - -map-automorphism-Π : - { l1 l2 : Level} {A : UU l1} {B : A → UU l2} - ( e : A ≃ A) (f : (a : A) → B a ≃ B (map-equiv e a)) → - ( (a : A) → B a) → ((a : A) → B a) -map-automorphism-Π {B = B} e f = - ( map-Π (λ a → (map-inv-is-equiv (is-equiv-map-equiv (f a))))) ∘ - ( precomp-Π (map-equiv e) B) - -abstract - is-equiv-map-automorphism-Π : - { l1 l2 : Level} {A : UU l1} {B : A → UU l2} - ( e : A ≃ A) (f : (a : A) → B a ≃ B (map-equiv e a)) → - is-equiv (map-automorphism-Π e f) - is-equiv-map-automorphism-Π {B = B} e f = - is-equiv-comp _ _ - ( is-equiv-precomp-Π-is-equiv (is-equiv-map-equiv e) B) - ( is-equiv-map-Π-is-fiberwise-equiv - ( λ a → is-equiv-map-inv-is-equiv (is-equiv-map-equiv (f a)))) - -automorphism-Π : - { l1 l2 : Level} {A : UU l1} {B : A → UU l2} - ( e : A ≃ A) (f : (a : A) → B a ≃ B (map-equiv e a)) → - ( (a : A) → B a) ≃ ((a : A) → B a) -pr1 (automorphism-Π e f) = map-automorphism-Π e f -pr2 (automorphism-Π e f) = is-equiv-map-automorphism-Π e f +module _ + {l1 l2 l3 l4 : Level} + {A' : UU l1} (B' : A' → UU l2) {A : UU l3} (B : A → UU l4) + where + + HTPY-map-equiv-Π : + (e e' : A' ≃ A) → htpy-equiv e e' → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) + HTPY-map-equiv-Π e e' H = + ( f : (a' : A') → B' a' ≃ B (map-equiv e a')) → + ( f' : (a' : A') → B' a' ≃ B (map-equiv e' a')) → + ( K : (a' : A') → + ((tr B (H a')) ∘ (map-equiv (f a'))) ~ (map-equiv (f' a'))) → + ( map-equiv-Π B e f) ~ (map-equiv-Π B e' f') + + htpy-map-equiv-Π-refl-htpy : + (e : A' ≃ A) → HTPY-map-equiv-Π e e (refl-htpy-equiv e) + htpy-map-equiv-Π-refl-htpy e f f' K = + ( htpy-map-Π + ( λ a → + ( tr B (is-section-map-inv-is-equiv (is-equiv-map-equiv e) a)) ·l + ( K (map-inv-is-equiv (is-equiv-map-equiv e) a)))) ·r + ( precomp-Π (map-inv-is-equiv (is-equiv-map-equiv e)) B') + + abstract + htpy-map-equiv-Π : + (e e' : A' ≃ A) (H : htpy-equiv e e') → HTPY-map-equiv-Π e e' H + htpy-map-equiv-Π e e' H f f' K = + ind-htpy-equiv e + ( HTPY-map-equiv-Π e) + ( htpy-map-equiv-Π-refl-htpy e) + e' H f f' K + + compute-htpy-map-equiv-Π : + ( e : A' ≃ A) → + htpy-map-equiv-Π e e (refl-htpy-equiv e) = htpy-map-equiv-Π-refl-htpy e + compute-htpy-map-equiv-Π e = + compute-ind-htpy-equiv e + ( HTPY-map-equiv-Π e) + ( htpy-map-equiv-Π-refl-htpy e) +``` + +### The action on automorphisms + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : A → UU l2} + (e : A ≃ A) (f : (a : A) → B a ≃ B (map-equiv e a)) + where + + map-automorphism-Π : ( (a : A) → B a) → ((a : A) → B a) + map-automorphism-Π = + ( map-Π (λ a → (map-inv-is-equiv (is-equiv-map-equiv (f a))))) ∘ + ( precomp-Π (map-equiv e) B) + + abstract + is-equiv-map-automorphism-Π : + is-equiv map-automorphism-Π + is-equiv-map-automorphism-Π = + is-equiv-comp _ _ + ( is-equiv-precomp-Π-is-equiv (is-equiv-map-equiv e) B) + ( is-equiv-map-Π-is-fiberwise-equiv + ( λ a → is-equiv-map-inv-is-equiv (is-equiv-map-equiv (f a)))) + + automorphism-Π : ((a : A) → B a) ≃ ((a : A) → B a) + automorphism-Π = (map-automorphism-Π , is-equiv-map-automorphism-Π) ``` ### Families of retracts induce retracts of dependent function types diff --git a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md index 1cd22d5acc..26edd4377a 100644 --- a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md @@ -167,7 +167,7 @@ module _ pr1 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) = unit-is-reflective-subuniverse 𝒫 is-reflective-𝒫 pr2 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) - X is-in-subuniverse-X = + ( X , is-in-subuniverse-X) = is-local-is-in-subuniverse-is-reflective-subuniverse 𝒫 is-reflective-𝒫 X A is-in-subuniverse-X From 8cfb546cff428d0a66b5816b1f4ff37ee33a8854 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 14:55:31 +0200 Subject: [PATCH 41/69] modulation! --- ...ctoriality-dependent-function-types.lagda.md | 17 ++++++++++++----- 1 file changed, 12 insertions(+), 5 deletions(-) diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md index 0df22316f6..8c0773ef43 100644 --- a/src/foundation/functoriality-dependent-function-types.lagda.md +++ b/src/foundation/functoriality-dependent-function-types.lagda.md @@ -92,9 +92,15 @@ module _ #### Computing `map-equiv-Π` ```agda +module _ + {l1 l2 l3 l4 : Level} + {A' : UU l1} {B' : A' → UU l2} {A : UU l3} (B : A → UU l4) + (e : A' ≃ A) (f : (a' : A') → B' a' ≃ B (map-equiv e a')) + where + compute-map-equiv-Π : (h : (a' : A') → B' a') (a' : A') → - map-equiv-Π h (map-equiv e a') = map-equiv (f a') (h a') + map-equiv-Π B e f h (map-equiv e a') = map-equiv (f a') (h a') compute-map-equiv-Π h a' = ( ap ( λ t → @@ -187,10 +193,11 @@ module _ ( equiv-Σ ( λ a → f i a = b) ( equiv-universal-property-unit (A i)) - ( λ h → equiv-ap - ( equiv-universal-property-unit (B i)) - ( map-Π (λ _ → f i) h) - ( point b))) + ( λ h → + equiv-ap + ( equiv-universal-property-unit (B i)) + ( map-Π (λ _ → f i) h) + ( point b))) ( H (λ _ → i) (point b)) is-emb-map-Π-is-emb' : From 479b9e4f9d774e1de4de7a42a2dee05ed7d7f6c8 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 3 Oct 2025 15:03:12 +0200 Subject: [PATCH 42/69] more edits --- ...oriality-dependent-function-types.lagda.md | 42 ++++++++----------- .../functoriality-function-types.lagda.md | 5 +-- .../reflective-subuniverses.lagda.md | 2 +- 3 files changed, 20 insertions(+), 29 deletions(-) diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md index 8c0773ef43..3ce0a35468 100644 --- a/src/foundation/functoriality-dependent-function-types.lagda.md +++ b/src/foundation/functoriality-dependent-function-types.lagda.md @@ -85,8 +85,7 @@ module _ ( is-section-map-inv-is-equiv (is-equiv-map-equiv e) a)))) equiv-Π : ((a' : A') → B' a') ≃ ((a : A) → B a) - pr1 equiv-Π = map-equiv-Π - pr2 equiv-Π = is-equiv-map-equiv-Π + equiv-Π = (map-equiv-Π , is-equiv-map-equiv-Π) ``` #### Computing `map-equiv-Π` @@ -119,7 +118,7 @@ module _ id-map-equiv-Π : {l1 l2 : Level} {A : UU l1} (B : A → UU l2) → - (map-equiv-Π B (id-equiv {A = A}) (λ a → id-equiv {A = B a})) ~ id + map-equiv-Π B (id-equiv {A = A}) (λ a → id-equiv {A = B a}) ~ id id-map-equiv-Π B h = eq-htpy (compute-map-equiv-Π B id-equiv (λ _ → id-equiv) h) ``` @@ -127,16 +126,14 @@ id-map-equiv-Π B h = eq-htpy (compute-map-equiv-Π B id-equiv (λ _ → id-equi ```agda module _ - {l1 l2 l3 : Level} {A : UU l1} + {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3} where equiv-htpy-map-Π-fam-equiv : - { B : A → UU l2} {C : A → UU l3} → - ( e : fam-equiv B C) (f g : (a : A) → B a) → - ( f ~ g) ≃ (map-Π (map-fam-equiv e) f ~ map-Π (map-fam-equiv e) g) + (e : fam-equiv B C) (f g : (a : A) → B a) → + (f ~ g) ≃ (map-Π (map-fam-equiv e) f ~ map-Π (map-fam-equiv e) g) equiv-htpy-map-Π-fam-equiv e f g = - equiv-Π-equiv-family - ( λ a → equiv-ap (e a) (f a) (g a)) + equiv-Π-equiv-family (λ a → equiv-ap (e a) (f a) (g a)) ``` ### Families of truncated maps induce truncated maps on dependent function types @@ -174,10 +171,10 @@ module _ {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} where - is-trunc-map-Π-is-trunc-map' : + is-trunc-map-Π' : (k : 𝕋) {l4 : Level} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π' α f) - is-trunc-map-Π-is-trunc-map' k {J = J} α f H h = + is-trunc-map-Π' k {J = J} α f H h = is-trunc-equiv' k ( (j : J) → fiber (f (α j)) (h j)) ( compute-fiber-map-Π' α f h) @@ -200,13 +197,12 @@ module _ ( point b))) ( H (λ _ → i) (point b)) - is-emb-map-Π-is-emb' : + is-emb-map-Π' : {l4 : Level} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → ((i : I) → is-emb (f i)) → is-emb (map-Π' α f) - is-emb-map-Π-is-emb' α f H = + is-emb-map-Π' α f H = is-emb-is-prop-map - ( is-trunc-map-Π-is-trunc-map' neg-one-𝕋 α f - ( λ i → is-prop-map-is-emb (H i))) + ( is-trunc-map-Π' neg-one-𝕋 α f (λ i → is-prop-map-is-emb (H i))) ``` ### The action on homotopies of equivalences @@ -220,11 +216,10 @@ module _ HTPY-map-equiv-Π : (e e' : A' ≃ A) → htpy-equiv e e' → UU (l1 ⊔ l2 ⊔ l3 ⊔ l4) HTPY-map-equiv-Π e e' H = - ( f : (a' : A') → B' a' ≃ B (map-equiv e a')) → - ( f' : (a' : A') → B' a' ≃ B (map-equiv e' a')) → - ( K : (a' : A') → - ((tr B (H a')) ∘ (map-equiv (f a'))) ~ (map-equiv (f' a'))) → - ( map-equiv-Π B e f) ~ (map-equiv-Π B e' f') + (f : (a' : A') → B' a' ≃ B (map-equiv e a')) → + (f' : (a' : A') → B' a' ≃ B (map-equiv e' a')) → + (K : (a' : A') → tr B (H a') ∘ map-equiv (f a') ~ map-equiv (f' a')) → + map-equiv-Π B e f ~ map-equiv-Π B e' f' htpy-map-equiv-Π-refl-htpy : (e : A' ≃ A) → HTPY-map-equiv-Π e e (refl-htpy-equiv e) @@ -238,14 +233,13 @@ module _ abstract htpy-map-equiv-Π : (e e' : A' ≃ A) (H : htpy-equiv e e') → HTPY-map-equiv-Π e e' H - htpy-map-equiv-Π e e' H f f' K = + htpy-map-equiv-Π e = ind-htpy-equiv e ( HTPY-map-equiv-Π e) ( htpy-map-equiv-Π-refl-htpy e) - e' H f f' K compute-htpy-map-equiv-Π : - ( e : A' ≃ A) → + (e : A' ≃ A) → htpy-map-equiv-Π e e (refl-htpy-equiv e) = htpy-map-equiv-Π-refl-htpy e compute-htpy-map-equiv-Π e = compute-ind-htpy-equiv e @@ -261,7 +255,7 @@ module _ (e : A ≃ A) (f : (a : A) → B a ≃ B (map-equiv e a)) where - map-automorphism-Π : ( (a : A) → B a) → ((a : A) → B a) + map-automorphism-Π : ((a : A) → B a) → ((a : A) → B a) map-automorphism-Π = ( map-Π (λ a → (map-inv-is-equiv (is-equiv-map-equiv (f a))))) ∘ ( precomp-Π (map-equiv e) B) diff --git a/src/foundation/functoriality-function-types.lagda.md b/src/foundation/functoriality-function-types.lagda.md index f8d8ab6ca1..c7fa12cb18 100644 --- a/src/foundation/functoriality-function-types.lagda.md +++ b/src/foundation/functoriality-function-types.lagda.md @@ -71,10 +71,7 @@ module _ is-trunc-map k f → {l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f) is-trunc-map-postcomp-is-trunc-map is-trunc-f A = - is-trunc-map-Π-is-trunc-map' k - ( terminal-map A) - ( point f) - ( point is-trunc-f) + is-trunc-map-Π' k (terminal-map A) (point f) (point is-trunc-f) is-trunc-map-is-trunc-map-postcomp : ({l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f)) → diff --git a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md index 26edd4377a..f66155010c 100644 --- a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md @@ -167,7 +167,7 @@ module _ pr1 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) = unit-is-reflective-subuniverse 𝒫 is-reflective-𝒫 pr2 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) - ( X , is-in-subuniverse-X) = + ( X , is-in-subuniverse-X) = is-local-is-in-subuniverse-is-reflective-subuniverse 𝒫 is-reflective-𝒫 X A is-in-subuniverse-X From 018c2deb682ac8a8a65b21b7dc70b29705944703 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Wed, 22 Oct 2025 19:27:54 +0200 Subject: [PATCH 43/69] another pushout characterization of smash --- references.bib | 14 ++++++++++++++ .../smash-products-of-pointed-types.lagda.md | 15 +++++++++++++++ 2 files changed, 29 insertions(+) diff --git a/references.bib b/references.bib index da0e6f65ee..b9460ff1c0 100644 --- a/references.bib +++ b/references.bib @@ -602,6 +602,20 @@ @misc{Lavenir23 primaryclass = {math.AT}, } +@article{Ljungström24, + title = {Symmetric monoidal smash products in homotopy type theory}, + author = {Ljungström, Axel}, + year = 2024, + month = oct, + journal = {Mathematical Structures in Computer Science}, + publisher = {Cambridge University Press (CUP)}, + volume = 34, + number = 9, + pages = {985–1007}, + doi = {10.1017/s0960129524000318}, + issn = {1469-8072}, +} + @book{Lurie09, title = {Higher Topos Theory}, author = {Jacob Lurie}, diff --git a/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md b/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md index b23c3d17b8..8b9b565eac 100644 --- a/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md +++ b/src/synthetic-homotopy-theory/smash-products-of-pointed-types.lagda.md @@ -475,6 +475,21 @@ the canonical maps `X → X ⋊∗ Y` and `Y → X ⋉∗ Y`, i.e., we have push > This remains to be formalized. +### The smash product as a bipointed pushout + +Given two pointed types `X` and `Y`, the smash product is the pushout +{{#cite Ljungström24}} + +```text + X + Y ------> X × Y + | | + | | + ∨ ⌜ ∨ + 1 + 1 ------> X ∧ Y. +``` + +> This remains to be formalized. + ## References {{#bibliography}} From 62797d1047d763b17fe5024a891ad97003bad9cb Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 23 Oct 2025 16:37:57 +0200 Subject: [PATCH 44/69] =?UTF-8?q?(=F0=9D=91=9B+1)-truncated=20pointed=20ma?= =?UTF-8?q?ps=20induce=20=F0=9D=91=9B-truncated=20maps=20on=20loop=20space?= =?UTF-8?q?s?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../functoriality-loop-spaces.lagda.md | 26 +++++++++++++++++++ 1 file changed, 26 insertions(+) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index d7cc51c5be..4b4254086e 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -14,6 +14,8 @@ open import foundation.equivalences open import foundation.faithful-maps open import foundation.function-types open import foundation.homotopies +open import foundation.truncation-levels +open import foundation.truncated-maps open import foundation.identity-types open import foundation.universe-levels @@ -77,6 +79,30 @@ module _ ( ap (map-pointed-map f) x)) ``` +### (𝑛+1)-truncated pointed maps induce 𝑛-truncated maps on loop spaces + +```agda +module _ + {l1 l2 : Level} (k : 𝕋) {A : Pointed-Type l1} {B : Pointed-Type l2} + where + + is-trunc-map-Ω : + (f : A →∗ B) → + is-trunc-map (succ-𝕋 k) (map-pointed-map f) → + is-trunc-map k (map-Ω f) + is-trunc-map-Ω f H = + is-trunc-map-comp k + ( tr-type-Ω (preserves-point-pointed-map f)) + ( ap (map-pointed-map f)) + ( is-trunc-map-is-equiv k + ( is-equiv-tr-type-Ω (preserves-point-pointed-map f))) + ( is-trunc-map-ap-is-trunc-map-succ k + ( map-pointed-map f) + ( H) + ( point-Pointed-Type A) + ( point-Pointed-Type A)) +``` + ### Faithful pointed maps induce embeddings on loop spaces ```agda From c999911f4cbdfdb5e92acfbc45c13548b10d84fc Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 23 Oct 2025 16:46:13 +0200 Subject: [PATCH 45/69] edits loop spaces --- .../loop-spaces.lagda.md | 44 +++++++++---------- 1 file changed, 20 insertions(+), 24 deletions(-) diff --git a/src/synthetic-homotopy-theory/loop-spaces.lagda.md b/src/synthetic-homotopy-theory/loop-spaces.lagda.md index be414f65db..533970b711 100644 --- a/src/synthetic-homotopy-theory/loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/loop-spaces.lagda.md @@ -48,7 +48,7 @@ module _ refl-Ω = refl Ω : Pointed-Type l - Ω = pair type-Ω refl-Ω + Ω = (type-Ω , refl-Ω) ``` ### The magma of loops on a pointed space @@ -84,11 +84,12 @@ module _ coherence-unit-laws-mul-Ω = refl Ω-H-Space : H-Space l - pr1 Ω-H-Space = Ω A - pr1 (pr2 Ω-H-Space) = mul-Ω A - pr1 (pr2 (pr2 Ω-H-Space)) = left-unit-law-mul-Ω - pr1 (pr2 (pr2 (pr2 Ω-H-Space))) = right-unit-law-mul-Ω - pr2 (pr2 (pr2 (pr2 Ω-H-Space))) = coherence-unit-laws-mul-Ω + pr1 Ω-H-Space = + ( Ω A , + mul-Ω A , + left-unit-law-mul-Ω , + right-unit-law-mul-Ω , + coherence-unit-laws-mul-Ω) ``` ### The wild quasigroup of loops on a pointed space @@ -134,14 +135,14 @@ module _ {l1 : Level} {A : UU l1} {x y : A} where - equiv-tr-Ω : x = y → Ω (pair A x) ≃∗ Ω (pair A y) - equiv-tr-Ω refl = pair id-equiv refl + equiv-tr-Ω : x = y → Ω (A , x) ≃∗ Ω (A , y) + equiv-tr-Ω refl = (id-equiv , refl) - equiv-tr-type-Ω : x = y → type-Ω (pair A x) ≃ type-Ω (pair A y) + equiv-tr-type-Ω : x = y → type-Ω (A , x) ≃ type-Ω (A , y) equiv-tr-type-Ω p = equiv-pointed-equiv (equiv-tr-Ω p) - tr-type-Ω : x = y → type-Ω (pair A x) → type-Ω (pair A y) + tr-type-Ω : x = y → type-Ω (A , x) → type-Ω (A , y) tr-type-Ω p = map-equiv (equiv-tr-type-Ω p) is-equiv-tr-type-Ω : (p : x = y) → is-equiv (tr-type-Ω p) @@ -151,21 +152,18 @@ module _ preserves-refl-tr-Ω refl = refl preserves-mul-tr-Ω : - (p : x = y) (u v : type-Ω (pair A x)) → - Id - ( tr-type-Ω p (mul-Ω (pair A x) u v)) - ( mul-Ω (pair A y) (tr-type-Ω p u) (tr-type-Ω p v)) + (p : x = y) (u v : type-Ω (A , x)) → + tr-type-Ω p (mul-Ω (A , x) u v) = + mul-Ω (A , y) (tr-type-Ω p u) (tr-type-Ω p v) preserves-mul-tr-Ω refl u v = refl preserves-inv-tr-Ω : - (p : x = y) (u : type-Ω (pair A x)) → - Id - ( tr-type-Ω p (inv-Ω (pair A x) u)) - ( inv-Ω (pair A y) (tr-type-Ω p u)) + (p : x = y) (u : type-Ω (A , x)) → + tr-type-Ω p (inv-Ω (A , x) u) = inv-Ω (A , y) (tr-type-Ω p u) preserves-inv-tr-Ω refl u = refl eq-tr-type-Ω : - (p : x = y) (q : type-Ω (pair A x)) → + (p : x = y) (q : type-Ω (A , x)) → tr-type-Ω p q = inv p ∙ (q ∙ p) eq-tr-type-Ω refl q = inv right-unit ``` @@ -181,9 +179,7 @@ module _ where pointed-equiv-loop-pointed-identity : - ( pair (point-Pointed-Type A = x) p) ≃∗ Ω A - pr1 pointed-equiv-loop-pointed-identity = - equiv-concat' (point-Pointed-Type A) (inv p) - pr2 pointed-equiv-loop-pointed-identity = - right-inv p + ( (point-Pointed-Type A = x) , p) ≃∗ Ω A + pointed-equiv-loop-pointed-identity = + ( equiv-concat' (point-Pointed-Type A) (inv p) , right-inv p) ``` From 632bd9268988ca7a7211c5a9620e0a9b7c509061 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 24 Oct 2025 15:35:40 +0200 Subject: [PATCH 46/69] =?UTF-8?q?The=20loop=20space=20of=20a=20(?= =?UTF-8?q?=F0=9D=91=98+1)-truncated=20type=20is=20=F0=9D=91=98-truncated?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../loop-spaces.lagda.md | 25 +++++++++++++++---- 1 file changed, 20 insertions(+), 5 deletions(-) diff --git a/src/synthetic-homotopy-theory/loop-spaces.lagda.md b/src/synthetic-homotopy-theory/loop-spaces.lagda.md index 533970b711..66bf86527a 100644 --- a/src/synthetic-homotopy-theory/loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/loop-spaces.lagda.md @@ -9,8 +9,10 @@ module synthetic-homotopy-theory.loop-spaces where ```agda open import foundation.dependent-pair-types open import foundation.equivalences +open import foundation.truncated-types open import foundation.identity-types open import foundation.universe-levels +open import foundation.truncation-levels open import structured-types.h-spaces open import structured-types.magmas @@ -23,10 +25,12 @@ open import structured-types.wild-quasigroups ## Idea -The **loop space** of a [pointed type](structured-types.pointed-types.md) `A` is -the pointed type of self-[identifications](foundation-core.identity-types.md) of -the base point of `A`. The loop space comes equipped with a group-like structure -induced by the groupoidal-like structure on identifications. +The +{{#concept "loop space" Disambiguation="of a pointed type" Agda=Ω WD="loop space" WDID=Q2066070}} +of a [pointed type](structured-types.pointed-types.md) `A` is the pointed type +of self-[identifications](foundation-core.identity-types.md) of the base point +of `A`. The loop space comes equipped with a group-like structure induced by the +groupoidal-like structure on identifications. ## Table of files directly related to loop spaces @@ -84,7 +88,7 @@ module _ coherence-unit-laws-mul-Ω = refl Ω-H-Space : H-Space l - pr1 Ω-H-Space = + Ω-H-Space = ( Ω A , mul-Ω A , left-unit-law-mul-Ω , @@ -183,3 +187,14 @@ module _ pointed-equiv-loop-pointed-identity = ( equiv-concat' (point-Pointed-Type A) (inv p) , right-inv p) ``` + +### The loop space of a (𝑘+1)-truncated type is 𝑘-truncated + +```agda +module _ + {l : Level} (k : 𝕋) (A : Pointed-Type l) + where + + is-trunc-Ω : is-trunc (succ-𝕋 k) (type-Pointed-Type A) → is-trunc k (type-Ω A) + is-trunc-Ω H = H (point-Pointed-Type A) (point-Pointed-Type A) +``` From c7347c98047a46675b20341fbd0eae6a2bc1937f Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 24 Oct 2025 15:55:21 +0200 Subject: [PATCH 47/69] =?UTF-8?q?If=20A=20is=20(=F0=9D=91=9B+=F0=9D=91=98)?= =?UTF-8?q?-truncated=20then=20=CE=A9=E2=81=BFA=20is=20=F0=9D=91=98-trunca?= =?UTF-8?q?ted?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../functoriality-loop-spaces.lagda.md | 4 ++-- .../iterated-loop-spaces.lagda.md | 18 ++++++++++++++++++ .../loop-spaces.lagda.md | 4 ++-- 3 files changed, 22 insertions(+), 4 deletions(-) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index 4b4254086e..6811336870 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -14,9 +14,9 @@ open import foundation.equivalences open import foundation.faithful-maps open import foundation.function-types open import foundation.homotopies -open import foundation.truncation-levels -open import foundation.truncated-maps open import foundation.identity-types +open import foundation.truncated-maps +open import foundation.truncation-levels open import foundation.universe-levels open import structured-types.faithful-pointed-maps diff --git a/src/synthetic-homotopy-theory/iterated-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/iterated-loop-spaces.lagda.md index 22f8b3630e..a91acd1752 100644 --- a/src/synthetic-homotopy-theory/iterated-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/iterated-loop-spaces.lagda.md @@ -9,7 +9,11 @@ module synthetic-homotopy-theory.iterated-loop-spaces where ```agda open import elementary-number-theory.natural-numbers +open import foundation.function-types +open import foundation.iterated-successors-truncation-levels open import foundation.iterating-functions +open import foundation.truncated-types +open import foundation.truncation-levels open import foundation.universe-levels open import structured-types.h-spaces @@ -61,6 +65,20 @@ module _ Ω-H-Space (iterated-loop-space n (pointed-type-H-Space X)) ``` +### If A is (𝑛+𝑘)-truncated then ΩⁿA is 𝑘-truncated + +```agda +is-trunc-iterated-loop-space : + {l : Level} (n : ℕ) (k : 𝕋) (A : Pointed-Type l) → + is-trunc (iterate-succ-𝕋 n k) (type-Pointed-Type A) → + is-trunc k (type-iterated-loop-space n A) +is-trunc-iterated-loop-space zero-ℕ k A H = H +is-trunc-iterated-loop-space (succ-ℕ n) k A H = + is-trunc-Ω k + ( iterated-loop-space n A) + ( is-trunc-iterated-loop-space n (succ-𝕋 k) A H) +``` + ## See also - [Double loop spaces](synthetic-homotopy-theory.double-loop-spaces.md) diff --git a/src/synthetic-homotopy-theory/loop-spaces.lagda.md b/src/synthetic-homotopy-theory/loop-spaces.lagda.md index 66bf86527a..a99e618f93 100644 --- a/src/synthetic-homotopy-theory/loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/loop-spaces.lagda.md @@ -9,10 +9,10 @@ module synthetic-homotopy-theory.loop-spaces where ```agda open import foundation.dependent-pair-types open import foundation.equivalences -open import foundation.truncated-types open import foundation.identity-types -open import foundation.universe-levels +open import foundation.truncated-types open import foundation.truncation-levels +open import foundation.universe-levels open import structured-types.h-spaces open import structured-types.magmas From f15dde70cdff4e78d675117fa88e51a851b087b0 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 24 Oct 2025 16:21:36 +0200 Subject: [PATCH 48/69] mildly simplify a proof --- src/foundation/surjective-maps.lagda.md | 58 +++++++++++++++---------- 1 file changed, 34 insertions(+), 24 deletions(-) diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md index 1e7120f2dd..560ebc0172 100644 --- a/src/foundation/surjective-maps.lagda.md +++ b/src/foundation/surjective-maps.lagda.md @@ -13,6 +13,7 @@ open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.diagonal-maps-of-types open import foundation.embeddings +open import foundation.injective-maps open import foundation.empty-types open import foundation.equality-cartesian-product-types open import foundation.functoriality-cartesian-product-types @@ -805,25 +806,25 @@ module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} where - is-surjective-postcomp-extension-surjective-map : + is-split-surjective-postcomp-extension-surjective-map : (f : A → B) (i : A → X) (g : X → Y) → - is-surjective f → is-emb g → - is-surjective (postcomp-extension f i g) - is-surjective-postcomp-extension-surjective-map f i g H K (h , L) = - unit-trunc-Prop - ( ( j , N) , - ( eq-htpy-extension f - ( g ∘ i) - ( postcomp-extension f i g (j , N)) - ( h , L) - ( M) - ( λ a → - ( ap - ( concat' (g (i a)) (M (f a))) - ( is-section-map-inv-is-equiv - ( K (i a) (j (f a))) - ( L a ∙ inv (M (f a))))) ∙ - ( is-section-inv-concat' (M (f a)) (L a))))) + is-surjective f → + is-emb g → + is-split-surjective (postcomp-extension f i g) + is-split-surjective-postcomp-extension-surjective-map f i g H K (h , L) = + ( ( j , N) , + ( eq-htpy-extension f + ( g ∘ i) + ( postcomp-extension f i g (j , N)) + ( h , L) + ( M) + ( λ a → + ( ap + ( concat' (g (i a)) (M (f a))) + ( is-section-map-inv-is-equiv + ( K (i a) (j (f a))) + ( L a ∙ inv (M (f a))))) ∙ + ( is-section-inv-concat' (M (f a)) (L a))))) where J : (b : B) → fiber g (h b) @@ -835,20 +836,29 @@ module _ j : B → X j b = pr1 (J b) - M : (g ∘ j) ~ h + M : g ∘ j ~ h M b = pr2 (J b) - N : i ~ (j ∘ f) - N a = map-inv-is-equiv (K (i a) (j (f a))) (L a ∙ inv (M (f a))) + N : i ~ j ∘ f + N a = is-injective-is-emb K (L a ∙ inv (M (f a))) + + is-surjective-postcomp-extension-surjective-map : + (f : A → B) (i : A → X) (g : X → Y) → + is-surjective f → is-emb g → + is-surjective (postcomp-extension f i g) + is-surjective-postcomp-extension-surjective-map f i g H K = + is-surjective-is-split-surjective + ( is-split-surjective-postcomp-extension-surjective-map f i g H K) is-equiv-postcomp-extension-is-surjective : (f : A → B) (i : A → X) (g : X → Y) → is-surjective f → is-emb g → is-equiv (postcomp-extension f i g) is-equiv-postcomp-extension-is-surjective f i g H K = - is-equiv-is-emb-is-surjective - ( is-surjective-postcomp-extension-surjective-map f i g H K) - ( is-emb-postcomp-extension f i g K) + is-equiv-is-split-surjective-is-injective + ( postcomp-extension f i g) + ( is-injective-is-emb (is-emb-postcomp-extension f i g K)) + ( is-split-surjective-postcomp-extension-surjective-map f i g H K) equiv-postcomp-extension-surjection : (f : A ↠ B) (i : A → X) (g : X ↪ Y) → From da3b38acd3133c5f032ef18b0ff990990f1128dc Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 24 Oct 2025 16:41:40 +0200 Subject: [PATCH 49/69] add a link --- src/foundation/surjective-maps.lagda.md | 2 +- .../postcomposition-extensions-maps.lagda.md | 6 ++++++ 2 files changed, 7 insertions(+), 1 deletion(-) diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md index 560ebc0172..8d8f6a7b44 100644 --- a/src/foundation/surjective-maps.lagda.md +++ b/src/foundation/surjective-maps.lagda.md @@ -13,7 +13,6 @@ open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.diagonal-maps-of-types open import foundation.embeddings -open import foundation.injective-maps open import foundation.empty-types open import foundation.equality-cartesian-product-types open import foundation.functoriality-cartesian-product-types @@ -22,6 +21,7 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopy-induction open import foundation.identity-types open import foundation.inhabited-types +open import foundation.injective-maps open import foundation.postcomposition-dependent-functions open import foundation.propositional-truncations open import foundation.split-surjective-maps diff --git a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md index 410be2dea8..c888561eed 100644 --- a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md @@ -142,3 +142,9 @@ module _ trunc-map-postcomp-extension f i (g , G) = ( postcomp-extension f i g , is-trunc-map-postcomp-extension f i g G) ``` + +## See also + +- In [`foundation.surjective-maps`](foundation.surjective-maps.md) it is shown + that postcomposition of extensions along surjective maps by an embedding is an + equivalence. From fe3ed62749968c40e5832ff1d70d66b993c6e8a1 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Sat, 25 Oct 2025 16:54:17 +0200 Subject: [PATCH 50/69] Computing extension types as a dependent product --- .../precomposition-functions.lagda.md | 2 +- ...property-family-of-fibers-of-maps.lagda.md | 6 +- .../extensions-maps.lagda.md | 77 ++++++++++++++++++- 3 files changed, 80 insertions(+), 5 deletions(-) diff --git a/src/foundation/precomposition-functions.lagda.md b/src/foundation/precomposition-functions.lagda.md index 74deeba4c9..51c0079846 100644 --- a/src/foundation/precomposition-functions.lagda.md +++ b/src/foundation/precomposition-functions.lagda.md @@ -131,7 +131,7 @@ commutes. ```agda module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (X : UU l3) + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) {X : UU l3} where compute-extension-fiber-precomp' : diff --git a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md index f70e51e4ec..1ae1eaa5e6 100644 --- a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md +++ b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md @@ -636,7 +636,7 @@ dependent products: ```agda module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) (U : UU l3) + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) {U : UU l3} (g : A → U) where @@ -652,7 +652,7 @@ module _ equivalence-reasoning fiber (precomp f U) g ≃ Σ (B → U) (λ h → (a : A) → g a = (h ∘ f) a) - by compute-extension-fiber-precomp f U g + by compute-extension-fiber-precomp f g ≃ Σ ( B → U) ( λ h → (b : B) ((a , _) : fiber f b) → g a = h b) by @@ -681,7 +681,7 @@ module _ equivalence-reasoning fiber (precomp f U) g ≃ Σ (B → U) (λ h → (h ∘ f) ~ g) - by compute-extension-fiber-precomp' f U g + by compute-extension-fiber-precomp' f g ≃ Σ ( B → U) ( λ h → (b : B) ((a , _) : fiber' f b) → h b = g a) by diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index 888385a62f..3289ea3d9f 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -10,6 +10,8 @@ module orthogonal-factorization-systems.extensions-maps where open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions open import foundation.contractible-types +open import foundation.precomposition-functions +open import foundation.universal-property-family-of-fibers-of-maps open import foundation.dependent-pair-types open import foundation.embeddings open import foundation.equivalences @@ -33,7 +35,7 @@ open import foundation.truncation-levels open import foundation.type-arithmetic-dependent-pair-types open import foundation.universe-levels open import foundation.whiskering-homotopies-composition - +open import foundation.precomposition-dependent-functions open import foundation-core.torsorial-type-families ``` @@ -272,6 +274,79 @@ module _ map-inv-equiv (extensionality-extension e e') (H , K) ``` +### Computing extension types as a dependent product + +Extension types are equivalent to fibers of precomposition maps, which in turn +have a Π-type characterization. Given `i : A → B` and `g : B → C`, then + +```text + extension i g ≃ ((y : B) → Σ (c : C), (x : X) → (i x = y) → (c = g y)). +``` + +We give 4 different formulations of this equivalence for both the nondependent +and dependent extension type. + +```agda +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (i : A → B) + {C : UU l3} (g : A → C) + where + + + equiv-fiber-Π-curry-precomp-extension : + extension i g ≃ ((b : B) → Σ C (λ u → (a : A) → i a = b → g a = u)) + equiv-fiber-Π-curry-precomp-extension = + ( compute-fiber-Π-curry-precomp i g) ∘e + ( inv-equiv (compute-extension-fiber-precomp i g)) + + equiv-fiber-Π-curry-precomp-extension' : + extension i g ≃ ((b : B) → Σ C (λ u → (a : A) → b = i a → u = g a)) + equiv-fiber-Π-curry-precomp-extension' = + ( compute-fiber-Π-curry-precomp' i g) ∘e + ( inv-equiv (compute-extension-fiber-precomp i g)) + + equiv-fiber-Π-precomp-extension : + extension i g ≃ fiber-Π-precomp i g + equiv-fiber-Π-precomp-extension = + ( compute-fiber-Π-precomp i g) ∘e + ( inv-equiv (compute-extension-fiber-precomp i g)) + + equiv-fiber-Π-precomp-extension' : + extension i g ≃ fiber-Π-precomp' i g + equiv-fiber-Π-precomp-extension' = + ( compute-fiber-Π-precomp' i g) ∘e + ( inv-equiv (compute-extension-fiber-precomp i g)) + +module _ + {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (i : A → B) + {P : B → UU l3} (g : (x : A) → P (i x)) + where + + equiv-fiber-Π-curry-precomp-Π-extension-dependent-type : + extension-dependent-type i P g ≃ fiber-Π-curry-precomp-Π i P g + equiv-fiber-Π-curry-precomp-Π-extension-dependent-type = + ( compute-fiber-Π-curry-precomp-Π i P g) ∘e + ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) + + equiv-fiber-Π-curry-precomp-Π-extension-dependent-type' : + extension-dependent-type i P g ≃ fiber-Π-curry-precomp-Π' i P g + equiv-fiber-Π-curry-precomp-Π-extension-dependent-type' = + ( compute-fiber-Π-curry-precomp-Π' i P g) ∘e + ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) + + equiv-fiber-Π-precomp-Π-extension-dependent-type : + extension-dependent-type i P g ≃ fiber-Π-precomp-Π i P g + equiv-fiber-Π-precomp-Π-extension-dependent-type = + ( compute-fiber-Π-precomp-Π i P g) ∘e + ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) + + equiv-fiber-Π-precomp-Π-extension-dependent-type' : + extension-dependent-type i P g ≃ fiber-Π-precomp-Π' i P g + equiv-fiber-Π-precomp-Π-extension-dependent-type' = + ( compute-fiber-Π-precomp-Π' i P g) ∘e + ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) +``` + ### The total type of extensions is equivalent to `(y : B) → P y` ```agda From dd2a0b7f0b1e445a3028c1ff465b81c22ff8dd60 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 31 Oct 2025 22:28:46 +0100 Subject: [PATCH 51/69] edits --- src/foundation/epimorphisms.lagda.md | 10 +++++----- .../eckmann-hilton-argument.lagda.md | 4 ++++ 2 files changed, 9 insertions(+), 5 deletions(-) diff --git a/src/foundation/epimorphisms.lagda.md b/src/foundation/epimorphisms.lagda.md index b7c5bb9ab4..3c78e79e8f 100644 --- a/src/foundation/epimorphisms.lagda.md +++ b/src/foundation/epimorphisms.lagda.md @@ -84,7 +84,7 @@ module _ where is-equiv-diagonal-into-fibers-precomp-is-epimorphism : - is-epimorphism f → is-equiv (diagonal-into-fibers-precomp f X) + is-epimorphism f → is-equiv (diagonal-into-fibers-precomp f) is-equiv-diagonal-into-fibers-precomp-is-epimorphism e = is-equiv-map-section-family ( λ g → (g , refl)) @@ -102,10 +102,10 @@ module _ universal-property-pushout f f (cocone-codiagonal-map f) universal-property-pushout-is-epimorphism e X = is-equiv-comp - ( map-equiv (compute-total-fiber-precomp f X)) - ( diagonal-into-fibers-precomp f X) + ( map-equiv (compute-total-fiber-precomp f)) + ( diagonal-into-fibers-precomp f) ( is-equiv-diagonal-into-fibers-precomp-is-epimorphism f X e) - ( is-equiv-map-equiv (compute-total-fiber-precomp f X)) + ( is-equiv-map-equiv (compute-total-fiber-precomp f)) ``` If the map `f : A → B` is epi, then its codiagonal is an equivalence. @@ -142,7 +142,7 @@ If the map `f : A → B` is epi, then its codiagonal is an equivalence. ( λ g → is-contr-equiv ( Σ (B → X) (λ h → coherence-square-maps f f h g)) - ( compute-fiber-precomp f X g) + ( compute-fiber-precomp f g) ( is-contr-fam-is-equiv-map-section-family ( λ h → ( vertical-map-cocone f f diff --git a/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md b/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md index bef3c627e1..127f438e83 100644 --- a/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md +++ b/src/synthetic-homotopy-theory/eckmann-hilton-argument.lagda.md @@ -782,6 +782,10 @@ module _ ( tr³-horizontal-concat-right-unit-law-right-whisker-left-unit-law-left-whisker-Ω²)) ``` +## See also + +- [Medial magmas](structured-types.medial-magmas.md) + ## External links - [The Eckmann-Hilton argument](https://1lab.dev/Algebra.Magma.Unital.EckmannHilton.html) From 542c7e084a290c4b01f3c46b7ef176a406755be3 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 14:52:52 +0100 Subject: [PATCH 52/69] revert stupid name change --- ...oriality-dependent-function-types.lagda.md | 6 ++-- ...unctoriality-dependent-pair-types.lagda.md | 6 ++-- src/foundation-core/subtypes.lagda.md | 2 +- src/foundation-core/truncated-maps.lagda.md | 4 +-- src/foundation/0-maps.lagda.md | 12 ++++---- src/foundation/decidable-embeddings.lagda.md | 10 +++---- src/foundation/embeddings.lagda.md | 4 +-- src/foundation/faithful-maps.lagda.md | 4 +-- ...oriality-dependent-function-types.lagda.md | 16 +++++------ ...unctoriality-dependent-pair-types.lagda.md | 20 ++++++------- .../functoriality-function-types.lagda.md | 4 +-- src/logic/de-morgan-embeddings.lagda.md | 26 ++++++++--------- src/logic/de-morgan-subtypes.lagda.md | 2 +- .../double-negation-eliminating-maps.lagda.md | 6 ++-- ...double-negation-stable-embeddings.lagda.md | 28 +++++++++---------- .../double-negation-stable-subtypes.lagda.md | 2 +- .../anodyne-maps.lagda.md | 8 +++--- .../extensions-maps.lagda.md | 8 +++--- .../postcomposition-extensions-maps.lagda.md | 4 +-- .../functoriality-loop-spaces.lagda.md | 4 +-- .../binomial-types.lagda.md | 4 +-- 21 files changed, 90 insertions(+), 90 deletions(-) diff --git a/src/foundation-core/functoriality-dependent-function-types.lagda.md b/src/foundation-core/functoriality-dependent-function-types.lagda.md index 1308e74fe2..cabf45322d 100644 --- a/src/foundation-core/functoriality-dependent-function-types.lagda.md +++ b/src/foundation-core/functoriality-dependent-function-types.lagda.md @@ -97,10 +97,10 @@ module _ where abstract - is-contr-map-Π-is-fiberwise-contr-map : + is-contr-map-map-Π-is-fiberwise-contr-map : {f : (i : I) → A i → B i} → ((i : I) → is-contr-map (f i)) → is-contr-map (map-Π f) - is-contr-map-Π-is-fiberwise-contr-map H g = + is-contr-map-map-Π-is-fiberwise-contr-map H g = is-contr-equiv' _ (compute-fiber-map-Π _ g) (is-contr-Π (λ i → H i (g i))) abstract @@ -108,7 +108,7 @@ module _ {f : (i : I) → A i → B i} → is-fiberwise-equiv f → is-equiv (map-Π f) is-equiv-map-Π-is-fiberwise-equiv is-equiv-f = is-equiv-is-contr-map - ( is-contr-map-Π-is-fiberwise-contr-map + ( is-contr-map-map-Π-is-fiberwise-contr-map ( is-contr-map-is-equiv ∘ is-equiv-f)) equiv-Π-equiv-family : diff --git a/src/foundation-core/functoriality-dependent-pair-types.lagda.md b/src/foundation-core/functoriality-dependent-pair-types.lagda.md index 5334f6737b..f3d941d326 100644 --- a/src/foundation-core/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation-core/functoriality-dependent-pair-types.lagda.md @@ -389,9 +389,9 @@ module _ where abstract - is-contr-map-Σ-map-base : + is-contr-map-map-Σ-map-base : is-contr-map f → is-contr-map (map-Σ-map-base f C) - is-contr-map-Σ-map-base is-contr-f (pair y z) = + is-contr-map-map-Σ-map-base is-contr-f (pair y z) = is-contr-equiv' ( fiber f y) ( compute-fiber-map-Σ-map-base f C (pair y z)) @@ -409,7 +409,7 @@ module _ is-equiv-map-Σ-map-base : is-equiv f → is-equiv (map-Σ-map-base f C) is-equiv-map-Σ-map-base is-equiv-f = is-equiv-is-contr-map - ( is-contr-map-Σ-map-base f C (is-contr-map-is-equiv is-equiv-f)) + ( is-contr-map-map-Σ-map-base f C (is-contr-map-is-equiv is-equiv-f)) equiv-Σ-equiv-base : {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3) (e : A ≃ B) → diff --git a/src/foundation-core/subtypes.lagda.md b/src/foundation-core/subtypes.lagda.md index 25af195532..5310ef728e 100644 --- a/src/foundation-core/subtypes.lagda.md +++ b/src/foundation-core/subtypes.lagda.md @@ -226,7 +226,7 @@ module _ pr1 (trunc-map-into-subtype f p) x = (map-trunc-map f x , p x) pr2 (trunc-map-into-subtype f p) = is-trunc-map-into-subtype - ( is-trunc-map-trunc-map f) + ( is-trunc-map-map-trunc-map f) ( p) ``` diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 47395ec194..cfbf8939e5 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -53,9 +53,9 @@ module _ map-trunc-map = pr1 abstract - is-trunc-map-trunc-map : + is-trunc-map-map-trunc-map : (f : trunc-map k A B) → is-trunc-map k (map-trunc-map f) - is-trunc-map-trunc-map = pr2 + is-trunc-map-map-trunc-map = pr2 ``` ## Properties diff --git a/src/foundation/0-maps.lagda.md b/src/foundation/0-maps.lagda.md index 4890707a44..365e87053d 100644 --- a/src/foundation/0-maps.lagda.md +++ b/src/foundation/0-maps.lagda.md @@ -43,9 +43,9 @@ module _ map-0-map : {A : UU l1} {B : UU l2} → 0-map A B → A → B map-0-map = pr1 - is-0-map-0-map : + is-0-map-map-0-map : {A : UU l1} {B : UU l2} (f : 0-map A B) → is-0-map (map-0-map f) - is-0-map-0-map = pr2 + is-0-map-map-0-map = pr2 ``` ## Properties @@ -140,8 +140,8 @@ module _ where abstract - is-0-map-Σ-map-base : is-0-map f → is-0-map (map-Σ-map-base f C) - is-0-map-Σ-map-base = is-trunc-map-Σ-map-base zero-𝕋 C + is-0-map-map-Σ-map-base : is-0-map f → is-0-map (map-Σ-map-base f C) + is-0-map-map-Σ-map-base = is-trunc-map-map-Σ-map-base zero-𝕋 C ``` ### The functorial action of `Σ` preserves `0`-maps @@ -152,7 +152,7 @@ module _ (D : B → UU l4) {f : A → B} {g : (x : A) → C x → D (f x)} where - is-0-map-Σ : + is-0-map-map-Σ : is-0-map f → ((x : A) → is-0-map (g x)) → is-0-map (map-Σ D f g) - is-0-map-Σ = is-trunc-map-Σ zero-𝕋 D + is-0-map-map-Σ = is-trunc-map-map-Σ zero-𝕋 D ``` diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md index f9d1d73e96..6f6a7725f7 100644 --- a/src/foundation/decidable-embeddings.lagda.md +++ b/src/foundation/decidable-embeddings.lagda.md @@ -132,17 +132,17 @@ module _ is-decidable-emb map-decidable-emb is-decidable-emb-decidable-emb = pr2 e - is-emb-decidable-emb : is-emb map-decidable-emb - is-emb-decidable-emb = + is-emb-map-decidable-emb : is-emb map-decidable-emb + is-emb-map-decidable-emb = is-emb-is-decidable-emb is-decidable-emb-decidable-emb - is-decidable-decidable-emb : + is-decidable-map-decidable-emb : is-decidable-map map-decidable-emb - is-decidable-decidable-emb = + is-decidable-map-decidable-emb = is-decidable-map-is-decidable-emb is-decidable-emb-decidable-emb emb-decidable-emb : X ↪ Y - emb-decidable-emb = map-decidable-emb , is-emb-decidable-emb + emb-decidable-emb = map-decidable-emb , is-emb-map-decidable-emb ``` ## Properties diff --git a/src/foundation/embeddings.lagda.md b/src/foundation/embeddings.lagda.md index 9001bb666b..9d0c224af3 100644 --- a/src/foundation/embeddings.lagda.md +++ b/src/foundation/embeddings.lagda.md @@ -223,7 +223,7 @@ module _ is-emb-map-Σ-map-base : {f : A → B} (C : B → UU l3) → is-emb f → is-emb (map-Σ-map-base f C) is-emb-map-Σ-map-base C H = - is-emb-is-prop-map (is-prop-map-Σ-map-base C (is-prop-map-is-emb H)) + is-emb-is-prop-map (is-prop-map-map-Σ-map-base C (is-prop-map-is-emb H)) emb-Σ-emb-base : (f : A ↪ B) (C : B → UU l3) → Σ A (λ a → C (map-emb f a)) ↪ Σ B C @@ -240,7 +240,7 @@ module _ is-emb f → ((x : A) → is-emb (g x)) → is-emb (map-Σ D f g) is-emb-map-Σ D H K = is-emb-is-prop-map - ( is-prop-map-Σ D + ( is-prop-map-map-Σ D ( is-prop-map-is-emb H) ( λ x → is-prop-map-is-emb (K x))) diff --git a/src/foundation/faithful-maps.lagda.md b/src/foundation/faithful-maps.lagda.md index 281041beb5..341c52fb7a 100644 --- a/src/foundation/faithful-maps.lagda.md +++ b/src/foundation/faithful-maps.lagda.md @@ -237,7 +237,7 @@ module _ is-faithful f → is-faithful (map-Σ-map-base f C) is-faithful-map-Σ-map-base H = is-faithful-is-0-map - ( is-0-map-Σ-map-base C (is-0-map-is-faithful H)) + ( is-0-map-map-Σ-map-base C (is-0-map-is-faithful H)) faithful-map-Σ-faithful-map-base : (f : faithful-map A B) (C : B → UU l3) → @@ -256,7 +256,7 @@ module _ is-faithful f → ((x : A) → is-faithful (g x)) → is-faithful (map-Σ D f g) is-faithful-map-Σ H K = is-faithful-is-0-map - ( is-0-map-Σ D + ( is-0-map-map-Σ D ( is-0-map-is-faithful H) ( λ x → is-0-map-is-faithful (K x))) ``` diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md index 3ce0a35468..cc9c75b565 100644 --- a/src/foundation/functoriality-dependent-function-types.lagda.md +++ b/src/foundation/functoriality-dependent-function-types.lagda.md @@ -144,10 +144,10 @@ module _ where abstract - is-trunc-map-Π : + is-trunc-map-map-Π : (k : 𝕋) (f : (i : I) → A i → B i) → ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π f) - is-trunc-map-Π k f H h = + is-trunc-map-map-Π k f H h = is-trunc-equiv' k ( (i : I) → fiber (f i) (h i)) ( compute-fiber-map-Π f h) @@ -158,7 +158,7 @@ module _ {f : (i : I) → A i → B i} → ((i : I) → is-emb (f i)) → is-emb (map-Π f) is-emb-map-Π {f} H = is-emb-is-prop-map - ( is-trunc-map-Π neg-one-𝕋 f (λ i → is-prop-map-is-emb (H i))) + ( is-trunc-map-map-Π neg-one-𝕋 f (λ i → is-prop-map-is-emb (H i))) emb-Π : ((i : I) → A i ↪ B i) → ((i : I) → A i) ↪ ((i : I) → B i) emb-Π f = (map-Π (map-emb ∘ f) , is-emb-map-Π (is-emb-map-emb ∘ f)) @@ -171,20 +171,20 @@ module _ {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3} where - is-trunc-map-Π' : + is-trunc-map-map-Π' : (k : 𝕋) {l4 : Level} {J : UU l4} (α : J → I) (f : (i : I) → A i → B i) → ((i : I) → is-trunc-map k (f i)) → is-trunc-map k (map-Π' α f) - is-trunc-map-Π' k {J = J} α f H h = + is-trunc-map-map-Π' k {J = J} α f H h = is-trunc-equiv' k ( (j : J) → fiber (f (α j)) (h j)) ( compute-fiber-map-Π' α f h) ( is-trunc-Π k (λ j → H (α j) (h j))) - is-trunc-map-is-trunc-map-Π' : + is-trunc-map-is-trunc-map-map-Π' : (k : 𝕋) (f : (i : I) → A i → B i) → ({l : Level} {J : UU l} (α : J → I) → is-trunc-map k (map-Π' α f)) → (i : I) → is-trunc-map k (f i) - is-trunc-map-is-trunc-map-Π' k f H i b = + is-trunc-map-is-trunc-map-map-Π' k f H i b = is-trunc-equiv' k ( fiber (map-Π (λ _ → f i)) (point b)) ( equiv-Σ @@ -202,7 +202,7 @@ module _ ((i : I) → is-emb (f i)) → is-emb (map-Π' α f) is-emb-map-Π' α f H = is-emb-is-prop-map - ( is-trunc-map-Π' neg-one-𝕋 α f (λ i → is-prop-map-is-emb (H i))) + ( is-trunc-map-map-Π' neg-one-𝕋 α f (λ i → is-prop-map-is-emb (H i))) ``` ### The action on homotopies of equivalences diff --git a/src/foundation/functoriality-dependent-pair-types.lagda.md b/src/foundation/functoriality-dependent-pair-types.lagda.md index b3dfd0b0fc..1f93fd022d 100644 --- a/src/foundation/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation/functoriality-dependent-pair-types.lagda.md @@ -181,37 +181,37 @@ module _ where abstract - is-trunc-map-Σ-map-base : + is-trunc-map-map-Σ-map-base : (k : 𝕋) {f : A → B} (C : B → UU l3) → is-trunc-map k f → is-trunc-map k (map-Σ-map-base f C) - is-trunc-map-Σ-map-base k {f} C H y = + is-trunc-map-map-Σ-map-base k {f} C H y = is-trunc-equiv' k ( fiber f (pr1 y)) ( compute-fiber-map-Σ-map-base f C y) ( H (pr1 y)) abstract - is-prop-map-Σ-map-base : + is-prop-map-map-Σ-map-base : {f : A → B} (C : B → UU l3) → is-prop-map f → is-prop-map (map-Σ-map-base f C) - is-prop-map-Σ-map-base C = is-trunc-map-Σ-map-base neg-one-𝕋 C + is-prop-map-map-Σ-map-base C = is-trunc-map-map-Σ-map-base neg-one-𝕋 C module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} where abstract - is-trunc-map-Σ : + is-trunc-map-map-Σ : (k : 𝕋) (D : B → UU l4) {f : A → B} {g : (x : A) → C x → D (f x)} → is-trunc-map k f → ((x : A) → is-trunc-map k (g x)) → is-trunc-map k (map-Σ D f g) - is-trunc-map-Σ k D {f} {g} H K = + is-trunc-map-map-Σ k D {f} {g} H K = is-trunc-map-left-map-triangle k ( map-Σ D f g) ( map-Σ-map-base f D) ( tot g) ( triangle-map-Σ D f g) - ( is-trunc-map-Σ-map-base k D H) + ( is-trunc-map-map-Σ-map-base k D H) ( is-trunc-map-tot k K) module _ @@ -222,13 +222,13 @@ module _ is-contr-map-Σ : is-contr-map f → ((x : A) → is-contr-map (g x)) → is-contr-map (map-Σ D f g) - is-contr-map-Σ = is-trunc-map-Σ neg-two-𝕋 D + is-contr-map-Σ = is-trunc-map-map-Σ neg-two-𝕋 D abstract - is-prop-map-Σ : + is-prop-map-map-Σ : is-prop-map f → ((x : A) → is-prop-map (g x)) → is-prop-map (map-Σ D f g) - is-prop-map-Σ = is-trunc-map-Σ neg-one-𝕋 D + is-prop-map-map-Σ = is-trunc-map-map-Σ neg-one-𝕋 D ``` ### Commuting squares of maps on total spaces diff --git a/src/foundation/functoriality-function-types.lagda.md b/src/foundation/functoriality-function-types.lagda.md index c7fa12cb18..7797a47622 100644 --- a/src/foundation/functoriality-function-types.lagda.md +++ b/src/foundation/functoriality-function-types.lagda.md @@ -71,13 +71,13 @@ module _ is-trunc-map k f → {l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f) is-trunc-map-postcomp-is-trunc-map is-trunc-f A = - is-trunc-map-Π' k (terminal-map A) (point f) (point is-trunc-f) + is-trunc-map-map-Π' k (terminal-map A) (point f) (point is-trunc-f) is-trunc-map-is-trunc-map-postcomp : ({l3 : Level} (A : UU l3) → is-trunc-map k (postcomp A f)) → is-trunc-map k f is-trunc-map-is-trunc-map-postcomp is-trunc-postcomp-f = - is-trunc-map-is-trunc-map-Π' k + is-trunc-map-is-trunc-map-map-Π' k ( point f) ( λ {l} {J} α → is-trunc-postcomp-f {l} J) ( star) diff --git a/src/logic/de-morgan-embeddings.lagda.md b/src/logic/de-morgan-embeddings.lagda.md index b590deac87..a65add4e2c 100644 --- a/src/logic/de-morgan-embeddings.lagda.md +++ b/src/logic/de-morgan-embeddings.lagda.md @@ -143,19 +143,19 @@ module _ map-de-morgan-emb : X → Y map-de-morgan-emb = pr1 e - is-de-morgan-emb-de-morgan-emb : is-de-morgan-emb map-de-morgan-emb - is-de-morgan-emb-de-morgan-emb = pr2 e + is-de-morgan-emb-map-de-morgan-emb : is-de-morgan-emb map-de-morgan-emb + is-de-morgan-emb-map-de-morgan-emb = pr2 e - is-emb-de-morgan-emb : is-emb map-de-morgan-emb - is-emb-de-morgan-emb = - is-emb-is-de-morgan-emb is-de-morgan-emb-de-morgan-emb + is-emb-map-de-morgan-emb : is-emb map-de-morgan-emb + is-emb-map-de-morgan-emb = + is-emb-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb - is-de-morgan-de-morgan-emb : is-de-morgan-map map-de-morgan-emb - is-de-morgan-de-morgan-emb = - is-de-morgan-map-is-de-morgan-emb is-de-morgan-emb-de-morgan-emb + is-de-morgan-map-de-morgan-emb : is-de-morgan-map map-de-morgan-emb + is-de-morgan-map-de-morgan-emb = + is-de-morgan-map-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb emb-de-morgan-emb : X ↪ Y - emb-de-morgan-emb = map-de-morgan-emb , is-emb-de-morgan-emb + emb-de-morgan-emb = map-de-morgan-emb , is-emb-map-de-morgan-emb ``` ## Properties @@ -425,7 +425,7 @@ abstract ( is-prop-is-de-morgan-emb) ( map-de-morgan-emb f) ( refl-htpy) - ( is-de-morgan-emb-de-morgan-emb f) + ( is-de-morgan-emb-map-de-morgan-emb f) abstract is-equiv-htpy-eq-de-morgan-emb : @@ -478,7 +478,7 @@ module _ de-morgan-emb-tot f = ( tot (map-de-morgan-emb ∘ f) , is-de-morgan-emb-tot - ( is-de-morgan-emb-de-morgan-emb ∘ f)) + ( is-de-morgan-emb-map-de-morgan-emb ∘ f)) ``` ### The map on total spaces induced by a De Morgan embedding on the base is a De Morgan embedding @@ -502,7 +502,7 @@ module _ de-morgan-emb-map-Σ-map-base f = ( map-Σ-map-base (map-de-morgan-emb f) C , is-de-morgan-emb-map-Σ-map-base - ( is-de-morgan-emb-de-morgan-emb f)) + ( is-de-morgan-emb-map-de-morgan-emb f)) ``` ### The functoriality of dependent pair types on De Morgan embeddings @@ -531,7 +531,7 @@ module _ ( ( map-Σ D (map-decidable-emb f) (map-de-morgan-emb ∘ g)) , ( is-de-morgan-emb-map-Σ ( is-decidable-emb-decidable-emb f) - ( is-de-morgan-emb-de-morgan-emb ∘ g))) + ( is-de-morgan-emb-map-de-morgan-emb ∘ g))) ``` ### Products of De Morgan embeddings are De Morgan embeddings diff --git a/src/logic/de-morgan-subtypes.lagda.md b/src/logic/de-morgan-subtypes.lagda.md index 0578cc921e..8981009f85 100644 --- a/src/logic/de-morgan-subtypes.lagda.md +++ b/src/logic/de-morgan-subtypes.lagda.md @@ -161,7 +161,7 @@ module _ fiber (map-de-morgan-emb f) y pr2 (de-morgan-subtype-de-morgan-emb y) = is-de-morgan-prop-map-is-de-morgan-emb - ( is-de-morgan-emb-de-morgan-emb f) + ( is-de-morgan-emb-map-de-morgan-emb f) ( y) compute-type-de-morgan-type-de-morgan-emb : diff --git a/src/logic/double-negation-eliminating-maps.lagda.md b/src/logic/double-negation-eliminating-maps.lagda.md index 1909dd860c..8b9c1f8e95 100644 --- a/src/logic/double-negation-eliminating-maps.lagda.md +++ b/src/logic/double-negation-eliminating-maps.lagda.md @@ -84,9 +84,9 @@ module _ map-double-negation-eliminating-map : A → B map-double-negation-eliminating-map = pr1 f - is-double-negation-eliminating-double-negation-eliminating-map : + is-double-negation-eliminating-map-double-negation-eliminating-map : is-double-negation-eliminating-map map-double-negation-eliminating-map - is-double-negation-eliminating-double-negation-eliminating-map = pr2 f + is-double-negation-eliminating-map-double-negation-eliminating-map = pr2 f ``` ## Properties @@ -348,7 +348,7 @@ module _ (f : A →¬¬ B) → ε-operator-map (map-double-negation-eliminating-map f) ε-operator-double-negation-eliminating-map f = ε-operator-map-is-double-negation-eliminating-map - ( is-double-negation-eliminating-double-negation-eliminating-map f) + ( is-double-negation-eliminating-map-double-negation-eliminating-map f) ``` ## See also diff --git a/src/logic/double-negation-stable-embeddings.lagda.md b/src/logic/double-negation-stable-embeddings.lagda.md index ff510de7fa..24bd420213 100644 --- a/src/logic/double-negation-stable-embeddings.lagda.md +++ b/src/logic/double-negation-stable-embeddings.lagda.md @@ -145,25 +145,25 @@ module _ map-double-negation-stable-emb : X → Y map-double-negation-stable-emb = pr1 e - is-double-negation-stable-emb-double-negation-stable-emb : + is-double-negation-stable-emb-map-double-negation-stable-emb : is-double-negation-stable-emb map-double-negation-stable-emb - is-double-negation-stable-emb-double-negation-stable-emb = pr2 e + is-double-negation-stable-emb-map-double-negation-stable-emb = pr2 e - is-emb-double-negation-stable-emb : + is-emb-map-double-negation-stable-emb : is-emb map-double-negation-stable-emb - is-emb-double-negation-stable-emb = + is-emb-map-double-negation-stable-emb = is-emb-is-double-negation-stable-emb - is-double-negation-stable-emb-double-negation-stable-emb + is-double-negation-stable-emb-map-double-negation-stable-emb - is-double-negation-eliminating-double-negation-stable-emb : + is-double-negation-eliminating-map-double-negation-stable-emb : is-double-negation-eliminating-map map-double-negation-stable-emb - is-double-negation-eliminating-double-negation-stable-emb = + is-double-negation-eliminating-map-double-negation-stable-emb = is-double-negation-eliminating-map-is-double-negation-stable-emb - is-double-negation-stable-emb-double-negation-stable-emb + is-double-negation-stable-emb-map-double-negation-stable-emb emb-double-negation-stable-emb : X ↪ Y emb-double-negation-stable-emb = - map-double-negation-stable-emb , is-emb-double-negation-stable-emb + map-double-negation-stable-emb , is-emb-map-double-negation-stable-emb ``` ## Properties @@ -450,7 +450,7 @@ abstract ( is-prop-is-double-negation-stable-emb) ( map-double-negation-stable-emb f) ( refl-htpy) - ( is-double-negation-stable-emb-double-negation-stable-emb f) + ( is-double-negation-stable-emb-map-double-negation-stable-emb f) abstract is-equiv-htpy-eq-double-negation-stable-emb : @@ -536,7 +536,7 @@ module _ double-negation-stable-emb-tot f = ( tot (map-double-negation-stable-emb ∘ f) , is-double-negation-stable-emb-tot - ( is-double-negation-stable-emb-double-negation-stable-emb ∘ f)) + ( is-double-negation-stable-emb-map-double-negation-stable-emb ∘ f)) ``` ### The map on total spaces induced by a double negation stable embedding on the base is a double negation stable embedding @@ -560,7 +560,7 @@ module _ double-negation-stable-emb-map-Σ-map-base f = ( map-Σ-map-base (map-double-negation-stable-emb f) C , is-double-negation-stable-emb-map-Σ-map-base - ( is-double-negation-stable-emb-double-negation-stable-emb f)) + ( is-double-negation-stable-emb-map-double-negation-stable-emb f)) ``` ### The functoriality of dependent pair types preserves double negation stable embeddings @@ -590,8 +590,8 @@ module _ ( map-double-negation-stable-emb f) ( map-double-negation-stable-emb ∘ g)) , ( is-double-negation-stable-emb-map-Σ - ( is-double-negation-stable-emb-double-negation-stable-emb f) - ( is-double-negation-stable-emb-double-negation-stable-emb ∘ g))) + ( is-double-negation-stable-emb-map-double-negation-stable-emb f) + ( is-double-negation-stable-emb-map-double-negation-stable-emb ∘ g))) ``` ### Products of double negation stable embeddings are double negation stable embeddings diff --git a/src/logic/double-negation-stable-subtypes.lagda.md b/src/logic/double-negation-stable-subtypes.lagda.md index 29a0af7ea3..9345efa5c3 100644 --- a/src/logic/double-negation-stable-subtypes.lagda.md +++ b/src/logic/double-negation-stable-subtypes.lagda.md @@ -172,7 +172,7 @@ module _ fiber (map-double-negation-stable-emb f) y pr2 (double-negation-stable-subtype-double-negation-stable-emb y) = is-double-negation-stable-prop-map-is-double-negation-stable-emb - ( is-double-negation-stable-emb-double-negation-stable-emb f) + ( is-double-negation-stable-emb-map-double-negation-stable-emb f) ( y) compute-type-double-negation-stable-type-double-negation-stable-emb : diff --git a/src/orthogonal-factorization-systems/anodyne-maps.lagda.md b/src/orthogonal-factorization-systems/anodyne-maps.lagda.md index ba5406b0bc..8f86a1c09d 100644 --- a/src/orthogonal-factorization-systems/anodyne-maps.lagda.md +++ b/src/orthogonal-factorization-systems/anodyne-maps.lagda.md @@ -136,9 +136,9 @@ module _ (f : A → B) (j : C → D) (i : E → F) where - is-anodyne-map-product : + is-anodyne-map-map-product : is-anodyne-map f j → is-anodyne-map f i → is-anodyne-map f (map-product j i) - is-anodyne-map-product J I g H = + is-anodyne-map-map-product J I g H = is-orthogonal-left-product j i g (J g H) (I g H) ``` @@ -151,11 +151,11 @@ module _ (f : A → B) (j : C → D) (i : E → F) where - is-anodyne-map-coproduct : + is-anodyne-map-map-coproduct : is-anodyne-map f j → is-anodyne-map f i → is-anodyne-map f (map-coproduct j i) - is-anodyne-map-coproduct J I g H = + is-anodyne-map-map-coproduct J I g H = is-orthogonal-left-coproduct j i g (J g H) (I g H) ``` diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index 3289ea3d9f..d5d16e3066 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -10,8 +10,6 @@ module orthogonal-factorization-systems.extensions-maps where open import foundation.action-on-identifications-dependent-functions open import foundation.action-on-identifications-functions open import foundation.contractible-types -open import foundation.precomposition-functions -open import foundation.universal-property-family-of-fibers-of-maps open import foundation.dependent-pair-types open import foundation.embeddings open import foundation.equivalences @@ -25,6 +23,8 @@ open import foundation.homotopy-induction open import foundation.identity-types open import foundation.monomorphisms open import foundation.postcomposition-functions +open import foundation.precomposition-dependent-functions +open import foundation.precomposition-functions open import foundation.propositions open import foundation.sets open import foundation.structure-identity-principle @@ -33,9 +33,10 @@ open import foundation.truncated-maps open import foundation.truncated-types open import foundation.truncation-levels open import foundation.type-arithmetic-dependent-pair-types +open import foundation.universal-property-family-of-fibers-of-maps open import foundation.universe-levels open import foundation.whiskering-homotopies-composition -open import foundation.precomposition-dependent-functions + open import foundation-core.torsorial-type-families ``` @@ -292,7 +293,6 @@ module _ {C : UU l3} (g : A → C) where - equiv-fiber-Π-curry-precomp-extension : extension i g ≃ ((b : B) → Σ C (λ u → (a : A) → i a = b → g a = u)) equiv-fiber-Π-curry-precomp-extension = diff --git a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md index c888561eed..828dfc2c7b 100644 --- a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md @@ -128,11 +128,11 @@ module _ (f : A → B) (i : A → X) (g : X → Y) → is-trunc-map k g → is-trunc-map k (postcomp-extension f i g) is-trunc-map-postcomp-extension f i g G = - is-trunc-map-Σ k + is-trunc-map-map-Σ k ( is-extension f (g ∘ i)) ( is-trunc-map-postcomp-is-trunc-map k g G B) ( λ j → - is-trunc-map-Π k + is-trunc-map-map-Π k ( λ a → ap g) ( λ a → is-trunc-map-ap-is-trunc-map k g G (i a) (j (f a)))) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index 6811336870..209d7620de 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -86,11 +86,11 @@ module _ {l1 l2 : Level} (k : 𝕋) {A : Pointed-Type l1} {B : Pointed-Type l2} where - is-trunc-map-Ω : + is-trunc-map-map-Ω : (f : A →∗ B) → is-trunc-map (succ-𝕋 k) (map-pointed-map f) → is-trunc-map k (map-Ω f) - is-trunc-map-Ω f H = + is-trunc-map-map-Ω f H = is-trunc-map-comp k ( tr-type-Ω (preserves-point-pointed-map f)) ( ap (map-pointed-map f)) diff --git a/src/univalent-combinatorics/binomial-types.lagda.md b/src/univalent-combinatorics/binomial-types.lagda.md index 88c619060f..8b2f68067a 100644 --- a/src/univalent-combinatorics/binomial-types.lagda.md +++ b/src/univalent-combinatorics/binomial-types.lagda.md @@ -97,7 +97,7 @@ module _ is-emb-map-emb-binomial-type-Level : is-emb map-decidable-emb-binomial-type-Level is-emb-map-emb-binomial-type-Level = - is-emb-decidable-emb decidable-emb-binomial-type-Level + is-emb-map-decidable-emb decidable-emb-binomial-type-Level ``` ### The standard binomial types @@ -127,7 +127,7 @@ module _ abstract is-emb-map-emb-binomial-type : is-emb map-decidable-emb-binomial-type is-emb-map-emb-binomial-type = - is-emb-decidable-emb decidable-emb-binomial-type + is-emb-map-decidable-emb decidable-emb-binomial-type ``` ### The type of decidable subtypes of `A` such that the total space is merely equivalent to a given finite type From 8e8c485e133c1288dac6ccd9eb77ee67895f24fa Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 15:08:54 +0100 Subject: [PATCH 53/69] revert stupid name change --- src/foundation/decidable-embeddings.lagda.md | 22 +++++++++---------- src/foundation/decidable-subtypes.lagda.md | 2 +- ...unctoriality-dependent-pair-types.lagda.md | 4 ++-- src/logic/de-morgan-embeddings.lagda.md | 6 ++--- ...double-negation-stable-embeddings.lagda.md | 4 ++-- 5 files changed, 19 insertions(+), 19 deletions(-) diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md index 6f6a7725f7..a58beedcbb 100644 --- a/src/foundation/decidable-embeddings.lagda.md +++ b/src/foundation/decidable-embeddings.lagda.md @@ -128,18 +128,18 @@ module _ map-decidable-emb : X → Y map-decidable-emb = pr1 e - is-decidable-emb-decidable-emb : + is-decidable-emb-map-decidable-emb : is-decidable-emb map-decidable-emb - is-decidable-emb-decidable-emb = pr2 e + is-decidable-emb-map-decidable-emb = pr2 e is-emb-map-decidable-emb : is-emb map-decidable-emb is-emb-map-decidable-emb = - is-emb-is-decidable-emb is-decidable-emb-decidable-emb + is-emb-is-decidable-emb is-decidable-emb-map-decidable-emb - is-decidable-map-decidable-emb : + is-decidable-map-map-decidable-emb : is-decidable-map map-decidable-emb - is-decidable-map-decidable-emb = - is-decidable-map-is-decidable-emb is-decidable-emb-decidable-emb + is-decidable-map-map-decidable-emb = + is-decidable-map-is-decidable-emb is-decidable-emb-map-decidable-emb emb-decidable-emb : X ↪ Y emb-decidable-emb = map-decidable-emb , is-emb-map-decidable-emb @@ -387,7 +387,7 @@ abstract ( is-prop-is-decidable-emb) ( map-decidable-emb f) ( refl-htpy) - ( is-decidable-emb-decidable-emb f) + ( is-decidable-emb-map-decidable-emb f) abstract is-equiv-htpy-eq-decidable-emb : @@ -465,7 +465,7 @@ module _ decidable-emb-tot : ((x : A) → B x ↪ᵈ C x) → Σ A B ↪ᵈ Σ A C decidable-emb-tot f = ( tot (λ x → map-decidable-emb (f x)) , - is-decidable-emb-tot (λ x → is-decidable-emb-decidable-emb (f x))) + is-decidable-emb-tot (λ x → is-decidable-emb-map-decidable-emb (f x))) ``` ### The map on total spaces induced by a decidable embedding on the base is a decidable embedding @@ -485,7 +485,7 @@ module _ (f : A ↪ᵈ B) → Σ A (C ∘ map-decidable-emb f) ↪ᵈ Σ B C decidable-emb-map-Σ-map-base f = ( map-Σ-map-base (map-decidable-emb f) C , - is-decidable-emb-map-Σ-map-base ((is-decidable-emb-decidable-emb f))) + is-decidable-emb-map-Σ-map-base ((is-decidable-emb-map-decidable-emb f))) ``` ### The functoriality of dependent pair types preserves decidable embeddings @@ -513,8 +513,8 @@ module _ decidable-emb-Σ f g = ( ( map-Σ D (map-decidable-emb f) (λ x → map-decidable-emb (g x))) , ( is-decidable-emb-map-Σ - ( is-decidable-emb-decidable-emb f) - ( λ x → is-decidable-emb-decidable-emb (g x)))) + ( is-decidable-emb-map-decidable-emb f) + ( λ x → is-decidable-emb-map-decidable-emb (g x)))) ``` ### Products of decidable embeddings are decidable embeddings diff --git a/src/foundation/decidable-subtypes.lagda.md b/src/foundation/decidable-subtypes.lagda.md index e2cee41424..d54451386b 100644 --- a/src/foundation/decidable-subtypes.lagda.md +++ b/src/foundation/decidable-subtypes.lagda.md @@ -168,7 +168,7 @@ module _ fiber (map-decidable-emb f) y pr2 (decidable-subtype-decidable-emb y) = is-decidable-prop-map-is-decidable-emb - ( is-decidable-emb-decidable-emb f) + ( is-decidable-emb-map-decidable-emb f) ( y) compute-type-decidable-subtype-decidable-emb : diff --git a/src/foundation/functoriality-dependent-pair-types.lagda.md b/src/foundation/functoriality-dependent-pair-types.lagda.md index 1f93fd022d..beffdd38f2 100644 --- a/src/foundation/functoriality-dependent-pair-types.lagda.md +++ b/src/foundation/functoriality-dependent-pair-types.lagda.md @@ -219,10 +219,10 @@ module _ where abstract - is-contr-map-Σ : + is-contr-map-map-Σ : is-contr-map f → ((x : A) → is-contr-map (g x)) → is-contr-map (map-Σ D f g) - is-contr-map-Σ = is-trunc-map-map-Σ neg-two-𝕋 D + is-contr-map-map-Σ = is-trunc-map-map-Σ neg-two-𝕋 D abstract is-prop-map-map-Σ : diff --git a/src/logic/de-morgan-embeddings.lagda.md b/src/logic/de-morgan-embeddings.lagda.md index a65add4e2c..3391dea43b 100644 --- a/src/logic/de-morgan-embeddings.lagda.md +++ b/src/logic/de-morgan-embeddings.lagda.md @@ -150,8 +150,8 @@ module _ is-emb-map-de-morgan-emb = is-emb-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb - is-de-morgan-map-de-morgan-emb : is-de-morgan-map map-de-morgan-emb - is-de-morgan-map-de-morgan-emb = + is-de-morgan-map-map-de-morgan-emb : is-de-morgan-map map-de-morgan-emb + is-de-morgan-map-map-de-morgan-emb = is-de-morgan-map-is-de-morgan-emb is-de-morgan-emb-map-de-morgan-emb emb-de-morgan-emb : X ↪ Y @@ -530,7 +530,7 @@ module _ de-morgan-emb-Σ f g = ( ( map-Σ D (map-decidable-emb f) (map-de-morgan-emb ∘ g)) , ( is-de-morgan-emb-map-Σ - ( is-decidable-emb-decidable-emb f) + ( is-decidable-emb-map-decidable-emb f) ( is-de-morgan-emb-map-de-morgan-emb ∘ g))) ``` diff --git a/src/logic/double-negation-stable-embeddings.lagda.md b/src/logic/double-negation-stable-embeddings.lagda.md index 24bd420213..f6f753758a 100644 --- a/src/logic/double-negation-stable-embeddings.lagda.md +++ b/src/logic/double-negation-stable-embeddings.lagda.md @@ -155,9 +155,9 @@ module _ is-emb-is-double-negation-stable-emb is-double-negation-stable-emb-map-double-negation-stable-emb - is-double-negation-eliminating-map-double-negation-stable-emb : + is-double-negation-eliminating-map-map-double-negation-stable-emb : is-double-negation-eliminating-map map-double-negation-stable-emb - is-double-negation-eliminating-map-double-negation-stable-emb = + is-double-negation-eliminating-map-map-double-negation-stable-emb = is-double-negation-eliminating-map-is-double-negation-stable-emb is-double-negation-stable-emb-map-double-negation-stable-emb From 282c71d7542f6e7fa12c3d5c11fd55d51caeaf99 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 15:09:42 +0100 Subject: [PATCH 54/69] revert stupid name change --- src/trees/functoriality-w-types.lagda.md | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/trees/functoriality-w-types.lagda.md b/src/trees/functoriality-w-types.lagda.md index ea23c2e517..e5a123d283 100644 --- a/src/trees/functoriality-w-types.lagda.md +++ b/src/trees/functoriality-w-types.lagda.md @@ -120,18 +120,18 @@ abstract ### For any family of equivalences `e` over `f`, if `f` is truncated then `map-𝕎 f e` is truncated ```agda -is-trunc-map-𝕎 : +is-trunc-map-map-𝕎 : {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) (f : A → C) (e : (x : A) → B x ≃ D (f x)) → is-trunc-map k f → is-trunc-map k (map-𝕎 D f e) -is-trunc-map-𝕎 k D f e H (tree-𝕎 c γ) = +is-trunc-map-map-𝕎 k D f e H (tree-𝕎 c γ) = is-trunc-equiv k ( fiber-map-𝕎 D f e (tree-𝕎 c γ)) ( equiv-fiber-map-𝕎 D f e (tree-𝕎 c γ)) ( is-trunc-Σ ( H c) - ( λ t → is-trunc-Π k (λ d → is-trunc-map-𝕎 k D f e H (γ d)))) + ( λ t → is-trunc-Π k (λ d → is-trunc-map-map-𝕎 k D f e H (γ d)))) ``` ### For any family of equivalences `e` over `f`, if `f` is an equivalence then `map-𝕎 f e` is an equivalence @@ -143,7 +143,7 @@ is-equiv-map-𝕎 : is-equiv f → is-equiv (map-𝕎 D f e) is-equiv-map-𝕎 D f e H = is-equiv-is-contr-map - ( is-trunc-map-𝕎 neg-two-𝕋 D f e (is-contr-map-is-equiv H)) + ( is-trunc-map-map-𝕎 neg-two-𝕋 D f e (is-contr-map-is-equiv H)) equiv-𝕎 : {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) @@ -164,7 +164,7 @@ is-emb-map-𝕎 : is-emb f → is-emb (map-𝕎 D f e) is-emb-map-𝕎 D f e H = is-emb-is-prop-map - (is-trunc-map-𝕎 neg-one-𝕋 D f e (is-prop-map-is-emb H)) + (is-trunc-map-map-𝕎 neg-one-𝕋 D f e (is-prop-map-is-emb H)) emb-𝕎 : {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} (D : C → UU l4) From ce1e7fda924c59faa89b1ac1ae2b84d1688b0cd6 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 15:14:24 +0100 Subject: [PATCH 55/69] fix a renaming mistake --- .../types-separated-at-maps.lagda.md | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md b/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md index c5e4c76aa2..a909c66dac 100644 --- a/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/types-separated-at-maps.lagda.md @@ -28,11 +28,11 @@ module _ {l1 l2 l3 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) where - is-rel-separated-family : (X → UU l3) → UU (l1 ⊔ l2 ⊔ l3) - is-rel-separated-family A = (x : X) (y z : A x) → is-local f (y = z) + is-map-separated-family : (X → UU l3) → UU (l1 ⊔ l2 ⊔ l3) + is-map-separated-family A = (x : X) (y z : A x) → is-local f (y = z) - is-rel-separated : UU l3 → UU (l1 ⊔ l2 ⊔ l3) - is-rel-separated A = is-rel-separated-family (λ _ → A) + is-map-separated : UU l3 → UU (l1 ⊔ l2 ⊔ l3) + is-map-separated A = is-map-separated-family (λ _ → A) ``` ## References From 3ed97cc4f33536d90c50427823106b2f9b9c154b Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 15:49:55 +0100 Subject: [PATCH 56/69] revert another naming change --- src/foundation-core/truncated-maps.lagda.md | 12 ++++++------ src/foundation/diagonals-of-maps.lagda.md | 2 +- src/foundation/faithful-maps.lagda.md | 4 ++-- src/foundation/mere-path-cosplit-maps.lagda.md | 2 +- src/foundation/path-cosplit-maps.lagda.md | 2 +- .../precomposition-dependent-functions.lagda.md | 2 +- .../functoriality-loop-spaces.lagda.md | 2 +- 7 files changed, 13 insertions(+), 13 deletions(-) diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 720b4c6286..23cc51fa58 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -114,18 +114,18 @@ module _ where abstract - is-trunc-map-succ-is-trunc-map-ap : + is-trunc-map-is-trunc-map-ap : ((x y : A) → is-trunc-map k (ap f {x} {y})) → is-trunc-map (succ-𝕋 k) f - is-trunc-map-succ-is-trunc-map-ap is-trunc-map-ap-f b (x , p) (x' , p') = + is-trunc-map-is-trunc-map-ap is-trunc-map-ap-f b (x , p) (x' , p') = is-trunc-equiv k ( fiber (ap f) (p ∙ inv p')) ( equiv-fiber-ap-eq-fiber f (x , p) (x' , p')) ( is-trunc-map-ap-f x x' (p ∙ inv p')) abstract - is-trunc-map-ap-is-trunc-map-succ : + is-trunc-map-ap-is-trunc-map : is-trunc-map (succ-𝕋 k) f → (x y : A) → is-trunc-map k (ap f {x} {y}) - is-trunc-map-ap-is-trunc-map-succ is-trunc-map-f x y p = + is-trunc-map-ap-is-trunc-map is-trunc-map-f x y p = is-trunc-is-equiv' k ( (x , p) = (y , refl)) ( eq-fiber-fiber-ap f x y p) @@ -144,7 +144,7 @@ module _ is-trunc-map-ap-is-trunc-map : is-trunc-map k f → (x y : A) → is-trunc-map k (ap f {x} {y}) is-trunc-map-ap-is-trunc-map H = - is-trunc-map-ap-is-trunc-map-succ k f (is-trunc-map-succ-is-trunc-map k H) + is-trunc-map-ap-is-trunc-map k f (is-trunc-map-succ-is-trunc-map k H) ``` ### The domain of any `k`-truncated map into a `k+1`-truncated type is `k+1`-truncated @@ -162,7 +162,7 @@ is-trunc-is-trunc-map-into-is-trunc ( k) ( ap f) ( is-trunc-B (f a) (f a')) - ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-map-f a a') + ( is-trunc-map-ap-is-trunc-map k f is-trunc-map-f a a') ``` ### A family of types is a family of `k`-truncated types if and only of the projection map is `k`-truncated diff --git a/src/foundation/diagonals-of-maps.lagda.md b/src/foundation/diagonals-of-maps.lagda.md index 8454997692..a965e62bc4 100644 --- a/src/foundation/diagonals-of-maps.lagda.md +++ b/src/foundation/diagonals-of-maps.lagda.md @@ -88,7 +88,7 @@ abstract is-trunc-is-equiv k (fiber (ap f) p) ( fiber-ap-fiber-diagonal-map f (triple x y p)) ( is-equiv-fiber-ap-fiber-diagonal-map f (triple x y p)) - ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y p) + ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y p) abstract is-trunc-map-is-trunc-diagonal-map : diff --git a/src/foundation/faithful-maps.lagda.md b/src/foundation/faithful-maps.lagda.md index 341c52fb7a..0e523f08b3 100644 --- a/src/foundation/faithful-maps.lagda.md +++ b/src/foundation/faithful-maps.lagda.md @@ -107,12 +107,12 @@ module _ is-0-map-is-faithful : is-faithful f → is-0-map f is-0-map-is-faithful H = - is-trunc-map-succ-is-trunc-map-ap neg-one-𝕋 f + is-trunc-map-is-trunc-map-ap neg-one-𝕋 f ( λ x y → is-prop-map-is-emb (H x y)) is-faithful-is-0-map : is-0-map f → is-faithful f is-faithful-is-0-map H x y = - is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map-succ neg-one-𝕋 f H x y) + is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map neg-one-𝕋 f H x y) ``` ## Properties diff --git a/src/foundation/mere-path-cosplit-maps.lagda.md b/src/foundation/mere-path-cosplit-maps.lagda.md index a32550ef06..94523eeebc 100644 --- a/src/foundation/mere-path-cosplit-maps.lagda.md +++ b/src/foundation/mere-path-cosplit-maps.lagda.md @@ -106,7 +106,7 @@ is-mere-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = unit-trunc-Prop (retraction-is-contr-map is-trunc-f) is-mere-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = is-mere-path-cosplit-is-trunc k - ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) + ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is merely `k+1`-path-cosplit diff --git a/src/foundation/path-cosplit-maps.lagda.md b/src/foundation/path-cosplit-maps.lagda.md index dbb0e02eff..9a52358e95 100644 --- a/src/foundation/path-cosplit-maps.lagda.md +++ b/src/foundation/path-cosplit-maps.lagda.md @@ -171,7 +171,7 @@ is-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = retraction-is-contr-map is-trunc-f is-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = is-path-cosplit-is-trunc k - ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) + ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is `k+1`-path-cosplit diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index 1afc8d566f..359702f52d 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -232,7 +232,7 @@ is-trunc-map-succ-precomp-Π : ((g h : (b : B) → C b) → is-trunc-map k (precomp-Π f (eq-value g h))) → is-trunc-map (succ-𝕋 k) (precomp-Π f C) is-trunc-map-succ-precomp-Π {k = k} {f = f} {C = C} H = - is-trunc-map-succ-is-trunc-map-ap k (precomp-Π f C) + is-trunc-map-is-trunc-map-ap k (precomp-Π f C) ( λ g h → is-trunc-map-top-is-trunc-map-bottom-is-equiv k ( ap (precomp-Π f C)) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index 209d7620de..680cb9792b 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -96,7 +96,7 @@ module _ ( ap (map-pointed-map f)) ( is-trunc-map-is-equiv k ( is-equiv-tr-type-Ω (preserves-point-pointed-map f))) - ( is-trunc-map-ap-is-trunc-map-succ k + ( is-trunc-map-ap-is-trunc-map k ( map-pointed-map f) ( H) ( point-Pointed-Type A) From e581125da0ac5a92444902dc6bda7de20fd309cc Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 15:55:17 +0100 Subject: [PATCH 57/69] Revert "revert another naming change" This reverts commit 3ed97cc4f33536d90c50427823106b2f9b9c154b. --- src/foundation-core/truncated-maps.lagda.md | 12 ++++++------ src/foundation/diagonals-of-maps.lagda.md | 2 +- src/foundation/faithful-maps.lagda.md | 4 ++-- src/foundation/mere-path-cosplit-maps.lagda.md | 2 +- src/foundation/path-cosplit-maps.lagda.md | 2 +- .../precomposition-dependent-functions.lagda.md | 2 +- .../functoriality-loop-spaces.lagda.md | 2 +- 7 files changed, 13 insertions(+), 13 deletions(-) diff --git a/src/foundation-core/truncated-maps.lagda.md b/src/foundation-core/truncated-maps.lagda.md index 23cc51fa58..720b4c6286 100644 --- a/src/foundation-core/truncated-maps.lagda.md +++ b/src/foundation-core/truncated-maps.lagda.md @@ -114,18 +114,18 @@ module _ where abstract - is-trunc-map-is-trunc-map-ap : + is-trunc-map-succ-is-trunc-map-ap : ((x y : A) → is-trunc-map k (ap f {x} {y})) → is-trunc-map (succ-𝕋 k) f - is-trunc-map-is-trunc-map-ap is-trunc-map-ap-f b (x , p) (x' , p') = + is-trunc-map-succ-is-trunc-map-ap is-trunc-map-ap-f b (x , p) (x' , p') = is-trunc-equiv k ( fiber (ap f) (p ∙ inv p')) ( equiv-fiber-ap-eq-fiber f (x , p) (x' , p')) ( is-trunc-map-ap-f x x' (p ∙ inv p')) abstract - is-trunc-map-ap-is-trunc-map : + is-trunc-map-ap-is-trunc-map-succ : is-trunc-map (succ-𝕋 k) f → (x y : A) → is-trunc-map k (ap f {x} {y}) - is-trunc-map-ap-is-trunc-map is-trunc-map-f x y p = + is-trunc-map-ap-is-trunc-map-succ is-trunc-map-f x y p = is-trunc-is-equiv' k ( (x , p) = (y , refl)) ( eq-fiber-fiber-ap f x y p) @@ -144,7 +144,7 @@ module _ is-trunc-map-ap-is-trunc-map : is-trunc-map k f → (x y : A) → is-trunc-map k (ap f {x} {y}) is-trunc-map-ap-is-trunc-map H = - is-trunc-map-ap-is-trunc-map k f (is-trunc-map-succ-is-trunc-map k H) + is-trunc-map-ap-is-trunc-map-succ k f (is-trunc-map-succ-is-trunc-map k H) ``` ### The domain of any `k`-truncated map into a `k+1`-truncated type is `k+1`-truncated @@ -162,7 +162,7 @@ is-trunc-is-trunc-map-into-is-trunc ( k) ( ap f) ( is-trunc-B (f a) (f a')) - ( is-trunc-map-ap-is-trunc-map k f is-trunc-map-f a a') + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-map-f a a') ``` ### A family of types is a family of `k`-truncated types if and only of the projection map is `k`-truncated diff --git a/src/foundation/diagonals-of-maps.lagda.md b/src/foundation/diagonals-of-maps.lagda.md index a965e62bc4..8454997692 100644 --- a/src/foundation/diagonals-of-maps.lagda.md +++ b/src/foundation/diagonals-of-maps.lagda.md @@ -88,7 +88,7 @@ abstract is-trunc-is-equiv k (fiber (ap f) p) ( fiber-ap-fiber-diagonal-map f (triple x y p)) ( is-equiv-fiber-ap-fiber-diagonal-map f (triple x y p)) - ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y p) + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y p) abstract is-trunc-map-is-trunc-diagonal-map : diff --git a/src/foundation/faithful-maps.lagda.md b/src/foundation/faithful-maps.lagda.md index 0e523f08b3..341c52fb7a 100644 --- a/src/foundation/faithful-maps.lagda.md +++ b/src/foundation/faithful-maps.lagda.md @@ -107,12 +107,12 @@ module _ is-0-map-is-faithful : is-faithful f → is-0-map f is-0-map-is-faithful H = - is-trunc-map-is-trunc-map-ap neg-one-𝕋 f + is-trunc-map-succ-is-trunc-map-ap neg-one-𝕋 f ( λ x y → is-prop-map-is-emb (H x y)) is-faithful-is-0-map : is-0-map f → is-faithful f is-faithful-is-0-map H x y = - is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map neg-one-𝕋 f H x y) + is-emb-is-prop-map (is-trunc-map-ap-is-trunc-map-succ neg-one-𝕋 f H x y) ``` ## Properties diff --git a/src/foundation/mere-path-cosplit-maps.lagda.md b/src/foundation/mere-path-cosplit-maps.lagda.md index 94523eeebc..a32550ef06 100644 --- a/src/foundation/mere-path-cosplit-maps.lagda.md +++ b/src/foundation/mere-path-cosplit-maps.lagda.md @@ -106,7 +106,7 @@ is-mere-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = unit-trunc-Prop (retraction-is-contr-map is-trunc-f) is-mere-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = is-mere-path-cosplit-is-trunc k - ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is merely `k+1`-path-cosplit diff --git a/src/foundation/path-cosplit-maps.lagda.md b/src/foundation/path-cosplit-maps.lagda.md index 9a52358e95..dbb0e02eff 100644 --- a/src/foundation/path-cosplit-maps.lagda.md +++ b/src/foundation/path-cosplit-maps.lagda.md @@ -171,7 +171,7 @@ is-path-cosplit-is-trunc neg-two-𝕋 is-trunc-f = retraction-is-contr-map is-trunc-f is-path-cosplit-is-trunc (succ-𝕋 k) {f = f} is-trunc-f x y = is-path-cosplit-is-trunc k - ( is-trunc-map-ap-is-trunc-map k f is-trunc-f x y) + ( is-trunc-map-ap-is-trunc-map-succ k f is-trunc-f x y) ``` ### If a map is `k`-path-cosplit then it is `k+1`-path-cosplit diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index 359702f52d..1afc8d566f 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -232,7 +232,7 @@ is-trunc-map-succ-precomp-Π : ((g h : (b : B) → C b) → is-trunc-map k (precomp-Π f (eq-value g h))) → is-trunc-map (succ-𝕋 k) (precomp-Π f C) is-trunc-map-succ-precomp-Π {k = k} {f = f} {C = C} H = - is-trunc-map-is-trunc-map-ap k (precomp-Π f C) + is-trunc-map-succ-is-trunc-map-ap k (precomp-Π f C) ( λ g h → is-trunc-map-top-is-trunc-map-bottom-is-equiv k ( ap (precomp-Π f C)) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index 680cb9792b..209d7620de 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -96,7 +96,7 @@ module _ ( ap (map-pointed-map f)) ( is-trunc-map-is-equiv k ( is-equiv-tr-type-Ω (preserves-point-pointed-map f))) - ( is-trunc-map-ap-is-trunc-map k + ( is-trunc-map-ap-is-trunc-map-succ k ( map-pointed-map f) ( H) ( point-Pointed-Type A) From 15517ad59d948b33a576e8b3cce7d3b14840c693 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 16:17:43 +0100 Subject: [PATCH 58/69] fixes --- ...s-with-respect-to-truncated-types.lagda.md | 11 +++++----- .../localizations-at-maps.lagda.md | 4 ++-- .../reflective-subuniverses.lagda.md | 5 +++-- .../subuniverse-connected-maps.lagda.md | 20 +++++++++---------- ...morphisms-polynomial-endofunctors.lagda.md | 3 +-- 5 files changed, 21 insertions(+), 22 deletions(-) diff --git a/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md b/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md index e5ff5b16a5..bec1d7a0cf 100644 --- a/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md +++ b/src/foundation/epimorphisms-with-respect-to-truncated-types.lagda.md @@ -187,7 +187,7 @@ module _ is-equiv-diagonal-into-fibers-precomp-is-epimorphism-Truncated-Type : is-epimorphism-Truncated-Type k f → {l : Level} (X : Truncated-Type l k) → - is-equiv (diagonal-into-fibers-precomp f (type-Truncated-Type X)) + is-equiv (diagonal-into-fibers-precomp f) is-equiv-diagonal-into-fibers-precomp-is-epimorphism-Truncated-Type e X = is-equiv-map-section-family ( λ g → g , refl) @@ -202,11 +202,10 @@ module _ is-equiv (diagonal-into-cocone f (type-Truncated-Type X)) is-equiv-diagonal-into-cocone-is-epimorphism-Truncated-Type e X = is-equiv-comp - ( map-equiv (compute-total-fiber-precomp f (type-Truncated-Type X))) - ( diagonal-into-fibers-precomp f (type-Truncated-Type X)) + ( map-equiv (compute-total-fiber-precomp f)) + ( diagonal-into-fibers-precomp f) ( is-equiv-diagonal-into-fibers-precomp-is-epimorphism-Truncated-Type e X) - ( is-equiv-map-equiv - ( compute-total-fiber-precomp f (type-Truncated-Type X))) + ( is-equiv-map-equiv (compute-total-fiber-precomp f)) is-equiv-horizontal-map-cocone-is-epimorphism-Truncated-Type : is-epimorphism-Truncated-Type k f → @@ -241,7 +240,7 @@ module _ is-contr-equiv ( Σ ( B → (type-Truncated-Type X)) ( λ h → coherence-square-maps f f h g)) - ( compute-fiber-precomp f (type-Truncated-Type X) g) + ( compute-fiber-precomp f g) ( is-contr-is-equiv-pr1 (h X) g)) is-epimorphism-is-equiv-vertical-map-cocone-Truncated-Type : diff --git a/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md b/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md index 1144d685d4..f7acd2f484 100644 --- a/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md @@ -81,8 +81,8 @@ module _ pr1 is-localization-Y pr1 (pr2 (is-subuniverse-localization-is-localization is-localization-Y)) = η pr2 (pr2 (is-subuniverse-localization-is-localization is-localization-Y)) - Z is-local-Z = - pr2 is-localization-Y Z is-local-Z + (Z , is-local-Z) = + pr2 is-localization-Y Z is-local-Z ``` It remains to construct a converse. diff --git a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md index f66155010c..f00277ed1d 100644 --- a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md @@ -186,8 +186,9 @@ module _ is-in-subuniverse-subuniverse-localization 𝒫 (L A) pr2 (pr2 (pr2 is-reflective-has-all-localizations-subuniverse)) A B is-in-subuniverse-A = - is-subuniverse-equiv-unit-subuniverse-localization - 𝒫 (L B) A is-in-subuniverse-A + is-subuniverse-equiv-unit-subuniverse-localization 𝒫 + ( L B) + ( A , is-in-subuniverse-A) ``` ## Recursion for reflective subuniverses diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index 93906ac736..a057e3a354 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -495,30 +495,30 @@ module _ ( refl-htpy-subuniverse-connected-map f) ( is-subuniverse-connected-map-subuniverse-connected-map K f) - htpy-eq-connected-map : + htpy-eq-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → f = g → htpy-subuniverse-connected-map f g - htpy-eq-connected-map f .f refl = refl-htpy-subuniverse-connected-map f + htpy-eq-subuniverse-connected-map f .f refl = refl-htpy-subuniverse-connected-map f - is-equiv-htpy-eq-connected-map : + is-equiv-htpy-eq-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → - is-equiv (htpy-eq-connected-map f g) - is-equiv-htpy-eq-connected-map f = + is-equiv (htpy-eq-subuniverse-connected-map f g) + is-equiv-htpy-eq-subuniverse-connected-map f = fundamental-theorem-id ( is-torsorial-htpy-subuniverse-connected-map f) - ( htpy-eq-connected-map f) + ( htpy-eq-subuniverse-connected-map f) - extensionality-connected-map : + extensionality-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → (f = g) ≃ htpy-subuniverse-connected-map f g - pr1 (extensionality-connected-map f g) = htpy-eq-connected-map f g - pr2 (extensionality-connected-map f g) = is-equiv-htpy-eq-connected-map f g + pr1 (extensionality-subuniverse-connected-map f g) = htpy-eq-subuniverse-connected-map f g + pr2 (extensionality-subuniverse-connected-map f g) = is-equiv-htpy-eq-subuniverse-connected-map f g eq-htpy-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → htpy-subuniverse-connected-map f g → (f = g) eq-htpy-subuniverse-connected-map f g = - map-inv-equiv (extensionality-connected-map f g) + map-inv-equiv (extensionality-subuniverse-connected-map f g) ``` ### All maps are `Contr`-connected diff --git a/src/trees/morphisms-polynomial-endofunctors.lagda.md b/src/trees/morphisms-polynomial-endofunctors.lagda.md index a30bed6f20..85cd8ee091 100644 --- a/src/trees/morphisms-polynomial-endofunctors.lagda.md +++ b/src/trees/morphisms-polynomial-endofunctors.lagda.md @@ -350,9 +350,8 @@ module _ by equiv-tot ( λ (a , p) → - compute-coherence-triangle-fiber-precomp' + compute-extension-fiber-precomp' ( α₁ a) - ( X) ( inv-tr (λ c' → Q₁ c' → X) p x)) ≃ Σ ( fiber α₀ c) ( λ (a , p) → From e9da00c5292e75e9ac1ad48cbefa5eaaaab0680d Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 18:37:41 +0100 Subject: [PATCH 59/69] remove relative separation work --- src/foundation/surjective-maps.lagda.md | 59 +++---- src/orthogonal-factorization-systems.lagda.md | 3 - .../extensions-maps.lagda.md | 119 +++++--------- .../postcomposition-extensions-maps.lagda.md | 150 ------------------ ...rations-types-global-subuniverses.lagda.md | 90 ----------- ...ve-separations-types-subuniverses.lagda.md | 55 ------- .../subuniverse-connected-maps.lagda.md | 9 +- 7 files changed, 68 insertions(+), 417 deletions(-) delete mode 100644 src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md delete mode 100644 src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md delete mode 100644 src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md index 6619b55f87..0f9dde78f7 100644 --- a/src/foundation/surjective-maps.lagda.md +++ b/src/foundation/surjective-maps.lagda.md @@ -22,7 +22,6 @@ open import foundation.fundamental-theorem-of-identity-types open import foundation.homotopy-induction open import foundation.identity-types open import foundation.inhabited-types -open import foundation.injective-maps open import foundation.postcomposition-dependent-functions open import foundation.propositional-truncations open import foundation.split-surjective-maps @@ -54,7 +53,6 @@ open import foundation-core.truncated-maps open import foundation-core.truncation-levels open import orthogonal-factorization-systems.extensions-maps -open import orthogonal-factorization-systems.postcomposition-extensions-maps ``` @@ -822,25 +820,25 @@ module _ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} where - is-split-surjective-postcomp-extension-surjective-map : + is-surjective-postcomp-extension-surjective-map : (f : A → B) (i : A → X) (g : X → Y) → - is-surjective f → - is-emb g → - is-split-surjective (postcomp-extension f i g) - is-split-surjective-postcomp-extension-surjective-map f i g H K (h , L) = - ( ( j , N) , - ( eq-htpy-extension f - ( g ∘ i) - ( postcomp-extension f i g (j , N)) - ( h , L) - ( M) - ( λ a → - ( ap - ( concat' (g (i a)) (M (f a))) - ( is-section-map-inv-is-equiv - ( K (i a) (j (f a))) - ( L a ∙ inv (M (f a))))) ∙ - ( is-section-inv-concat' (M (f a)) (L a))))) + is-surjective f → is-emb g → + is-surjective (postcomp-extension f i g) + is-surjective-postcomp-extension-surjective-map f i g H K (h , L) = + unit-trunc-Prop + ( ( j , N) , + ( eq-htpy-extension f + ( g ∘ i) + ( postcomp-extension f i g (j , N)) + ( h , L) + ( M) + ( λ a → + ( ap + ( concat' (g (i a)) (M (f a))) + ( is-section-map-inv-is-equiv + ( K (i a) (j (f a))) + ( L a ∙ inv (M (f a))))) ∙ + ( is-section-inv-concat' (M (f a)) (L a))))) where J : (b : B) → fiber g (h b) @@ -852,29 +850,20 @@ module _ j : B → X j b = pr1 (J b) - M : g ∘ j ~ h + M : (g ∘ j) ~ h M b = pr2 (J b) - N : i ~ j ∘ f - N a = is-injective-is-emb K (L a ∙ inv (M (f a))) - - is-surjective-postcomp-extension-surjective-map : - (f : A → B) (i : A → X) (g : X → Y) → - is-surjective f → is-emb g → - is-surjective (postcomp-extension f i g) - is-surjective-postcomp-extension-surjective-map f i g H K = - is-surjective-is-split-surjective - ( is-split-surjective-postcomp-extension-surjective-map f i g H K) + N : i ~ (j ∘ f) + N a = map-inv-is-equiv (K (i a) (j (f a))) (L a ∙ inv (M (f a))) is-equiv-postcomp-extension-is-surjective : (f : A → B) (i : A → X) (g : X → Y) → is-surjective f → is-emb g → is-equiv (postcomp-extension f i g) is-equiv-postcomp-extension-is-surjective f i g H K = - is-equiv-is-split-surjective-is-injective - ( postcomp-extension f i g) - ( is-injective-is-emb (is-emb-postcomp-extension f i g K)) - ( is-split-surjective-postcomp-extension-surjective-map f i g H K) + is-equiv-is-emb-is-surjective + ( is-surjective-postcomp-extension-surjective-map f i g H K) + ( is-emb-postcomp-extension f i g K) equiv-postcomp-extension-surjection : (f : A ↠ B) (i : A → X) (g : X ↪ Y) → diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index 21205c2716..a5d4e14027 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -56,7 +56,6 @@ open import orthogonal-factorization-systems.null-types public open import orthogonal-factorization-systems.open-modalities public open import orthogonal-factorization-systems.orthogonal-factorization-systems public open import orthogonal-factorization-systems.orthogonal-maps public -open import orthogonal-factorization-systems.postcomposition-extensions-maps public open import orthogonal-factorization-systems.precomposition-lifts-families-of-elements public open import orthogonal-factorization-systems.pullback-hom public open import orthogonal-factorization-systems.raise-modalities public @@ -64,8 +63,6 @@ open import orthogonal-factorization-systems.reflective-global-subuniverses publ open import orthogonal-factorization-systems.reflective-modalities public open import orthogonal-factorization-systems.reflective-subuniverses public open import orthogonal-factorization-systems.regular-cd-structures public -open import orthogonal-factorization-systems.relative-separations-types-global-subuniverses public -open import orthogonal-factorization-systems.relative-separations-types-subuniverses public open import orthogonal-factorization-systems.sigma-closed-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-subuniverses public diff --git a/src/orthogonal-factorization-systems/extensions-maps.lagda.md b/src/orthogonal-factorization-systems/extensions-maps.lagda.md index d5d16e3066..ccdc296bd0 100644 --- a/src/orthogonal-factorization-systems/extensions-maps.lagda.md +++ b/src/orthogonal-factorization-systems/extensions-maps.lagda.md @@ -16,24 +16,19 @@ open import foundation.equivalences open import foundation.function-types open import foundation.functoriality-dependent-function-types open import foundation.functoriality-dependent-pair-types -open import foundation.functoriality-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.monomorphisms open import foundation.postcomposition-functions -open import foundation.precomposition-dependent-functions -open import foundation.precomposition-functions open import foundation.propositions open import foundation.sets open import foundation.structure-identity-principle open import foundation.transport-along-identifications -open import foundation.truncated-maps open import foundation.truncated-types open import foundation.truncation-levels open import foundation.type-arithmetic-dependent-pair-types -open import foundation.universal-property-family-of-fibers-of-maps open import foundation.universe-levels open import foundation.whiskering-homotopies-composition @@ -216,6 +211,20 @@ module _ is-extension-right-whisker F h = F ∘ h ``` +### Postcomposition of extensions + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → + extension f i → extension f (g ∘ i) + postcomp-extension f i g = + map-Σ (is-extension f (g ∘ i)) (postcomp B g) (λ j H → g ·l H) +``` + ## Properties ### Characterizing identifications of extensions of maps @@ -275,78 +284,6 @@ module _ map-inv-equiv (extensionality-extension e e') (H , K) ``` -### Computing extension types as a dependent product - -Extension types are equivalent to fibers of precomposition maps, which in turn -have a Π-type characterization. Given `i : A → B` and `g : B → C`, then - -```text - extension i g ≃ ((y : B) → Σ (c : C), (x : X) → (i x = y) → (c = g y)). -``` - -We give 4 different formulations of this equivalence for both the nondependent -and dependent extension type. - -```agda -module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (i : A → B) - {C : UU l3} (g : A → C) - where - - equiv-fiber-Π-curry-precomp-extension : - extension i g ≃ ((b : B) → Σ C (λ u → (a : A) → i a = b → g a = u)) - equiv-fiber-Π-curry-precomp-extension = - ( compute-fiber-Π-curry-precomp i g) ∘e - ( inv-equiv (compute-extension-fiber-precomp i g)) - - equiv-fiber-Π-curry-precomp-extension' : - extension i g ≃ ((b : B) → Σ C (λ u → (a : A) → b = i a → u = g a)) - equiv-fiber-Π-curry-precomp-extension' = - ( compute-fiber-Π-curry-precomp' i g) ∘e - ( inv-equiv (compute-extension-fiber-precomp i g)) - - equiv-fiber-Π-precomp-extension : - extension i g ≃ fiber-Π-precomp i g - equiv-fiber-Π-precomp-extension = - ( compute-fiber-Π-precomp i g) ∘e - ( inv-equiv (compute-extension-fiber-precomp i g)) - - equiv-fiber-Π-precomp-extension' : - extension i g ≃ fiber-Π-precomp' i g - equiv-fiber-Π-precomp-extension' = - ( compute-fiber-Π-precomp' i g) ∘e - ( inv-equiv (compute-extension-fiber-precomp i g)) - -module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (i : A → B) - {P : B → UU l3} (g : (x : A) → P (i x)) - where - - equiv-fiber-Π-curry-precomp-Π-extension-dependent-type : - extension-dependent-type i P g ≃ fiber-Π-curry-precomp-Π i P g - equiv-fiber-Π-curry-precomp-Π-extension-dependent-type = - ( compute-fiber-Π-curry-precomp-Π i P g) ∘e - ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) - - equiv-fiber-Π-curry-precomp-Π-extension-dependent-type' : - extension-dependent-type i P g ≃ fiber-Π-curry-precomp-Π' i P g - equiv-fiber-Π-curry-precomp-Π-extension-dependent-type' = - ( compute-fiber-Π-curry-precomp-Π' i P g) ∘e - ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) - - equiv-fiber-Π-precomp-Π-extension-dependent-type : - extension-dependent-type i P g ≃ fiber-Π-precomp-Π i P g - equiv-fiber-Π-precomp-Π-extension-dependent-type = - ( compute-fiber-Π-precomp-Π i P g) ∘e - ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) - - equiv-fiber-Π-precomp-Π-extension-dependent-type' : - extension-dependent-type i P g ≃ fiber-Π-precomp-Π' i P g - equiv-fiber-Π-precomp-Π-extension-dependent-type' = - ( compute-fiber-Π-precomp-Π' i P g) ∘e - ( inv-equiv (compute-extension-fiber-precomp-Π i P g)) -``` - ### The total type of extensions is equivalent to `(y : B) → P y` ```agda @@ -386,7 +323,7 @@ module _ ((x : A) → is-trunc (succ-𝕋 k) (P (i x))) → (g : (x : B) → P x) → is-trunc k (is-extension i f g) is-trunc-is-extension-dependent-type f is-trunc-P g = - is-trunc-Π k (λ x → is-trunc-P x (f x) (g (i x))) + is-trunc-Π k λ x → is-trunc-P x (f x) (g (i x)) is-trunc-extension-dependent-type : {P : B → UU l3} (f : (x : A) → P (i x)) → @@ -416,7 +353,7 @@ module _ ((x : A) → is-prop (P (i x))) → (g : (x : B) → P x) → is-contr (is-extension i f g) is-contr-is-extension f is-prop-P g = - is-contr-Π (λ x → is-prop-P x (f x) (g (i x))) + is-contr-Π λ x → is-prop-P x (f x) (g (i x)) is-prop-is-extension : {P : B → UU l3} (f : (x : A) → P (i x)) → @@ -439,7 +376,7 @@ module _ is-extension-self = refl-htpy extension-self : extension-dependent-type id P f - extension-self = (f , is-extension-self) + extension-self = f , is-extension-self ``` ### The identity is an extension of every map along themselves @@ -453,7 +390,27 @@ module _ is-extension-along-self = refl-htpy extension-along-self : extension f f - extension-along-self = (id , is-extension-along-self) + extension-along-self = id , is-extension-along-self +``` + +### Postcomposition of extensions by an embedding is an embedding + +```agda +module _ + {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} + where + + is-emb-postcomp-extension : + (f : A → B) (i : A → X) (g : X → Y) → is-emb g → + is-emb (postcomp-extension f i g) + is-emb-postcomp-extension f i g H = + is-emb-map-Σ + ( is-extension f (g ∘ i)) + ( is-mono-is-emb g H B) + ( λ j → + is-emb-is-equiv + ( is-equiv-map-Π-is-fiberwise-equiv + ( λ x → H (i x) (j (f x))))) ``` ## See also diff --git a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md b/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md deleted file mode 100644 index 828dfc2c7b..0000000000 --- a/src/orthogonal-factorization-systems/postcomposition-extensions-maps.lagda.md +++ /dev/null @@ -1,150 +0,0 @@ -# Postcomposition of extensions of maps - -```agda -module orthogonal-factorization-systems.postcomposition-extensions-maps where -``` - -
Imports - -```agda -open import foundation.action-on-identifications-dependent-functions -open import foundation.action-on-identifications-functions -open import foundation.contractible-types -open import foundation.dependent-pair-types -open import foundation.embeddings -open import foundation.equivalences -open import foundation.function-types -open import foundation.functoriality-dependent-function-types -open import foundation.functoriality-dependent-pair-types -open import foundation.functoriality-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.monomorphisms -open import foundation.postcomposition-functions -open import foundation.propositions -open import foundation.sets -open import foundation.structure-identity-principle -open import foundation.transport-along-identifications -open import foundation.truncated-maps -open import foundation.truncated-types -open import foundation.truncation-levels -open import foundation.type-arithmetic-dependent-pair-types -open import foundation.universe-levels -open import foundation.whiskering-homotopies-composition - -open import foundation-core.torsorial-type-families - -open import orthogonal-factorization-systems.extensions-maps -``` - -
- -## Idea - -Given a map `i : A → B` and a map `f : A → X`, then we may -{{#concept "postcompose" Disambiguation="extension of map by map" Agda=postcomp-extension}} -any [extension](orthogonal-factorization-systems.extensions-maps.md) -`α : extension i f` by `g` to obtain an extension `g ∘ α : extension i (g ∘ f)`. - -## Definition - -### Postcomposition of extensions - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - postcomp-extension : - (i : A → B) (f : A → X) (g : X → Y) → - extension i f → extension i (g ∘ f) - postcomp-extension i f g = - map-Σ (is-extension i (g ∘ f)) (postcomp B g) (λ j H → g ·l H) -``` - -## Properties - -### Postcomposition of extensions by an equivalence is an equivalence - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-equiv-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-equiv g → - is-equiv (postcomp-extension f i g) - is-equiv-postcomp-extension f i g G = - is-equiv-map-Σ - ( is-extension f (g ∘ i)) - ( is-equiv-postcomp-is-equiv g G B) - ( λ j → - is-equiv-map-Π-is-fiberwise-equiv - ( λ x → is-emb-is-equiv G (i x) (j (f x)))) - - equiv-postcomp-extension : - (f : A → B) (i : A → X) (g : X ≃ Y) → - extension f i ≃ extension f (map-equiv g ∘ i) - equiv-postcomp-extension f i (g , G) = - ( postcomp-extension f i g , is-equiv-postcomp-extension f i g G) -``` - -### Postcomposition of extensions by an embedding is an embedding - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-emb-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-emb g → - is-emb (postcomp-extension f i g) - is-emb-postcomp-extension f i g H = - is-emb-map-Σ - ( is-extension f (g ∘ i)) - ( is-mono-is-emb g H B) - ( λ j → - is-emb-is-equiv - ( is-equiv-map-Π-is-fiberwise-equiv - ( λ x → H (i x) (j (f x))))) - - emb-postcomp-extension : - (f : A → B) (i : A → X) (g : X ↪ Y) → - extension f i ↪ extension f (map-emb g ∘ i) - emb-postcomp-extension f i (g , G) = - postcomp-extension f i g , is-emb-postcomp-extension f i g G -``` - -### Postcomposition of extensions by a `k`-truncated map is `k`-truncated - -```agda -module _ - {l1 l2 l3 l4 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4} - where - - is-trunc-map-postcomp-extension : - (f : A → B) (i : A → X) (g : X → Y) → is-trunc-map k g → - is-trunc-map k (postcomp-extension f i g) - is-trunc-map-postcomp-extension f i g G = - is-trunc-map-map-Σ k - ( is-extension f (g ∘ i)) - ( is-trunc-map-postcomp-is-trunc-map k g G B) - ( λ j → - is-trunc-map-map-Π k - ( λ a → ap g) - ( λ a → is-trunc-map-ap-is-trunc-map k g G (i a) (j (f a)))) - - trunc-map-postcomp-extension : - (f : A → B) (i : A → X) (g : trunc-map k X Y) → - trunc-map k (extension f i) (extension f (map-trunc-map g ∘ i)) - trunc-map-postcomp-extension f i (g , G) = - ( postcomp-extension f i g , is-trunc-map-postcomp-extension f i g G) -``` - -## See also - -- In [`foundation.surjective-maps`](foundation.surjective-maps.md) it is shown - that postcomposition of extensions along surjective maps by an embedding is an - equivalence. diff --git a/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md deleted file mode 100644 index b68f408e34..0000000000 --- a/src/orthogonal-factorization-systems/relative-separations-types-global-subuniverses.lagda.md +++ /dev/null @@ -1,90 +0,0 @@ -# Relative separations of types at global subuniverses - -```agda -module orthogonal-factorization-systems.relative-separations-types-global-subuniverses where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.equivalences -open import foundation.global-subuniverses -open import foundation.universe-levels - -open import foundation-core.function-types -open import foundation-core.identity-types -open import foundation-core.propositions - -open import orthogonal-factorization-systems.extensions-maps -open import orthogonal-factorization-systems.postcomposition-extensions-maps -``` - -
- -## Idea - -Consider a [global subuniverse](foundation.global-subuniverses.md) `K` and a map -`δ : X → Y`. A type `A` is said to be -{{#concept "`K`-separated relative to `δ`" Agda=is-rel-separated}} if the type -of extensions of any map `f : X → A` along `δ` is in `K`. - -As a special case, a type `A` is _`K`-separated_ if it is `K`-separated relative -to the terminal projection at a two-element type. - -## Definitions - -### The predicate of being separated - -```agda -module _ - {α : Level → Level} {l1 l2 l3 : Level} (K : global-subuniverse α) - {X : UU l1} {Y : UU l2} (δ : X → Y) - where - - is-rel-separated : (A : UU l3) → UU (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - is-rel-separated A = - (f : X → A) → is-in-global-subuniverse K (extension δ f) - - is-prop-is-rel-separated : - (A : UU l3) → is-prop (is-rel-separated A) - is-prop-is-rel-separated A = - is-prop-Π (λ f → is-prop-is-in-global-subuniverse K (extension δ f)) - - is-rel-separated-Prop : - (A : UU l3) → Prop (α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - is-rel-separated-Prop A = - ( is-rel-separated A , is-prop-is-rel-separated A) -``` - -### The global subuniverse of `K`-separated types relative to `δ` - -```agda -module _ - {α : Level → Level} {l1 l2 : Level} (K : global-subuniverse α) - {X : UU l1} {Y : UU l2} (δ : X → Y) - where - - is-closed-under-equiv-rel-separated-global-subuniverse : - {l4 l5 : Level} → - is-closed-under-equiv-subuniverses - ( λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - ( λ l3 → is-rel-separated-Prop {l3 = l3} K δ) - ( l4) - ( l5) - is-closed-under-equiv-rel-separated-global-subuniverse A B e H f = - is-closed-under-equiv-global-subuniverse K - ( extension δ (map-inv-equiv e ∘ f)) - ( extension δ f) - ( inv-equiv (equiv-postcomp-extension δ f (inv-equiv e))) - ( H (map-inv-equiv e ∘ f)) - - rel-separated-global-subuniverse : - global-subuniverse (λ l3 → α (l1 ⊔ l2 ⊔ l3) ⊔ l1 ⊔ l3) - rel-separated-global-subuniverse = - λ where - .subuniverse-global-subuniverse l3 → - is-rel-separated-Prop K δ - .is-closed-under-equiv-global-subuniverse → - is-closed-under-equiv-rel-separated-global-subuniverse -``` diff --git a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md b/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md deleted file mode 100644 index 56d9cee15d..0000000000 --- a/src/orthogonal-factorization-systems/relative-separations-types-subuniverses.lagda.md +++ /dev/null @@ -1,55 +0,0 @@ -# Relative separations of types at subuniverses - -```agda -module orthogonal-factorization-systems.relative-separations-types-subuniverses where -``` - -
Imports - -```agda -open import foundation.dependent-pair-types -open import foundation.subuniverses -open import foundation.universe-levels - -open import foundation-core.identity-types -open import foundation-core.propositions - -open import orthogonal-factorization-systems.extensions-maps -``` - -
- -## Idea - -Consider a [subuniverse](foundation.subuniverses.md) `K` and a map `δ : X → Y`. -A type `A` is said to be -{{#concept "`K`-separated relative to `δ`" Agda=is-rel-separated'}} if the type -of extensions of any map `f : X → A` along `δ` is in `K`. - -## Definitions - -### The predicate of being rel-separated - -```agda -module _ - {l1 l2 l3 : Level} (K : subuniverse l1 l2) - {X : UU l1} {Y : UU l1} (δ : X → Y) - where - - is-rel-separated' : (A : UU l1) → UU (l1 ⊔ l2) - is-rel-separated' A = (f : X → A) → is-in-subuniverse K (extension δ f) - - is-prop-is-rel-separated' : - (A : UU l1) → is-prop (is-rel-separated' A) - is-prop-is-rel-separated' A = - is-prop-Π (λ f → is-prop-is-in-subuniverse K (extension δ f)) - - is-rel-separated'-Prop : - (A : UU l1) → Prop (l1 ⊔ l2) - is-rel-separated'-Prop A = - ( is-rel-separated' A , is-prop-is-rel-separated' A) -``` - -## See also - -- [Relative separations of types at global subuniverses](orthogonal-factorization-systems.relative-separations-types-global-subuniverses.md) diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md index a057e3a354..02a4e9d74c 100644 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md @@ -498,7 +498,8 @@ module _ htpy-eq-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → f = g → htpy-subuniverse-connected-map f g - htpy-eq-subuniverse-connected-map f .f refl = refl-htpy-subuniverse-connected-map f + htpy-eq-subuniverse-connected-map f .f refl = + refl-htpy-subuniverse-connected-map f is-equiv-htpy-eq-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → @@ -511,8 +512,10 @@ module _ extensionality-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → (f = g) ≃ htpy-subuniverse-connected-map f g - pr1 (extensionality-subuniverse-connected-map f g) = htpy-eq-subuniverse-connected-map f g - pr2 (extensionality-subuniverse-connected-map f g) = is-equiv-htpy-eq-subuniverse-connected-map f g + pr1 (extensionality-subuniverse-connected-map f g) = + htpy-eq-subuniverse-connected-map f g + pr2 (extensionality-subuniverse-connected-map f g) = + is-equiv-htpy-eq-subuniverse-connected-map f g eq-htpy-subuniverse-connected-map : (f g : subuniverse-connected-map K A B) → From 7ed317c29c1118b49c7d0e4d4d4a43480d7a9320 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 18:37:44 +0100 Subject: [PATCH 60/69] a header --- .../subuniverse-equivalences.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md index 3ac2ccf951..340ebe99e9 100644 --- a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md +++ b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md @@ -1,4 +1,4 @@ -# `K`-Equivalences +# Equivalences with respect to a subuniverse ```agda module orthogonal-factorization-systems.subuniverse-equivalences where From 75d66ea03a7971ac6d86d7f64522ea586e048b04 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 18:42:23 +0100 Subject: [PATCH 61/69] remove subuniverse equivalence work --- src/orthogonal-factorization-systems.lagda.md | 2 - .../localizations-at-maps.lagda.md | 4 +- .../localizations-at-subuniverses.lagda.md | 41 +- .../reflective-subuniverses.lagda.md | 7 +- .../subuniverse-connected-maps.lagda.md | 906 ------------------ .../subuniverse-equivalences.lagda.md | 528 ---------- .../types-local-at-maps.lagda.md | 35 +- .../functoriality-loop-spaces.lagda.md | 2 +- 8 files changed, 32 insertions(+), 1493 deletions(-) delete mode 100644 src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md delete mode 100644 src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md diff --git a/src/orthogonal-factorization-systems.lagda.md b/src/orthogonal-factorization-systems.lagda.md index a5d4e14027..e2a8bf7286 100644 --- a/src/orthogonal-factorization-systems.lagda.md +++ b/src/orthogonal-factorization-systems.lagda.md @@ -67,8 +67,6 @@ open import orthogonal-factorization-systems.sigma-closed-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-modalities public open import orthogonal-factorization-systems.sigma-closed-reflective-subuniverses public open import orthogonal-factorization-systems.stable-orthogonal-factorization-systems public -open import orthogonal-factorization-systems.subuniverse-connected-maps public -open import orthogonal-factorization-systems.subuniverse-equivalences public open import orthogonal-factorization-systems.types-colocal-at-maps public open import orthogonal-factorization-systems.types-local-at-maps public open import orthogonal-factorization-systems.types-separated-at-maps public diff --git a/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md b/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md index f7acd2f484..1144d685d4 100644 --- a/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/localizations-at-maps.lagda.md @@ -81,8 +81,8 @@ module _ pr1 is-localization-Y pr1 (pr2 (is-subuniverse-localization-is-localization is-localization-Y)) = η pr2 (pr2 (is-subuniverse-localization-is-localization is-localization-Y)) - (Z , is-local-Z) = - pr2 is-localization-Y Z is-local-Z + Z is-local-Z = + pr2 is-localization-Y Z is-local-Z ``` It remains to construct a converse. diff --git a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md index 97f0b6d507..c12ced2306 100644 --- a/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/localizations-at-subuniverses.lagda.md @@ -12,7 +12,6 @@ open import foundation.dependent-pair-types open import foundation.subuniverses open import foundation.universe-levels -open import orthogonal-factorization-systems.subuniverse-equivalences open import orthogonal-factorization-systems.types-local-at-maps ``` @@ -21,30 +20,19 @@ open import orthogonal-factorization-systems.types-local-at-maps ## Idea Let `P` be a [subuniverse](foundation.subuniverses.md). Given a type `X`, its -{{#concept "localization" Disambiguation="of a type at a subuniverse" Agda=subuniverse-localization}} -at `P`, or **`P`-localization**, is a type `Y` in `P` and a `P`-equivalence -`η : X → Y`, i.e., a map such that for every `Z` in `P` the -[precomposition map](foundation-core.function-types.md) +**localization** at `P`, or **`P`-localization**, is a type `Y` in `P` and a map +`η : X → Y` such that every type in `P` is +`η`[-local](orthogonal-factorization-systems.types-local-at-maps.md). I.e., for +every `Z` in `P`, the [precomposition map](foundation-core.function-types.md) ```text - ∘ η : (Y → Z) → (X → Z) ``` -is an [equivalence](foundation-core.equivalences.md). In other words, every type -in `P` is `η`[-local](orthogonal-factorization-systems.types-local-at-maps.md). +is an [equivalence](foundation-core.equivalences.md). ## Definition -### The predicate on a map of being a localization at a subuniverse - -```agda -is-subuniverse-localization-map : - {l1 l2 lP : Level} (P : subuniverse l1 lP) {A : UU l2} {PA : UU l1} - (η : A → PA) → UU (lsuc l1 ⊔ l2 ⊔ lP) -is-subuniverse-localization-map P {A} {PA} η = - is-in-subuniverse P PA × is-subuniverse-equiv P η -``` - ### The predicate of being a localization at a subuniverse ```agda @@ -52,7 +40,9 @@ is-subuniverse-localization : {l1 l2 lP : Level} (P : subuniverse l1 lP) → UU l2 → UU l1 → UU (lsuc l1 ⊔ l2 ⊔ lP) is-subuniverse-localization {l1} {l2} P X Y = - (is-in-subuniverse P Y) × (subuniverse-equiv P X Y) + ( is-in-subuniverse P Y) × + ( Σ ( X → Y) + ( λ η → (Z : UU l1) → is-in-subuniverse P Z → is-local η Z)) ``` ```agda @@ -68,7 +58,8 @@ module _ unit-is-subuniverse-localization = pr1 (pr2 is-localization-Y) is-local-at-unit-is-in-subuniverse-is-subuniverse-localization : - (Z : type-subuniverse P) → is-local unit-is-subuniverse-localization (pr1 Z) + (Z : UU l1) → is-in-subuniverse P Z → + is-local unit-is-subuniverse-localization Z is-local-at-unit-is-in-subuniverse-is-subuniverse-localization = pr2 (pr2 is-localization-Y) ``` @@ -100,19 +91,15 @@ module _ is-in-subuniverse-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) - type-subuniverse-subuniverse-localization : type-subuniverse P - type-subuniverse-subuniverse-localization = - ( type-subuniverse-localization , - is-in-subuniverse-subuniverse-localization) - unit-subuniverse-localization : X → type-subuniverse-localization unit-subuniverse-localization = unit-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) - is-subuniverse-equiv-unit-subuniverse-localization : - is-subuniverse-equiv P unit-subuniverse-localization - is-subuniverse-equiv-unit-subuniverse-localization = + is-local-at-unit-is-in-subuniverse-subuniverse-localization : + (Z : UU l1) → + is-in-subuniverse P Z → is-local unit-subuniverse-localization Z + is-local-at-unit-is-in-subuniverse-subuniverse-localization = is-local-at-unit-is-in-subuniverse-is-subuniverse-localization P ( is-subuniverse-localization-subuniverse-localization) ``` diff --git a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md index f00277ed1d..41a30148b0 100644 --- a/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md +++ b/src/orthogonal-factorization-systems/reflective-subuniverses.lagda.md @@ -167,7 +167,7 @@ module _ pr1 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) = unit-is-reflective-subuniverse 𝒫 is-reflective-𝒫 pr2 (pr2 (pr2 (has-all-localizations-is-reflective-subuniverse A))) - ( X , is-in-subuniverse-X) = + X is-in-subuniverse-X = is-local-is-in-subuniverse-is-reflective-subuniverse 𝒫 is-reflective-𝒫 X A is-in-subuniverse-X @@ -186,9 +186,8 @@ module _ is-in-subuniverse-subuniverse-localization 𝒫 (L A) pr2 (pr2 (pr2 is-reflective-has-all-localizations-subuniverse)) A B is-in-subuniverse-A = - is-subuniverse-equiv-unit-subuniverse-localization 𝒫 - ( L B) - ( A , is-in-subuniverse-A) + is-local-at-unit-is-in-subuniverse-subuniverse-localization + 𝒫 (L B) A is-in-subuniverse-A ``` ## Recursion for reflective subuniverses diff --git a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md b/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md deleted file mode 100644 index 02a4e9d74c..0000000000 --- a/src/orthogonal-factorization-systems/subuniverse-connected-maps.lagda.md +++ /dev/null @@ -1,906 +0,0 @@ -# Connected maps with respect to a subuniverse - -```agda -module orthogonal-factorization-systems.subuniverse-connected-maps where -``` - -
Imports - -```agda -open import foundation.connected-types -open import foundation.constant-maps -open import foundation.contractible-types -open import foundation.coproduct-types -open import foundation.dependent-identifications -open import foundation.dependent-pair-types -open import foundation.dependent-universal-property-equivalences -open import foundation.diagonal-maps-of-types -open import foundation.equivalences -open import foundation.equivalences-arrows -open import foundation.function-extensionality -open import foundation.functoriality-cartesian-product-types -open import foundation.functoriality-coproduct-types -open import foundation.functoriality-dependent-function-types -open import foundation.fundamental-theorem-of-identity-types -open import foundation.homotopy-induction -open import foundation.inhabited-types -open import foundation.iterated-successors-truncation-levels -open import foundation.postcomposition-functions -open import foundation.precomposition-dependent-functions -open import foundation.precomposition-functions -open import foundation.precomposition-type-families -open import foundation.propositional-truncations -open import foundation.sections -open import foundation.split-surjective-maps -open import foundation.structure-identity-principle -open import foundation.subtype-identity-principle -open import foundation.subuniverses -open import foundation.surjective-maps -open import foundation.truncated-types -open import foundation.truncation-levels -open import foundation.truncations -open import foundation.unit-type -open import foundation.univalence -open import foundation.universal-property-coproduct-types -open import foundation.universal-property-dependent-pair-types -open import foundation.universal-property-family-of-fibers-of-maps -open import foundation.universal-property-unit-type -open import foundation.universe-levels - -open import foundation-core.contractible-maps -open import foundation-core.embeddings -open import foundation-core.fibers-of-maps -open import foundation-core.function-types -open import foundation-core.functoriality-dependent-pair-types -open import foundation-core.homotopies -open import foundation-core.identity-types -open import foundation-core.injective-maps -open import foundation-core.postcomposition-dependent-functions -open import foundation-core.propositions -open import foundation-core.subtypes -open import foundation-core.torsorial-type-families -open import foundation-core.truncated-maps - -open import orthogonal-factorization-systems.localizations-at-subuniverses -open import orthogonal-factorization-systems.orthogonal-maps -open import orthogonal-factorization-systems.relative-separations-types-subuniverses -open import orthogonal-factorization-systems.subuniverse-equivalences - -open import synthetic-homotopy-theory.cocones-under-spans -open import synthetic-homotopy-theory.dependent-pullback-property-pushouts -open import synthetic-homotopy-theory.pushouts -``` - -
- -## Idea - -Given a [subuniverse](foundation.subuniverses.md) `K`, A map `f : A → B` is said -to be -{{#concept "`K`-connected" Disambiguation="map of types" Agda=is-subuniverse-connected-map}} -if it satisfies the -{{#concept "universal property" Disambiguation="subuniverse connected map of types"}} -of `K`-connected maps: - -For every `K`-valued family `U` over `B`, the -[dependent precomposition map](foundation-core.precomposition-dependent-functions.md) - -```text - - ∘ f : ((b : B) → U b) → ((a : A) → U (f a)) -``` - -is an [equivalence](foundation-core.equivalences.md). - -Equivalently, a `K`-connected map `f : A → B` is a map that is -[left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to maps -`h : C → B` whose fibers are in `K`. - -## Definitions - -### The predicate on maps of being `K`-connected - -```agda -is-subuniverse-connected-map-Prop : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → - (A → B) → Prop (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -is-subuniverse-connected-map-Prop K {B = B} f = - Π-Prop (B → type-subuniverse K) λ U → is-equiv-Prop (precomp-Π f (pr1 ∘ U)) - -is-subuniverse-connected-map : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → - (A → B) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -is-subuniverse-connected-map K {B = B} f = - (U : B → type-subuniverse K) → is-equiv (precomp-Π f (pr1 ∘ U)) - -is-prop-is-subuniverse-connected-map : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) → - is-prop (is-subuniverse-connected-map K f) -is-prop-is-subuniverse-connected-map K f = - is-prop-type-Prop (is-subuniverse-connected-map-Prop K f) -``` - -### The type of `K`-connected maps between two types - -```agda -subuniverse-connected-map : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) → - UU l1 → UU l2 → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -subuniverse-connected-map K A B = - type-subtype (is-subuniverse-connected-map-Prop K {A} {B}) - -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} - where - - map-subuniverse-connected-map : subuniverse-connected-map K A B → A → B - map-subuniverse-connected-map = - inclusion-subtype (is-subuniverse-connected-map-Prop K) - - is-subuniverse-connected-map-subuniverse-connected-map : - (f : subuniverse-connected-map K A B) → - is-subuniverse-connected-map K (map-subuniverse-connected-map f) - is-subuniverse-connected-map-subuniverse-connected-map = - is-in-subtype-inclusion-subtype (is-subuniverse-connected-map-Prop K) - - emb-inclusion-subuniverse-connected-map : - subuniverse-connected-map K A B ↪ (A → B) - emb-inclusion-subuniverse-connected-map = - emb-subtype (is-subuniverse-connected-map-Prop K) -``` - -### The type of `K`-connected maps into a type - -```agda -Subuniverse-Connected-Map : - {l1 l3 l4 : Level} (l2 : Level) (K : subuniverse l3 l4) (A : UU l1) → - UU (l1 ⊔ lsuc l3 ⊔ l4 ⊔ lsuc l2) -Subuniverse-Connected-Map l2 K A = Σ (UU l2) (subuniverse-connected-map K A) - -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} (f : Subuniverse-Connected-Map l2 K A) - where - - type-Subuniverse-Connected-Map : UU l2 - type-Subuniverse-Connected-Map = pr1 f - - subuniverse-connected-map-Subuniverse-Connected-Map : - subuniverse-connected-map K A type-Subuniverse-Connected-Map - subuniverse-connected-map-Subuniverse-Connected-Map = pr2 f - - map-Subuniverse-Connected-Map : A → type-Subuniverse-Connected-Map - map-Subuniverse-Connected-Map = - map-subuniverse-connected-map K - ( subuniverse-connected-map-Subuniverse-Connected-Map) - - is-subuniverse-connected-map-Subuniverse-Connected-Map : - is-subuniverse-connected-map K map-Subuniverse-Connected-Map - is-subuniverse-connected-map-Subuniverse-Connected-Map = - is-subuniverse-connected-map-subuniverse-connected-map K - ( subuniverse-connected-map-Subuniverse-Connected-Map) -``` - -## Properties - -### Equivalent characterizations of `K`-connected maps - -#### Contractible fiber-localization condition - -Given a subuniverse `K` then a map `f` is `K`-connected if if the fibers are -`K`-equivalent to contractible types. - -**Proof.** We have an equivalence of arrows - -```text - ~ - ((b : B) → unit → U b) --------> ((b : B) → U b) - | | - | | precomp-Π f U - | | - ∨ ~ ∨ - ((b : B) → fiber f b → U b) ---> ((a : A) → U (f a)) -``` - -where the left-hand map is -`map-Π (λ b → precomp (terminal-map (fiber f b)) (U b))`, hence if each terminal -map is a `K`-equivalence then this is an equivalence as well. ∎ - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - where - - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers : - ((b : B) → is-subuniverse-equiv K (terminal-map (fiber f b))) → - is-subuniverse-connected-map K f - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers H U = - is-equiv-target-is-equiv-source-equiv-arrow - ( map-Π (λ b → precomp (terminal-map (fiber f b)) (pr1 (U b)))) - ( precomp-Π f (pr1 ∘ U)) - ( ( equiv-Π-equiv-family - ( λ b → equiv-universal-property-unit (pr1 (U b)))) , - ( equiv-universal-property-family-of-fibers f (pr1 ∘ U)) , - ( refl-htpy)) - ( is-equiv-map-Π-is-fiberwise-equiv (λ b → H b (U b))) -``` - -#### Constant map condition on fibers - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - where - - is-subuniverse-connected-map-is-equiv-diagonal-exponential-fibers : - ( (b : B) (U : type-subuniverse K) → - is-equiv (diagonal-exponential (pr1 U) (fiber f b))) → - is-subuniverse-connected-map K f - is-subuniverse-connected-map-is-equiv-diagonal-exponential-fibers H = - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f - ( λ b U → - is-equiv-precomp-terminal-map-is-equiv-diagonal-exponential (H b U)) -``` - -#### Section condition - -A map is `K`-connected if and only if its dependent precomposition maps admit -sections and the fibers have `K`-localizations. - -```agda -module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) - (Kfib : B → UU l3) (η : (b : B) → fiber f b → Kfib b) - ( is-htpy-injective-precomp-η-Kfib : - (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) - where - - is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' : - ( fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl))) → - ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π'' Fη b = - ( pr1 (Fη b) , - is-htpy-injective-precomp-η-Kfib b - ( λ where (a , refl) → pr2 (Fη b) (a , refl))) - - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' : - fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl)) → - ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' (s , H) b = - ( s b , - is-htpy-injective-precomp-η-Kfib b (λ where (a , refl) → htpy-eq H a)) - -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) - (s : (U : B → type-subuniverse K) → section (precomp-Π f (pr1 ∘ U))) - where - - is-contr-subuniverse-localization-fiber-has-section-precomp-Π : - ((b : B) → is-contr (pr1 (Kfib b))) - is-contr-subuniverse-localization-fiber-has-section-precomp-Π = - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' f - ( type-subuniverse-localization K ∘ Kfib) - ( unit-subuniverse-localization K ∘ Kfib) - ( λ b H → - htpy-eq - ( is-injective-is-equiv - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) - ( type-subuniverse-subuniverse-localization K (Kfib b))) - ( eq-htpy H))) - ( is-split-surjective-section - ( precomp-Π f (type-subuniverse-localization K ∘ Kfib)) - ( s (type-subuniverse-subuniverse-localization K ∘ Kfib)) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) - - is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π : - (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) - is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π b = - is-subuniverse-equiv-comp K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-is-equiv K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( is-equiv-terminal-map-is-contr - ( is-contr-subuniverse-localization-fiber-has-section-precomp-Π b))) - - is-subuniverse-connected-map-has-section-precomp-Π : - is-subuniverse-connected-map K f - is-subuniverse-connected-map-has-section-precomp-Π = - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f - ( is-subuniverse-equiv-terminal-map-fibers-has-section-precomp-Π) -``` - -#### Surjection condition - -A map is `K`-connected if and only if its dependent precomposition maps are -surjective and the fibers have `K`-localizations. - -In fact, it suffices that the fibers have `K`-localizations and the family - -```text - b ↦ - Σ ( u : K(fiber f b)), - ( ((a , p) : fiber f b) → - dependent-identification (b ↦ K(fiber f b)) p u (η (f a) (a , refl))) -``` - -is inhabited, which is a slightly weaker condition than inhabitedness of the -fiber of `precomp-Π f` over the map `a ↦ η (f a) (a , refl)`. - -```agda -module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) - (Kfib : B → UU l3) - (η : (b : B) → fiber f b → Kfib b) - ( is-htpy-injective-precomp-η-Kfib : - (b : B) {g h : Kfib b → Kfib b} → - precomp (η b) (Kfib b) g ~ precomp (η b) (Kfib b) h → g ~ h) - where - - abstract - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' : - ( (b : B) → - is-inhabited - ( family-fiber-Π-precomp-Π' f Kfib (λ a → η (f a) (a , refl)) b)) → - ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' - Fη b = - rec-trunc-Prop - ( is-contr-Prop (Kfib b)) - ( λ (sb , Hb) → - ( sb , - is-htpy-injective-precomp-η-Kfib b - ( λ where (a , refl) → Hb (a , refl)))) - ( Fη b) - - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π : - is-inhabited (fiber (precomp-Π f Kfib) (λ a → η (f a) (a , refl))) → - ((b : B) → is-contr (Kfib b)) - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π - s b = - rec-trunc-Prop - ( is-contr-Prop (Kfib b)) - ( λ s → - is-contr-subuniverse-localization-fiber-has-section-precomp-Π' - f Kfib η is-htpy-injective-precomp-η-Kfib s b) - ( s) - -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - (Kfib : (b : B) → subuniverse-localization K (fiber f b)) - where - - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π : - ( (b : B) → - is-inhabited - ( family-fiber-Π-precomp-Π' f - ( type-subuniverse-localization K ∘ Kfib) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) - ( b))) → - ((b : B) → is-contr (type-subuniverse-localization K (Kfib b))) - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π = - is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π' - ( f) - ( type-subuniverse-localization K ∘ Kfib) - ( unit-subuniverse-localization K ∘ Kfib) - ( λ b H → - htpy-eq - ( is-injective-is-equiv - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) - ( type-subuniverse-subuniverse-localization K (Kfib b))) - ( eq-htpy H))) - - is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π : - ( (b : B) → - is-inhabited - ( family-fiber-Π-precomp-Π' f - ( type-subuniverse-localization K ∘ Kfib) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) - ( b))) → - (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) - is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π - s b = - is-subuniverse-equiv-comp K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-is-equiv K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( is-equiv-terminal-map-is-contr - ( is-contr-subuniverse-localization-fiber-is-inhabited-family-fiber-Π-precomp-Π - s b))) - - is-subuniverse-connected-map-is-inhabited-family-fiber-Π-precomp-Π : - ( (b : B) → - is-inhabited - ( family-fiber-Π-precomp-Π' f - ( type-subuniverse-localization K ∘ Kfib) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl)) - ( b))) → - is-subuniverse-connected-map K f - is-subuniverse-connected-map-is-inhabited-family-fiber-Π-precomp-Π Fη = - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f - ( is-subuniverse-equiv-terminal-map-fibers-is-inhabited-family-fiber-Π-precomp-Π - ( Fη)) - - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π : - ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → - ((b : B) → is-contr (pr1 (Kfib b))) - is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π H = - is-contr-subuniverse-localization-fiber-is-inhabited-fiber-precomp-Π f - ( type-subuniverse-localization K ∘ Kfib) - ( unit-subuniverse-localization K ∘ Kfib) - ( λ b H → - htpy-eq - ( is-injective-is-equiv - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b) - ( type-subuniverse-subuniverse-localization K (Kfib b))) - ( eq-htpy H))) - ( H ( type-subuniverse-subuniverse-localization K ∘ Kfib) - ( λ a → unit-subuniverse-localization K (Kfib (f a)) (a , refl))) - - is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π : - ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → - (b : B) → is-subuniverse-equiv K (terminal-map (fiber f b)) - is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π H b = - is-subuniverse-equiv-comp K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-unit-subuniverse-localization K (Kfib b)) - ( is-subuniverse-equiv-is-equiv K - ( terminal-map (type-subuniverse-localization K (Kfib b))) - ( is-equiv-terminal-map-is-contr - ( is-contr-subuniverse-localization-fiber-is-surjective-precomp-Π - H b))) - - is-subuniverse-connected-map-is-surjective-precomp-Π : - ((U : B → type-subuniverse K) → is-surjective (precomp-Π f (pr1 ∘ U))) → - is-subuniverse-connected-map K f - is-subuniverse-connected-map-is-surjective-precomp-Π H = - is-subuniverse-connected-map-is-subuniverse-equiv-terminal-map-fibers K f - ( is-subuniverse-equiv-terminal-map-fibers-is-surjective-precomp-Π H) -``` - -### Characterizing equality of `K`-connected maps - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} - where - - htpy-subuniverse-connected-map : - (f g : subuniverse-connected-map K A B) → UU (l1 ⊔ l2) - htpy-subuniverse-connected-map f g = - map-subuniverse-connected-map K f ~ map-subuniverse-connected-map K g - - refl-htpy-subuniverse-connected-map : - (f : subuniverse-connected-map K A B) → htpy-subuniverse-connected-map f f - refl-htpy-subuniverse-connected-map f = refl-htpy - - is-torsorial-htpy-subuniverse-connected-map : - (f : subuniverse-connected-map K A B) → - is-torsorial (htpy-subuniverse-connected-map f) - is-torsorial-htpy-subuniverse-connected-map f = - is-torsorial-Eq-subtype - ( is-torsorial-htpy (map-subuniverse-connected-map K f)) - ( is-prop-is-subuniverse-connected-map K) - ( map-subuniverse-connected-map K f) - ( refl-htpy-subuniverse-connected-map f) - ( is-subuniverse-connected-map-subuniverse-connected-map K f) - - htpy-eq-subuniverse-connected-map : - (f g : subuniverse-connected-map K A B) → - f = g → htpy-subuniverse-connected-map f g - htpy-eq-subuniverse-connected-map f .f refl = - refl-htpy-subuniverse-connected-map f - - is-equiv-htpy-eq-subuniverse-connected-map : - (f g : subuniverse-connected-map K A B) → - is-equiv (htpy-eq-subuniverse-connected-map f g) - is-equiv-htpy-eq-subuniverse-connected-map f = - fundamental-theorem-id - ( is-torsorial-htpy-subuniverse-connected-map f) - ( htpy-eq-subuniverse-connected-map f) - - extensionality-subuniverse-connected-map : - (f g : subuniverse-connected-map K A B) → - (f = g) ≃ htpy-subuniverse-connected-map f g - pr1 (extensionality-subuniverse-connected-map f g) = - htpy-eq-subuniverse-connected-map f g - pr2 (extensionality-subuniverse-connected-map f g) = - is-equiv-htpy-eq-subuniverse-connected-map f g - - eq-htpy-subuniverse-connected-map : - (f g : subuniverse-connected-map K A B) → - htpy-subuniverse-connected-map f g → (f = g) - eq-htpy-subuniverse-connected-map f g = - map-inv-equiv (extensionality-subuniverse-connected-map f g) -``` - -### All maps are `Contr`-connected - -```agda -module _ - {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) - where - - is-Contr-connected-map : is-subuniverse-connected-map (is-contr-Prop {l3}) f - is-Contr-connected-map U = - is-equiv-is-contr - ( precomp-Π f (pr1 ∘ U)) - ( is-contr-Π (pr2 ∘ U)) - ( is-contr-Π (pr2 ∘ U ∘ f)) -``` - -### Equivalences are `K`-connected for any `K` - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} - where - - is-subuniverse-connected-map-is-equiv : - {f : A → B} → is-equiv f → is-subuniverse-connected-map K f - is-subuniverse-connected-map-is-equiv H U = - is-equiv-precomp-Π-is-equiv H (pr1 ∘ U) - - is-subuniverse-connected-map-equiv : - (e : A ≃ B) → is-subuniverse-connected-map K (map-equiv e) - is-subuniverse-connected-map-equiv e = - is-subuniverse-connected-map-is-equiv (is-equiv-map-equiv e) - - subuniverse-connected-map-equiv : - (A ≃ B) → subuniverse-connected-map K A B - subuniverse-connected-map-equiv e = - ( map-equiv e , is-subuniverse-connected-map-equiv e) -``` - -### `K`-connected maps are `K`-equivalences - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - where - - is-subuniverse-equiv-is-subuniverse-connected-map : - is-subuniverse-connected-map K f → is-subuniverse-equiv K f - is-subuniverse-equiv-is-subuniverse-connected-map F U = F (λ _ → U) -``` - -### The composition of two `K`-connected maps is `K`-connected - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {C : UU l3} - where - - is-subuniverse-connected-map-comp : - {g : B → C} {f : A → B} → - is-subuniverse-connected-map K g → - is-subuniverse-connected-map K f → - is-subuniverse-connected-map K (g ∘ f) - is-subuniverse-connected-map-comp {g} {f} G F U = - is-equiv-comp - ( precomp-Π f (pr1 ∘ U ∘ g)) - ( precomp-Π g (pr1 ∘ U)) - ( G U) - ( F (U ∘ g)) - - comp-subuniverse-connected-map : - subuniverse-connected-map K B C → - subuniverse-connected-map K A B → - subuniverse-connected-map K A C - pr1 (comp-subuniverse-connected-map g f) = - map-subuniverse-connected-map K g ∘ map-subuniverse-connected-map K f - pr2 (comp-subuniverse-connected-map g f) = - is-subuniverse-connected-map-comp - ( is-subuniverse-connected-map-subuniverse-connected-map K g) - ( is-subuniverse-connected-map-subuniverse-connected-map K f) -``` - -### Right cancellation of `K`-connected maps - -```agda -is-subuniverse-connected-map-left-factor : - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {C : UU l3} - {g : B → C} {h : A → B} → - is-subuniverse-connected-map K h → - is-subuniverse-connected-map K (g ∘ h) → - is-subuniverse-connected-map K g -is-subuniverse-connected-map-left-factor K {g = g} {h} H GH U = - is-equiv-right-factor - ( precomp-Π h (pr1 ∘ U ∘ g)) - ( precomp-Π g (pr1 ∘ U)) - ( H (U ∘ g)) - ( GH U) -``` - -### Characterization of the identity type of `Subuniverse-Connected-Map l2 K A` - -```agda -equiv-Subuniverse-Connected-Map : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} → - (f g : Subuniverse-Connected-Map l2 K A) → UU (l1 ⊔ l2) -equiv-Subuniverse-Connected-Map K f g = - Σ ( type-Subuniverse-Connected-Map K f ≃ type-Subuniverse-Connected-Map K g) - ( λ e → - map-equiv e ∘ map-Subuniverse-Connected-Map K f ~ - map-Subuniverse-Connected-Map K g) - -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} - (f : Subuniverse-Connected-Map l2 K A) - where - - id-equiv-Subuniverse-Connected-Map : equiv-Subuniverse-Connected-Map K f f - id-equiv-Subuniverse-Connected-Map = (id-equiv , refl-htpy) - - is-torsorial-equiv-Subuniverse-Connected-Map : - is-torsorial (equiv-Subuniverse-Connected-Map K f) - is-torsorial-equiv-Subuniverse-Connected-Map = - is-torsorial-Eq-structure - ( is-torsorial-equiv (type-Subuniverse-Connected-Map K f)) - ( type-Subuniverse-Connected-Map K f , id-equiv) - ( is-torsorial-htpy-subuniverse-connected-map K - ( subuniverse-connected-map-Subuniverse-Connected-Map K f)) - - equiv-eq-Subuniverse-Connected-Map : - (g : Subuniverse-Connected-Map l2 K A) → - f = g → equiv-Subuniverse-Connected-Map K f g - equiv-eq-Subuniverse-Connected-Map .f refl = - id-equiv-Subuniverse-Connected-Map - - is-equiv-equiv-eq-Subuniverse-Connected-Map : - (g : Subuniverse-Connected-Map l2 K A) → - is-equiv (equiv-eq-Subuniverse-Connected-Map g) - is-equiv-equiv-eq-Subuniverse-Connected-Map = - fundamental-theorem-id - ( is-torsorial-equiv-Subuniverse-Connected-Map) - ( equiv-eq-Subuniverse-Connected-Map) - - extensionality-Subuniverse-Connected-Map : - (g : Subuniverse-Connected-Map l2 K A) → - (f = g) ≃ equiv-Subuniverse-Connected-Map K f g - extensionality-Subuniverse-Connected-Map g = - ( equiv-eq-Subuniverse-Connected-Map g , - is-equiv-equiv-eq-Subuniverse-Connected-Map g) - - eq-equiv-Subuniverse-Connected-Map : - (g : Subuniverse-Connected-Map l2 K A) → - equiv-Subuniverse-Connected-Map K f g → f = g - eq-equiv-Subuniverse-Connected-Map g = - map-inv-equiv (extensionality-Subuniverse-Connected-Map g) -``` - -### `K`-connected maps are closed under cobase change - -```agda -module _ - {l1 l2 l3 l4 l5 l6 : Level} (K : subuniverse l5 l6) - {S : UU l1} {A : UU l2} {B : UU l3} - (f : S → A) (g : S → B) {X : UU l4} (c : cocone f g X) - where - - is-subuniverse-connected-map-cobase-change' : - dependent-pullback-property-pushout f g c → - is-subuniverse-connected-map K g → - is-subuniverse-connected-map K (horizontal-map-cocone f g c) - is-subuniverse-connected-map-cobase-change' H G U = - is-equiv-vertical-map-is-pullback _ _ - ( cone-dependent-pullback-property-pushout f g c (pr1 ∘ U)) - ( G (U ∘ vertical-map-cocone f g c)) - ( H (pr1 ∘ U)) - - is-subuniverse-connected-map-cobase-change : - is-pushout f g c → - is-subuniverse-connected-map K g → - is-subuniverse-connected-map K (horizontal-map-cocone f g c) - is-subuniverse-connected-map-cobase-change H = - is-subuniverse-connected-map-cobase-change' - ( dependent-pullback-property-pushout-is-pushout f g c H) -``` - -### `K`-connected maps are closed under retracts of maps - -**Proof.** Given a retract of maps - -```text - i r - A' ------> A ------> A' - | | | - f'| I f R | f' - ∨ ∨ ∨ - B' ------> B ------> B' - i' r' -``` - -with higher coherence `α`, and a `K`-valued family `U` over `B'` there is by -functoriality an induced retract of dependent precomposition maps - -```text - Π(A'),(U∘f') <--- Π(A'),(U∘r'∘i'∘f) <--- Π(A),(U∘r'∘f) <--- Π(A'),(U∘f') - ∧ ∧ ∧ - | α* □ Π(I),(U∘r') | Π(R),U | - Π(f'),U | Π(f),(U∘r') | Π(f'),U - | | | - Π(B'),U <--------- Π(B'),(U∘r'∘i') <----- Π(B),(U∘r') <--- Π(B'),(U), -``` - -and since equivalences are closed under retracts of maps, if `f` is -`K`-connected then so is `f'`. ∎ - -Note that, since equivalences are already closed under noncoherent retracts of -maps, we are not obligated to produce the higher coherence of this retract. - -> This remains to be formalized. - -### The total map induced by a family of maps is `K`-connected if and only if all maps in the family are `K`-connected - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : A → UU l2} {C : A → UU l3} - (f : (x : A) → B x → C x) - where - - is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map : - ((x : A) → is-subuniverse-connected-map K (f x)) → - is-subuniverse-connected-map K (tot f) - is-subuniverse-connected-map-tot-is-fiberwise-subuniverse-connected-map H U = - is-equiv-source-is-equiv-target-equiv-arrow - ( precomp-Π (tot f) (pr1 ∘ U)) - ( map-Π (λ i → precomp-Π (f i) (pr1 ∘ U ∘ (i ,_)))) - ( equiv-ev-pair , equiv-ev-pair , refl-htpy) - ( is-equiv-map-Π-is-fiberwise-equiv (λ i → H i (U ∘ (i ,_)))) - - -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot : - -- is-subuniverse-connected-map K (tot f) → - -- (x : A) → is-subuniverse-connected-map K (f x) - -- is-fiberwise-subuniverse-connected-map-is-subuniverse-connected-map-tot H = {! !} - -- -- is-subuniverse-connected-equiv (inv-compute-fiber-tot f (x , y)) (H (x , y)) -``` - -### Coproducts of `K`-connected maps - -```agda -module _ - {l1 l2 l3 l4 l5 l6 : Level} (K : subuniverse l5 l6) - {A : UU l1} {B : UU l2} {A' : UU l3} {B' : UU l4} - {f : A → B} {f' : A' → B'} - where - - is-subuniverse-connected-map-coproduct : - is-subuniverse-connected-map K f → - is-subuniverse-connected-map K f' → - is-subuniverse-connected-map K (map-coproduct f f') - is-subuniverse-connected-map-coproduct F F' U = - is-equiv-source-is-equiv-target-equiv-arrow - ( precomp-Π (map-coproduct f f') (pr1 ∘ U)) - ( map-product - ( precomp-Π f (pr1 ∘ U ∘ inl)) - ( precomp-Π f' (pr1 ∘ U ∘ inr))) - ( equiv-dependent-universal-property-coproduct - ( pr1 ∘ U) , - equiv-dependent-universal-property-coproduct - ( pr1 ∘ U ∘ map-coproduct f f') , - refl-htpy) - ( is-equiv-map-product _ _ (F (U ∘ inl)) (F' (U ∘ inr))) -``` - -### The map on dependent pair types induced by a `K`-connected map is a `K`-equivalence - -This requires the added assumption that `K` is closed under exponentiation by -arbitrary types. - -This is a generalization of Lemma 2.27 in {{#cite CORS20}}, since the unit map -of the $ΣK$-localization is a $ΣK$-equivalence, and $ΣK$-equivalences are in -particular $K$-connected. - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : subuniverse-connected-map K A B) (P : B → UU l3) - where - - map-Σ-map-base-subuniverse-connected-map : - Σ A (P ∘ map-subuniverse-connected-map K f) → Σ B P - map-Σ-map-base-subuniverse-connected-map = - map-Σ-map-base (map-subuniverse-connected-map K f) P - - is-subuniverse-equiv-map-Σ-map-base-subuniverse-connected-map : - ((V : UU l3) (U : type-subuniverse K) → is-in-subuniverse K (V → pr1 U)) → - is-subuniverse-equiv K map-Σ-map-base-subuniverse-connected-map - is-subuniverse-equiv-map-Σ-map-base-subuniverse-connected-map H U = - is-equiv-source-is-equiv-target-equiv-arrow - ( precomp map-Σ-map-base-subuniverse-connected-map (pr1 U)) - ( precomp-Π (map-subuniverse-connected-map K f) (λ y → (b : P y) → pr1 U)) - ( equiv-ev-pair , equiv-ev-pair , refl-htpy) - ( is-subuniverse-connected-map-subuniverse-connected-map K f - ( λ y → ((P y → pr1 U) , H (P y) U))) -``` - -### There is a `K`-equivalence between the fiber of a map and the fiber of its `ΣK`-localization - -This is generalization of Corollary 2.29 in {{#cite CORS20}}. - -Assume given arbitrary `f : A → B` and `Σ(δ)K`-equivalences `A → A'` and -`B → B'`. Then there is - -```text - fiber f b = Σ A (λ a → f a = b) - → Σ A (λ a → ║f a = b║) - ≃ Σ A (λ a → |f a| = |b|) - ≃ Σ A (λ a → ║f║ |a| = |b|) - → Σ ║A║ (λ t → ║f║ t = |b|) - = fiber ║f║ |b| -``` - -where the first and last maps are `K`-equivalences. - -```text -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) (b : B) - where - - fiber-map-trunc-fiber : - fiber f b → fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b) - fiber-map-trunc-fiber = - ( map-Σ-map-base-subuniverse-equiv - ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) ∘ - ( tot - ( λ a → - ( concat (naturality-unit-trunc (succ-𝕋 K) f a) (unit-trunc b)) ∘ - ( map-effectiveness-trunc K (f a) b) ∘ - ( unit-trunc))) - - abstract - is-subuniverse-equiv-fiber-map-trunc-fiber : - is-subuniverse-equiv K fiber-map-trunc-fiber - is-subuniverse-equiv-fiber-map-trunc-fiber = - is-subuniverse-equiv-comp - ( map-Σ-map-base-subuniverse-equiv - ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) - ( tot - ( λ a → - ( concat (naturality-unit-trunc (succ-𝕋 K) f a) (unit-trunc b)) ∘ - ( map-effectiveness-trunc K (f a) b) ∘ - ( unit-trunc))) - ( is-subuniverse-equiv-is-subuniverse-equiv-equiv - ( equiv-tot - ( λ a → - ( equiv-concat - ( naturality-unit-trunc (succ-𝕋 K) f a) - ( unit-trunc b)) ∘e - ( effectiveness-trunc K (f a) b))) - ( λ (a , p) → a , unit-trunc p) - ( is-equiv-map-equiv (equiv-trunc-Σ K))) - ( is-subuniverse-equiv-map-Σ-map-base-unit-trunc - ( λ t → map-trunc (succ-𝕋 K) f t = unit-trunc b)) - - subuniverse-equiv-fiber-map-trunc-fiber : - subuniverse-equiv K - ( fiber f b) - ( fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b)) - pr1 subuniverse-equiv-fiber-map-trunc-fiber = - fiber-map-trunc-fiber - pr2 subuniverse-equiv-fiber-map-trunc-fiber = - is-subuniverse-equiv-fiber-map-trunc-fiber - - equiv-trunc-fiber-map-trunc-fiber : - type-trunc K (fiber f b) ≃ - type-trunc K (fiber (map-trunc (succ-𝕋 K) f) (unit-trunc b)) - equiv-trunc-fiber-map-trunc-fiber = - equiv-trunc-subuniverse-equiv K - ( subuniverse-equiv-fiber-map-trunc-fiber) -``` diff --git a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md b/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md deleted file mode 100644 index 340ebe99e9..0000000000 --- a/src/orthogonal-factorization-systems/subuniverse-equivalences.lagda.md +++ /dev/null @@ -1,528 +0,0 @@ -# Equivalences with respect to a subuniverse - -```agda -module orthogonal-factorization-systems.subuniverse-equivalences where -``` - -
Imports - -```agda -open import foundation.commuting-squares-of-maps -open import foundation.connected-maps -open import foundation.connected-types -open import foundation.contractible-types -open import foundation.dependent-pair-types -open import foundation.functoriality-truncation -open import foundation.identity-types -open import foundation.precomposition-functions -open import foundation.precomposition-functions-into-subuniverses -open import foundation.propositional-truncations -open import foundation.subuniverses -open import foundation.truncations -open import foundation.type-arithmetic-dependent-pair-types -open import foundation.universal-property-dependent-pair-types -open import foundation.universal-property-equivalences -open import foundation.universal-property-truncation -open import foundation.universe-levels - -open import foundation-core.contractible-maps -open import foundation-core.equivalences -open import foundation-core.fibers-of-maps -open import foundation-core.function-types -open import foundation-core.functoriality-dependent-pair-types -open import foundation-core.homotopies -open import foundation-core.retractions -open import foundation-core.sections -open import foundation-core.transport-along-identifications -open import foundation-core.truncated-maps -open import foundation-core.truncated-types -open import foundation-core.truncation-levels -``` - -
- -## Idea - -Given a [subuniverse](foundation.subuniverses.md) `K`, A map `f : A → B` is said -to be a -{{#concept "`K`-equivalence" Disambiguation="map of types" Agda=is-subuniverse-equiv}} -if it satisfies the -{{#concept "universal property" Disambiguation="subuniverse connected map of types"}} -of `K`-equivalences: - -For every `K`-type `U`, the -[precomposition map](foundation-core.precomposition-functions.md) - -```text - - ∘ f : (B → U) → (A → U) -``` - -is an [equivalence](foundation-core.equivalences.md). - -## Definition - -```agda -is-subuniverse-equiv : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} → - (A → B) → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -is-subuniverse-equiv K f = - (U : type-subuniverse K) → is-equiv (precomp f (pr1 U)) - -subuniverse-equiv : - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) → - UU l1 → UU l2 → UU (l1 ⊔ l2 ⊔ lsuc l3 ⊔ l4) -subuniverse-equiv K A B = Σ (A → B) (is-subuniverse-equiv K) - -module _ - {l1 l2 l3 l4 : Level} - (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} - (f : subuniverse-equiv K A B) - where - - map-subuniverse-equiv : A → B - map-subuniverse-equiv = pr1 f - - is-subuniverse-equiv-subuniverse-equiv : - is-subuniverse-equiv K map-subuniverse-equiv - is-subuniverse-equiv-subuniverse-equiv = pr2 f -``` - -## Properties - -### Equivalences are `K`-equivalences for all `K` - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} (f : A → B) - where - - is-subuniverse-equiv-is-equiv : is-equiv f → is-subuniverse-equiv K f - is-subuniverse-equiv-is-equiv H U = is-equiv-precomp-is-equiv f H (pr1 U) -``` - -### The identity map is a `K`-equivalence for all `K` - -```agda -is-subuniverse-equiv-id : - {l1 l2 l3 : Level} (K : subuniverse l2 l3) {A : UU l1} → - is-subuniverse-equiv K (id' A) -is-subuniverse-equiv-id K = is-subuniverse-equiv-is-equiv K id is-equiv-id -``` - -### The `K`-equivalences are closed under homotopies - -```agda -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) - {A : UU l1} {B : UU l2} {f g : A → B} - where - - is-subuniverse-equiv-htpy : - f ~ g → is-subuniverse-equiv K g → is-subuniverse-equiv K f - is-subuniverse-equiv-htpy H G U = - is-equiv-htpy (precomp g (pr1 U)) (htpy-precomp H (pr1 U)) (G U) - - is-subuniverse-equiv-htpy' : - f ~ g → is-subuniverse-equiv K f → is-subuniverse-equiv K g - is-subuniverse-equiv-htpy' H G U = - is-equiv-htpy' (precomp f (pr1 U)) (htpy-precomp H (pr1 U)) (G U) -``` - -### The `K`-equivalences are closed under composition - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {C : UU l3} - where - - is-subuniverse-equiv-comp : - (g : B → C) (f : A → B) → - is-subuniverse-equiv K f → - is-subuniverse-equiv K g → - is-subuniverse-equiv K (g ∘ f) - is-subuniverse-equiv-comp g f F G U = - is-equiv-comp (precomp f (pr1 U)) (precomp g (pr1 U)) (G U) (F U) - - subuniverse-equiv-comp : - subuniverse-equiv K B C → - subuniverse-equiv K A B → - subuniverse-equiv K A C - pr1 (subuniverse-equiv-comp g f) = - map-subuniverse-equiv K g ∘ map-subuniverse-equiv K f - pr2 (subuniverse-equiv-comp g f) = - is-subuniverse-equiv-comp - ( map-subuniverse-equiv K g) - ( map-subuniverse-equiv K f) - ( is-subuniverse-equiv-subuniverse-equiv K f) - ( is-subuniverse-equiv-subuniverse-equiv K g) -``` - -### The class of `K`-equivalences has the left and right cancellation property - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {C : UU l3} - (g : B → C) (f : A → B) (GF : is-subuniverse-equiv K (g ∘ f)) - where - - is-subuniverse-equiv-left-factor : - is-subuniverse-equiv K f → is-subuniverse-equiv K g - is-subuniverse-equiv-left-factor F U = - is-equiv-right-factor (precomp f (pr1 U)) (precomp g (pr1 U)) (F U) (GF U) - - is-subuniverse-equiv-right-factor : - is-subuniverse-equiv K g → is-subuniverse-equiv K f - is-subuniverse-equiv-right-factor G U = - is-equiv-left-factor (precomp f (pr1 U)) (precomp g (pr1 U)) (GF U) (G U) -``` - -### The class of `K`-equivalences satisfies the 3-for-2 property - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {X : UU l3} - (f : A → X) (g : B → X) (h : A → B) (p : f ~ g ∘ h) - where - - is-subuniverse-equiv-top-map-triangle : - is-subuniverse-equiv K g → - is-subuniverse-equiv K f → - is-subuniverse-equiv K h - is-subuniverse-equiv-top-map-triangle G F = - is-subuniverse-equiv-right-factor K g h - ( is-subuniverse-equiv-htpy' K p F) - ( G) - - is-subuniverse-equiv-right-map-triangle : - is-subuniverse-equiv K f → - is-subuniverse-equiv K h → - is-subuniverse-equiv K g - is-subuniverse-equiv-right-map-triangle F = - is-subuniverse-equiv-left-factor K g h (is-subuniverse-equiv-htpy' K p F) - - is-subuniverse-equiv-left-map-triangle : - is-subuniverse-equiv K h → - is-subuniverse-equiv K g → - is-subuniverse-equiv K f - is-subuniverse-equiv-left-map-triangle H G = - is-subuniverse-equiv-htpy K p (is-subuniverse-equiv-comp K g h H G) -``` - -### Sections of `K`-equivalences are `K`-equivalences - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {f : A → B} - where - - is-subuniverse-equiv-map-section : - (K : subuniverse l3 l4) (s : section f) → - is-subuniverse-equiv K f → - is-subuniverse-equiv K (map-section f s) - is-subuniverse-equiv-map-section K (s , h) = - is-subuniverse-equiv-right-factor K f s - ( is-subuniverse-equiv-is-equiv K (f ∘ s) (is-equiv-htpy-id h)) -``` - -### Retractions of `K`-equivalences are `K`-equivalences - -```agda -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {f : A → B} - where - - is-subuniverse-equiv-map-retraction : - (K : subuniverse l3 l4) (r : retraction f) → - is-subuniverse-equiv K f → - is-subuniverse-equiv K (map-retraction f r) - is-subuniverse-equiv-map-retraction K (r , h) = - is-subuniverse-equiv-left-factor K r f - ( is-subuniverse-equiv-is-equiv K (r ∘ f) (is-equiv-htpy-id h)) -``` - -### Composing `K`-equivalences with equivalences - -```agda -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) - {A : UU l1} {B : UU l2} {C : UU l3} - where - - is-subuniverse-equiv-is-equiv-is-subuniverse-equiv : - (g : B → C) (f : A → B) → - is-subuniverse-equiv K g → - is-equiv f → - is-subuniverse-equiv K (g ∘ f) - is-subuniverse-equiv-is-equiv-is-subuniverse-equiv g f eg ef = - is-subuniverse-equiv-comp K g f - ( is-subuniverse-equiv-is-equiv K f ef) - ( eg) - - is-subuniverse-equiv-is-subuniverse-equiv-is-equiv : - (g : B → C) (f : A → B) → - is-equiv g → - is-subuniverse-equiv K f → - is-subuniverse-equiv K (g ∘ f) - is-subuniverse-equiv-is-subuniverse-equiv-is-equiv g f eg ef = - is-subuniverse-equiv-comp K g f - ( ef) - ( is-subuniverse-equiv-is-equiv K g eg) - - is-subuniverse-equiv-equiv-is-subuniverse-equiv : - (g : B → C) (f : A ≃ B) → - is-subuniverse-equiv K g → - is-subuniverse-equiv K (g ∘ map-equiv f) - is-subuniverse-equiv-equiv-is-subuniverse-equiv g (f , ef) eg = - is-subuniverse-equiv-is-equiv-is-subuniverse-equiv g f eg ef - - is-subuniverse-equiv-is-subuniverse-equiv-equiv : - (g : B ≃ C) (f : A → B) → - is-subuniverse-equiv K f → - is-subuniverse-equiv K (map-equiv g ∘ f) - is-subuniverse-equiv-is-subuniverse-equiv-equiv (g , eg) f ef = - is-subuniverse-equiv-is-subuniverse-equiv-is-equiv g f eg ef -``` - -### Being `K`-connected is invariant under `K`-equivalences - -```text -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} - where - - is-connected-is-subuniverse-equiv-is-connected : - (f : A → B) → is-subuniverse-equiv K f → - is-connected K B → is-connected K A - is-connected-is-subuniverse-equiv-is-connected f e = - is-contr-equiv (type-trunc K B) (map-trunc K f , e) - - is-connected-is-subuniverse-equiv-is-connected' : - (f : A → B) → is-subuniverse-equiv K f → - is-connected K A → is-connected K B - is-connected-is-subuniverse-equiv-is-connected' f e = - is-contr-equiv' (type-trunc K A) (map-trunc K f , e) - - is-connected-subuniverse-equiv-is-connected : - subuniverse-equiv K A B → is-connected K B → is-connected K A - is-connected-subuniverse-equiv-is-connected f = - is-connected-is-subuniverse-equiv-is-connected - ( map-subuniverse-equiv K f) - ( is-subuniverse-equiv-subuniverse-equiv K f) - - is-connected-subuniverse-equiv-is-connected' : - subuniverse-equiv K A B → is-connected K A → is-connected K B - is-connected-subuniverse-equiv-is-connected' f = - is-connected-is-subuniverse-equiv-is-connected' - ( map-subuniverse-equiv K f) - ( is-subuniverse-equiv-subuniverse-equiv K f) -``` - -### Every `ΣK`-equivalence is `K`-connected - -This is a generalization of Proposition 2.30 in {{#cite CORS20}}. - -```text -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} (f : A → B) - where - - is-connected-map-is-succ-subuniverse-equiv : - is-subuniverse-equiv (succ-𝕋 K) f → is-connected-map K f - is-connected-map-is-succ-subuniverse-equiv e b = - is-connected-subuniverse-equiv-is-connected - ( subuniverse-equiv-fiber-map-trunc-fiber f b) - ( is-connected-map-is-equiv e (unit-trunc b)) -``` - -### A map is `K`-connected if and only if its `ΣK`-localization is - -```text -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} {f : A → B} - where - - is-connected-map-trunc-succ-is-succ-connected-domain : - is-connected-map K f → - is-connected-map K (map-trunc (succ-𝕋 K) f) - is-connected-map-trunc-succ-is-succ-connected-domain cf t = - apply-universal-property-trunc-Prop - ( is-surjective-unit-trunc-succ t) - ( is-connected-Prop K (fiber (map-trunc (succ-𝕋 K) f) t)) - ( λ (b , p) → - tr - ( λ s → is-connected K (fiber (map-trunc (succ-𝕋 K) f) s)) - ( p) - ( is-connected-subuniverse-equiv-is-connected' - ( subuniverse-equiv-fiber-map-trunc-fiber f b) - ( cf b))) - - is-connected-map-is-connected-map-trunc-succ : - is-connected-map K (map-trunc (succ-𝕋 K) f) → - is-connected-map K f - is-connected-map-is-connected-map-trunc-succ cf' b = - is-connected-subuniverse-equiv-is-connected - ( subuniverse-equiv-fiber-map-trunc-fiber f b) - ( cf' (unit-trunc b)) -``` - -### The codomain of a `K`-connected map is `ΣK`-connected if its domain is `ΣK`-connected - -This follows part of the proof of Proposition 2.31 in {{#cite CORS20}}. - -**Proof.** Let $f : A → B$ be a $K$-connected map on a $K+1$-connected domain. -To show that the codomain is $K+1$-connected it is enough to show that $f$ is a -$K+1$-equivalence, in other words, that $║f║ₖ₊₁$ is an equivalence. By previous -computations we know that $║f║ₖ₊₁$ is $K$-truncated since the domain is -$K+1$-connected, and that $║f║ₖ₊₁$ is $K$-connected since $f$ is $K$-connected, -so we are done. ∎ - -```text -module _ - {l1 l2 l3 l4 : Level} (K : subuniverse l3 l4) {A : UU l1} {B : UU l2} (f : A → B) - where - - is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map : - is-connected-map K f → - is-connected (succ-𝕋 K) A → - is-subuniverse-equiv (succ-𝕋 K) f - is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map - cf cA = - is-equiv-is-connected-map-is-trunc-map - ( is-trunc-map-trunc-succ-is-succ-connected-domain f cA) - ( is-connected-map-trunc-succ-is-succ-connected-domain cf) - - is-succ-connected-codomain-is-succ-connected-domain-is-connected-map : - is-connected-map K f → - is-connected (succ-𝕋 K) A → - is-connected (succ-𝕋 K) B - is-succ-connected-codomain-is-succ-connected-domain-is-connected-map cf cA = - is-connected-is-subuniverse-equiv-is-connected' f - ( is-subuniverse-equiv-succ-is-succ-connected-domain-is-connected-map - ( cf) - ( cA)) - ( cA) -``` - -### If `g ∘ f` is `ΣK`-connected, then `f` is `K`-connected if and only if `g` is `ΣK`-connected - -This is an instance of Proposition 2.31 in {{#cite CORS20}}. - -**Proof.** If $g$ is $ΣK$-connected then by the cancellation property of -$ΣK$-equivalences, $f$ is a $K+1$-equivalence, and so in particular -$K$-connected. - -Conversely, assume $f$ is $K$-connected. We want to show that the fibers of $g$ -are $K+1$-connected, so let $c$ be an element of the codomain of $g$. The fibers -of the composite $g ∘ f$ compute as - -$$ - \operatorname{fiber}_{g\circ f}(c) ≃ - \sum_{(b , p) : \operatorname{fiber}_{g}(c)}{\operatorname{fiber}_{f}(b)}. -$$ - -By the previous lemma, since $\operatorname{fiber}_{g\circ f}(c)$ is -$K+1$-connected, $\operatorname{fiber}_{g}(c)$ is $K+1$-connected if the first -projection map of this type is $K$-connected, and its fibers compute to the -fibers of $f$. ∎ - -```text -module _ - {l1 l2 l3 l4 l5 : Level} (K : subuniverse l4 l5) {A : UU l1} {B : UU l2} {C : UU l3} - (g : B → C) (f : A → B) (cgf : is-connected-map (succ-𝕋 K) (g ∘ f)) - where - - is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor : - is-connected-map (succ-𝕋 K) g → is-subuniverse-equiv (succ-𝕋 K) f - is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor - cg = - is-subuniverse-equiv-right-factor g f - ( is-subuniverse-equiv-is-connected-map (g ∘ f) cgf) - ( is-subuniverse-equiv-is-connected-map g cg) - - is-connected-map-right-factor-is-succ-connected-map-left-factor : - is-connected-map (succ-𝕋 K) g → is-connected-map K f - is-connected-map-right-factor-is-succ-connected-map-left-factor cg = - is-connected-map-is-succ-subuniverse-equiv f - ( is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor - ( cg)) - - is-connected-map-right-factor-is-succ-connected-map-right-factor : - is-connected-map K f → is-connected-map (succ-𝕋 K) g - is-connected-map-right-factor-is-succ-connected-map-right-factor cf c = - is-succ-connected-codomain-is-succ-connected-domain-is-connected-map - ( pr1) - ( λ p → - is-connected-equiv - ( equiv-fiber-pr1 (fiber f ∘ pr1) p) - ( cf (pr1 p))) - ( is-connected-equiv' (compute-fiber-comp g f c) (cgf c)) -``` - -As a corollary, if $g ∘ f$ is $(K + 1)$-connected for some $g$, and $f$ is -$K$-connected, then $f$ is a $K+1$-equivalence. - -```text - is-succ-truncation-equiv-is-succ-connected-comp : - is-connected-map K f → is-subuniverse-equiv (succ-𝕋 K) f - is-succ-truncation-equiv-is-succ-connected-comp cf = - is-succ-subuniverse-equiv-right-factor-is-succ-connected-map-left-factor - ( is-connected-map-right-factor-is-succ-connected-map-right-factor cf) -``` - -### A `K`-equivalence with a section is `K`-connected - -**Proof.** If $K ≐ -2$ notice that every map is $-2$-connected. So let -$K ≐ n + 1$ for some truncation level $n$ and let $f$ be our $K$-equivalence -with a section $s$. By assumption, we have a commuting triangle of maps - -```text - A - ∧ \ - s / \ f - / ∨ - B ======== B. -``` - -By the previous lemma, since the identity map is $K$-connected, it thus suffices -to show that $s$ is $n$-connected. But by the cancellation property of -$n+1$-equivalences $s$ is an $n+1$-equivalence and $n+1$-equivalences are in -particular $n$-connected. ∎ - -```text -module _ - {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} (f : A → B) - where - - is-connected-map-section-is-subuniverse-equiv-succ : - (K : subuniverse l3 l4) (s : section f) → - is-subuniverse-equiv (succ-𝕋 K) f → - is-connected-map K (map-section f s) - is-connected-map-section-is-subuniverse-equiv-succ K (s , h) e = - is-connected-map-is-succ-subuniverse-equiv s - ( is-subuniverse-equiv-map-section (succ-𝕋 K) (s , h) e) - - is-connected-map-is-subuniverse-equiv-section : - (K : subuniverse l3 l4) → - section f → is-subuniverse-equiv K f → is-connected-map K f - is-connected-map-is-subuniverse-equiv-section neg-two-𝕋 (s , h) e = - is-neg-two-connected-map f - is-connected-map-is-subuniverse-equiv-section (succ-𝕋 K) (s , h) e = - is-connected-map-right-factor-is-succ-connected-map-right-factor f s - ( is-connected-map-htpy-id h) - ( is-connected-map-section-is-subuniverse-equiv-succ K (s , h) e) -``` - -## References - -- The notion of `K`-equivalence is a special case of the notion of - `L`-equivalence, where `L` is a reflective subuniverse. These were studied in - {{#cite CORS20}}. -- The class of `K`-equivalences is - [left orthogonal](orthogonal-factorization-systems.orthogonal-maps.md) to the - class of `K`-étale maps. This was shown in {{#cite CR21}}. - -{{#bibliography}} diff --git a/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md b/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md index 076c7343f3..1ebbc9e3de 100644 --- a/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md +++ b/src/orthogonal-factorization-systems/types-local-at-maps.lagda.md @@ -14,7 +14,6 @@ open import foundation.contractible-maps open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.dependent-universal-property-equivalences -open import foundation.embeddings open import foundation.empty-types open import foundation.equivalences open import foundation.families-of-equivalences @@ -29,7 +28,6 @@ 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-maps open import foundation.propositions open import foundation.retractions open import foundation.retracts-of-maps @@ -318,36 +316,27 @@ module _ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) where - section-is-local-domains'' : - section (precomp f X) → is-emb (precomp f Y) → section f - pr1 (section-is-local-domains'' section-precomp-X _) = + section-is-local-domains' : section (precomp f X) → is-local f Y → section f + pr1 (section-is-local-domains' section-precomp-X is-local-Y) = pr1 section-precomp-X id - pr2 (section-is-local-domains'' section-precomp-X is-emb-precomp-Y) = + pr2 (section-is-local-domains' section-precomp-X is-local-Y) = htpy-eq ( ap ( pr1) - { ( f ∘ pr1 section-precomp-X id) , + { ( f ∘ pr1 (section-is-local-domains' section-precomp-X is-local-Y)) , ( ap (postcomp X f) (pr2 section-precomp-X id))} { id , refl} - ( eq-is-prop (is-prop-map-is-emb is-emb-precomp-Y f))) - - section-is-local-domains' : - section (precomp f X) → is-local f Y → section f - section-is-local-domains' section-precomp-X is-local-Y = - section-is-local-domains'' section-precomp-X (is-emb-is-equiv is-local-Y) - - is-equiv-is-local-domains' : - section (precomp f X) → is-emb (precomp f Y) → is-equiv f - pr1 (is-equiv-is-local-domains' section-precomp-X is-emb-precomp-Y) = - section-is-local-domains'' section-precomp-X is-emb-precomp-Y - pr2 (is-equiv-is-local-domains' section-precomp-X is-emb-precomp-Y) = + ( eq-is-contr (is-contr-map-is-equiv is-local-Y f))) + + is-equiv-is-local-domains' : section (precomp f X) → is-local f Y → is-equiv f + pr1 (is-equiv-is-local-domains' section-precomp-X is-local-Y) = + section-is-local-domains' section-precomp-X is-local-Y + pr2 (is-equiv-is-local-domains' section-precomp-X is-local-Y) = retraction-map-section-precomp f section-precomp-X is-equiv-is-local-domains : is-local f X → is-local f Y → is-equiv f - is-equiv-is-local-domains is-local-X is-local-Y = - is-equiv-is-local-domains' - ( section-is-equiv is-local-X) - ( is-emb-is-equiv is-local-Y) + is-equiv-is-local-domains is-local-X = + is-equiv-is-local-domains' (pr1 is-local-X) ``` ### If `f` has a retraction and the codomain of `f` is `f`-local, then `f` is an equivalence diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index 1a47382283..b3d7af4af0 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -365,7 +365,7 @@ module _ ``` ### The operator `pointed-map-Ω` preserves equivalences - + ```agda module _ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} From 38deb3bd35d683f6295929e39f08510622fc0aec Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Thu, 6 Nov 2025 18:44:54 +0100 Subject: [PATCH 62/69] pre-commit --- .../functoriality-loop-spaces.lagda.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md index b3d7af4af0..1a47382283 100644 --- a/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md +++ b/src/synthetic-homotopy-theory/functoriality-loop-spaces.lagda.md @@ -365,7 +365,7 @@ module _ ``` ### The operator `pointed-map-Ω` preserves equivalences - + ```agda module _ {l1 l2 : Level} {A : Pointed-Type l1} {B : Pointed-Type l2} From 7e874629e3f19c09dd0a4b0c2236286d5a83844a Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 14:06:36 +0100 Subject: [PATCH 63/69] reorg `connected-types` --- src/foundation/connected-types.lagda.md | 77 +++++++++++++------------ 1 file changed, 40 insertions(+), 37 deletions(-) diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index 637cc49594..c3371d5606 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -36,7 +36,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 Agda=Connected-Type}} +if its `k`-[truncation](foundation.truncations.md) is +[contractible](foundation-core.contractible-types.md). ## Definition @@ -197,42 +200,6 @@ is-connected-is-connected-succ-𝕋 k H = ( H)) ``` -### The total space of a family of `k`-connected types over a `k`-connected type is `k`-connected - -```agda -is-connected-Σ : - {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} → - is-connected k A → ((x : A) → is-connected k (B x)) → - is-connected k (Σ A B) -is-connected-Σ k H K = - is-contr-equiv _ (equiv-trunc k (equiv-pr1 K) ∘e equiv-trunc-Σ k) H -``` - -### If the total space of a family of `k`-connected types is `k`-connected so is the base - -**Proof.** We compute - -```text - ║Σ (x : A), B x║ₖ ≃ ║Σ (x : A), ║B x║ₖ║ₖ by equiv-trunc-Σ - ≃ ║Σ (x : A), 1 ║ₖ by k-connectedness of B - ≃ ║A║ₖ by the right unit law of Σ -``` - -and so, in particular, if the total space is `k`-connected so is the base. ∎ - -```agda -is-connected-base : - {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} → - ((x : A) → is-connected k (B x)) → is-connected k (Σ A B) → is-connected k A -is-connected-base k {A} {B} K = - is-contr-equiv' - ( type-trunc k (Σ A B)) - ( equivalence-reasoning - type-trunc k (Σ A B) - ≃ type-trunc k (Σ A (type-trunc k ∘ B)) by equiv-trunc-Σ k - ≃ type-trunc k A by equiv-trunc k (right-unit-law-Σ-is-contr K)) -``` - ### An inhabited type `A` is `k + 1`-connected if and only if its identity types are `k`-connected ```agda @@ -283,6 +250,42 @@ module _ ( center (K a x))))) ``` +### The total space of a family of `k`-connected types over a `k`-connected type is `k`-connected + +```agda +is-connected-Σ : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} → + is-connected k A → ((x : A) → is-connected k (B x)) → + is-connected k (Σ A B) +is-connected-Σ k H K = + is-contr-equiv _ (equiv-trunc k (equiv-pr1 K) ∘e equiv-trunc-Σ k) H +``` + +### If the total space of a family of `k`-connected types is `k`-connected so is the base + +**Proof.** We compute + +```text + ║Σ (x : A), B x║ₖ ≃ ║Σ (x : A), ║B x║ₖ║ₖ by equiv-trunc-Σ + ≃ ║Σ (x : A), 1 ║ₖ by k-connectedness of B + ≃ ║A║ₖ by the right unit law of Σ +``` + +and so, in particular, if the total space is `k`-connected so is the base. ∎ + +```agda +is-connected-base : + {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : A → UU l2} → + ((x : A) → is-connected k (B x)) → is-connected k (Σ A B) → is-connected k A +is-connected-base k {A} {B} K = + is-contr-equiv' + ( type-trunc k (Σ A B)) + ( equivalence-reasoning + type-trunc k (Σ A B) + ≃ type-trunc k (Σ A (type-trunc k ∘ B)) by equiv-trunc-Σ k + ≃ type-trunc k A by equiv-trunc k (right-unit-law-Σ-is-contr K)) +``` + ### If the domain of `f` is `k+1`-connected, then the `k+1`-truncation of `f` is `k`-truncated ```agda From 4f5546bf398a24e1f11bf800026e26d71c7c6e29 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 14:25:22 +0100 Subject: [PATCH 64/69] Coproducts of -1-connected types are not 0-connected --- src/foundation/0-connected-types.lagda.md | 83 +++++++++++++++++------ 1 file changed, 62 insertions(+), 21 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index 756f782225..d1384c6b67 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -16,11 +16,16 @@ open import foundation.functoriality-set-truncation open import foundation.images open import foundation.inhabited-types open import foundation.mere-equality +open import foundation.empty-types open import foundation.propositional-truncations open import foundation.set-truncations open import foundation.sets +open import foundation.negation +open import foundation.coproduct-types +open import foundation.equality-coproduct-types open import foundation.surjective-maps open import foundation.unit-type +open import foundation.mere-equality open import foundation.universal-property-contractible-types open import foundation.universal-property-unit-type open import foundation.universe-levels @@ -222,11 +227,35 @@ is-0-connected-is-contr X p = ### The image of a function with a `0`-connected domain is `0`-connected ```agda -abstract +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) + where abstract + + is-0-connected-im-is-0-connected-domain' : + A → all-elements-merely-equal A → is-contr (im (map-trunc-Set f)) + is-0-connected-im-is-0-connected-domain' a C = + is-contr-im + ( trunc-Set B) + ( unit-trunc-Set a) + ( apply-dependent-universal-property-trunc-Set' + ( λ x → + set-Prop + ( Id-Prop + ( trunc-Set B) + ( map-trunc-Set f x) + ( map-trunc-Set f (unit-trunc-Set a)))) + ( λ a' → + apply-universal-property-trunc-Prop + ( C a' a) + ( Id-Prop + ( trunc-Set B) + ( map-trunc-Set f (unit-trunc-Set a')) + ( map-trunc-Set f (unit-trunc-Set a))) + ( λ where refl → refl))) + is-0-connected-im-is-0-connected-domain : - {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) → is-0-connected A → is-0-connected (im f) - is-0-connected-im-is-0-connected-domain {A = A} {B} f C = + is-0-connected-im-is-0-connected-domain C = apply-universal-property-trunc-Prop ( is-inhabited-is-0-connected C) ( is-contr-Prop _) @@ -234,28 +263,40 @@ abstract is-contr-equiv' ( im (map-trunc-Set f)) ( equiv-trunc-im-Set f) - ( is-contr-im - ( trunc-Set B) - ( unit-trunc-Set a) - ( apply-dependent-universal-property-trunc-Set' - ( λ x → - set-Prop - ( Id-Prop - ( trunc-Set B) - ( map-trunc-Set f x) - ( map-trunc-Set f (unit-trunc-Set a)))) - ( λ a' → - apply-universal-property-trunc-Prop - ( mere-eq-is-0-connected C a' a) - ( Id-Prop - ( trunc-Set B) - ( map-trunc-Set f (unit-trunc-Set a')) - ( map-trunc-Set f (unit-trunc-Set a))) - ( λ where refl → refl))))) + ( is-0-connected-im-is-0-connected-domain' + ( a) + ( mere-eq-is-0-connected C))) abstract is-0-connected-im-point' : {l1 : Level} {A : UU l1} (f : unit → A) → is-0-connected (im f) is-0-connected-im-point' f = is-0-connected-im-is-0-connected-domain f is-0-connected-unit + + is-0-connected-im-point : + {l1 : Level} {A : UU l1} (a : A) → is-0-connected (im (point a)) + is-0-connected-im-point a = is-0-connected-im-point' (point a) +``` + +### Coproducts of -1-connected types are not 0-connected + +```agda +module _ + {l1 l2 : Level} {A : UU l1} {B : UU l2} + where abstract + + is-not-0-connected-coproduct-has-element : + A → B → ¬ is-0-connected (A + B) + is-not-0-connected-coproduct-has-element a b H = + apply-universal-property-trunc-Prop + ( mere-eq-is-0-connected H (inl a) (inr b)) + ( empty-Prop) + ( is-empty-eq-coproduct-inl-inr a b) + + is-not-0-connected-coproduct-is-inhabited : + is-inhabited A → is-inhabited B → ¬ is-0-connected (A + B) + is-not-0-connected-coproduct-is-inhabited |a| |b| = + apply-twice-universal-property-trunc-Prop |a| |b| + ( neg-type-Prop (is-0-connected (A + B))) + ( is-not-0-connected-coproduct-has-element) ``` From b318d76c1445c09bf8ce85805cb880e39e2576a1 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 15:02:26 +0100 Subject: [PATCH 65/69] edits 0-connected types --- src/foundation/0-connected-types.lagda.md | 97 +++++++++++++------ .../homomorphisms-concrete-groups.lagda.md | 2 +- 2 files changed, 70 insertions(+), 29 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index d1384c6b67..64569b70a9 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -1,4 +1,4 @@ -# `0`-Connected types +# 0-Connected types ```agda module foundation.0-connected-types where @@ -12,6 +12,7 @@ open import foundation.constant-maps open import foundation.contractible-types open import foundation.dependent-pair-types open import foundation.fiber-inclusions +open import foundation.evaluation-functions open import foundation.functoriality-set-truncation open import foundation.images open import foundation.inhabited-types @@ -61,7 +62,13 @@ is-0-connected A = type-Prop (is-0-connected-Prop A) is-prop-is-0-connected : {l : Level} (A : UU l) → is-prop (is-0-connected A) is-prop-is-0-connected A = is-prop-type-Prop (is-0-connected-Prop A) +``` + +## Properties + +### 0-connected types are inhabited +```agda abstract is-inhabited-is-0-connected : {l : Level} {A : UU l} → is-0-connected A → is-inhabited A @@ -70,13 +77,21 @@ abstract ( center C) ( set-Prop (trunc-Prop A)) ( unit-trunc-Prop) +``` +### All elements of 0-connected types are merely equal + +```agda abstract mere-eq-is-0-connected : {l : Level} {A : UU l} → is-0-connected A → all-elements-merely-equal A mere-eq-is-0-connected {A = A} H x y = apply-effectiveness-unit-trunc-Set (eq-is-contr H) +``` + +### A type is 0-connected if and only if it is inhabited and all elements are merely equal +```agda abstract is-0-connected-mere-eq : {l : Level} {A : UU l} (a : A) → @@ -96,17 +111,22 @@ abstract apply-universal-property-trunc-Prop H ( is-0-connected-Prop _) ( λ a → is-0-connected-mere-eq a (K a)) +``` -is-0-connected-is-surjective-point : - {l1 : Level} {A : UU l1} (a : A) → - is-surjective (point a) → is-0-connected A -is-0-connected-is-surjective-point a H = - is-0-connected-mere-eq a - ( λ x → - apply-universal-property-trunc-Prop - ( H x) - ( mere-eq-Prop a x) - ( λ u → unit-trunc-Prop (pr2 u))) +### A type is 0-connected iff any point inclusion is surjective + +```agda +abstract + is-0-connected-is-surjective-point : + {l1 : Level} {A : UU l1} (a : A) → + is-surjective (point a) → is-0-connected A + is-0-connected-is-surjective-point a H = + is-0-connected-mere-eq a + ( λ x → + apply-universal-property-trunc-Prop + ( H x) + ( mere-eq-Prop a x) + ( λ u → unit-trunc-Prop (pr2 u))) abstract is-surjective-point-is-0-connected : @@ -117,21 +137,29 @@ abstract ( mere-eq-is-0-connected H a x) ( trunc-Prop (fiber (point a) x)) ( λ where refl → unit-trunc-Prop (star , refl)) +``` + +### If `A` is 0-connected and `B` is `k+1`-truncated then every evaluation map `(A → B) → B` is `k`-truncated -is-trunc-map-ev-point-is-connected : +```agda +is-trunc-map-ev-is-connected : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (a : A) → - 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-0-connected A → is-trunc (succ-𝕋 k) B → + is-trunc-map k (ev-function B a) +is-trunc-map-ev-is-connected k {A} {B} a H K = is-trunc-map-comp k - ( ev-point' star {B}) + ( ev-function B star) ( precomp (point a) B) ( is-trunc-map-is-equiv k ( universal-property-contr-is-contr star is-contr-unit B)) ( is-trunc-map-precomp-Π-is-surjective k ( is-surjective-point-is-0-connected a H) ( λ _ → (B , K))) +``` + +### 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) → @@ -148,7 +176,11 @@ apply-dependent-universal-property-is-0-connected : {l : Level} (P : A → Prop l) → type-Prop (P a) → (x : A) → type-Prop (P x) apply-dependent-universal-property-is-0-connected a H P = map-inv-equiv (equiv-dependent-universal-property-is-0-connected a H P) +``` +### A type `A` is 0-connected if and only if every fiber inclusion `B a → Σ A B` is surjective + +```agda abstract is-surjective-fiber-inclusion : {l1 l2 : Level} {A : UU l1} {B : A → UU l2} → @@ -177,7 +209,11 @@ 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 @@ -189,7 +225,7 @@ is-0-connected-equiv' : 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 _ @@ -205,7 +241,7 @@ module _ ( is-contr-product p1 p2) ``` -### The unit type is `0`-connected +### The unit type is 0-connected ```agda abstract @@ -214,7 +250,7 @@ abstract is-contr-equiv' unit equiv-unit-trunc-unit-Set is-contr-unit ``` -### A contractible type is `0`-connected +### Contractible types are 0-connected ```agda is-0-connected-is-contr : @@ -224,16 +260,16 @@ is-0-connected-is-contr X p = is-contr-equiv X (inv-equiv (equiv-unit-trunc-Set (X , is-set-is-contr p))) p ``` -### The image of a function with a `0`-connected domain is `0`-connected +### The image of a function with a 0-connected domain is 0-connected ```agda module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) where abstract - is-0-connected-im-is-0-connected-domain' : + is-contr-im-map-trunc-Set-is-0-connected-domain' : A → all-elements-merely-equal A → is-contr (im (map-trunc-Set f)) - is-0-connected-im-is-0-connected-domain' a C = + is-contr-im-map-trunc-Set-is-0-connected-domain' a C = is-contr-im ( trunc-Set B) ( unit-trunc-Set a) @@ -263,18 +299,23 @@ module _ is-contr-equiv' ( im (map-trunc-Set f)) ( equiv-trunc-im-Set f) - ( is-0-connected-im-is-0-connected-domain' + ( is-contr-im-map-trunc-Set-is-0-connected-domain' ( a) ( mere-eq-is-0-connected C))) +``` -abstract - is-0-connected-im-point' : - {l1 : Level} {A : UU l1} (f : unit → A) → is-0-connected (im f) +### The image of a point is 0-connected + +```agda +module _ + {l1 : Level} {A : UU l1} + where abstract + + is-0-connected-im-point' : (f : unit → A) → is-0-connected (im f) is-0-connected-im-point' f = is-0-connected-im-is-0-connected-domain f is-0-connected-unit - is-0-connected-im-point : - {l1 : Level} {A : UU l1} (a : A) → is-0-connected (im (point a)) + is-0-connected-im-point : (a : A) → is-0-connected (im (point a)) is-0-connected-im-point a = is-0-connected-im-point' (point a) ``` diff --git a/src/group-theory/homomorphisms-concrete-groups.lagda.md b/src/group-theory/homomorphisms-concrete-groups.lagda.md index 4b12ff6251..2283144a3f 100644 --- a/src/group-theory/homomorphisms-concrete-groups.lagda.md +++ b/src/group-theory/homomorphisms-concrete-groups.lagda.md @@ -34,7 +34,7 @@ module _ is-set-hom-Concrete-Group : is-set hom-Concrete-Group is-set-hom-Concrete-Group = - is-trunc-map-ev-point-is-connected + is-trunc-map-ev-is-connected ( zero-𝕋) ( shape-Concrete-Group G) ( is-0-connected-classifying-type-Concrete-Group G) From 3b4efed16f5aa2d8c484027b125af831b28c9284 Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 17:46:43 +0100 Subject: [PATCH 66/69] edit --- src/foundation/0-connected-types.lagda.md | 15 +++++++++------ 1 file changed, 9 insertions(+), 6 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index 64569b70a9..e175458208 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -289,6 +289,14 @@ module _ ( map-trunc-Set f (unit-trunc-Set a))) ( λ where refl → refl))) + is-0-connected-im-is-0-connected-domain' : + A → all-elements-merely-equal A → is-0-connected (im f) + is-0-connected-im-is-0-connected-domain' a C = + is-contr-equiv' + ( im (map-trunc-Set f)) + ( equiv-trunc-im-Set f) + ( is-contr-im-map-trunc-Set-is-0-connected-domain' a C) + is-0-connected-im-is-0-connected-domain : is-0-connected A → is-0-connected (im f) is-0-connected-im-is-0-connected-domain C = @@ -296,12 +304,7 @@ module _ ( is-inhabited-is-0-connected C) ( is-contr-Prop _) ( λ a → - is-contr-equiv' - ( im (map-trunc-Set f)) - ( equiv-trunc-im-Set f) - ( is-contr-im-map-trunc-Set-is-0-connected-domain' - ( a) - ( mere-eq-is-0-connected C))) + is-0-connected-im-is-0-connected-domain' a (mere-eq-is-0-connected C)) ``` ### The image of a point is 0-connected From 9bb5004491db7da6073cdc59fcd40a538c7c07de Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 17:49:59 +0100 Subject: [PATCH 67/69] edit --- src/foundation/0-connected-types.lagda.md | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index e175458208..d69f2ed5ea 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -79,7 +79,7 @@ abstract ( unit-trunc-Prop) ``` -### All elements of 0-connected types are merely equal +### Elements of 0-connected types are all merely equal ```agda abstract @@ -89,7 +89,7 @@ abstract apply-effectiveness-unit-trunc-Set (eq-is-contr H) ``` -### A type is 0-connected if and only if it is inhabited and all elements are merely equal +### A type is 0-connected if it is inhabited and all elements are merely equal ```agda abstract @@ -144,8 +144,8 @@ abstract ```agda is-trunc-map-ev-is-connected : {l1 l2 : Level} (k : 𝕋) {A : UU l1} {B : UU l2} (a : A) → - is-0-connected A → is-trunc (succ-𝕋 k) B → - is-trunc-map k (ev-function B a) + is-0-connected A → is-trunc (succ-𝕋 k) B → + is-trunc-map k (ev-function B a) is-trunc-map-ev-is-connected k {A} {B} a H K = is-trunc-map-comp k ( ev-function B star) @@ -322,7 +322,7 @@ module _ is-0-connected-im-point a = is-0-connected-im-point' (point a) ``` -### Coproducts of -1-connected types are not 0-connected +### Coproducts of inhabited types are not 0-connected ```agda module _ From b978a5eb3ffeb8fedb31c5f07b03f23d7add08ea Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Fri, 7 Nov 2025 17:50:17 +0100 Subject: [PATCH 68/69] pre-commit --- src/foundation/0-connected-types.lagda.md | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index d69f2ed5ea..e2cf2b6486 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -10,23 +10,22 @@ module foundation.0-connected-types where open import foundation.action-on-identifications-functions open import foundation.constant-maps open import foundation.contractible-types +open import foundation.coproduct-types open import foundation.dependent-pair-types -open import foundation.fiber-inclusions +open import foundation.empty-types +open import foundation.equality-coproduct-types open import foundation.evaluation-functions +open import foundation.fiber-inclusions open import foundation.functoriality-set-truncation open import foundation.images open import foundation.inhabited-types open import foundation.mere-equality -open import foundation.empty-types +open import foundation.negation open import foundation.propositional-truncations open import foundation.set-truncations open import foundation.sets -open import foundation.negation -open import foundation.coproduct-types -open import foundation.equality-coproduct-types open import foundation.surjective-maps open import foundation.unit-type -open import foundation.mere-equality open import foundation.universal-property-contractible-types open import foundation.universal-property-unit-type open import foundation.universe-levels From 176187e3da951214c109d48916209a2aef12504b Mon Sep 17 00:00:00 2001 From: Fredrik Bakke Date: Mon, 10 Nov 2025 15:17:19 +0100 Subject: [PATCH 69/69] self-review --- src/foundation/0-connected-types.lagda.md | 127 +++++++++--------- src/foundation/connected-maps.lagda.md | 3 +- src/foundation/connected-types.lagda.md | 3 +- ...recomposition-dependent-functions.lagda.md | 2 +- ...property-family-of-fibers-of-maps.lagda.md | 104 +++++++------- .../universal-property-unit-type.lagda.md | 2 +- 6 files changed, 130 insertions(+), 111 deletions(-) diff --git a/src/foundation/0-connected-types.lagda.md b/src/foundation/0-connected-types.lagda.md index e2cf2b6486..a9f50ba0fe 100644 --- a/src/foundation/0-connected-types.lagda.md +++ b/src/foundation/0-connected-types.lagda.md @@ -264,46 +264,49 @@ is-0-connected-is-contr X p = ```agda module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) - where abstract + where - is-contr-im-map-trunc-Set-is-0-connected-domain' : - A → all-elements-merely-equal A → is-contr (im (map-trunc-Set f)) - is-contr-im-map-trunc-Set-is-0-connected-domain' a C = - is-contr-im - ( trunc-Set B) - ( unit-trunc-Set a) - ( apply-dependent-universal-property-trunc-Set' - ( λ x → - set-Prop - ( Id-Prop - ( trunc-Set B) - ( map-trunc-Set f x) - ( map-trunc-Set f (unit-trunc-Set a)))) - ( λ a' → - apply-universal-property-trunc-Prop - ( C a' a) - ( Id-Prop - ( trunc-Set B) - ( map-trunc-Set f (unit-trunc-Set a')) - ( map-trunc-Set f (unit-trunc-Set a))) - ( λ where refl → refl))) - - is-0-connected-im-is-0-connected-domain' : - A → all-elements-merely-equal A → is-0-connected (im f) - is-0-connected-im-is-0-connected-domain' a C = - is-contr-equiv' - ( im (map-trunc-Set f)) - ( equiv-trunc-im-Set f) - ( is-contr-im-map-trunc-Set-is-0-connected-domain' a C) - - is-0-connected-im-is-0-connected-domain : - is-0-connected A → is-0-connected (im f) - is-0-connected-im-is-0-connected-domain C = - apply-universal-property-trunc-Prop - ( is-inhabited-is-0-connected C) - ( is-contr-Prop _) - ( λ a → - is-0-connected-im-is-0-connected-domain' a (mere-eq-is-0-connected C)) + abstract + is-contr-im-map-trunc-Set-is-0-connected-domain' : + A → all-elements-merely-equal A → is-contr (im (map-trunc-Set f)) + is-contr-im-map-trunc-Set-is-0-connected-domain' a C = + is-contr-im + ( trunc-Set B) + ( unit-trunc-Set a) + ( apply-dependent-universal-property-trunc-Set' + ( λ x → + set-Prop + ( Id-Prop + ( trunc-Set B) + ( map-trunc-Set f x) + ( map-trunc-Set f (unit-trunc-Set a)))) + ( λ a' → + apply-universal-property-trunc-Prop + ( C a' a) + ( Id-Prop + ( trunc-Set B) + ( map-trunc-Set f (unit-trunc-Set a')) + ( map-trunc-Set f (unit-trunc-Set a))) + ( λ where refl → refl))) + + abstract + is-0-connected-im-is-0-connected-domain' : + A → all-elements-merely-equal A → is-0-connected (im f) + is-0-connected-im-is-0-connected-domain' a C = + is-contr-equiv' + ( im (map-trunc-Set f)) + ( equiv-trunc-im-Set f) + ( is-contr-im-map-trunc-Set-is-0-connected-domain' a C) + + abstract + is-0-connected-im-is-0-connected-domain : + is-0-connected A → is-0-connected (im f) + is-0-connected-im-is-0-connected-domain C = + apply-universal-property-trunc-Prop + ( is-inhabited-is-0-connected C) + ( is-contr-Prop _) + ( λ a → + is-0-connected-im-is-0-connected-domain' a (mere-eq-is-0-connected C)) ``` ### The image of a point is 0-connected @@ -311,14 +314,16 @@ module _ ```agda module _ {l1 : Level} {A : UU l1} - where abstract + where - is-0-connected-im-point' : (f : unit → A) → is-0-connected (im f) - is-0-connected-im-point' f = - is-0-connected-im-is-0-connected-domain f is-0-connected-unit + abstract + is-0-connected-im-point' : (f : unit → A) → is-0-connected (im f) + is-0-connected-im-point' f = + is-0-connected-im-is-0-connected-domain f is-0-connected-unit - is-0-connected-im-point : (a : A) → is-0-connected (im (point a)) - is-0-connected-im-point a = is-0-connected-im-point' (point a) + abstract + is-0-connected-im-point : (a : A) → is-0-connected (im (point a)) + is-0-connected-im-point a = is-0-connected-im-point' (point a) ``` ### Coproducts of inhabited types are not 0-connected @@ -326,20 +331,22 @@ module _ ```agda module _ {l1 l2 : Level} {A : UU l1} {B : UU l2} - where abstract + where - is-not-0-connected-coproduct-has-element : - A → B → ¬ is-0-connected (A + B) - is-not-0-connected-coproduct-has-element a b H = - apply-universal-property-trunc-Prop - ( mere-eq-is-0-connected H (inl a) (inr b)) - ( empty-Prop) - ( is-empty-eq-coproduct-inl-inr a b) - - is-not-0-connected-coproduct-is-inhabited : - is-inhabited A → is-inhabited B → ¬ is-0-connected (A + B) - is-not-0-connected-coproduct-is-inhabited |a| |b| = - apply-twice-universal-property-trunc-Prop |a| |b| - ( neg-type-Prop (is-0-connected (A + B))) - ( is-not-0-connected-coproduct-has-element) + abstract + is-not-0-connected-coproduct-has-element : + A → B → ¬ is-0-connected (A + B) + is-not-0-connected-coproduct-has-element a b H = + apply-universal-property-trunc-Prop + ( mere-eq-is-0-connected H (inl a) (inr b)) + ( empty-Prop) + ( is-empty-eq-coproduct-inl-inr a b) + + abstract + is-not-0-connected-coproduct-is-inhabited : + is-inhabited A → is-inhabited B → ¬ is-0-connected (A + B) + is-not-0-connected-coproduct-is-inhabited |a| |b| = + apply-twice-universal-property-trunc-Prop |a| |b| + ( neg-type-Prop (is-0-connected (A + B))) + ( is-not-0-connected-coproduct-has-element) ``` diff --git a/src/foundation/connected-maps.lagda.md b/src/foundation/connected-maps.lagda.md index 7843126958..7613972118 100644 --- a/src/foundation/connected-maps.lagda.md +++ b/src/foundation/connected-maps.lagda.md @@ -7,6 +7,7 @@ module foundation.connected-maps where
Imports ```agda +open import foundation.action-on-identifications-functions open import foundation.connected-types open import foundation.dependent-pair-types open import foundation.function-extensionality @@ -184,7 +185,7 @@ module _ htpy-eq-connected-map : (f g : connected-map k A B) → f = g → htpy-connected-map f g - htpy-eq-connected-map f .f refl = refl-htpy-connected-map f + htpy-eq-connected-map f g H = htpy-eq (ap pr1 H) is-equiv-htpy-eq-connected-map : (f g : connected-map k A B) → is-equiv (htpy-eq-connected-map f g) diff --git a/src/foundation/connected-types.lagda.md b/src/foundation/connected-types.lagda.md index c3371d5606..75ef75f011 100644 --- a/src/foundation/connected-types.lagda.md +++ b/src/foundation/connected-types.lagda.md @@ -271,7 +271,8 @@ is-connected-Σ k H K = ≃ ║A║ₖ by the right unit law of Σ ``` -and so, in particular, if the total space is `k`-connected so is the base. ∎ +and so, in particular, if the total space is `k`-connected then so is the base. +∎ ```agda is-connected-base : diff --git a/src/foundation/precomposition-dependent-functions.lagda.md b/src/foundation/precomposition-dependent-functions.lagda.md index 1afc8d566f..d1bcc3dbb4 100644 --- a/src/foundation/precomposition-dependent-functions.lagda.md +++ b/src/foundation/precomposition-dependent-functions.lagda.md @@ -94,7 +94,7 @@ module _ we compute the fiber family of dependent precomposition maps as a dependent product ```text - compute-fiber-Π-precomp-Π : + dependent-product-characterization-fiber-precomp-Π : fiber (precomp-Π f U) g ≃ ( (b : B) → Σ (u : U b), diff --git a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md index 1ae1eaa5e6..8c4000e1bd 100644 --- a/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md +++ b/src/foundation/universal-property-family-of-fibers-of-maps.lagda.md @@ -538,15 +538,17 @@ module _ (g : (a : A) → U (f a)) where - family-fiber-Π-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp-Π b = + family-dependent-product-fiber-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) + family-dependent-product-fiber-precomp-Π b = Σ (U b) (λ u → ((a , p) : fiber f b) → dependent-identification U p (g a) u) - fiber-Π-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp-Π = (b : B) → family-fiber-Π-precomp-Π b + dependent-product-fiber-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) + dependent-product-fiber-precomp-Π = + (b : B) → family-dependent-product-fiber-precomp-Π b - compute-fiber-Π-precomp-Π : fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π - compute-fiber-Π-precomp-Π = + dependent-product-characterization-fiber-precomp-Π : + fiber (precomp-Π f U) g ≃ dependent-product-fiber-precomp-Π + dependent-product-characterization-fiber-precomp-Π = equivalence-reasoning fiber (precomp-Π f U) g ≃ Σ ((b : B) → U b) (λ h → (a : A) → g a = (h ∘ f) a) @@ -565,30 +567,33 @@ module _ ((a , p) : fiber f b) → dependent-identification U p (g a) u)) by inv-distributive-Π-Σ - family-fiber-Π-curry-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-curry-precomp-Π b = + family-fiber-dependent-product-curry-precomp-Π : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-dependent-product-curry-precomp-Π b = Σ ( U b) ( λ u → (a : A) (p : f a = b) → dependent-identification U p (g a) u) - fiber-Π-curry-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-curry-precomp-Π = (b : B) → family-fiber-Π-curry-precomp-Π b + fiber-dependent-product-curry-precomp-Π : UU (l1 ⊔ l2 ⊔ l3) + fiber-dependent-product-curry-precomp-Π = + (b : B) → family-fiber-dependent-product-curry-precomp-Π b - compute-fiber-Π-curry-precomp-Π : - fiber (precomp-Π f U) g ≃ fiber-Π-curry-precomp-Π - compute-fiber-Π-curry-precomp-Π = + curried-dependent-product-characterization-fiber-precomp-Π : + fiber (precomp-Π f U) g ≃ fiber-dependent-product-curry-precomp-Π + curried-dependent-product-characterization-fiber-precomp-Π = ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e - ( compute-fiber-Π-precomp-Π) + ( dependent-product-characterization-fiber-precomp-Π) - family-fiber-Π-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp-Π' b = + family-dependent-product-fiber-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) + family-dependent-product-fiber-precomp-Π' b = Σ ( U b) ( λ u → ((a , p) : fiber' f b) → dependent-identification U p u (g a)) - fiber-Π-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp-Π' = (b : B) → family-fiber-Π-precomp-Π' b + dependent-product-fiber-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) + dependent-product-fiber-precomp-Π' = + (b : B) → family-dependent-product-fiber-precomp-Π' b - compute-fiber-Π-precomp-Π' : fiber (precomp-Π f U) g ≃ fiber-Π-precomp-Π' - compute-fiber-Π-precomp-Π' = + dependent-product-characterization-fiber-precomp-Π' : + fiber (precomp-Π f U) g ≃ dependent-product-fiber-precomp-Π' + dependent-product-characterization-fiber-precomp-Π' = equivalence-reasoning fiber (precomp-Π f U) g ≃ Σ ((b : B) → U b) (λ h → (a : A) → (h ∘ f) a = g a) @@ -607,21 +612,22 @@ module _ ((a , p) : fiber' f b) → dependent-identification U p u (g a))) by inv-distributive-Π-Σ - family-fiber-Π-curry-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-curry-precomp-Π' b = + family-fiber-dependent-product-curry-precomp-Π' : B → UU (l1 ⊔ l2 ⊔ l3) + family-fiber-dependent-product-curry-precomp-Π' b = Σ (U b) (λ u → (a : A) (p : b = f a) → dependent-identification U p u (g a)) - fiber-Π-curry-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-curry-precomp-Π' = (b : B) → family-fiber-Π-curry-precomp-Π' b + fiber-dependent-product-curry-precomp-Π' : UU (l1 ⊔ l2 ⊔ l3) + fiber-dependent-product-curry-precomp-Π' = + (b : B) → family-fiber-dependent-product-curry-precomp-Π' b - compute-fiber-Π-curry-precomp-Π' : - fiber (precomp-Π f U) g ≃ fiber-Π-curry-precomp-Π' - compute-fiber-Π-curry-precomp-Π' = + curried-dependent-product-characterization-fiber-precomp-Π' : + fiber (precomp-Π f U) g ≃ fiber-dependent-product-curry-precomp-Π' + curried-dependent-product-characterization-fiber-precomp-Π' = ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e - ( compute-fiber-Π-precomp-Π') + ( dependent-product-characterization-fiber-precomp-Π') ``` -### Computing the fibers of precomposition functions as dependent products +### Fibers of precomposition functions as dependent products We give four equivalences for the fibers of precomposition functions as dependent products: @@ -640,15 +646,17 @@ module _ (g : A → U) where - family-fiber-Π-precomp : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp b = + family-dependent-product-fiber-precomp : B → UU (l1 ⊔ l2 ⊔ l3) + family-dependent-product-fiber-precomp b = Σ U (λ u → ((a , _) : fiber f b) → g a = u) - fiber-Π-precomp : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp = (b : B) → family-fiber-Π-precomp b + dependent-product-fiber-precomp : UU (l1 ⊔ l2 ⊔ l3) + dependent-product-fiber-precomp = + (b : B) → family-dependent-product-fiber-precomp b - compute-fiber-Π-precomp : fiber (precomp f U) g ≃ fiber-Π-precomp - compute-fiber-Π-precomp = + dependent-product-characterization-fiber-precomp : + fiber (precomp f U) g ≃ dependent-product-fiber-precomp + dependent-product-characterization-fiber-precomp = equivalence-reasoning fiber (precomp f U) g ≃ Σ (B → U) (λ h → (a : A) → g a = (h ∘ f) a) @@ -663,21 +671,23 @@ module _ ≃ ( (b : B) → Σ U (λ u → ((a , _) : fiber f b) → g a = u)) by inv-distributive-Π-Σ - compute-fiber-Π-curry-precomp : + curried-dependent-product-characterization-fiber-precomp : fiber (precomp f U) g ≃ ((b : B) → Σ U (λ u → (a : A) → f a = b → g a = u)) - compute-fiber-Π-curry-precomp = + curried-dependent-product-characterization-fiber-precomp = ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e - ( compute-fiber-Π-precomp) + ( dependent-product-characterization-fiber-precomp) - family-fiber-Π-precomp' : B → UU (l1 ⊔ l2 ⊔ l3) - family-fiber-Π-precomp' b = + family-dependent-product-fiber-precomp' : B → UU (l1 ⊔ l2 ⊔ l3) + family-dependent-product-fiber-precomp' b = Σ U (λ u → ((a , _) : fiber' f b) → u = g a) - fiber-Π-precomp' : UU (l1 ⊔ l2 ⊔ l3) - fiber-Π-precomp' = (b : B) → family-fiber-Π-precomp' b + dependent-product-fiber-precomp' : UU (l1 ⊔ l2 ⊔ l3) + dependent-product-fiber-precomp' = + (b : B) → family-dependent-product-fiber-precomp' b - compute-fiber-Π-precomp' : fiber (precomp f U) g ≃ fiber-Π-precomp' - compute-fiber-Π-precomp' = + dependent-product-characterization-fiber-precomp' : + fiber (precomp f U) g ≃ dependent-product-fiber-precomp' + dependent-product-characterization-fiber-precomp' = equivalence-reasoning fiber (precomp f U) g ≃ Σ (B → U) (λ h → (h ∘ f) ~ g) @@ -692,9 +702,9 @@ module _ ≃ ( (b : B) → Σ U (λ u → ((a , _) : fiber' f b) → u = g a)) by inv-distributive-Π-Σ - compute-fiber-Π-curry-precomp' : + curried-dependent-product-characterization-fiber-precomp' : fiber (precomp f U) g ≃ ((b : B) → Σ U (λ u → (a : A) → b = f a → u = g a)) - compute-fiber-Π-curry-precomp' = + curried-dependent-product-characterization-fiber-precomp' = ( equiv-Π-equiv-family (λ b → equiv-tot (λ u → equiv-ev-pair))) ∘e - ( compute-fiber-Π-precomp') + ( dependent-product-characterization-fiber-precomp') ``` diff --git a/src/foundation/universal-property-unit-type.lagda.md b/src/foundation/universal-property-unit-type.lagda.md index bf7b958180..8579a1bb27 100644 --- a/src/foundation/universal-property-unit-type.lagda.md +++ b/src/foundation/universal-property-unit-type.lagda.md @@ -142,7 +142,7 @@ abstract ( refl-htpy) ``` -### Precomposition with the terminal map is equivalent to the diagonal exponential +### Precomposition with the terminal map is an equivalence if and only if the diagonal exponential is ```agda is-equiv-precomp-terminal-map-is-equiv-diagonal-exponential :