Skip to content

Commit

Permalink
resolve merge conflicts
Browse files Browse the repository at this point in the history
  • Loading branch information
EgbertRijke committed Nov 20, 2024
2 parents ef70c90 + 76c5d87 commit 3013c3f
Show file tree
Hide file tree
Showing 42 changed files with 532 additions and 172 deletions.
14 changes: 8 additions & 6 deletions scripts/preprocessor_git_metadata.py
Original file line number Diff line number Diff line change
Expand Up @@ -84,12 +84,14 @@ def get_author_element_for_file(filename, include_contributors, contributors):
'HEAD', '--', filename
], capture_output=True, text=True, check=True).stdout.splitlines()

# Collect authors and sort by number of commits
author_names = [
author['displayName']
for author in sorted_authors_from_raw_shortlog_lines(raw_authors_git_output, contributors)
]
attribution_text = f'<p><i>Content created by {format_multiple_authors_attribution(author_names)}.</i></p>'
# If all commits to a file are chore commits, then there are no authors
if raw_authors_git_output:
# Collect authors and sort by number of commits
author_names = [
author['displayName']
for author in sorted_authors_from_raw_shortlog_lines(raw_authors_git_output, contributors)
]
attribution_text = f'<p><i>Content created by {format_multiple_authors_attribution(author_names)}.</i></p>'

file_log_output = subprocess.run([
'git', 'log',
Expand Down
18 changes: 18 additions & 0 deletions src/foundation-core/contractible-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -133,6 +133,24 @@ module _
is-equiv-is-contr _ is-contr-A is-contr-B
```

### Contractibility of the base of a contractible dependent sum

Given a type `A` and a type family over it `B`, then if the dependent sum
`Σ A B` is contractible, it follows that if there is a section `(x : A) B x`
then `A` is contractible.

```agda
module _
{l1 l2 : Level} (A : UU l1) (B : A UU l2)
where

abstract
is-contr-base-is-contr-Σ' :
((x : A) B x) is-contr (Σ A B) is-contr A
is-contr-base-is-contr-Σ' s =
is-contr-retract-of (Σ A B) ((λ a a , s a) , pr1 , refl-htpy)
```

### Contractibility of cartesian product types

Given two types `A` and `B`, the following are equivalent:
Expand Down
95 changes: 93 additions & 2 deletions src/foundation-core/functoriality-dependent-pair-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -7,8 +7,10 @@ module foundation-core.functoriality-dependent-pair-types where
<details><summary>Imports</summary>

```agda
open import foundation.action-on-identifications-dependent-functions
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.transport-along-identifications
open import foundation.universe-levels

open import foundation-core.contractible-maps
Expand All @@ -23,7 +25,7 @@ open import foundation-core.homotopies
open import foundation-core.identity-types
open import foundation-core.retractions
open import foundation-core.retracts-of-types
open import foundation-core.transport-along-identifications
open import foundation-core.sections
```

</details>
Expand Down Expand Up @@ -307,6 +309,78 @@ module _
is-equiv-fiber-map-Σ-map-base-fiber t
```

### The fibers of `map-Σ`

We compute the fibers of `map-Σ` first in terms of `fiber'` and then in terms of
`fiber`.

```agda
module _
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A UU l3} (D : B UU l4)
(f : A B) (g : (x : A) C x D (f x)) (t : Σ B D)
where

fiber-map-Σ' : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
fiber-map-Σ' =
Σ (fiber' f (pr1 t)) (λ s fiber' (g (pr1 s)) (tr D (pr2 s) (pr2 t)))

map-fiber-map-Σ' : fiber' (map-Σ D f g) t fiber-map-Σ'
map-fiber-map-Σ' ((a , c) , refl) = (a , refl) , (c , refl)

map-inv-fiber-map-Σ' : fiber-map-Σ' fiber' (map-Σ D f g) t
map-inv-fiber-map-Σ' ((a , p) , (c , q)) = ((a , c) , eq-pair-Σ p q)

abstract
is-section-map-inv-fiber-map-Σ' :
is-section map-fiber-map-Σ' map-inv-fiber-map-Σ'
is-section-map-inv-fiber-map-Σ' ((a , refl) , (c , refl)) = refl

abstract
is-retraction-map-inv-fiber-map-Σ' :
is-retraction map-fiber-map-Σ' map-inv-fiber-map-Σ'
is-retraction-map-inv-fiber-map-Σ' ((a , c) , refl) = refl

is-equiv-map-fiber-map-Σ' : is-equiv map-fiber-map-Σ'
is-equiv-map-fiber-map-Σ' =
is-equiv-is-invertible
map-inv-fiber-map-Σ'
is-section-map-inv-fiber-map-Σ'
is-retraction-map-inv-fiber-map-Σ'

compute-fiber-map-Σ' : fiber' (map-Σ D f g) t ≃ fiber-map-Σ'
compute-fiber-map-Σ' = (map-fiber-map-Σ' , is-equiv-map-fiber-map-Σ')

fiber-map-Σ : UU (l1 ⊔ l2 ⊔ l3 ⊔ l4)
fiber-map-Σ =
Σ (fiber f (pr1 t)) (λ s fiber (g (pr1 s)) (inv-tr D (pr2 s) (pr2 t)))

map-fiber-map-Σ : fiber (map-Σ D f g) t fiber-map-Σ
map-fiber-map-Σ ((a , c) , refl) = (a , refl) , (c , refl)

map-inv-fiber-map-Σ : fiber-map-Σ fiber (map-Σ D f g) t
map-inv-fiber-map-Σ ((a , refl) , c , refl) = ((a , c) , refl)

abstract
is-section-map-inv-fiber-map-Σ :
is-section map-fiber-map-Σ map-inv-fiber-map-Σ
is-section-map-inv-fiber-map-Σ ((a , refl) , (c , refl)) = refl

abstract
is-retraction-map-inv-fiber-map-Σ :
is-retraction map-fiber-map-Σ map-inv-fiber-map-Σ
is-retraction-map-inv-fiber-map-Σ ((a , c) , refl) = refl

is-equiv-map-fiber-map-Σ : is-equiv map-fiber-map-Σ
is-equiv-map-fiber-map-Σ =
is-equiv-is-invertible
map-inv-fiber-map-Σ
is-section-map-inv-fiber-map-Σ
is-retraction-map-inv-fiber-map-Σ

compute-fiber-map-Σ : fiber (map-Σ D f g) t ≃ fiber-map-Σ
compute-fiber-map-Σ = (map-fiber-map-Σ , is-equiv-map-fiber-map-Σ)
```

### If `f : A B` is a contractible map, then so is `map-Σ-map-base f C`

```agda
Expand Down Expand Up @@ -457,10 +531,27 @@ module _
{l1 l2 : Level} {A : UU l1} {B : A UU l2}
where

compute-map-Σ-id : map-Σ B id (λ x id) ~ id
compute-map-Σ-id : map-Σ B id (λ _ id) ~ id
compute-map-Σ-id = refl-htpy
```

### `map-Σ` preserves composition of maps

```agda
module _
{l1 l2 l3 l4 l5 l6 : Level}
{A : UU l1} {B : UU l2} {C : UU l3}
{A' : A UU l4} {B' : B UU l5} {C' : C UU l6}
{f : A B} {f' : (x : A) A' x B' (f x)}
{g : B C} {g' : (y : B) B' y C' (g y)}
where

preserves-comp-map-Σ :
map-Σ C' (g ∘ f) (λ x x' g' (f x) (f' x x')) ~
map-Σ C' g g' ∘ map-Σ B' f f'
preserves-comp-map-Σ = refl-htpy
```

### Computing the action on identifications of `tot`

```agda
Expand Down
19 changes: 19 additions & 0 deletions src/foundation-core/propositions.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -183,6 +183,25 @@ pr2 (Σ-Prop P Q) =
( λ p is-prop-type-Prop (Q p))
```

### If `Σ A B` is a proposition and there is a section `(x : A) B x` then `A` is a proposition

```agda
module _
{l1 l2 : Level} {A : UU l1} {B : A UU l2} (s : (x : A) B x)
where

is-proof-irrelevant-base-is-proof-irrelevant-Σ' :
is-proof-irrelevant (Σ A B) is-proof-irrelevant A
is-proof-irrelevant-base-is-proof-irrelevant-Σ' H a =
is-contr-base-is-contr-Σ' A B s (H (a , s a))

is-prop-base-is-prop-Σ' : is-prop (Σ A B) is-prop A
is-prop-base-is-prop-Σ' H =
is-prop-is-proof-irrelevant
( is-proof-irrelevant-base-is-proof-irrelevant-Σ'
( is-proof-irrelevant-is-prop H))
```

### Propositions are closed under cartesian product types

```agda
Expand Down
19 changes: 19 additions & 0 deletions src/foundation-core/truncated-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,7 @@ module foundation-core.truncated-types where
<details><summary>Imports</summary>

```agda
open import foundation.action-on-identifications-dependent-functions
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.equality-cartesian-product-types
Expand Down Expand Up @@ -217,6 +218,24 @@ fiber-Truncated-Type A B f b =
Σ-Truncated-Type A (λ a Id-Truncated-Type' B (f a) b)
```

### Truncatedness of the base of a truncated dependent sum

```agda
abstract
is-trunc-base-is-trunc-Σ' :
{l1 l2 : Level} {k : 𝕋} {A : UU l1} {B : A UU l2}
((x : A) B x) is-trunc k (Σ A B) is-trunc k A
is-trunc-base-is-trunc-Σ' {k = neg-two-𝕋} {A} {B} =
is-contr-base-is-contr-Σ' A B
is-trunc-base-is-trunc-Σ' {k = succ-𝕋 k} s is-trunc-ΣAB x y =
is-trunc-base-is-trunc-Σ'
( apd s)
( is-trunc-equiv' k
( (x , s x) = (y , s y))
( equiv-pair-eq-Σ (x , s x) (y , s y))
( is-trunc-ΣAB (x , s x) (y , s y)))
```

### Products of truncated types are truncated

```agda
Expand Down
2 changes: 1 addition & 1 deletion src/foundation.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -354,7 +354,7 @@ open import foundation.retractions public
open import foundation.retracts-of-maps public
open import foundation.retracts-of-types public
open import foundation.sections public
open import foundation.separated-types public
open import foundation.separated-types-subuniverses public
open import foundation.sequences public
open import foundation.sequential-limits public
open import foundation.set-presented-types public
Expand Down
2 changes: 1 addition & 1 deletion src/foundation/continuations.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -35,8 +35,8 @@ open import foundation-core.sections
open import foundation-core.transport-along-identifications

open import orthogonal-factorization-systems.extensions-of-maps
open import orthogonal-factorization-systems.local-types
open import orthogonal-factorization-systems.modal-operators
open import orthogonal-factorization-systems.types-local-at-maps
open import orthogonal-factorization-systems.uniquely-eliminating-modalities
```

Expand Down
8 changes: 3 additions & 5 deletions src/foundation/contractible-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -157,11 +157,9 @@ module _

is-contr-Σ-is-prop :
((x : A) is-prop (B x)) ((x : A) B x a = x) is-contr (Σ A B)
pr1 (is-contr-Σ-is-prop p f) = pair a b
pr2 (is-contr-Σ-is-prop p f) (pair x y) =
eq-type-subtype
( λ x' pair (B x') (p x'))
( f x y)
pr1 (is-contr-Σ-is-prop p f) = a , b
pr2 (is-contr-Σ-is-prop p f) (x , y) =
eq-type-subtype (λ x' B x' , p x') (f x y)
```

### The diagonal of contractible types
Expand Down
4 changes: 4 additions & 0 deletions src/foundation/double-negation-modality.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -16,10 +16,14 @@ open import foundation.propositions
open import foundation.unit-type
open import foundation.universe-levels

open import foundation-core.function-types
open import foundation-core.transport-along-identifications

open import orthogonal-factorization-systems.continuation-modalities
open import orthogonal-factorization-systems.large-lawvere-tierney-topologies
open import orthogonal-factorization-systems.lawvere-tierney-topologies
open import orthogonal-factorization-systems.modal-operators
open import orthogonal-factorization-systems.types-local-at-maps
open import orthogonal-factorization-systems.uniquely-eliminating-modalities
```

Expand Down
Original file line number Diff line number Diff line change
Expand Up @@ -35,7 +35,7 @@ The **functorial action** of the
`f : A → B` and `g : C D` and returns a map

```text
f × g : A × B C × D`
f × g : A × B C × D
```

between the cartesian product types. This functorial action is _bifunctorial_ in
Expand Down
10 changes: 5 additions & 5 deletions src/foundation/global-subuniverses.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -116,7 +116,7 @@ module _

## Properties

### Global subuniverses are closed under homogenous equivalences
### Global subuniverses are closed under equivalences between types in a single universe

This is true for any family of subuniverses indexed by universe levels.

Expand All @@ -126,14 +126,14 @@ module _
{l : Level} {X Y : UU l}
where

is-in-global-subuniverse-homogenous-equiv :
is-in-global-subuniverse-equiv-Level :
X ≃ Y is-in-global-subuniverse P X is-in-global-subuniverse P Y
is-in-global-subuniverse-homogenous-equiv =
is-in-global-subuniverse-equiv-Level =
is-in-subuniverse-equiv (subuniverse-global-subuniverse P l)

is-in-global-subuniverse-homogenous-equiv' :
is-in-global-subuniverse-equiv-Level' :
X ≃ Y is-in-global-subuniverse P Y is-in-global-subuniverse P X
is-in-global-subuniverse-homogenous-equiv' =
is-in-global-subuniverse-equiv-Level' =
is-in-subuniverse-equiv' (subuniverse-global-subuniverse P l)
```

Expand Down
24 changes: 24 additions & 0 deletions src/foundation/inhabited-types.lagda.md
Original file line number Diff line number Diff line change
Expand Up @@ -7,9 +7,11 @@ module foundation.inhabited-types where
<details><summary>Imports</summary>

```agda
open import foundation.action-on-identifications-functions
open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equality-dependent-function-types
open import foundation.function-extensionality
open import foundation.functoriality-propositional-truncation
open import foundation.fundamental-theorem-of-identity-types
open import foundation.propositional-truncations
Expand All @@ -18,6 +20,7 @@ open import foundation.univalence
open import foundation.universe-levels

open import foundation-core.equivalences
open import foundation-core.homotopies
open import foundation-core.identity-types
open import foundation-core.propositions
open import foundation-core.torsorial-type-families
Expand Down Expand Up @@ -220,6 +223,27 @@ is-contr-is-inhabited-is-prop {A = A} p h =
( λ a a , eq-is-prop' p a)
```

### Contractibility of the base of a dependent sum

Given a type `A` and a type family over it `B`, then if the dependent sum
`Σ A B` is contractible, it follows that if there merely exists a section
`(x : A) B x`, then `A` is contractible.

```agda
module _
{l1 l2 : Level} (A : UU l1) (B : A UU l2)
where

abstract
is-contr-base-is-contr-Σ :
is-inhabited ((x : A) B x) is-contr (Σ A B) is-contr A
is-contr-base-is-contr-Σ s is-contr-ΣAB =
rec-trunc-Prop
( is-contr-Prop A)
( λ s is-contr-base-is-contr-Σ' A B s is-contr-ΣAB)
( s)
```

## See also

- The notion of _nonempty types_ is treated in
Expand Down
Loading

0 comments on commit 3013c3f

Please sign in to comment.