diff --git a/references.bib b/references.bib
index aef4fe98c3..2c7232900e 100644
--- a/references.bib
+++ b/references.bib
@@ -285,6 +285,25 @@ @article{EKMS93
keywords = {axiality,closure operation,Galois connection,interior operation,polarity}
}
+@article{Esc08,
+ author = {Escardó, Martín Hötzel},
+ title = {Exhaustible sets in higher-type computation},
+ journal = {Logical Methods in Computer Science},
+ volume = {4},
+ year = {2008},
+ month = {8},
+ number = {3},
+ issue = {3},
+ publisher = {Centre pour la Communication Scientifique Directe (CCSD)},
+ pages = {3:3, 37},
+ issn = {1860-5974},
+ doi = {10.2168/LMCS-4(3:3)2008},
+ eprint = {0808.0441},
+ eprinttype = {arxiv},
+ eprintclass = {cs},
+ primaryclass = {cs.LO}
+}
+
@online{Esc19DefinitionsEquivalence,
title = {Definitions of Equivalence Satisfying Judgmental/Strict Groupoid Laws?},
author = {Escardó, Martín Hötzel},
@@ -501,6 +520,22 @@ @online{MR23
keywords = {20B30 03B15,Mathematics - Group Theory,Mathematics - Logic}
}
+@article{Esc21b,
+ author = {Escardó, Martín Hötzel},
+ title = {Injective types in univalent mathematics},
+ journal = {Mathematical Structures in Computer Science},
+ volume = {31},
+ year = {2021},
+ number = {1},
+ pages = {89--111},
+ issn = {0960-1295,1469-8072},
+ doi = {10.1017/S0960129520000225},
+ eprint = {1903.01211},
+ archiveprefix = {arXiv},
+ primaryclass = {math.CT},
+ url = {https://martinescardo.github.io/TypeTopology/InjectiveTypes.Article.html}
+}
+
@book{MRR88,
title = {A {{Course}} in {{Constructive Algebra}}},
author = {Mines, Ray and Richman, Fred and Ruitenburg, Wim},
diff --git a/src/foundation-core/decidable-propositions.lagda.md b/src/foundation-core/decidable-propositions.lagda.md
index 3f6eb2e2ce..eb48a26b7d 100644
--- a/src/foundation-core/decidable-propositions.lagda.md
+++ b/src/foundation-core/decidable-propositions.lagda.md
@@ -13,10 +13,12 @@ open import foundation.dependent-pair-types
open import foundation.double-negation
open import foundation.negation
open import foundation.propositional-truncations
+open import foundation.transport-along-identifications
open import foundation.unit-type
open import foundation.universe-levels
open import foundation-core.cartesian-product-types
+open import foundation-core.contractible-types
open import foundation-core.empty-types
open import foundation-core.function-types
open import foundation-core.functoriality-dependent-pair-types
@@ -32,14 +34,11 @@ A {{#concept "decidable proposition" Agda=is-decidable-Prop}} is a
[proposition](foundation-core.propositions.md) that has a
[decidable](foundation.decidable-types.md) underlying type.
-## Definition
+## Definitions
-### The subtype of decidable propositions
+### The property of a proposition of being decidable
```agda
-is-decidable-prop : {l : Level} → UU l → UU l
-is-decidable-prop A = is-prop A × is-decidable A
-
is-prop-is-decidable :
{l : Level} {A : UU l} → is-prop A → is-prop (is-decidable A)
is-prop-is-decidable is-prop-A =
@@ -50,6 +49,16 @@ is-decidable-Prop :
pr1 (is-decidable-Prop P) = is-decidable (type-Prop P)
pr2 (is-decidable-Prop P) = is-prop-is-decidable (is-prop-type-Prop P)
+is-decidable-type-Prop : {l : Level} → Prop l → UU l
+is-decidable-type-Prop P = is-decidable (type-Prop P)
+```
+
+### The subuniverse of decidable propositions
+
+```agda
+is-decidable-prop : {l : Level} → UU l → UU l
+is-decidable-prop A = is-prop A × is-decidable A
+
is-prop-is-decidable-prop :
{l : Level} (X : UU l) → is-prop (is-decidable-prop X)
is-prop-is-decidable-prop X =
@@ -63,6 +72,16 @@ is-decidable-prop-Prop :
{l : Level} (A : UU l) → Prop l
pr1 (is-decidable-prop-Prop A) = is-decidable-prop A
pr2 (is-decidable-prop-Prop A) = is-prop-is-decidable-prop A
+
+module _
+ {l : Level} {A : UU l} (H : is-decidable-prop A)
+ where
+
+ is-prop-type-is-decidable-prop : is-prop A
+ is-prop-type-is-decidable-prop = pr1 H
+
+ is-decidable-type-is-decidable-prop : is-decidable A
+ is-decidable-type-is-decidable-prop = pr2 H
```
### Decidable propositions
@@ -93,7 +112,8 @@ module _
is-decidable-prop-type-Decidable-Prop = pr2 P
is-decidable-prop-Decidable-Prop : Prop l
- pr1 is-decidable-prop-Decidable-Prop = is-decidable type-Decidable-Prop
+ pr1 is-decidable-prop-Decidable-Prop =
+ is-decidable type-Decidable-Prop
pr2 is-decidable-prop-Decidable-Prop =
is-prop-is-decidable is-prop-type-Decidable-Prop
```
@@ -110,6 +130,14 @@ pr1 empty-Decidable-Prop = empty
pr2 empty-Decidable-Prop = is-decidable-prop-empty
```
+### Empty types are decidable propositions
+
+```agda
+is-decidable-prop-is-empty :
+ {l : Level} {A : UU l} → is-empty A → is-decidable-prop A
+is-decidable-prop-is-empty H = is-prop-is-empty H , inr H
+```
+
### The unit type is a decidable proposition
```agda
@@ -122,6 +150,14 @@ pr1 unit-Decidable-Prop = unit
pr2 unit-Decidable-Prop = is-decidable-prop-unit
```
+### Contractible types are decidable propositions
+
+```agda
+is-decidable-prop-is-contr :
+ {l : Level} {A : UU l} → is-contr A → is-decidable-prop A
+is-decidable-prop-is-contr H = is-prop-is-contr H , inl (center H)
+```
+
### The product of two decidable propositions is a decidable proposition
```agda
@@ -154,6 +190,42 @@ module _
pr2 product-Decidable-Prop = is-decidable-prop-product-Decidable-Prop
```
+### The dependent sum of a family of decidable propositions over a decidable proposition
+
+```agda
+module _
+ {l1 l2 : Level} {P : UU l1} {Q : P → UU l2}
+ (H : is-decidable-prop P) (K : (x : P) → is-decidable-prop (Q x))
+ where
+
+ is-prop-is-decidable-prop-Σ : is-prop (Σ P Q)
+ is-prop-is-decidable-prop-Σ =
+ is-prop-Σ
+ ( is-prop-type-is-decidable-prop H)
+ ( is-prop-type-is-decidable-prop ∘ K)
+
+ is-decidable-is-decidable-prop-Σ : is-decidable (Σ P Q)
+ is-decidable-is-decidable-prop-Σ =
+ rec-coproduct
+ ( λ x →
+ rec-coproduct
+ ( λ y → inl (x , y))
+ ( λ ny →
+ inr
+ ( λ xy →
+ ny
+ ( tr Q
+ ( eq-is-prop (is-prop-type-is-decidable-prop H))
+ ( pr2 xy))))
+ ( is-decidable-type-is-decidable-prop (K x)))
+ ( λ nx → inr (λ xy → nx (pr1 xy)))
+ ( is-decidable-type-is-decidable-prop H)
+
+ is-decidable-prop-Σ : is-decidable-prop (Σ P Q)
+ is-decidable-prop-Σ =
+ ( is-prop-is-decidable-prop-Σ , is-decidable-is-decidable-prop-Σ)
+```
+
### The negation operation on decidable propositions
```agda
diff --git a/src/foundation-core/embeddings.lagda.md b/src/foundation-core/embeddings.lagda.md
index 3eac0587bb..6fe6dabc72 100644
--- a/src/foundation-core/embeddings.lagda.md
+++ b/src/foundation-core/embeddings.lagda.md
@@ -30,7 +30,9 @@ better behaved homotopically than
being an equivalence is a [property](foundation-core.propositions.md) under
[function extensionality](foundation.function-extensionality.md).
-## Definition
+## Definitions
+
+### The predicate on maps of being embeddings
```agda
module _
@@ -48,7 +50,11 @@ module _
inv-equiv-ap-is-emb :
{f : A → B} (e : is-emb f) {x y : A} → (f x = f y) ≃ (x = y)
inv-equiv-ap-is-emb e = inv-equiv (equiv-ap-is-emb e)
+```
+### The type of embeddings
+
+```agda
infix 5 _↪_
_↪_ :
{l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2)
@@ -73,6 +79,29 @@ module _
inv-equiv-ap-emb e = inv-equiv (equiv-ap-emb e)
```
+### The type of embeddings into a type
+
+```agda
+Emb : (l1 : Level) {l2 : Level} (B : UU l2) → UU (lsuc l1 ⊔ l2)
+Emb l1 B = Σ (UU l1) (λ A → A ↪ B)
+
+module _
+ {l1 l2 : Level} {B : UU l2} (f : Emb l1 B)
+ where
+
+ type-Emb : UU l1
+ type-Emb = pr1 f
+
+ emb-Emb : type-Emb ↪ B
+ emb-Emb = pr2 f
+
+ map-Emb : type-Emb → B
+ map-Emb = map-emb emb-Emb
+
+ is-emb-map-Emb : is-emb map-Emb
+ is-emb-map-Emb = is-emb-map-emb emb-Emb
+```
+
## Examples
### The identity map is an embedding
diff --git a/src/foundation-core/equality-dependent-pair-types.lagda.md b/src/foundation-core/equality-dependent-pair-types.lagda.md
index 6b6374433d..9dacb2ce72 100644
--- a/src/foundation-core/equality-dependent-pair-types.lagda.md
+++ b/src/foundation-core/equality-dependent-pair-types.lagda.md
@@ -45,11 +45,18 @@ module _
```agda
refl-Eq-Σ : (s : Σ A B) → Eq-Σ s s
- pr1 (refl-Eq-Σ (pair a b)) = refl
- pr2 (refl-Eq-Σ (pair a b)) = refl
+ refl-Eq-Σ s = refl , refl
+
+ eq-base-eq-pair : {s t : Σ A B} → s = t → pr1 s = pr1 t
+ eq-base-eq-pair = ap pr1
+
+ dependent-identification-eq-pair :
+ {s t : Σ A B} (p : s = t) →
+ dependent-identification B (eq-base-eq-pair p) (pr2 s) (pr2 t)
+ dependent-identification-eq-pair {s} p = tr-ap pr1 (λ x _ → pr2 x) p (pr1 s)
pair-eq-Σ : {s t : Σ A B} → s = t → Eq-Σ s t
- pair-eq-Σ {s} refl = refl-Eq-Σ s
+ pair-eq-Σ p = eq-base-eq-pair p , dependent-identification-eq-pair p
eq-pair-eq-base :
{x y : A} {s : B x} (p : x = y) → (x , s) = (y , tr B p s)
@@ -78,11 +85,11 @@ module _
is-retraction-pair-eq-Σ :
(s t : Σ A B) → pair-eq-Σ {s} {t} ∘ eq-pair-Σ' {s} {t} ~ id {A = Eq-Σ s t}
- is-retraction-pair-eq-Σ (pair x y) (pair .x .y) (pair refl refl) = refl
+ is-retraction-pair-eq-Σ (x , y) (.x , .y) (refl , refl) = refl
is-section-pair-eq-Σ :
- (s t : Σ A B) → ((eq-pair-Σ' {s} {t}) ∘ (pair-eq-Σ {s} {t})) ~ id
- is-section-pair-eq-Σ (pair x y) .(pair x y) refl = refl
+ (s t : Σ A B) → eq-pair-Σ' {s} {t} ∘ pair-eq-Σ {s} {t} ~ id
+ is-section-pair-eq-Σ (x , y) .(x , y) refl = refl
abstract
is-equiv-eq-pair-Σ : (s t : Σ A B) → is-equiv (eq-pair-Σ' {s} {t})
@@ -110,14 +117,6 @@ module _
η-pair : (t : Σ A B) → (pair (pr1 t) (pr2 t)) = t
η-pair t = refl
-
- eq-base-eq-pair-Σ : {s t : Σ A B} → (s = t) → (pr1 s = pr1 t)
- eq-base-eq-pair-Σ p = pr1 (pair-eq-Σ p)
-
- dependent-eq-family-eq-pair-Σ :
- {s t : Σ A B} → (p : s = t) →
- dependent-identification B (eq-base-eq-pair-Σ p) (pr2 s) (pr2 t)
- dependent-eq-family-eq-pair-Σ p = pr2 (pair-eq-Σ p)
```
### Lifting equality to the total space
@@ -128,7 +127,7 @@ module _
where
lift-eq-Σ :
- {x y : A} (p : x = y) (b : B x) → (pair x b) = (pair y (tr B p b))
+ {x y : A} (p : x = y) (b : B x) → (x , b) = (y , tr B p b)
lift-eq-Σ refl b = refl
```
@@ -148,7 +147,7 @@ tr-Σ C refl z = refl
```agda
tr-eq-pair-Σ :
{l1 l2 l3 : Level} {A : UU l1} {a0 a1 : A}
- {B : A → UU l2} {b0 : B a0} {b1 : B a1} (C : (Σ A B) → UU l3)
+ {B : A → UU l2} {b0 : B a0} {b1 : B a1} (C : Σ A B → UU l3)
(p : a0 = a1) (q : dependent-identification B p b0 b1) (u : C (a0 , b0)) →
tr C (eq-pair-Σ p q) u =
tr (λ x → C (a1 , x)) q (tr C (eq-pair-Σ p refl) u)
diff --git a/src/foundation-core/fibers-of-maps.lagda.md b/src/foundation-core/fibers-of-maps.lagda.md
index 066a033e3c..72320d9a12 100644
--- a/src/foundation-core/fibers-of-maps.lagda.md
+++ b/src/foundation-core/fibers-of-maps.lagda.md
@@ -9,12 +9,14 @@ module foundation-core.fibers-of-maps where
```agda
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
+open import foundation.strictly-right-unital-concatenation-identifications
open import foundation.universe-levels
open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.homotopies
open import foundation-core.identity-types
+open import foundation-core.postcomposition-functions
open import foundation-core.retractions
open import foundation-core.sections
open import foundation-core.transport-along-identifications
@@ -373,6 +375,76 @@ module _
pr2 inv-compute-fiber-comp = is-equiv-map-inv-compute-fiber-comp
```
+### Fibers of homotopic maps are equivalent
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ {f g : A → B} (H : g ~ f) (y : B)
+ where
+
+ map-equiv-fiber-htpy : fiber f y → fiber g y
+ map-equiv-fiber-htpy (x , p) = x , H x ∙ᵣ p
+
+ map-inv-equiv-fiber-htpy : fiber g y → fiber f y
+ map-inv-equiv-fiber-htpy (x , p) = x , inv (H x) ∙ᵣ p
+
+ is-section-map-inv-equiv-fiber-htpy :
+ is-section map-equiv-fiber-htpy map-inv-equiv-fiber-htpy
+ is-section-map-inv-equiv-fiber-htpy (x , refl) =
+ eq-Eq-fiber g (g x) refl (inv (right-inv-right-strict-concat (H x)))
+
+ is-retraction-map-inv-equiv-fiber-htpy :
+ is-retraction map-equiv-fiber-htpy map-inv-equiv-fiber-htpy
+ is-retraction-map-inv-equiv-fiber-htpy (x , refl) =
+ eq-Eq-fiber f (f x) refl (inv (left-inv-right-strict-concat (H x)))
+
+ is-equiv-map-equiv-fiber-htpy : is-equiv map-equiv-fiber-htpy
+ is-equiv-map-equiv-fiber-htpy =
+ is-equiv-is-invertible
+ map-inv-equiv-fiber-htpy
+ is-section-map-inv-equiv-fiber-htpy
+ 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
+```
+
+We repeat the construction for `fiber'`.
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ {f g : A → B} (H : g ~ f) (y : B)
+ where
+
+ map-equiv-fiber-htpy' : fiber' f y → fiber' g y
+ map-equiv-fiber-htpy' (x , p) = (x , p ∙ inv (H x))
+
+ map-inv-equiv-fiber-htpy' : fiber' g y → fiber' f y
+ map-inv-equiv-fiber-htpy' (x , p) = (x , p ∙ H x)
+
+ is-section-map-inv-equiv-fiber-htpy' :
+ is-section map-equiv-fiber-htpy' map-inv-equiv-fiber-htpy'
+ is-section-map-inv-equiv-fiber-htpy' (x , p) =
+ ap (pair x) (is-retraction-inv-concat' (H x) p)
+
+ is-retraction-map-inv-equiv-fiber-htpy' :
+ is-retraction map-equiv-fiber-htpy' map-inv-equiv-fiber-htpy'
+ is-retraction-map-inv-equiv-fiber-htpy' (x , p) =
+ ap (pair x) (is-section-inv-concat' (H x) p)
+
+ is-equiv-map-equiv-fiber-htpy' : is-equiv map-equiv-fiber-htpy'
+ is-equiv-map-equiv-fiber-htpy' =
+ is-equiv-is-invertible
+ map-inv-equiv-fiber-htpy'
+ is-section-map-inv-equiv-fiber-htpy'
+ 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'
+```
+
## Table of files about fibers of maps
The following table lists files that are about fibers of maps as a general
diff --git a/src/foundation-core/path-split-maps.lagda.md b/src/foundation-core/path-split-maps.lagda.md
index 2d4e324279..ef41e36e50 100644
--- a/src/foundation-core/path-split-maps.lagda.md
+++ b/src/foundation-core/path-split-maps.lagda.md
@@ -23,8 +23,8 @@ open import foundation-core.sections
## Idea
-A map `f : A → B` is said to be **path split** if it has a
-[section](foundation-core.sections.md) and its
+A map `f : A → B` is said to be {{#concept "path split" Agda=is-path-split}} if
+it has a [section](foundation-core.sections.md) and its
[action on identifications](foundation.action-on-identifications-functions.md)
`x = y → f x = f y` has a section for each `x y : A`. By the
[fundamental theorem of identity types](foundation.fundamental-theorem-of-identity-types.md),
diff --git a/src/foundation-core/propositions.lagda.md b/src/foundation-core/propositions.lagda.md
index 005680a831..5aeafc52d4 100644
--- a/src/foundation-core/propositions.lagda.md
+++ b/src/foundation-core/propositions.lagda.md
@@ -126,6 +126,16 @@ module _
abstract
eq-is-proof-irrelevant : is-proof-irrelevant A → all-elements-equal A
eq-is-proof-irrelevant = eq-is-prop' ∘ is-prop-is-proof-irrelevant
+
+abstract
+ eq-type-Prop : {l : Level} (P : Prop l) → {x y : type-Prop P} → x = y
+ eq-type-Prop P = eq-is-prop (is-prop-type-Prop P)
+
+abstract
+ is-proof-irrelevant-type-Prop :
+ {l : Level} (P : Prop l) → is-proof-irrelevant (type-Prop P)
+ is-proof-irrelevant-type-Prop P =
+ is-proof-irrelevant-is-prop (is-prop-type-Prop P)
```
### Propositions are closed under equivalences
diff --git a/src/foundation-core/transport-along-identifications.lagda.md b/src/foundation-core/transport-along-identifications.lagda.md
index 911b7b0255..49b511e7e0 100644
--- a/src/foundation-core/transport-along-identifications.lagda.md
+++ b/src/foundation-core/transport-along-identifications.lagda.md
@@ -39,6 +39,47 @@ tr B refl b = b
## Properties
+### The dependent action of a family of maps over a map on identifications in the base type
+
+```agda
+tr-ap :
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} {D : C → UU l4}
+ (f : A → C) (g : (x : A) → B x → D (f x))
+ {x y : A} (p : x = y) (z : B x) →
+ tr D (ap f p) (g x z) = g y (tr B p z)
+tr-ap f g refl z = refl
+```
+
+### Every family of maps preserves transport
+
+```agda
+module _
+ {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3}
+ (f : (i : I) → A i → B i)
+ where
+
+ preserves-tr :
+ {i j : I} (p : i = j) (x : A i) →
+ f j (tr A p x) = tr B p (f i x)
+ preserves-tr refl x = refl
+
+ compute-preserves-tr :
+ {i j : I} (p : i = j) (x : A i) →
+ preserves-tr p x =
+ inv (tr-ap id f p x) ∙ ap (λ q → tr B q (f i x)) (ap-id p)
+ compute-preserves-tr refl x = refl
+```
+
+### Substitution law for transport
+
+```agda
+substitution-law-tr :
+ {l1 l2 l3 : Level} {X : UU l1} {A : UU l2} (B : A → UU l3) (f : X → A)
+ {x y : X} (p : x = y) {x' : B (f x)} →
+ tr B (ap f p) x' = tr (B ∘ f) p x'
+substitution-law-tr B f p {x'} = tr-ap f (λ _ → id) p x'
+```
+
### Transport preserves concatenation of identifications
```agda
@@ -70,34 +111,12 @@ module _
eq-transpose-tr' refl q = q
```
-### Every family of maps preserves transport
-
-```agda
-preserves-tr :
- {l1 l2 l3 : Level} {I : UU l1} {A : I → UU l2} {B : I → UU l3}
- (f : (i : I) → A i → B i)
- {i j : I} (p : i = j) (x : A i) →
- f j (tr A p x) = tr B p (f i x)
-preserves-tr f refl x = refl
-```
-
-### Transporting along the action on identifications of a function
-
-```agda
-tr-ap :
- {l1 l2 l3 l4 : Level} {A : UU l1} {B : A → UU l2} {C : UU l3} {D : C → UU l4}
- (f : A → C) (g : (x : A) → B x → D (f x))
- {x y : A} (p : x = y) (z : B x) →
- tr D (ap f p) (g x z) = g y (tr B p z)
-tr-ap f g refl z = refl
-```
-
### Computing maps out of identity types as transports
```agda
module _
{l1 l2 : Level} {A : UU l1} {B : A → UU l2} {a : A}
- (f : (x : A) → (a = x) → B x)
+ (f : (x : A) → a = x → B x)
where
compute-map-out-of-identity-type :
@@ -110,7 +129,7 @@ module _
```agda
tr-Id-left :
{l : Level} {A : UU l} {a b c : A} (q : b = c) (p : b = a) →
- tr (_= a) q p = (inv q ∙ p)
+ tr (_= a) q p = inv q ∙ p
tr-Id-left refl p = refl
```
@@ -119,16 +138,6 @@ tr-Id-left refl p = refl
```agda
tr-Id-right :
{l : Level} {A : UU l} {a b c : A} (q : b = c) (p : a = b) →
- tr (a =_) q p = (p ∙ q)
+ tr (a =_) q p = p ∙ q
tr-Id-right refl p = inv right-unit
```
-
-### Substitution law for transport
-
-```agda
-substitution-law-tr :
- {l1 l2 l3 : Level} {X : UU l1} {A : UU l2} (B : A → UU l3) (f : X → A)
- {x y : X} (p : x = y) {x' : B (f x)} →
- tr B (ap f p) x' = tr (B ∘ f) p x'
-substitution-law-tr B f refl = refl
-```
diff --git a/src/foundation.lagda.md b/src/foundation.lagda.md
index 0460bc24cd..4e72fe1c04 100644
--- a/src/foundation.lagda.md
+++ b/src/foundation.lagda.md
@@ -129,6 +129,7 @@ open import foundation.descent-coproduct-types public
open import foundation.descent-dependent-pair-types public
open import foundation.descent-empty-types public
open import foundation.descent-equivalences public
+open import foundation.diaconescus-theorem public
open import foundation.diagonal-maps-cartesian-products-of-types public
open import foundation.diagonal-maps-of-types public
open import foundation.diagonal-span-diagrams public
@@ -264,6 +265,7 @@ open import foundation.lawveres-fixed-point-theorem public
open import foundation.lesser-limited-principle-of-omniscience public
open import foundation.lifts-types public
open import foundation.limited-principle-of-omniscience public
+open import foundation.locale-of-propositions public
open import foundation.locally-small-types public
open import foundation.logical-equivalences public
open import foundation.maps-in-global-subuniverses public
@@ -311,6 +313,7 @@ open import foundation.partitions public
open import foundation.path-algebra public
open import foundation.path-cosplit-maps public
open import foundation.path-split-maps public
+open import foundation.path-split-type-families public
open import foundation.perfect-images public
open import foundation.permutations-spans-families-of-types public
open import foundation.pi-decompositions public
@@ -324,7 +327,6 @@ open import foundation.precomposition-dependent-functions public
open import foundation.precomposition-functions public
open import foundation.precomposition-functions-into-subuniverses public
open import foundation.precomposition-type-families public
-open import foundation.preimages-of-subtypes public
open import foundation.preunivalence public
open import foundation.preunivalent-type-families public
open import foundation.principle-of-omniscience public
@@ -443,6 +445,7 @@ open import foundation.type-arithmetic-standard-pullbacks public
open import foundation.type-arithmetic-unit-type public
open import foundation.type-duality public
open import foundation.type-theoretic-principle-of-choice public
+open import foundation.uniformly-decidable-type-families public
open import foundation.unions-subtypes public
open import foundation.uniqueness-image public
open import foundation.uniqueness-quantification public
diff --git a/src/foundation/axiom-of-choice.lagda.md b/src/foundation/axiom-of-choice.lagda.md
index 319776bde7..26774379b7 100644
--- a/src/foundation/axiom-of-choice.lagda.md
+++ b/src/foundation/axiom-of-choice.lagda.md
@@ -32,31 +32,53 @@ open import foundation-core.sets
## Idea
-The {{#concept "axiom of choice" WD="axiom of choice" WDID=Q179692 Agda=AC-0}}
+The {{#concept "axiom of choice" WD="axiom of choice" WDID=Q179692 Agda=AC0}}
asserts that for every family of [inhabited](foundation.inhabited-types.md)
types `B` indexed by a [set](foundation-core.sets.md) `A`, the type of sections
of that family `(x : A) → B x` is inhabited.
## Definition
+### Instances of choice
+
+```agda
+instance-choice : {l1 l2 : Level} (A : UU l1) → (A → UU l2) → UU (l1 ⊔ l2)
+instance-choice A B =
+ ((x : A) → is-inhabited (B x)) → is-inhabited ((x : A) → B x)
+```
+
+Note that the above statement, were it true for all indexing types `A`, is
+inconsistent with [univalence](foundation.univalence.md). For a concrete
+counterexample see Lemma 3.8.5 in {{#cite UF13}}.
+
### The axiom of choice restricted to sets
```agda
-AC-Set : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
-AC-Set l1 l2 =
- (A : Set l1) (B : type-Set A → Set l2) →
- ((x : type-Set A) → is-inhabited (type-Set (B x))) →
- is-inhabited ((x : type-Set A) → type-Set (B x))
+instance-choice-Set :
+ {l1 l2 : Level} (A : Set l1) → (type-Set A → Set l2) → UU (l1 ⊔ l2)
+instance-choice-Set A B = instance-choice (type-Set A) (type-Set ∘ B)
+
+level-AC-Set : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
+level-AC-Set l1 l2 =
+ (A : Set l1) (B : type-Set A → Set l2) → instance-choice-Set A B
+
+AC-Set : UUω
+AC-Set = {l1 l2 : Level} → level-AC-Set l1 l2
```
### The axiom of choice
```agda
-AC-0 : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
-AC-0 l1 l2 =
- (A : Set l1) (B : type-Set A → UU l2) →
- ((x : type-Set A) → is-inhabited (B x)) →
- is-inhabited ((x : type-Set A) → B x)
+instance-choice₀ :
+ {l1 l2 : Level} (A : Set l1) → (type-Set A → UU l2) → UU (l1 ⊔ l2)
+instance-choice₀ A = instance-choice (type-Set A)
+
+level-AC0 : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
+level-AC0 l1 l2 =
+ (A : Set l1) (B : type-Set A → UU l2) → instance-choice₀ A B
+
+AC0 : UUω
+AC0 = {l1 l2 : Level} → level-AC0 l1 l2
```
## Properties
@@ -64,10 +86,10 @@ AC-0 l1 l2 =
### Every type is set-projective if and only if the axiom of choice holds
```agda
-is-set-projective-AC-0 :
- {l1 l2 l3 : Level} → AC-0 l2 (l1 ⊔ l2) →
+is-set-projective-AC0 :
+ {l1 l2 l3 : Level} → level-AC0 l2 (l1 ⊔ l2) →
(X : UU l3) → is-set-projective l1 l2 X
-is-set-projective-AC-0 ac X A B f h =
+is-set-projective-AC0 ac X A B f h =
map-trunc-Prop
( ( map-Σ
( λ g → ((map-surjection f) ∘ g) = h)
@@ -76,11 +98,11 @@ is-set-projective-AC-0 ac X A B f h =
( section-is-split-surjective (map-surjection f)))
( ac B (fiber (map-surjection f)) (is-surjective-map-surjection f))
-AC-0-is-set-projective :
+AC0-is-set-projective :
{l1 l2 : Level} →
({l : Level} (X : UU l) → is-set-projective (l1 ⊔ l2) l1 X) →
- AC-0 l1 l2
-AC-0-is-set-projective H A B K =
+ level-AC0 l1 l2
+AC0-is-set-projective H A B K =
map-trunc-Prop
( map-equiv (equiv-Π-section-pr1 {B = B}) ∘ tot (λ g → htpy-eq))
( H ( type-Set A)
@@ -89,3 +111,12 @@ AC-0-is-set-projective H A B K =
( pr1 , (λ a → map-trunc-Prop (map-inv-fiber-pr1 B a) (K a)))
( id))
```
+
+## See also
+
+- [Diaconescu's theorem](foundation.diaconescus-theorem.md), which states that
+ the axiom of choice implies the law of excluded middle.
+
+## References
+
+{{#bibliography}}
diff --git a/src/foundation/binary-transport.lagda.md b/src/foundation/binary-transport.lagda.md
index 5ce9abd175..1bf6c1e593 100644
--- a/src/foundation/binary-transport.lagda.md
+++ b/src/foundation/binary-transport.lagda.md
@@ -37,7 +37,7 @@ module _
where
binary-tr : (p : x = x') (q : y = y') → C x y → C x' y'
- binary-tr p q c = tr (C _) q (tr (λ u → C u _) p c)
+ binary-tr p q c = tr (C x') q (tr (λ u → C u y) p c)
is-equiv-binary-tr : (p : x = x') (q : y = y') → is-equiv (binary-tr p q)
is-equiv-binary-tr refl refl = is-equiv-id
diff --git a/src/foundation/booleans.lagda.md b/src/foundation/booleans.lagda.md
index 82e07d5dfc..4c8195a978 100644
--- a/src/foundation/booleans.lagda.md
+++ b/src/foundation/booleans.lagda.md
@@ -7,7 +7,10 @@ module foundation.booleans where
Imports
```agda
+open import foundation.decidable-equality
+open import foundation.decidable-types
open import foundation.dependent-pair-types
+open import foundation.discrete-types
open import foundation.involutions
open import foundation.negated-equality
open import foundation.raising-universe-levels
@@ -193,18 +196,30 @@ abstract
( λ x y → eq-Eq-bool)
bool-Set : Set lzero
-pr1 bool-Set = bool
-pr2 bool-Set = is-set-bool
+bool-Set = bool , is-set-bool
+```
+
+### The booleans have decidable equality
+
+```agda
+has-decidable-equality-bool : has-decidable-equality bool
+has-decidable-equality-bool true true = inl refl
+has-decidable-equality-bool true false = inr neq-true-false-bool
+has-decidable-equality-bool false true = inr neq-false-true-bool
+has-decidable-equality-bool false false = inl refl
+
+bool-Discrete-Type : Discrete-Type lzero
+bool-Discrete-Type = bool , has-decidable-equality-bool
```
### The "is true" predicate on booleans
```agda
is-true : bool → UU lzero
-is-true = Eq-bool true
+is-true = _= true
is-prop-is-true : (b : bool) → is-prop (is-true b)
-is-prop-is-true = is-prop-Eq-bool true
+is-prop-is-true b = is-set-bool b true
is-true-Prop : bool → Prop lzero
is-true-Prop b = is-true b , is-prop-is-true b
@@ -214,15 +229,27 @@ is-true-Prop b = is-true b , is-prop-is-true b
```agda
is-false : bool → UU lzero
-is-false = Eq-bool false
+is-false = _= false
is-prop-is-false : (b : bool) → is-prop (is-false b)
-is-prop-is-false = is-prop-Eq-bool false
+is-prop-is-false b = is-set-bool b false
is-false-Prop : bool → Prop lzero
is-false-Prop b = is-false b , is-prop-is-false b
```
+### A boolean cannot be both true and false
+
+```agda
+not-is-false-is-true : (x : bool) → is-true x → ¬ (is-false x)
+not-is-false-is-true true t ()
+not-is-false-is-true false () f
+
+not-is-true-is-false : (x : bool) → is-false x → ¬ (is-true x)
+not-is-true-is-false true () f
+not-is-true-is-false false t ()
+```
+
### The type of booleans is equivalent to `Fin 2`
```agda
diff --git a/src/foundation/cantors-theorem.lagda.md b/src/foundation/cantors-theorem.lagda.md
index a00b09caeb..96593b0ef3 100644
--- a/src/foundation/cantors-theorem.lagda.md
+++ b/src/foundation/cantors-theorem.lagda.md
@@ -11,6 +11,7 @@ open import foundation.action-on-identifications-functions
open import foundation.decidable-propositions
open import foundation.decidable-subtypes
open import foundation.dependent-pair-types
+open import foundation.double-negation-stable-propositions
open import foundation.function-extensionality
open import foundation.logical-equivalences
open import foundation.negation
@@ -22,6 +23,10 @@ open import foundation.universe-levels
open import foundation-core.empty-types
open import foundation-core.fibers-of-maps
open import foundation-core.propositions
+
+open import logic.de-morgan-propositions
+open import logic.de-morgan-subtypes
+open import logic.double-negation-stable-subtypes
```
@@ -85,6 +90,9 @@ module _
### Cantor's theorem for the set of decidable subtypes
+**Statement.** There is no surjective map from a type `X` to its set of
+decidable subtypes `𝒫ᵈ(X)`.
+
```agda
module _
{l1 l2 : Level} {X : UU l1} (f : X → decidable-subtype l2 X)
@@ -110,6 +118,68 @@ module _
( not-in-image-map-theorem-decidable-Cantor)
```
+### Cantor's theorem for the set of double negation stable subtypes
+
+**Statement.** There is no surjective map from a type `X` to its set of double
+negation stable subtypes `𝒫^¬¬(X)`.
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} (f : X → double-negation-stable-subtype l2 X)
+ where
+
+ map-theorem-double-negation-stable-Cantor :
+ double-negation-stable-subtype l2 X
+ map-theorem-double-negation-stable-Cantor x =
+ neg-Double-Negation-Stable-Prop (f x x)
+
+ abstract
+ not-in-image-map-theorem-double-negation-stable-Cantor :
+ ¬ (fiber f map-theorem-double-negation-stable-Cantor)
+ not-in-image-map-theorem-double-negation-stable-Cantor (x , α) =
+ no-fixed-points-neg-Double-Negation-Stable-Prop
+ ( f x x)
+ ( iff-eq (ap prop-Double-Negation-Stable-Prop (htpy-eq α x)))
+
+ abstract
+ theorem-double-negation-stable-Cantor : ¬ (is-surjective f)
+ theorem-double-negation-stable-Cantor H =
+ apply-universal-property-trunc-Prop
+ ( H map-theorem-double-negation-stable-Cantor)
+ ( empty-Prop)
+ ( not-in-image-map-theorem-double-negation-stable-Cantor)
+```
+
+### Cantor's theorem for the set of De Morgan subtypes
+
+**Statement.** There is no surjective map from a type `X` to its set of De
+Morgan subtypes `𝒫ᵈᵐ(X)`.
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} (f : X → de-morgan-subtype l2 X)
+ where
+
+ map-theorem-de-morgan-Cantor : de-morgan-subtype l2 X
+ map-theorem-de-morgan-Cantor x = neg-De-Morgan-Prop (f x x)
+
+ abstract
+ not-in-image-map-theorem-de-morgan-Cantor :
+ ¬ (fiber f map-theorem-de-morgan-Cantor)
+ not-in-image-map-theorem-de-morgan-Cantor (x , α) =
+ no-fixed-points-neg-De-Morgan-Prop
+ ( f x x)
+ ( iff-eq (ap prop-De-Morgan-Prop (htpy-eq α x)))
+
+ abstract
+ theorem-de-morgan-Cantor : ¬ (is-surjective f)
+ theorem-de-morgan-Cantor H =
+ apply-universal-property-trunc-Prop
+ ( H map-theorem-de-morgan-Cantor)
+ ( empty-Prop)
+ ( not-in-image-map-theorem-de-morgan-Cantor)
+```
+
## References
A proof of Cantor's theorem first appeared in {{#cite Cantor1890/91}} where it
diff --git a/src/foundation/complements-subtypes.lagda.md b/src/foundation/complements-subtypes.lagda.md
index f105addb16..df47f07a3b 100644
--- a/src/foundation/complements-subtypes.lagda.md
+++ b/src/foundation/complements-subtypes.lagda.md
@@ -9,22 +9,37 @@ module foundation.complements-subtypes where
```agda
open import foundation.decidable-propositions
open import foundation.decidable-subtypes
+open import foundation.double-negation-stable-propositions
open import foundation.full-subtypes
open import foundation.negation
+open import foundation.postcomposition-functions
+open import foundation.powersets
open import foundation.propositional-truncations
open import foundation.unions-subtypes
open import foundation.universe-levels
open import foundation-core.function-types
open import foundation-core.subtypes
+
+open import logic.double-negation-stable-subtypes
+
+open import order-theory.large-posets
+open import order-theory.opposite-large-posets
+open import order-theory.order-preserving-maps-large-posets
+open import order-theory.order-preserving-maps-large-preorders
+open import order-theory.order-preserving-maps-posets
+open import order-theory.order-preserving-maps-preorders
+open import order-theory.posets
```
## Idea
-The **complement** of a [subtype](foundation-core.subtypes.md) `P` of `A`
-consists of the elements that are not in `P`.
+The
+{{#concept "complement" Disambiguation="of a subtype" Agda=complement-subtype}}
+of a [subtype](foundation-core.subtypes.md) `P ⊆ A` consists of the elements
+that are [not](foundation-core.negation.md) in `P`.
## Definition
@@ -36,45 +51,29 @@ complement-subtype :
complement-subtype P x = neg-Prop (P x)
```
-### Complements of decidable subtypes
+## Properties
+
+### Complements of subtypes are double negation stable
```agda
-complement-decidable-subtype :
- {l1 l2 : Level} {A : UU l1} → decidable-subtype l2 A → decidable-subtype l2 A
-complement-decidable-subtype P x = neg-Decidable-Prop (P x)
+complement-double-negation-stable-subtype' :
+ {l1 l2 : Level} {A : UU l1} →
+ subtype l2 A → double-negation-stable-subtype l2 A
+complement-double-negation-stable-subtype' P x =
+ neg-type-Double-Negation-Stable-Prop (is-in-subtype P x)
```
-## Properties
-
-### The union of a subtype `P` with its complement is the full subtype if and only if `P` is a decidable subtype
+### Taking complements gives a contravariant endooperator on the powerset posets
```agda
-module _
- {l1 l2 : Level} {A : UU l1}
- where
-
- is-full-union-subtype-complement-subtype :
- (P : subtype l2 A) → is-decidable-subtype P →
- is-full-subtype (union-subtype P (complement-subtype P))
- is-full-union-subtype-complement-subtype P d x =
- unit-trunc-Prop (d x)
-
- is-decidable-subtype-is-full-union-subtype-complement-subtype :
- (P : subtype l2 A) →
- is-full-subtype (union-subtype P (complement-subtype P)) →
- is-decidable-subtype P
- is-decidable-subtype-is-full-union-subtype-complement-subtype P H x =
- apply-universal-property-trunc-Prop
- ( H x)
- ( is-decidable-Prop (P x))
- ( id)
-
- is-full-union-subtype-complement-decidable-subtype :
- (P : decidable-subtype l2 A) →
- is-full-decidable-subtype
- ( union-decidable-subtype P (complement-decidable-subtype P))
- is-full-union-subtype-complement-decidable-subtype P =
- is-full-union-subtype-complement-subtype
- ( subtype-decidable-subtype P)
- ( is-decidable-decidable-subtype P)
+neg-hom-powerset :
+ {l1 : Level} {A : UU l1} →
+ hom-Large-Poset
+ ( λ l → l)
+ ( powerset-Large-Poset A)
+ ( opposite-Large-Poset (powerset-Large-Poset A))
+neg-hom-powerset =
+ make-hom-Large-Preorder
+ ( λ P x → neg-Prop (P x))
+ ( λ P Q f x → map-neg (f x))
```
diff --git a/src/foundation/coproduct-decompositions-subuniverse.lagda.md b/src/foundation/coproduct-decompositions-subuniverse.lagda.md
index 1662dfbdc9..3d8495b9a1 100644
--- a/src/foundation/coproduct-decompositions-subuniverse.lagda.md
+++ b/src/foundation/coproduct-decompositions-subuniverse.lagda.md
@@ -503,7 +503,7 @@ module _
eq-pair-Σ
( eq-pair-Σ
( eq-equiv (equiv-is-empty is-empty-raise-empty (pr2 x)))
- ( eq-is-prop (is-prop-type-Prop (P _))))
+ ( eq-type-Prop (P _)))
( eq-is-prop is-property-is-empty)))
( ( raise-empty l1 , C1) , is-empty-raise-empty)) ∘e
( ( inv-associative-Σ _ _ _) ∘e
diff --git a/src/foundation/coproduct-types.lagda.md b/src/foundation/coproduct-types.lagda.md
index 14b74377c5..aaadc73d63 100644
--- a/src/foundation/coproduct-types.lagda.md
+++ b/src/foundation/coproduct-types.lagda.md
@@ -176,11 +176,17 @@ module _
{l1 l2 : Level} {A : UU l1} {B : UU l2}
where
+ noncontractibility-coproduct-is-contr' :
+ is-contr A → is-contr B → noncontractibility' (A + B) 1
+ noncontractibility-coproduct-is-contr' HA HB =
+ inl (center HA) , inr (center HB) , neq-inl-inr
+
abstract
is-not-contractible-coproduct-is-contr :
is-contr A → is-contr B → is-not-contractible (A + B)
- is-not-contractible-coproduct-is-contr HA HB HAB =
- neq-inl-inr {x = center HA} {y = center HB} (eq-is-contr HAB)
+ is-not-contractible-coproduct-is-contr HA HB =
+ is-not-contractible-noncontractibility
+ ( 1 , noncontractibility-coproduct-is-contr' HA HB)
```
### Coproducts of mutually exclusive propositions are propositions
diff --git a/src/foundation/decidable-dependent-function-types.lagda.md b/src/foundation/decidable-dependent-function-types.lagda.md
index f0045ab035..c63c6ddeea 100644
--- a/src/foundation/decidable-dependent-function-types.lagda.md
+++ b/src/foundation/decidable-dependent-function-types.lagda.md
@@ -10,26 +10,51 @@ module foundation.decidable-dependent-function-types where
open import foundation.decidable-types
open import foundation.functoriality-dependent-function-types
open import foundation.maybe
+open import foundation.uniformly-decidable-type-families
open import foundation.universal-property-coproduct-types
open import foundation.universal-property-maybe
open import foundation.universe-levels
open import foundation-core.coproduct-types
+open import foundation-core.empty-types
open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.negation
```
## Idea
-We describe conditions under which dependent products are decidable.
+We describe conditions under which
+[dependent function types](foundation.dependent-function-types.md) are
+[decidable](foundation.decidable-types.md).
+
+## Properties
+
+### Decidability of dependent products of uniformly decidable families
+
+```agda
+is-decidable-Π-uniformly-decidable-family :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ is-decidable A →
+ is-uniformly-decidable-family B →
+ is-decidable ((a : A) → (B a))
+is-decidable-Π-uniformly-decidable-family (inl a) (inl b) =
+ inl b
+is-decidable-Π-uniformly-decidable-family (inl a) (inr b) =
+ inr (λ f → b a (f a))
+is-decidable-Π-uniformly-decidable-family (inr na) _ =
+ inl (ex-falso ∘ na)
+```
### Decidablitilty of dependent products over coproducts
```agda
is-decidable-Π-coproduct :
{l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : A + B → UU l3} →
- is-decidable ((a : A) → C (inl a)) → is-decidable ((b : B) → C (inr b)) →
+ is-decidable ((a : A) → C (inl a)) →
+ is-decidable ((b : B) → C (inr b)) →
is-decidable ((x : A + B) → C x)
is-decidable-Π-coproduct {C = C} dA dB =
is-decidable-equiv
@@ -42,7 +67,8 @@ is-decidable-Π-coproduct {C = C} dA dB =
```agda
is-decidable-Π-Maybe :
{l1 l2 : Level} {A : UU l1} {B : Maybe A → UU l2} →
- is-decidable ((x : A) → B (unit-Maybe x)) → is-decidable (B exception-Maybe) →
+ is-decidable ((x : A) → B (unit-Maybe x)) →
+ is-decidable (B exception-Maybe) →
is-decidable ((x : Maybe A) → B x)
is-decidable-Π-Maybe {B = B} du de =
is-decidable-equiv
@@ -53,15 +79,16 @@ is-decidable-Π-Maybe {B = B} du de =
### Decidability of dependent products over an equivalence
```agda
-is-decidable-Π-equiv :
+module _
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} {D : B → UU l4}
- (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x)) →
- is-decidable ((x : A) → C x) → is-decidable ((y : B) → D y)
-is-decidable-Π-equiv {D = D} e f = is-decidable-equiv' (equiv-Π D e f)
+ (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x))
+ where
-is-decidable-Π-equiv' :
- {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} {D : B → UU l4}
- (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x)) →
- is-decidable ((y : B) → D y) → is-decidable ((x : A) → C x)
-is-decidable-Π-equiv' {D = D} e f = is-decidable-equiv (equiv-Π D e f)
+ is-decidable-Π-equiv :
+ is-decidable ((x : A) → C x) → is-decidable ((y : B) → D y)
+ is-decidable-Π-equiv = is-decidable-equiv' (equiv-Π D e f)
+
+ is-decidable-Π-equiv' :
+ is-decidable ((y : B) → D y) → is-decidable ((x : A) → C x)
+ is-decidable-Π-equiv' = is-decidable-equiv (equiv-Π D e f)
```
diff --git a/src/foundation/decidable-dependent-pair-types.lagda.md b/src/foundation/decidable-dependent-pair-types.lagda.md
index 1632117634..4af4b5e568 100644
--- a/src/foundation/decidable-dependent-pair-types.lagda.md
+++ b/src/foundation/decidable-dependent-pair-types.lagda.md
@@ -12,51 +12,78 @@ open import foundation.dependent-pair-types
open import foundation.maybe
open import foundation.type-arithmetic-coproduct-types
open import foundation.type-arithmetic-unit-type
+open import foundation.uniformly-decidable-type-families
open import foundation.universe-levels
open import foundation-core.coproduct-types
open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.functoriality-dependent-pair-types
+open import foundation-core.negation
```
## Idea
-We describe conditions under which dependent sums are decidable.
+We describe conditions under which
+[dependent sums](foundation.dependent-pair-types.md) are
+[decidable](foundation.decidable-types.md)
+
+## Properites
+
+### Dependent sums of a uniformly decidable family of types
+
+```agda
+is-decidable-Σ-uniformly-decidable-family :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ is-decidable A →
+ is-uniformly-decidable-family B →
+ is-decidable (Σ A B)
+is-decidable-Σ-uniformly-decidable-family (inl a) (inl b) =
+ inl (a , b a)
+is-decidable-Σ-uniformly-decidable-family (inl a) (inr b) =
+ inr (λ x → b (pr1 x) (pr2 x))
+is-decidable-Σ-uniformly-decidable-family (inr a) _ =
+ inr (λ x → a (pr1 x))
+```
+
+### Decidability of dependent sums over coproducts
```agda
is-decidable-Σ-coproduct :
{l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : A + B → UU l3) →
is-decidable (Σ A (C ∘ inl)) → is-decidable (Σ B (C ∘ inr)) →
is-decidable (Σ (A + B) C)
-is-decidable-Σ-coproduct {l1} {l2} {l3} {A} {B} C dA dB =
+is-decidable-Σ-coproduct {A = A} {B} C dA dB =
is-decidable-equiv
( right-distributive-Σ-coproduct A B C)
( is-decidable-coproduct dA dB)
+```
+### Decidability of dependent sums over `Maybe`
+
+```agda
is-decidable-Σ-Maybe :
{l1 l2 : Level} {A : UU l1} {B : Maybe A → UU l2} →
is-decidable (Σ A (B ∘ unit-Maybe)) → is-decidable (B exception-Maybe) →
is-decidable (Σ (Maybe A) B)
-is-decidable-Σ-Maybe {l1} {l2} {A} {B} dA de =
+is-decidable-Σ-Maybe {A = A} {B} dA de =
is-decidable-Σ-coproduct B dA
- ( is-decidable-equiv
- ( left-unit-law-Σ (B ∘ inr))
- ( de))
+ ( is-decidable-equiv (left-unit-law-Σ (B ∘ inr)) de)
+```
-is-decidable-Σ-equiv :
- {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} {D : B → UU l4}
- (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x)) →
- is-decidable (Σ A C) → is-decidable (Σ B D)
-is-decidable-Σ-equiv {D = D} e f =
- is-decidable-equiv' (equiv-Σ D e f)
+### Decidability of dependent sums over equivalences
-is-decidable-Σ-equiv' :
+```agda
+module _
{l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} {D : B → UU l4}
- (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x)) →
- is-decidable (Σ B D) → is-decidable (Σ A C)
-is-decidable-Σ-equiv' {D = D} e f =
- is-decidable-equiv (equiv-Σ D e f)
+ (e : A ≃ B) (f : (x : A) → C x ≃ D (map-equiv e x))
+ where
+
+ is-decidable-Σ-equiv : is-decidable (Σ A C) → is-decidable (Σ B D)
+ is-decidable-Σ-equiv = is-decidable-equiv' (equiv-Σ D e f)
+
+ is-decidable-Σ-equiv' : is-decidable (Σ B D) → is-decidable (Σ A C)
+ is-decidable-Σ-equiv' = is-decidable-equiv (equiv-Σ D e f)
```
diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md
index fc452b45c0..f4780285bd 100644
--- a/src/foundation/decidable-embeddings.lagda.md
+++ b/src/foundation/decidable-embeddings.lagda.md
@@ -14,6 +14,7 @@ open import foundation.decidable-propositions
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.embeddings
+open import foundation.fibers-of-maps
open import foundation.functoriality-cartesian-product-types
open import foundation.functoriality-coproduct-types
open import foundation.fundamental-theorem-of-identity-types
@@ -21,9 +22,11 @@ open import foundation.homotopy-induction
open import foundation.identity-types
open import foundation.logical-equivalences
open import foundation.propositional-maps
+open import foundation.propositions
open import foundation.retracts-of-maps
open import foundation.subtype-identity-principle
open import foundation.type-arithmetic-dependent-pair-types
+open import foundation.unit-type
open import foundation.universal-property-equivalences
open import foundation.universe-levels
@@ -31,12 +34,10 @@ open import foundation-core.cartesian-product-types
open import foundation-core.coproduct-types
open import foundation-core.empty-types
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.injective-maps
-open import foundation-core.propositions
open import foundation-core.torsorial-type-families
```
@@ -64,88 +65,87 @@ is-decidable-emb :
{l1 l2 : Level} {X : UU l1} {Y : UU l2} → (X → Y) → UU (l1 ⊔ l2)
is-decidable-emb f = is-emb f × is-decidable-map f
-abstract
- is-emb-is-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
- is-decidable-emb f → is-emb f
- is-emb-is-decidable-emb = pr1
+is-emb-is-decidable-emb :
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
+ is-decidable-emb f → is-emb f
+is-emb-is-decidable-emb = pr1
is-decidable-map-is-decidable-emb :
{l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
is-decidable-emb f → is-decidable-map f
is-decidable-map-is-decidable-emb = pr2
+
+is-injective-is-decidable-emb :
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
+ is-decidable-emb f → is-injective f
+is-injective-is-decidable-emb = is-injective-is-emb ∘ is-emb-is-decidable-emb
```
### Decidably propositional maps
```agda
-is-decidable-prop-map :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} → (X → Y) → UU (l1 ⊔ l2)
-is-decidable-prop-map {Y = Y} f = (y : Y) → is-decidable-prop (fiber f y)
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ where
-abstract
- is-prop-map-is-decidable-prop-map :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
- is-decidable-prop-map f → is-prop-map f
- is-prop-map-is-decidable-prop-map H y = pr1 (H y)
+ is-decidable-prop-map : (X → Y) → UU (l1 ⊔ l2)
+ is-decidable-prop-map f = (y : Y) → is-decidable-prop (fiber f y)
-is-decidable-map-is-decidable-prop-map :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y} →
- is-decidable-prop-map f → is-decidable-map f
-is-decidable-map-is-decidable-prop-map H y = pr2 (H y)
+ is-prop-is-decidable-prop-map :
+ (f : X → Y) → is-prop (is-decidable-prop-map f)
+ is-prop-is-decidable-prop-map f =
+ is-prop-Π (λ y → is-prop-is-decidable-prop (fiber f y))
+
+ is-decidable-prop-map-Prop : (X → Y) → Prop (l1 ⊔ l2)
+ is-decidable-prop-map-Prop f =
+ ( is-decidable-prop-map f , is-prop-is-decidable-prop-map f)
+
+ abstract
+ is-prop-map-is-decidable-prop-map :
+ {f : X → Y} → is-decidable-prop-map f → is-prop-map f
+ is-prop-map-is-decidable-prop-map H y = pr1 (H y)
+
+ is-decidable-map-is-decidable-prop-map :
+ {f : X → Y} → is-decidable-prop-map f → is-decidable-map f
+ is-decidable-map-is-decidable-prop-map H y = pr2 (H y)
```
### The type of decidable embeddings
```agda
infix 5 _↪ᵈ_
-_↪ᵈ_ :
- {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+_↪ᵈ_ : {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
X ↪ᵈ Y = Σ (X → Y) is-decidable-emb
-map-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} → X ↪ᵈ Y → X → Y
-map-decidable-emb e = pr1 e
+decidable-emb : {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+decidable-emb = _↪ᵈ_
+
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪ᵈ Y)
+ where
+
+ map-decidable-emb : X → Y
+ map-decidable-emb = pr1 e
-abstract
is-decidable-emb-map-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪ᵈ Y) →
- is-decidable-emb (map-decidable-emb e)
- is-decidable-emb-map-decidable-emb e = pr2 e
+ is-decidable-emb map-decidable-emb
+ is-decidable-emb-map-decidable-emb = pr2 e
-abstract
- is-emb-map-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪ᵈ Y) →
- is-emb (map-decidable-emb e)
- is-emb-map-decidable-emb e =
- is-emb-is-decidable-emb (is-decidable-emb-map-decidable-emb 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
-abstract
is-decidable-map-map-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪ᵈ Y) →
- is-decidable-map (map-decidable-emb e)
- is-decidable-map-map-decidable-emb e =
- is-decidable-map-is-decidable-emb (is-decidable-emb-map-decidable-emb e)
-
-emb-decidable-emb :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} → X ↪ᵈ Y → X ↪ Y
-pr1 (emb-decidable-emb e) = map-decidable-emb e
-pr2 (emb-decidable-emb e) = is-emb-map-decidable-emb e
+ is-decidable-map map-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
```
## Properties
-### Being a decidably propositional map is a proposition
-
-```agda
-abstract
- is-prop-is-decidable-prop-map :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X → Y) →
- is-prop (is-decidable-prop-map f)
- is-prop-is-decidable-prop-map f =
- is-prop-Π (λ y → is-prop-is-decidable-prop (fiber f y))
-```
-
### Any map of which the fibers are decidable propositions is a decidable embedding
```agda
@@ -175,6 +175,26 @@ module _
is-decidable-map-is-decidable-emb H y
```
+### The first projection map of a dependent sum of decidable propositions is a decidable embedding
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (Q : A → Decidable-Prop l2)
+ where
+
+ is-decidable-prop-map-pr1 :
+ is-decidable-prop-map (pr1 {B = type-Decidable-Prop ∘ Q})
+ is-decidable-prop-map-pr1 y =
+ is-decidable-prop-equiv
+ ( equiv-fiber-pr1 (type-Decidable-Prop ∘ Q) y)
+ ( is-decidable-prop-type-Decidable-Prop (Q y))
+
+ is-decidable-emb-pr1 :
+ is-decidable-emb (pr1 {B = type-Decidable-Prop ∘ Q})
+ is-decidable-emb-pr1 =
+ is-decidable-emb-is-decidable-prop-map is-decidable-prop-map-pr1
+```
+
### Equivalences are decidable embeddings
```agda
@@ -189,13 +209,16 @@ abstract
### Identity maps are decidable embeddings
```agda
-abstract
- is-decidable-emb-id :
- {l1 : Level} {A : UU l1} → is-decidable-emb (id {A = A})
- is-decidable-emb-id = (is-emb-id , is-decidable-map-id)
+is-decidable-emb-id :
+ {l : Level} {A : UU l} → is-decidable-emb (id {A = A})
+is-decidable-emb-id = (is-emb-id , is-decidable-map-id)
-decidable-emb-id : {l1 : Level} {A : UU l1} → A ↪ᵈ A
+decidable-emb-id : {l : Level} {A : UU l} → A ↪ᵈ A
decidable-emb-id = (id , is-decidable-emb-id)
+
+is-decidable-prop-map-id :
+ {l : Level} {A : UU l} → is-decidable-prop-map (id {A = A})
+is-decidable-prop-map-id y = is-decidable-prop-is-contr (is-torsorial-Id' y)
```
### Being a decidable embedding is a property
@@ -214,6 +237,18 @@ abstract
( λ y → is-prop-is-decidable (is-prop-map-is-emb (pr1 H) y))))
```
+### Decidable embeddings are closed under homotopies
+
+```agda
+abstract
+ is-decidable-emb-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
+ f ~ g → is-decidable-emb g → is-decidable-emb f
+ is-decidable-emb-htpy {f = f} {g} H K =
+ ( is-emb-htpy H (is-emb-is-decidable-emb K) ,
+ is-decidable-map-htpy H (is-decidable-map-is-decidable-emb K))
+```
+
### Decidable embeddings are closed under composition
```agda
@@ -225,18 +260,10 @@ module _
abstract
is-decidable-map-comp-is-decidable-emb' :
is-decidable-emb g → is-decidable-map f → is-decidable-map (g ∘ f)
- is-decidable-map-comp-is-decidable-emb' K H x =
- rec-coproduct
- ( λ u →
- is-decidable-equiv
- ( ( left-unit-law-Σ-is-contr
- ( is-proof-irrelevant-is-prop
- ( is-prop-map-is-emb (is-emb-is-decidable-emb K) x) u)
- ( u)) ∘e
- ( compute-fiber-comp g f x))
- ( H (pr1 u)))
- ( λ α → inr (λ t → α (f (pr1 t) , pr2 t)))
- ( is-decidable-map-is-decidable-emb K x)
+ is-decidable-map-comp-is-decidable-emb' K =
+ is-decidable-map-comp
+ ( is-injective-is-decidable-emb K)
+ ( is-decidable-map-is-decidable-emb K)
is-decidable-map-comp-is-decidable-emb :
is-decidable-emb g → is-decidable-emb f → is-decidable-map (g ∘ f)
@@ -250,6 +277,29 @@ module _
is-decidable-emb-comp K H =
( is-emb-comp _ _ (pr1 K) (pr1 H) ,
is-decidable-map-comp-is-decidable-emb K H)
+
+ abstract
+ is-decidable-prop-map-comp :
+ is-decidable-prop-map g →
+ is-decidable-prop-map f →
+ is-decidable-prop-map (g ∘ f)
+ is-decidable-prop-map-comp K H =
+ is-decidable-prop-map-is-decidable-emb
+ ( is-decidable-emb-comp
+ ( is-decidable-emb-is-decidable-prop-map K)
+ ( is-decidable-emb-is-decidable-prop-map H))
+
+comp-decidable-emb :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ B ↪ᵈ C → A ↪ᵈ B → A ↪ᵈ C
+comp-decidable-emb (g , G) (f , F) =
+ ( g ∘ f , is-decidable-emb-comp G F)
+
+infixr 15 _∘ᵈ_
+_∘ᵈ_ :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ B ↪ᵈ C → A ↪ᵈ B → A ↪ᵈ C
+_∘ᵈ_ = comp-decidable-emb
```
### Left cancellation for decidable embeddings
@@ -273,18 +323,6 @@ module _
is-decidable-emb-right-factor' GH (is-emb-is-decidable-emb G)
```
-### Decidable embeddings are closed under homotopies
-
-```agda
-abstract
- is-decidable-emb-htpy :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
- f ~ g → is-decidable-emb g → is-decidable-emb f
- is-decidable-emb-htpy {f = f} {g} H K =
- ( is-emb-htpy H (is-emb-is-decidable-emb K) ,
- is-decidable-map-htpy H (is-decidable-map-is-decidable-emb K))
-```
-
### In a commuting triangle of maps, if the top and right maps are decidable embeddings so is the left map
```agda
@@ -360,9 +398,9 @@ abstract
( htpy-eq-decidable-emb f)
eq-htpy-decidable-emb :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A ↪ᵈ B} →
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪ᵈ B) →
htpy-decidable-emb f g → f = g
-eq-htpy-decidable-emb {f = f} {g} =
+eq-htpy-decidable-emb f g =
map-inv-is-equiv (is-equiv-htpy-eq-decidable-emb f g)
```
@@ -561,3 +599,45 @@ module _
( is-decidable-prop-map-retract-map R
( is-decidable-prop-map-is-decidable-emb G))
```
+
+### A type is a decidable proposition if and only if its terminal map is a decidable embedding
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ is-decidable-prop-is-decidable-emb-terminal-map :
+ is-decidable-emb (terminal-map A) → is-decidable-prop A
+ is-decidable-prop-is-decidable-emb-terminal-map H =
+ is-decidable-prop-equiv'
+ ( equiv-fiber-terminal-map star)
+ ( is-decidable-prop-map-is-decidable-emb H star)
+
+ is-decidable-emb-terminal-map-is-decidable-prop :
+ is-decidable-prop A → is-decidable-emb (terminal-map A)
+ is-decidable-emb-terminal-map-is-decidable-prop H =
+ is-decidable-emb-is-decidable-prop-map
+ ( λ y → is-decidable-prop-equiv (equiv-fiber-terminal-map y) H)
+```
+
+### If a dependent sum of propositions over a proposition is decidable, then the family is a family of decidable propositions
+
+```agda
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : type-Prop P → Prop l2)
+ where
+
+ is-decidable-prop-family-is-decidable-Σ :
+ is-decidable (Σ (type-Prop P) (type-Prop ∘ Q)) →
+ (p : type-Prop P) → is-decidable (type-Prop (Q p))
+ is-decidable-prop-family-is-decidable-Σ H p =
+ is-decidable-equiv'
+ ( equiv-fiber-pr1 (type-Prop ∘ Q) p)
+ ( is-decidable-map-is-decidable-emb
+ ( is-decidable-emb-right-factor'
+ ( is-decidable-emb-terminal-map-is-decidable-prop
+ ( is-prop-Σ (is-prop-type-Prop P) (is-prop-type-Prop ∘ Q) , H))
+ ( is-emb-terminal-map-is-prop (is-prop-type-Prop P)))
+ ( p))
+```
diff --git a/src/foundation/decidable-maps.lagda.md b/src/foundation/decidable-maps.lagda.md
index 893f11da73..926253c0dc 100644
--- a/src/foundation/decidable-maps.lagda.md
+++ b/src/foundation/decidable-maps.lagda.md
@@ -29,17 +29,22 @@ open import foundation-core.functoriality-dependent-pair-types
open import foundation-core.homotopies
open import foundation-core.injective-maps
open import foundation-core.retractions
+open import foundation-core.sections
```
-## Definition
+## Idea
A [map](foundation-core.function-types.md) is said to be
{{#concept "decidable" Disambiguation="map of types" Agda=is-decidable-map}} if
its [fibers](foundation-core.fibers-of-maps.md) are
[decidable types](foundation.decidable-types.md).
+## Definition
+
+### The structure on a map of decidability
+
```agda
module _
{l1 l2 : Level} {A : UU l1} {B : UU l2}
@@ -49,6 +54,28 @@ module _
is-decidable-map f = (y : B) → is-decidable (fiber f y)
```
+### The type of decidable maps
+
+```agda
+infix 5 _→ᵈ_
+
+_→ᵈ_ : {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+A →ᵈ B = Σ (A → B) (is-decidable-map)
+
+decidable-map : {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+decidable-map = _→ᵈ_
+
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A →ᵈ B)
+ where
+
+ map-decidable-map : A → B
+ map-decidable-map = pr1 f
+
+ is-decidable-decidable-map : is-decidable-map map-decidable-map
+ is-decidable-decidable-map = pr2 f
+```
+
## Properties
### Decidable maps are closed under homotopy
@@ -64,10 +91,36 @@ abstract
( K b)
```
+### Composition of decidable maps
+
+The composite `g ∘ f` of two decidable maps is decidable if `g` is injective.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {g : B → C} {f : A → B}
+ where
+
+ abstract
+ is-decidable-map-comp :
+ is-injective g →
+ is-decidable-map g →
+ is-decidable-map f →
+ is-decidable-map (g ∘ f)
+ is-decidable-map-comp H G F x =
+ rec-coproduct
+ ( λ u →
+ is-decidable-iff
+ ( λ v → (pr1 v) , ap g (pr2 v) ∙ pr2 u)
+ ( λ w → pr1 w , H (pr2 w ∙ inv (pr2 u)))
+ ( F (pr1 u)))
+ ( λ α → inr (λ t → α (f (pr1 t) , pr2 t)))
+ ( G x)
+```
+
### Left cancellation for decidable maps
-If a composite `g ∘ f` is decidable and the left factor `g` is injective, then
-the right factor `f` is decidable.
+If a composite `g ∘ f` is decidable and `g` is injective then `f` is decidable.
```agda
module _
@@ -77,9 +130,11 @@ module _
abstract
is-decidable-map-right-factor' :
is-decidable-map (g ∘ f) → is-injective g → is-decidable-map f
- is-decidable-map-right-factor' GF G y with (GF (g y))
- ... | inl q = inl (pr1 q , G (pr2 q))
- ... | inr q = inr (λ x → q ((pr1 x) , ap g (pr2 x)))
+ is-decidable-map-right-factor' GF G y =
+ rec-coproduct
+ ( λ q → inl (pr1 q , G (pr2 q)))
+ ( λ q → inr (λ x → q ((pr1 x) , ap g (pr2 x))))
+ ( GF (g y))
```
### Retracts into types with decidable equality are decidable
@@ -91,10 +146,19 @@ is-decidable-map-retraction :
is-decidable-map-retraction d i (r , R) b =
is-decidable-iff
( λ (p : i (r b) = b) → r b , p)
- ( λ t → ap (i ∘ r) (inv (pr2 t)) ∙ (ap i (R (pr1 t)) ∙ pr2 t))
+ ( λ t → ap (i ∘ r) (inv (pr2 t)) ∙ ap i (R (pr1 t)) ∙ pr2 t)
( d (i (r b)) b)
```
+### Maps with sections are decidable
+
+```agda
+is-decidable-map-section :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} →
+ (i : A → B) → section i → is-decidable-map i
+is-decidable-map-section i (s , S) b = inl (s b , S b)
+```
+
### Any map out of the empty type is decidable
```agda
diff --git a/src/foundation/decidable-propositions.lagda.md b/src/foundation/decidable-propositions.lagda.md
index 4315eab9c8..c6c5bfb814 100644
--- a/src/foundation/decidable-propositions.lagda.md
+++ b/src/foundation/decidable-propositions.lagda.md
@@ -115,8 +115,7 @@ module _
is-retraction-map-inv-equiv-bool-Decidable-Prop'
equiv-bool-Decidable-Prop' :
- ((Σ (Prop l) type-Prop) + (Σ (Prop l) (λ Q → ¬ (type-Prop Q)))) ≃
- bool
+ ((Σ (Prop l) type-Prop) + (Σ (Prop l) (λ Q → ¬ (type-Prop Q)))) ≃ bool
pr1 equiv-bool-Decidable-Prop' = map-equiv-bool-Decidable-Prop'
pr2 equiv-bool-Decidable-Prop' = is-equiv-map-equiv-bool-Decidable-Prop'
diff --git a/src/foundation/decidable-subtypes.lagda.md b/src/foundation/decidable-subtypes.lagda.md
index 9a162ba9f7..0bd5ecec7f 100644
--- a/src/foundation/decidable-subtypes.lagda.md
+++ b/src/foundation/decidable-subtypes.lagda.md
@@ -36,6 +36,8 @@ open import foundation-core.propositions
open import foundation-core.transport-along-identifications
open import foundation-core.truncated-types
open import foundation-core.truncation-levels
+
+open import logic.double-negation-stable-subtypes
```
@@ -227,6 +229,16 @@ module _
( iff-universes-Decidable-Prop l l' (S x))
```
+### Decidable subtypes are double negation stable
+
+```agda
+is-double-negation-stable-decicable-subtype :
+ {l1 l2 : Level} {A : UU l1} (P : decidable-subtype l2 A) →
+ is-double-negation-stable-subtype (subtype-decidable-subtype P)
+is-double-negation-stable-decicable-subtype P x =
+ double-negation-elim-is-decidable (is-decidable-decidable-subtype P x)
+```
+
### A decidable subtype of a `k+1`-truncated type is `k+1`-truncated
```agda
diff --git a/src/foundation/decidable-types.lagda.md b/src/foundation/decidable-types.lagda.md
index 2f34bdc31f..93bdd1707e 100644
--- a/src/foundation/decidable-types.lagda.md
+++ b/src/foundation/decidable-types.lagda.md
@@ -7,24 +7,27 @@ module foundation.decidable-types where
Imports
```agda
+open import foundation.action-on-identifications-functions
open import foundation.coproduct-types
open import foundation.dependent-pair-types
open import foundation.double-negation
open import foundation.empty-types
+open import foundation.equivalences
open import foundation.hilberts-epsilon-operators
open import foundation.logical-equivalences
open import foundation.negation
open import foundation.propositional-truncations
open import foundation.raising-universe-levels
+open import foundation.retracts-of-types
open import foundation.type-arithmetic-empty-type
open import foundation.unit-type
open import foundation.universe-levels
open import foundation-core.cartesian-product-types
-open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.propositions
-open import foundation-core.retracts-of-types
+open import foundation-core.retractions
+open import foundation-core.sections
```
@@ -154,36 +157,6 @@ is-decidable-function-type' (inl a) d with d a
is-decidable-function-type' (inr na) d = inl (ex-falso ∘ na)
```
-### Dependent sums of a uniformly decidable family of types over a decidable base is decidable
-
-```agda
-is-decidable-Σ-uniformly-decidable-family :
- {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
- is-decidable A → (((a : A) → B a) + ((a : A) → ¬ B a)) → is-decidable (Σ A B)
-is-decidable-Σ-uniformly-decidable-family (inl a) (inl b) =
- inl (a , b a)
-is-decidable-Σ-uniformly-decidable-family (inl a) (inr b) =
- inr (λ x → b (pr1 x) (pr2 x))
-is-decidable-Σ-uniformly-decidable-family (inr a) _ =
- inr (λ x → a (pr1 x))
-```
-
-### Dependent products of uniformly decidable families over decidable bases are decidable
-
-```agda
-is-decidable-Π-uniformly-decidable-family :
- {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
- is-decidable A →
- (((a : A) → B a) + ((a : A) → ¬ (B a))) →
- is-decidable ((a : A) → (B a))
-is-decidable-Π-uniformly-decidable-family (inl a) (inl b) =
- inl b
-is-decidable-Π-uniformly-decidable-family (inl a) (inr b) =
- inr (λ f → b a (f a))
-is-decidable-Π-uniformly-decidable-family (inr na) _ =
- inl (ex-falso ∘ na)
-```
-
### The negation of a decidable type is decidable
```agda
@@ -202,7 +175,18 @@ module _
is-decidable-iff :
(A → B) → (B → A) → is-decidable A → is-decidable B
is-decidable-iff f g (inl a) = inl (f a)
- is-decidable-iff f g (inr na) = inr (λ b → na (g b))
+ is-decidable-iff f g (inr na) = inr (na ∘ g)
+
+ is-decidable-iff' :
+ A ↔ B → is-decidable A → is-decidable B
+ is-decidable-iff' (f , g) = is-decidable-iff f g
+
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ iff-is-decidable : A ↔ B → is-decidable A ↔ is-decidable B
+ iff-is-decidable e = is-decidable-iff' e , is-decidable-iff' (inv-iff e)
```
### Decidable types are closed under retracts
@@ -214,26 +198,70 @@ module _
is-decidable-retract-of :
A retract-of B → is-decidable B → is-decidable A
- is-decidable-retract-of (pair i (pair r H)) (inl b) = inl (r b)
- is-decidable-retract-of (pair i (pair r H)) (inr f) = inr (f ∘ i)
+ is-decidable-retract-of R = is-decidable-iff' (iff-retract' R)
+
+ is-decidable-retract-of' :
+ A retract-of B → is-decidable A → is-decidable B
+ is-decidable-retract-of' R = is-decidable-iff' (inv-iff (iff-retract' R))
```
### Decidable types are closed under equivalences
```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
is-decidable-is-equiv :
{f : A → B} → is-equiv f → is-decidable B → is-decidable A
- is-decidable-is-equiv {f} (pair (pair g G) (pair h H)) =
- is-decidable-retract-of (pair f (pair h H))
+ is-decidable-is-equiv {f} H =
+ is-decidable-retract-of (retract-equiv (f , H))
is-decidable-equiv :
- (e : A ≃ B) → is-decidable B → is-decidable A
+ A ≃ B → is-decidable B → is-decidable A
is-decidable-equiv e = is-decidable-iff (map-inv-equiv e) (map-equiv e)
-is-decidable-equiv' :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} (e : A ≃ B) →
- is-decidable A → is-decidable B
-is-decidable-equiv' e = is-decidable-equiv (inv-equiv e)
+ is-decidable-equiv' :
+ A ≃ B → is-decidable A → is-decidable B
+ is-decidable-equiv' e = is-decidable-iff (map-equiv e) (map-inv-equiv e)
+```
+
+### Equivalent types have equivalent decidability predicates
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (e : A ≃ B)
+ where
+
+ map-equiv-is-decidable : is-decidable A → is-decidable B
+ map-equiv-is-decidable = is-decidable-equiv' e
+
+ map-inv-equiv-is-decidable : is-decidable B → is-decidable A
+ map-inv-equiv-is-decidable = is-decidable-equiv e
+
+ is-section-map-inv-equiv-is-decidable :
+ is-section map-equiv-is-decidable map-inv-equiv-is-decidable
+ is-section-map-inv-equiv-is-decidable (inl x) =
+ ap inl (is-section-map-inv-equiv e x)
+ is-section-map-inv-equiv-is-decidable (inr x) =
+ ap inr eq-neg
+
+ is-retraction-map-inv-equiv-is-decidable :
+ is-retraction map-equiv-is-decidable map-inv-equiv-is-decidable
+ is-retraction-map-inv-equiv-is-decidable (inl x) =
+ ap inl (is-retraction-map-inv-equiv e x)
+ is-retraction-map-inv-equiv-is-decidable (inr x) =
+ ap inr eq-neg
+
+ is-equiv-map-equiv-is-decidable : is-equiv map-equiv-is-decidable
+ is-equiv-map-equiv-is-decidable =
+ is-equiv-is-invertible
+ map-inv-equiv-is-decidable
+ is-section-map-inv-equiv-is-decidable
+ is-retraction-map-inv-equiv-is-decidable
+
+ equiv-is-decidable : is-decidable A ≃ is-decidable B
+ equiv-is-decidable = map-equiv-is-decidable , is-equiv-map-equiv-is-decidable
```
### Decidability implies double negation elimination
diff --git a/src/foundation/diaconescus-theorem.lagda.md b/src/foundation/diaconescus-theorem.lagda.md
new file mode 100644
index 0000000000..adcacce8cc
--- /dev/null
+++ b/src/foundation/diaconescus-theorem.lagda.md
@@ -0,0 +1,91 @@
+# Diaconescu's theorem
+
+```agda
+module foundation.diaconescus-theorem where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.axiom-of-choice
+open import foundation.booleans
+open import foundation.decidable-propositions
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.law-of-excluded-middle
+open import foundation.logical-equivalences
+open import foundation.propositional-truncations
+open import foundation.propositions
+open import foundation.universe-levels
+
+open import foundation-core.fibers-of-maps
+open import foundation-core.function-types
+open import foundation-core.identity-types
+
+open import synthetic-homotopy-theory.suspensions-of-propositions
+open import synthetic-homotopy-theory.suspensions-of-types
+```
+
+
+
+## Idea
+
+The [axiom of choice](foundation.axiom-of-choice.md) implies the
+[law of excluded middle](foundation.law-of-excluded-middle.md). This is often
+referred to as
+{{#concept "Diaconescu's theorem" WD="Diaconescu's theorem" WDID=Q3527059 Agda=theorem-Diaconescu}}.
+
+## Theorem
+
+We follow the proof given under Theorem 10.1.14 in {{#cite UF13}}.
+
+**Proof.** Given a [proposition](foundation-core.propositions.md) `P`, then its
+[suspension](synthetic-homotopy-theory.suspensions-of-propositions.md) `ΣP` is a
+[set](foundation-core.sets.md) with the property that `(N = S) ≃ ΣP`, where `N`
+and `S` are the _poles_ of `ΣP`. There is a surjection from the
+[booleans](foundation.booleans.md) onto the suspension `f : bool ↠ ΣP` such that
+`f true ≐ N` and `f false ≐ S`. Its
+[fiber family](foundation-core.fibers-of-maps.md) is in other words an
+[inhabited](foundation.inhabited-types.md) family over `ΣP`. Applying the axiom
+of choice to this family, we obtain a
+[mere](foundation.propositional-truncations.md)
+[section](foundation-core.sections.md) `s` of `f` which thus exhibits `P` as a
+[logical equivalent](foundation.logical-equivalences.md) to `f⁻¹ N = f⁻¹ S`.
+The latter is an [equation](foundation-core.identity-types.md) of booleans, and
+the booleans have [decidable equality](foundation.decidable-equality.md) so `P`
+must also be [decidable](foundation.decidable-propositions.md).
+
+```agda
+instance-theorem-Diaconescu :
+ {l : Level} (P : Prop l) →
+ instance-choice₀
+ ( suspension-set-Prop P)
+ ( fiber map-surjection-bool-suspension) →
+ is-decidable-type-Prop P
+instance-theorem-Diaconescu P ac-P =
+ rec-trunc-Prop
+ ( is-decidable-Prop P)
+ ( λ s →
+ is-decidable-iff'
+ ( ( iff-equiv (compute-eq-north-south-suspension-Prop P)) ∘iff
+ ( ( λ p →
+ ( inv (pr2 (s north-suspension))) ∙
+ ( ap map-surjection-bool-suspension p) ∙
+ ( pr2 (s south-suspension))) ,
+ ( ap (pr1 ∘ s))))
+ ( has-decidable-equality-bool
+ ( pr1 (s north-suspension))
+ ( pr1 (s south-suspension))))
+ ( ac-P is-surjective-map-surjection-bool-suspension)
+
+theorem-Diaconescu :
+ {l : Level} → level-AC0 l l → LEM l
+theorem-Diaconescu ac P =
+ instance-theorem-Diaconescu P
+ ( ac (suspension-set-Prop P) (fiber map-surjection-bool-suspension))
+```
+
+## References
+
+{{#bibliography}}
diff --git a/src/foundation/discrete-relaxed-sigma-decompositions.lagda.md b/src/foundation/discrete-relaxed-sigma-decompositions.lagda.md
index e77ad9f312..76cdac43a7 100644
--- a/src/foundation/discrete-relaxed-sigma-decompositions.lagda.md
+++ b/src/foundation/discrete-relaxed-sigma-decompositions.lagda.md
@@ -85,13 +85,7 @@ module _
( right-unit-law-Σ-is-contr is-discrete ∘e
matching-correspondence-Relaxed-Σ-Decomposition D))
pr1 (pr2 equiv-discrete-is-discrete-Relaxed-Σ-Decomposition) x =
- ( map-equiv (compute-raise-unit l4) ∘
- terminal-map (cotype-Relaxed-Σ-Decomposition D x) ,
- is-equiv-comp
- ( map-equiv (compute-raise-unit l4))
- ( terminal-map (cotype-Relaxed-Σ-Decomposition D x))
- ( is-equiv-terminal-map-is-contr (is-discrete x))
- ( is-equiv-map-equiv ( compute-raise-unit l4)))
+ equiv-raise-unit-is-contr (is-discrete x)
pr2 (pr2 equiv-discrete-is-discrete-Relaxed-Σ-Decomposition) a =
eq-pair-Σ
( ap ( λ f → map-equiv f a)
diff --git a/src/foundation/discrete-sigma-decompositions.lagda.md b/src/foundation/discrete-sigma-decompositions.lagda.md
index 19316103ef..3ed376b244 100644
--- a/src/foundation/discrete-sigma-decompositions.lagda.md
+++ b/src/foundation/discrete-sigma-decompositions.lagda.md
@@ -87,13 +87,7 @@ module _
( right-unit-law-Σ-is-contr is-discrete ∘e
matching-correspondence-Σ-Decomposition D))
pr1 (pr2 equiv-discrete-is-discrete-Σ-Decomposition) x =
- ( map-equiv (compute-raise-unit l4) ∘
- terminal-map (cotype-Σ-Decomposition D x) ,
- is-equiv-comp
- ( map-equiv (compute-raise-unit l4))
- ( terminal-map (cotype-Σ-Decomposition D x))
- ( is-equiv-terminal-map-is-contr (is-discrete x))
- ( is-equiv-map-equiv ( compute-raise-unit l4)))
+ equiv-raise-unit-is-contr (is-discrete x)
pr2 (pr2 equiv-discrete-is-discrete-Σ-Decomposition) a =
eq-pair-Σ
( ap ( λ f → map-equiv f a)
diff --git a/src/foundation/double-negation-modality.lagda.md b/src/foundation/double-negation-modality.lagda.md
index a3158f9e97..89f91c3ee0 100644
--- a/src/foundation/double-negation-modality.lagda.md
+++ b/src/foundation/double-negation-modality.lagda.md
@@ -19,6 +19,8 @@ open import foundation.universe-levels
open import foundation-core.function-types
open import foundation-core.transport-along-identifications
+open import logic.double-negation-elimination
+
open import orthogonal-factorization-systems.continuation-modalities
open import orthogonal-factorization-systems.large-lawvere-tierney-topologies
open import orthogonal-factorization-systems.lawvere-tierney-topologies
diff --git a/src/foundation/double-negation-stable-propositions.lagda.md b/src/foundation/double-negation-stable-propositions.lagda.md
index 760c158524..0ba325bf24 100644
--- a/src/foundation/double-negation-stable-propositions.lagda.md
+++ b/src/foundation/double-negation-stable-propositions.lagda.md
@@ -7,14 +7,36 @@ module foundation.double-negation-stable-propositions where
Imports
```agda
+open import foundation.cartesian-product-types
+open import foundation.conjunction
+open import foundation.coproduct-types
+open import foundation.decidable-propositions
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.disjunction
open import foundation.double-negation
+open import foundation.embeddings
open import foundation.empty-types
+open import foundation.equivalences
+open import foundation.existential-quantification
+open import foundation.logical-equivalences
open import foundation.negation
+open import foundation.propositional-extensionality
+open import foundation.propositions
+open import foundation.sets
+open import foundation.subtypes
+open import foundation.transport-along-identifications
+open import foundation.type-arithmetic-dependent-pair-types
open import foundation.unit-type
+open import foundation.universal-quantification
open import foundation.universe-levels
+open import foundation-core.contractible-types
open import foundation-core.function-types
-open import foundation-core.propositions
+open import foundation-core.identity-types
+open import foundation-core.retracts-of-types
+
+open import logic.double-negation-elimination
```
@@ -52,32 +74,327 @@ module _
is-prop-type-Prop is-double-negation-stable-Prop
```
+### The predicate on a type of being a double negation stable proposition
+
+```agda
+is-double-negation-stable-prop : {l : Level} → UU l → UU l
+is-double-negation-stable-prop X = (is-prop X) × (¬¬ X → X)
+
+is-prop-is-double-negation-stable-prop :
+ {l : Level} (X : UU l) → is-prop (is-double-negation-stable-prop X)
+is-prop-is-double-negation-stable-prop X =
+ is-prop-Σ
+ ( is-prop-is-prop X)
+ ( λ is-prop-X → is-prop-is-double-negation-stable (X , is-prop-X))
+
+is-double-negation-stable-prop-Prop : {l : Level} → UU l → Prop l
+is-double-negation-stable-prop-Prop X =
+ ( is-double-negation-stable-prop X , is-prop-is-double-negation-stable-prop X)
+
+module _
+ {l : Level} {A : UU l} (H : is-double-negation-stable-prop A)
+ where
+
+ is-prop-type-is-double-negation-stable-prop : is-prop A
+ is-prop-type-is-double-negation-stable-prop = pr1 H
+
+ has-double-negation-elim-is-double-negation-stable-prop :
+ has-double-negation-elim A
+ has-double-negation-elim-is-double-negation-stable-prop = pr2 H
+```
+
+### The subuniverse of double negation stable propositions
+
+```agda
+Double-Negation-Stable-Prop : (l : Level) → UU (lsuc l)
+Double-Negation-Stable-Prop l = Σ (UU l) (is-double-negation-stable-prop)
+
+module _
+ {l : Level} (P : Double-Negation-Stable-Prop l)
+ where
+
+ type-Double-Negation-Stable-Prop : UU l
+ type-Double-Negation-Stable-Prop = pr1 P
+
+ is-double-negation-stable-prop-type-Double-Negation-Stable-Prop :
+ is-double-negation-stable-prop type-Double-Negation-Stable-Prop
+ is-double-negation-stable-prop-type-Double-Negation-Stable-Prop = pr2 P
+
+ is-prop-type-Double-Negation-Stable-Prop :
+ is-prop type-Double-Negation-Stable-Prop
+ is-prop-type-Double-Negation-Stable-Prop =
+ is-prop-type-is-double-negation-stable-prop
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop)
+
+ prop-Double-Negation-Stable-Prop : Prop l
+ prop-Double-Negation-Stable-Prop =
+ ( type-Double-Negation-Stable-Prop ,
+ is-prop-type-Double-Negation-Stable-Prop)
+
+ has-double-negation-elim-type-Double-Negation-Stable-Prop :
+ has-double-negation-elim type-Double-Negation-Stable-Prop
+ has-double-negation-elim-type-Double-Negation-Stable-Prop =
+ has-double-negation-elim-is-double-negation-stable-prop
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop)
+```
+
## Properties
+### The forgetful map from double negation stable propositions to propositions is an embedding
+
+```agda
+is-emb-prop-Double-Negation-Stable-Prop :
+ {l : Level} → is-emb (prop-Double-Negation-Stable-Prop {l})
+is-emb-prop-Double-Negation-Stable-Prop =
+ is-emb-tot
+ ( λ X →
+ is-emb-inclusion-subtype
+ ( λ H →
+ has-double-negation-elim X , is-prop-has-double-negation-elim H))
+
+emb-prop-Double-Negation-Stable-Prop :
+ {l : Level} → Double-Negation-Stable-Prop l ↪ Prop l
+emb-prop-Double-Negation-Stable-Prop =
+ ( prop-Double-Negation-Stable-Prop , is-emb-prop-Double-Negation-Stable-Prop)
+```
+
+### The subuniverse of double negation stable propositions is a set
+
+```agda
+is-set-Double-Negation-Stable-Prop :
+ {l : Level} → is-set (Double-Negation-Stable-Prop l)
+is-set-Double-Negation-Stable-Prop {l} =
+ is-set-emb emb-prop-Double-Negation-Stable-Prop is-set-type-Prop
+
+set-Double-Negation-Stable-Prop : (l : Level) → Set (lsuc l)
+set-Double-Negation-Stable-Prop l =
+ ( Double-Negation-Stable-Prop l , is-set-Double-Negation-Stable-Prop)
+```
+
+### Extensionality of double negation stable propositions
+
+```agda
+module _
+ {l : Level} (P Q : Double-Negation-Stable-Prop l)
+ where
+
+ extensionality-Double-Negation-Stable-Prop :
+ ( P = Q) ≃
+ ( type-Double-Negation-Stable-Prop P ↔ type-Double-Negation-Stable-Prop Q)
+ extensionality-Double-Negation-Stable-Prop =
+ ( propositional-extensionality
+ ( prop-Double-Negation-Stable-Prop P)
+ ( prop-Double-Negation-Stable-Prop Q)) ∘e
+ ( equiv-ap-emb emb-prop-Double-Negation-Stable-Prop)
+
+ iff-eq-Double-Negation-Stable-Prop :
+ P = Q →
+ type-Double-Negation-Stable-Prop P ↔ type-Double-Negation-Stable-Prop Q
+ iff-eq-Double-Negation-Stable-Prop =
+ map-equiv extensionality-Double-Negation-Stable-Prop
+
+ eq-iff-Double-Negation-Stable-Prop :
+ (type-Double-Negation-Stable-Prop P → type-Double-Negation-Stable-Prop Q) →
+ (type-Double-Negation-Stable-Prop Q → type-Double-Negation-Stable-Prop P) →
+ P = Q
+ eq-iff-Double-Negation-Stable-Prop f g =
+ map-inv-equiv extensionality-Double-Negation-Stable-Prop (pair f g)
+```
+
+### Double negation stable propositions are closed under retracts
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-double-negation-stable-prop-retract :
+ A retract-of B →
+ is-double-negation-stable-prop B →
+ is-double-negation-stable-prop A
+ is-double-negation-stable-prop-retract e H =
+ ( is-prop-retract-of e
+ ( is-prop-type-is-double-negation-stable-prop H)) ,
+ ( has-double-negation-elim-retract e
+ ( has-double-negation-elim-is-double-negation-stable-prop H))
+```
+
+### Double negation stable propositions are closed under equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-double-negation-stable-prop-equiv :
+ A ≃ B → is-double-negation-stable-prop B → is-double-negation-stable-prop A
+ is-double-negation-stable-prop-equiv e =
+ is-double-negation-stable-prop-retract (retract-equiv e)
+
+ is-double-negation-stable-prop-equiv' :
+ B ≃ A → is-double-negation-stable-prop B → is-double-negation-stable-prop A
+ is-double-negation-stable-prop-equiv' e =
+ is-double-negation-stable-prop-retract (retract-inv-equiv e)
+```
+
### The empty proposition is double negation stable
```agda
-is-double-negation-stable-empty : is-double-negation-stable empty-Prop
-is-double-negation-stable-empty e = e id
+empty-Double-Negation-Stable-Prop : Double-Negation-Stable-Prop lzero
+empty-Double-Negation-Stable-Prop =
+ ( empty , is-prop-empty , double-negation-elim-empty)
```
### The unit proposition is double negation stable
```agda
-is-double-negation-stable-unit : is-double-negation-stable unit-Prop
-is-double-negation-stable-unit _ = star
+unit-Double-Negation-Stable-Prop : Double-Negation-Stable-Prop lzero
+unit-Double-Negation-Stable-Prop =
+ ( unit , is-prop-unit , double-negation-elim-unit)
+```
+
+```agda
+is-double-negation-stable-prop-is-contr :
+ {l : Level} {A : UU l} → is-contr A → is-double-negation-stable-prop A
+is-double-negation-stable-prop-is-contr H =
+ (is-prop-is-contr H , double-negation-elim-is-contr H)
+```
+
+### Decidable propositions are double negation stable
+
+```agda
+double-negation-stable-prop-Decidable-Prop :
+ {l : Level} → Decidable-Prop l → Double-Negation-Stable-Prop l
+double-negation-stable-prop-Decidable-Prop (A , H , d) =
+ ( A , H , double-negation-elim-is-decidable d)
+```
+
+### Negations of types are double negation stable propositions
+
+```agda
+neg-type-Double-Negation-Stable-Prop :
+ {l : Level} → UU l → Double-Negation-Stable-Prop l
+neg-type-Double-Negation-Stable-Prop A =
+ ( ¬ A , is-prop-neg , double-negation-elim-neg A)
+
+neg-Double-Negation-Stable-Prop :
+ {l : Level} → Double-Negation-Stable-Prop l → Double-Negation-Stable-Prop l
+neg-Double-Negation-Stable-Prop P =
+ neg-type-Double-Negation-Stable-Prop (type-Double-Negation-Stable-Prop P)
+```
+
+### Universal quantification over double negation stable propositions is double negation stable
+
+```agda
+is-double-negation-stable-prop-Π :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ ((a : A) → is-double-negation-stable-prop (B a)) →
+ is-double-negation-stable-prop ((a : A) → B a)
+is-double-negation-stable-prop-Π b =
+ ( is-prop-Π (is-prop-type-is-double-negation-stable-prop ∘ b)) ,
+ ( double-negation-elim-for-all
+ ( has-double-negation-elim-is-double-negation-stable-prop ∘ b))
+
+Π-Double-Negation-Stable-Prop :
+ {l1 l2 : Level}
+ (A : UU l1) (B : A → Double-Negation-Stable-Prop l2) →
+ Double-Negation-Stable-Prop (l1 ⊔ l2)
+Π-Double-Negation-Stable-Prop A B =
+ ( (a : A) → type-Double-Negation-Stable-Prop (B a)) ,
+ ( is-double-negation-stable-prop-Π
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop ∘ B))
+```
+
+### Implication into double negation stable propositions is double negation stable
+
+```agda
+is-double-negation-stable-prop-exp :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} →
+ is-double-negation-stable-prop B →
+ is-double-negation-stable-prop (A → B)
+is-double-negation-stable-prop-exp b =
+ is-double-negation-stable-prop-Π (λ _ → b)
+
+exp-Double-Negation-Stable-Prop :
+ {l1 l2 : Level}
+ (A : UU l1) (B : Double-Negation-Stable-Prop l2) →
+ Double-Negation-Stable-Prop (l1 ⊔ l2)
+exp-Double-Negation-Stable-Prop A B = Π-Double-Negation-Stable-Prop A (λ _ → B)
+```
+
+### Dependent sums of double negation stable propositions over double negation stable propositions are double negation stable
+
+```agda
+is-double-negation-stable-prop-Σ :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ is-double-negation-stable-prop A →
+ ((a : A) → is-double-negation-stable-prop (B a)) →
+ is-double-negation-stable-prop (Σ A B)
+is-double-negation-stable-prop-Σ a b =
+ ( is-prop-Σ
+ ( is-prop-type-is-double-negation-stable-prop a)
+ ( is-prop-type-is-double-negation-stable-prop ∘ b)) ,
+ ( double-negation-elim-Σ-is-prop-base
+ ( is-prop-type-is-double-negation-stable-prop a)
+ ( has-double-negation-elim-is-double-negation-stable-prop a)
+ ( has-double-negation-elim-is-double-negation-stable-prop ∘ b))
+
+Σ-Double-Negation-Stable-Prop :
+ {l1 l2 : Level}
+ (A : Double-Negation-Stable-Prop l1)
+ (B : type-Double-Negation-Stable-Prop A → Double-Negation-Stable-Prop l2) →
+ Double-Negation-Stable-Prop (l1 ⊔ l2)
+Σ-Double-Negation-Stable-Prop A B =
+ ( Σ ( type-Double-Negation-Stable-Prop A)
+ ( type-Double-Negation-Stable-Prop ∘ B)) ,
+ ( is-double-negation-stable-prop-Σ
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop A)
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop ∘ B))
+```
+
+### The conjunction of two double negation stable propositions is double negation stable
+
+```agda
+is-double-negation-stable-prop-product :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} →
+ is-double-negation-stable-prop A →
+ is-double-negation-stable-prop B →
+ is-double-negation-stable-prop (A × B)
+is-double-negation-stable-prop-product a b =
+ ( is-prop-product
+ ( is-prop-type-is-double-negation-stable-prop a)
+ ( is-prop-type-is-double-negation-stable-prop b)) ,
+ ( double-negation-elim-product
+ ( has-double-negation-elim-is-double-negation-stable-prop a)
+ ( has-double-negation-elim-is-double-negation-stable-prop b))
+
+product-Double-Negation-Stable-Prop :
+ {l1 l2 : Level} →
+ Double-Negation-Stable-Prop l1 →
+ Double-Negation-Stable-Prop l2 →
+ Double-Negation-Stable-Prop (l1 ⊔ l2)
+product-Double-Negation-Stable-Prop A B =
+ ( ( type-Double-Negation-Stable-Prop A) ×
+ ( type-Double-Negation-Stable-Prop B)) ,
+ ( is-double-negation-stable-prop-product
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop A)
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop B))
```
-### The negation of a type is double negation stable
+### Negation has no fixed points on double negation stable propositions
```agda
-is-double-negation-stable-neg :
- {l : Level} (A : UU l) → is-double-negation-stable (neg-type-Prop A)
-is-double-negation-stable-neg = double-negation-elim-neg
+abstract
+ no-fixed-points-neg-Double-Negation-Stable-Prop :
+ {l : Level} (P : Double-Negation-Stable-Prop l) →
+ ¬ (type-Double-Negation-Stable-Prop P ↔
+ ¬ (type-Double-Negation-Stable-Prop P))
+ no-fixed-points-neg-Double-Negation-Stable-Prop P =
+ no-fixed-points-neg (type-Double-Negation-Stable-Prop P)
```
## See also
- [The double negation modality](foundation.double-negation-modality.md)
- [Irrefutable propositions](foundation.irrefutable-propositions.md) are double
- negation stable.
+ negation connected types.
diff --git a/src/foundation/double-negation.lagda.md b/src/foundation/double-negation.lagda.md
index 4529283ca7..ee813f6384 100644
--- a/src/foundation/double-negation.lagda.md
+++ b/src/foundation/double-negation.lagda.md
@@ -7,15 +7,12 @@ module foundation.double-negation where
Imports
```agda
-open import foundation.dependent-pair-types
open import foundation.negation
open import foundation.propositional-truncations
open import foundation.universe-levels
-open import foundation-core.cartesian-product-types
open import foundation-core.coproduct-types
open import foundation-core.empty-types
-open import foundation-core.function-types
open import foundation-core.propositions
```
@@ -94,45 +91,13 @@ double-negation-linearity-implication {P = P} {Q = Q} f =
( λ (p : P) → map-neg (inr {A = P → Q} {B = Q → P}) f (λ _ → p))
```
-### Cases of double negation elimination
-
-```agda
-double-negation-elim-neg : {l : Level} (P : UU l) → ¬¬¬ P → ¬ P
-double-negation-elim-neg P f p = f (λ g → g p)
-
-double-negation-elim-product :
- {l1 l2 : Level} {P : UU l1} {Q : UU l2} →
- ¬¬ ((¬¬ P) × (¬¬ Q)) → (¬¬ P) × (¬¬ Q)
-pr1 (double-negation-elim-product {P = P} {Q = Q} f) =
- double-negation-elim-neg (¬ P) (map-double-negation pr1 f)
-pr2 (double-negation-elim-product {P = P} {Q = Q} f) =
- double-negation-elim-neg (¬ Q) (map-double-negation pr2 f)
-
-double-negation-elim-exp :
- {l1 l2 : Level} {P : UU l1} {Q : UU l2} →
- ¬¬ (P → ¬¬ Q) → (P → ¬¬ Q)
-double-negation-elim-exp {P = P} {Q = Q} f p =
- double-negation-elim-neg
- ( ¬ Q)
- ( map-double-negation (λ (g : P → ¬¬ Q) → g p) f)
-
-double-negation-elim-for-all :
- {l1 l2 : Level} {P : UU l1} {Q : P → UU l2} →
- ¬¬ ((p : P) → ¬¬ (Q p)) → (p : P) → ¬¬ (Q p)
-double-negation-elim-for-all {P = P} {Q = Q} f p =
- double-negation-elim-neg
- ( ¬ (Q p))
- ( map-double-negation (λ (g : (u : P) → ¬¬ (Q u)) → g p) f)
-```
-
### Maps into double negations extend along `intro-double-negation`
```agda
double-negation-extend :
{l1 l2 : Level} {P : UU l1} {Q : UU l2} →
(P → ¬¬ Q) → (¬¬ P → ¬¬ Q)
-double-negation-extend {P = P} {Q = Q} f =
- double-negation-elim-neg (¬ Q) ∘ (map-double-negation f)
+double-negation-extend {P = P} {Q = Q} f nnp nq = nnp (λ p → f p nq)
```
### The double negation of a type is logically equivalent to the double negation of its propositional truncation
@@ -154,6 +119,11 @@ abstract
map-double-negation unit-trunc-Prop
```
+## See also
+
+- [Double negation elimination](logic.double-negation-elimination.md)
+- [Irrefutable propositions](foundation.irrefutable-propositions.md)
+
## Table of files about propositional logic
The following table gives an overview of basic constructions in propositional
diff --git a/src/foundation/equality-cartesian-product-types.lagda.md b/src/foundation/equality-cartesian-product-types.lagda.md
index eb42462766..b500eb9d43 100644
--- a/src/foundation/equality-cartesian-product-types.lagda.md
+++ b/src/foundation/equality-cartesian-product-types.lagda.md
@@ -7,11 +7,13 @@ module foundation.equality-cartesian-product-types where
Imports
```agda
+open import foundation.action-on-identifications-binary-functions
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.universe-levels
open import foundation-core.cartesian-product-types
+open import foundation-core.equality-dependent-pair-types
open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.homotopies
@@ -36,7 +38,7 @@ module _
where
Eq-product : (s t : A × B) → UU (l1 ⊔ l2)
- Eq-product s t = ((pr1 s) = (pr1 t)) × ((pr2 s) = (pr2 t))
+ Eq-product s t = (pr1 s = pr1 t) × (pr2 s = pr2 t)
```
## Properties
@@ -49,23 +51,22 @@ module _
where
eq-pair' : {s t : A × B} → Eq-product s t → s = t
- eq-pair' {pair x y} {pair .x .y} (pair refl refl) = refl
+ eq-pair' (p , q) = ap-binary pair p q
- eq-pair :
- {s t : A × B} → (pr1 s) = (pr1 t) → (pr2 s) = (pr2 t) → s = t
- eq-pair p q = eq-pair' (pair p q)
+ eq-pair : {s t : A × B} → pr1 s = pr1 t → pr2 s = pr2 t → s = t
+ eq-pair p q = eq-pair' (p , q)
pair-eq : {s t : A × B} → s = t → Eq-product s t
pr1 (pair-eq α) = ap pr1 α
pr2 (pair-eq α) = ap pr2 α
is-retraction-pair-eq :
- {s t : A × B} → ((pair-eq {s} {t}) ∘ (eq-pair' {s} {t})) ~ id
- is-retraction-pair-eq {pair x y} {pair .x .y} (pair refl refl) = refl
+ {s t : A × B} → pair-eq {s} {t} ∘ eq-pair' {s} {t} ~ id
+ is-retraction-pair-eq (refl , refl) = refl
is-section-pair-eq :
- {s t : A × B} → ((eq-pair' {s} {t}) ∘ (pair-eq {s} {t})) ~ id
- is-section-pair-eq {pair x y} {pair .x .y} refl = refl
+ {s t : A × B} → eq-pair' {s} {t} ∘ pair-eq {s} {t} ~ id
+ is-section-pair-eq refl = refl
abstract
is-equiv-eq-pair :
@@ -99,13 +100,13 @@ module _
triangle-eq-pair :
{a0 a1 : A} {b0 b1 : B} (p : a0 = a1) (q : b0 = b1) →
- eq-pair p q = ((eq-pair p refl) ∙ (eq-pair refl q))
- triangle-eq-pair refl refl = refl
+ eq-pair p q = eq-pair p refl ∙ eq-pair refl q
+ triangle-eq-pair refl q = refl
triangle-eq-pair' :
{a0 a1 : A} {b0 b1 : B} (p : a0 = a1) (q : b0 = b1) →
- eq-pair p q = ((eq-pair refl q) ∙ (eq-pair p refl))
- triangle-eq-pair' refl refl = refl
+ eq-pair p q = eq-pair refl q ∙ eq-pair p refl
+ triangle-eq-pair' p refl = refl
```
### `eq-pair` preserves concatenation
@@ -114,9 +115,7 @@ module _
eq-pair-concat :
{l1 l2 : Level} {A : UU l1} {B : UU l2} {x x' x'' : A} {y y' y'' : B}
(p : x = x') (p' : x' = x'') (q : y = y') (q' : y' = y'') →
- ( eq-pair {s = pair x y} {t = pair x'' y''} (p ∙ p') (q ∙ q')) =
- ( ( eq-pair {s = pair x y} {t = pair x' y'} p q) ∙
- ( eq-pair p' q'))
+ eq-pair (p ∙ p') (q ∙ q') = eq-pair p q ∙ eq-pair p' q'
eq-pair-concat refl p' refl q' = refl
```
@@ -126,13 +125,13 @@ eq-pair-concat refl p' refl q' = refl
ap-pr1-eq-pair :
{l1 l2 : Level} {A : UU l1} {B : UU l2}
{x x' : A} (p : x = x') {y y' : B} (q : y = y') →
- ap pr1 (eq-pair {s = pair x y} {pair x' y'} p q) = p
+ ap pr1 (eq-pair p q) = p
ap-pr1-eq-pair refl refl = refl
ap-pr2-eq-pair :
{l1 l2 : Level} {A : UU l1} {B : UU l2}
{x x' : A} (p : x = x') {y y' : B} (q : y = y') →
- ap pr2 (eq-pair {s = pair x y} {pair x' y'} p q) = q
+ ap pr2 (eq-pair p q) = q
ap-pr2-eq-pair refl refl = refl
```
@@ -155,12 +154,12 @@ module _
```agda
left-unit-law-tr-eq-pair :
(C : A × B → UU l3) (q : b0 = b1) (u : C (a0 , b0)) →
- (tr C (eq-pair refl q) u) = tr (λ x → C (a0 , x)) q u
+ tr C (eq-pair refl q) u = tr (λ x → C (a0 , x)) q u
left-unit-law-tr-eq-pair C refl u = refl
right-unit-law-tr-eq-pair :
(C : A × B → UU l3) (p : a0 = a1) (u : C (a0 , b0)) →
- (tr C (eq-pair p refl) u) = tr (λ x → C (x , b0)) p u
+ tr C (eq-pair p refl) u = tr (λ x → C (x , b0)) p u
right-unit-law-tr-eq-pair C refl u = refl
```
diff --git a/src/foundation/equality-dependent-pair-types.lagda.md b/src/foundation/equality-dependent-pair-types.lagda.md
index 5126671f54..6ab7a7bbb3 100644
--- a/src/foundation/equality-dependent-pair-types.lagda.md
+++ b/src/foundation/equality-dependent-pair-types.lagda.md
@@ -192,13 +192,13 @@ module _
coh-eq-base-Σ² :
{s t : Σ A (λ x → Σ (B x) λ y → C x y)} (p : s = t) →
- eq-base-eq-pair-Σ p =
- eq-base-eq-pair-Σ (eq-base-eq-pair-Σ (ap (map-inv-associative-Σ' A B C) p))
+ eq-base-eq-pair p =
+ eq-base-eq-pair (eq-base-eq-pair (ap (map-inv-associative-Σ' A B C) p))
coh-eq-base-Σ² refl = refl
dependent-eq-second-component-eq-Σ² :
{s t : Σ A (λ x → Σ (B x) λ y → C x y)} (p : s = t) →
- dependent-identification B (eq-base-eq-pair-Σ p) (pr1 (pr2 s)) (pr1 (pr2 t))
+ dependent-identification B (eq-base-eq-pair p) (pr1 (pr2 s)) (pr1 (pr2 t))
dependent-eq-second-component-eq-Σ² {s = s} {t = t} p =
( ap (λ q → tr B q (pr1 (pr2 s))) (coh-eq-base-Σ² p)) ∙
( pr2
@@ -219,14 +219,14 @@ module _
coh-eq-base-Σ³ :
{ s t : Σ A (λ x → Σ (B x) (λ y → Σ (C x) (D x y)))} (p : s = t) →
- eq-base-eq-pair-Σ p =
- eq-base-eq-pair-Σ (ap (map-equiv (interchange-Σ-Σ-Σ D)) p)
+ eq-base-eq-pair p =
+ eq-base-eq-pair (ap (map-equiv (interchange-Σ-Σ-Σ D)) p)
coh-eq-base-Σ³ refl = refl
dependent-eq-third-component-eq-Σ³ :
{ s t : Σ A (λ x → Σ (B x) (λ y → Σ (C x) (D x y)))} (p : s = t) →
dependent-identification C
- ( eq-base-eq-pair-Σ p)
+ ( eq-base-eq-pair p)
( pr1 (pr2 (pr2 s)))
( pr1 (pr2 (pr2 t)))
dependent-eq-third-component-eq-Σ³ {s = s} {t = t} p =
diff --git a/src/foundation/equivalence-injective-type-families.lagda.md b/src/foundation/equivalence-injective-type-families.lagda.md
index d565eb4308..5b1ea1cb5f 100644
--- a/src/foundation/equivalence-injective-type-families.lagda.md
+++ b/src/foundation/equivalence-injective-type-families.lagda.md
@@ -37,9 +37,7 @@ as a map.
**Note.** The concept of equivalence injective type family as considered here is
unrelated to the concept of "injective type" as studied by Martín Escardó in
-_Injective types in univalent mathematics_
-([arXiv:1903.01211](https://arxiv.org/abs/1903.01211),
-[TypeTopology](https://www.cs.bham.ac.uk/~mhe/TypeTopology/InjectiveTypes.index.html)).
+_Injective types in univalent mathematics_ {{#cite Esc21b}}.
## Definition
@@ -108,3 +106,7 @@ module _
pr1 is-equivalence-injective-Prop = is-equivalence-injective P
pr2 is-equivalence-injective-Prop = is-prop-is-equivalence-injective
```
+
+## References
+
+{{#bibliography}}
diff --git a/src/foundation/fibers-of-maps.lagda.md b/src/foundation/fibers-of-maps.lagda.md
index e0995d367b..01e4a3f77b 100644
--- a/src/foundation/fibers-of-maps.lagda.md
+++ b/src/foundation/fibers-of-maps.lagda.md
@@ -12,6 +12,7 @@ open import foundation-core.fibers-of-maps public
open import foundation.action-on-identifications-functions
open import foundation.cones-over-cospan-diagrams
open import foundation.dependent-pair-types
+open import foundation.postcomposition-functions
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.type-arithmetic-unit-type
open import foundation.unit-type
@@ -20,6 +21,7 @@ open import foundation.universe-levels
open import foundation-core.contractible-types
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
@@ -109,6 +111,25 @@ module _
inv-equiv equiv-total-fiber-terminal-map
```
+### Computing the fibers of postcomposition by a projection map
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (B : A → UU l2)
+ where
+
+ compute-fiber-postcomp-pr1 :
+ {l3 : Level} {X : UU l3} (h : X → A) →
+ ((i : X) → B (h i)) ≃ fiber (postcomp X (pr1 {B = B})) h
+ compute-fiber-postcomp-pr1 {X = X} h =
+ compute-Π-fiber-postcomp X pr1 h ∘e
+ equiv-Π-equiv-family (λ i → inv-equiv-fiber-pr1 B (h i))
+
+ compute-fiber-id-postcomp-pr1 :
+ ((a : A) → B a) ≃ fiber (postcomp A (pr1 {B = B})) id
+ compute-fiber-id-postcomp-pr1 = compute-fiber-postcomp-pr1 id
+```
+
### Transport in fibers
```agda
diff --git a/src/foundation/function-types.lagda.md b/src/foundation/function-types.lagda.md
index 9af9c22baf..d6fb4ed16f 100644
--- a/src/foundation/function-types.lagda.md
+++ b/src/foundation/function-types.lagda.md
@@ -111,6 +111,11 @@ module _
{l1 l2 l3 : Level} {A : UU l1} {x y : A} (B : A → UU l2) (C : UU l3)
where
+ tr-function-type-fixed-codomain :
+ (p : x = y) (f : B x → C) →
+ tr (λ a → B a → C) p f = f ∘ tr B (inv p)
+ tr-function-type-fixed-codomain refl f = refl
+
compute-dependent-identification-function-type-fixed-codomain :
(p : x = y) (f : B x → C) (g : B y → C) →
((b : B x) → f b = g (tr B p b)) ≃
@@ -127,6 +132,19 @@ module _
( compute-dependent-identification-function-type-fixed-codomain p f g)
```
+### Transport in a family of function types with fixed domain
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {x y : A} (B : UU l2) (C : A → UU l3)
+ where
+
+ tr-function-type-fixed-domain :
+ (p : x = y) (f : B → C x) →
+ tr (λ a → B → C a) p f = tr C p ∘ f
+ tr-function-type-fixed-domain refl f = refl
+```
+
### Relation between `compute-dependent-identification-function-type` and `preserves-tr`
```agda
@@ -216,7 +234,7 @@ module _
( H s)
( k s)
( l s)))
- ( λ k l s → inv-equiv (equiv-funext)))) ∙
+ ( λ k l s → inv-equiv equiv-funext))) ∙
( eq-htpy-refl-htpy (h (i s))))
```
diff --git a/src/foundation/functoriality-dependent-function-types.lagda.md b/src/foundation/functoriality-dependent-function-types.lagda.md
index bb38a2d469..5993ac6c2a 100644
--- a/src/foundation/functoriality-dependent-function-types.lagda.md
+++ b/src/foundation/functoriality-dependent-function-types.lagda.md
@@ -103,15 +103,8 @@ module _
( f (map-inv-equiv e (map-equiv e a')))
( h (map-inv-equiv e (map-equiv e a')))))
( coherence-map-inv-equiv e a')) ∙
- ( ( tr-ap
- ( map-equiv e)
- ( λ _ → id)
- ( is-retraction-map-inv-equiv e a')
- ( map-equiv
- ( f (map-inv-equiv e (map-equiv e a')))
- ( h (map-inv-equiv e (map-equiv e a'))))) ∙
- ( α ( map-inv-equiv e (map-equiv e a'))
- ( is-retraction-map-inv-equiv e a')))
+ ( substitution-law-tr B (map-equiv e) (is-retraction-map-inv-equiv e a')) ∙
+ ( α (map-inv-equiv e (map-equiv e a')) (is-retraction-map-inv-equiv e a'))
where
α :
(x : A') (p : x = a') →
diff --git a/src/foundation/hilberts-epsilon-operators.lagda.md b/src/foundation/hilberts-epsilon-operators.lagda.md
index 6ad95cd052..1c23904a8b 100644
--- a/src/foundation/hilberts-epsilon-operators.lagda.md
+++ b/src/foundation/hilberts-epsilon-operators.lagda.md
@@ -19,8 +19,14 @@ open import foundation-core.function-types
## Idea
-Hilbert's $ε$-operator at a type `A` is a map `type-trunc-Prop A → A`. Contrary
-to Hilbert, we will not assume that such an operator exists for each type `A`.
+{{#concept "Hilbert's $ε$-operator"}} at a type `A` is a map
+
+```text
+ ε : ║A║₋₁ → A
+```
+
+Some authors also refer to this as _split support_ {{#cite KECA17}}. Contrary to
+Hilbert, we will not assume that such an operator exists for each type `A`.
## Definition
@@ -46,3 +52,12 @@ to Hilbert, we will not assume that such an operator exists for each type `A`.
ε-operator-equiv' e f =
(map-inv-equiv e ∘ f) ∘ (map-trunc-Prop (map-equiv e))
```
+
+## References
+
+{{#bibliography}}
+
+## External links
+
+- [Epsilon calculus](https://en.wikipedia.org/wiki/Epsilon_calculus) at
+ Wikipedia
diff --git a/src/foundation/images.lagda.md b/src/foundation/images.lagda.md
index f678e73231..4ffd596c45 100644
--- a/src/foundation/images.lagda.md
+++ b/src/foundation/images.lagda.md
@@ -36,7 +36,9 @@ open import foundation-core.truncation-levels
## Idea
-The **image** of a map is a type that satisfies the
+The
+{{#concept "image" Disambiguation="of a map" WD="image" WDID=Q860623 Agda=im}}
+of a map is a type that satisfies the
[universal property of the image](foundation.universal-property-image.md) of a
map.
@@ -73,7 +75,7 @@ module _
## Properties
-### We characterize the identity type of im f
+### We characterize the identity type of `im f`
```agda
module _
@@ -152,7 +154,7 @@ abstract
( trunc-Prop (fiber (map-unit-im f) (y , z)))
( α)
where
- α : fiber f y → type-Prop (trunc-Prop (fiber (map-unit-im f) (y , z)))
+ α : fiber f y → type-trunc-Prop (fiber (map-unit-im f) (y , z))
α (x , p) = unit-trunc-Prop (x , eq-type-subtype (trunc-Prop ∘ fiber f) p)
```
@@ -216,3 +218,8 @@ im-1-Type :
(f : A → type-1-Type X) → 1-Type (l1 ⊔ l2)
im-1-Type X f = im-Truncated-Type zero-𝕋 X f
```
+
+## External links
+
+- [Image (mathematics)]() at
+ Wikipedia
diff --git a/src/foundation/irrefutable-propositions.lagda.md b/src/foundation/irrefutable-propositions.lagda.md
index 7e0dd013dc..a6a0d7fcde 100644
--- a/src/foundation/irrefutable-propositions.lagda.md
+++ b/src/foundation/irrefutable-propositions.lagda.md
@@ -7,18 +7,22 @@ module foundation.irrefutable-propositions where
Imports
```agda
+open import foundation.cartesian-product-types
+open import foundation.contractible-types
open import foundation.coproduct-types
open import foundation.decidable-propositions
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.double-negation
+open import foundation.empty-types
open import foundation.function-types
open import foundation.negation
+open import foundation.propositions
open import foundation.subuniverses
open import foundation.unit-type
open import foundation.universe-levels
-open import foundation-core.propositions
+open import logic.double-negation-elimination
```
@@ -35,26 +39,57 @@ The [subuniverse](foundation.subuniverses.md) of
### The predicate on a proposition of being irrefutable
```agda
-is-irrefutable : {l : Level} → Prop l → UU l
-is-irrefutable P = ¬¬ (type-Prop P)
+module _
+ {l : Level} (P : Prop l)
+ where
+
+ is-irrefutable : UU l
+ is-irrefutable = ¬¬ (type-Prop P)
-is-prop-is-irrefutable : {l : Level} (P : Prop l) → is-prop (is-irrefutable P)
-is-prop-is-irrefutable P = is-prop-double-negation
+ is-prop-is-irrefutable : is-prop is-irrefutable
+ is-prop-is-irrefutable = is-prop-double-negation
-is-irrefutable-Prop : {l : Level} → Prop l → Prop l
-is-irrefutable-Prop = double-negation-Prop
+ is-irrefutable-Prop : Prop l
+ is-irrefutable-Prop = double-negation-Prop P
```
-### The subuniverse of irrefutable propositions
+### The predicate on a type of being an irrefutable proposition
```agda
-subuniverse-Irrefutable-Prop : (l : Level) → subuniverse l l
-subuniverse-Irrefutable-Prop l A =
- product-Prop (is-prop-Prop A) (double-negation-type-Prop A)
+module _
+ {l : Level} (P : UU l)
+ where
+
+ is-irrefutable-prop : UU l
+ is-irrefutable-prop = is-prop P × (¬¬ P)
+
+ is-prop-is-irrefutable-prop : is-prop is-irrefutable-prop
+ is-prop-is-irrefutable-prop =
+ is-prop-product (is-prop-is-prop P) is-prop-double-negation
+
+ is-irrefutable-prop-Prop : Prop l
+ is-irrefutable-prop-Prop = is-irrefutable-prop , is-prop-is-irrefutable-prop
+
+module _
+ {l : Level} {P : UU l} (H : is-irrefutable-prop P)
+ where
+
+ is-prop-type-is-irrefutable-prop : is-prop P
+ is-prop-type-is-irrefutable-prop = pr1 H
+
+ prop-is-irrefutable-prop : Prop l
+ prop-is-irrefutable-prop = P , is-prop-type-is-irrefutable-prop
+ is-irrefutable-is-irrefutable-prop :
+ is-irrefutable (P , is-prop-type-is-irrefutable-prop)
+ is-irrefutable-is-irrefutable-prop = pr2 H
+```
+
+### The subuniverse of irrefutable propositions
+
+```agda
Irrefutable-Prop : (l : Level) → UU (lsuc l)
-Irrefutable-Prop l =
- type-subuniverse (subuniverse-Irrefutable-Prop l)
+Irrefutable-Prop l = type-subuniverse is-irrefutable-prop-Prop
make-Irrefutable-Prop :
{l : Level} (P : Prop l) → is-irrefutable P → Irrefutable-Prop l
@@ -68,18 +103,47 @@ module _
type-Irrefutable-Prop : UU l
type-Irrefutable-Prop = pr1 P
+ is-irrefutable-prop-type-Irrefutable-Prop :
+ is-irrefutable-prop type-Irrefutable-Prop
+ is-irrefutable-prop-type-Irrefutable-Prop = pr2 P
+
is-prop-type-Irrefutable-Prop : is-prop type-Irrefutable-Prop
- is-prop-type-Irrefutable-Prop = pr1 (pr2 P)
+ is-prop-type-Irrefutable-Prop =
+ is-prop-type-is-irrefutable-prop is-irrefutable-prop-type-Irrefutable-Prop
prop-Irrefutable-Prop : Prop l
prop-Irrefutable-Prop = type-Irrefutable-Prop , is-prop-type-Irrefutable-Prop
is-irrefutable-Irrefutable-Prop : is-irrefutable prop-Irrefutable-Prop
- is-irrefutable-Irrefutable-Prop = pr2 (pr2 P)
+ is-irrefutable-Irrefutable-Prop =
+ is-irrefutable-is-irrefutable-prop is-irrefutable-prop-type-Irrefutable-Prop
```
## Properties
+### Provable propositions are irrefutable
+
+```agda
+module _
+ {l : Level} (P : Prop l)
+ where
+
+ is-irrefutable-has-element : type-Prop P → is-irrefutable P
+ is-irrefutable-has-element = intro-double-negation
+
+is-irrefutable-unit : is-irrefutable unit-Prop
+is-irrefutable-unit = is-irrefutable-has-element unit-Prop star
+```
+
+### Contractible types are irrefutable propositions
+
+```agda
+is-irrefutable-prop-is-contr :
+ {l : Level} {P : UU l} → is-contr P → is-irrefutable-prop P
+is-irrefutable-prop-is-contr H =
+ ( is-prop-is-contr H , intro-double-negation (center H))
+```
+
### If it is irrefutable that a proposition is irrefutable, then the proposition is irrefutable
```agda
@@ -87,13 +151,13 @@ module _
{l : Level} (P : Prop l)
where
- is-irrefutable-is-irrefutable-is-irrefutable :
+ is-idempotent-is-irrefutable :
is-irrefutable (is-irrefutable-Prop P) → is-irrefutable P
- is-irrefutable-is-irrefutable-is-irrefutable =
+ is-idempotent-is-irrefutable =
double-negation-elim-neg (¬ (type-Prop P))
```
-### The decidability of a proposition is irrefutable
+### Decidability is irrefutable
```agda
is-irrefutable-is-decidable : {l : Level} {A : UU l} → ¬¬ (is-decidable A)
@@ -111,16 +175,67 @@ module _
make-Irrefutable-Prop (is-decidable-Prop P) is-irrefutable-is-decidable-Prop
```
-### Provable propositions are irrefutable
+### Double negation elimination is irrefutable
+
+```agda
+is-irrefutable-double-negation-elim :
+ {l : Level} {A : UU l} → ¬¬ (has-double-negation-elim A)
+is-irrefutable-double-negation-elim H =
+ H (λ f → ex-falso (f (λ a → H (λ _ → a))))
+```
+
+### Dependent sums of irrefutable propositions
```agda
module _
- {l : Level} (P : Prop l)
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
where
- is-irrefutable-is-inhabited : type-Prop P → is-irrefutable P
- is-irrefutable-is-inhabited = intro-double-negation
+ is-irrefutable-Σ :
+ ¬¬ A → ((x : A) → ¬¬ B x) → ¬¬ (Σ A B)
+ is-irrefutable-Σ nna nnb nab = nna (λ a → nnb a (λ b → nab (a , b)))
+
+ is-irrefutable-prop-Σ :
+ is-irrefutable-prop A → ((x : A) → is-irrefutable-prop (B x)) →
+ is-irrefutable-prop (Σ A B)
+ is-irrefutable-prop-Σ a b =
+ ( is-prop-Σ
+ ( is-prop-type-is-irrefutable-prop a)
+ ( is-prop-type-is-irrefutable-prop ∘ b) ,
+ is-irrefutable-Σ
+ ( is-irrefutable-is-irrefutable-prop a)
+ ( is-irrefutable-is-irrefutable-prop ∘ b))
+```
-is-irrefutable-unit : is-irrefutable unit-Prop
-is-irrefutable-unit = is-irrefutable-is-inhabited unit-Prop star
+### Products of irrefutable propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-irrefutable-product : ¬¬ A → ¬¬ B → ¬¬ (A × B)
+ is-irrefutable-product nna nnb = is-irrefutable-Σ nna (λ _ → nnb)
+
+ is-irrefutable-prop-product :
+ is-irrefutable-prop A → is-irrefutable-prop B → is-irrefutable-prop (A × B)
+ is-irrefutable-prop-product a b = is-irrefutable-prop-Σ a (λ _ → b)
```
+
+### Coproducts of irrefutable propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-irrefutable-coproduct-inl : ¬¬ A → ¬¬ (A + B)
+ is-irrefutable-coproduct-inl nna x = nna (x ∘ inl)
+
+ is-irrefutable-coproduct-inr : ¬¬ B → ¬¬ (A + B)
+ is-irrefutable-coproduct-inr nnb x = nnb (x ∘ inr)
+```
+
+## See also
+
+- [De Morgan's law](logic.de-morgans-law.md) is irrefutable.
diff --git a/src/foundation/large-locale-of-propositions.lagda.md b/src/foundation/large-locale-of-propositions.lagda.md
index 816cd18c3d..d5ef3ea50a 100644
--- a/src/foundation/large-locale-of-propositions.lagda.md
+++ b/src/foundation/large-locale-of-propositions.lagda.md
@@ -8,6 +8,7 @@ module foundation.large-locale-of-propositions where
```agda
open import foundation.conjunction
+open import foundation.empty-types
open import foundation.existential-quantification
open import foundation.logical-equivalences
open import foundation.propositional-extensionality
@@ -18,6 +19,7 @@ open import foundation.universe-levels
open import foundation-core.function-types
open import foundation-core.propositions
+open import order-theory.bottom-elements-large-posets
open import order-theory.large-frames
open import order-theory.large-locales
open import order-theory.large-meet-semilattices
@@ -87,6 +89,17 @@ is-top-element-top-has-top-element-Large-Poset
star
```
+### The smallest element in the large poset of propositions
+
+```agda
+has-bottom-element-Prop-Large-Locale :
+ has-bottom-element-Large-Poset Prop-Large-Poset
+bottom-has-bottom-element-Large-Poset
+ has-bottom-element-Prop-Large-Locale = empty-Prop
+is-bottom-element-bottom-has-bottom-element-Large-Poset
+ has-bottom-element-Prop-Large-Locale P = ex-falso
+```
+
### The large poset of propositions is a large meet-semilattice
```agda
@@ -98,6 +111,12 @@ has-meets-is-large-meet-semilattice-Large-Poset
has-top-element-is-large-meet-semilattice-Large-Poset
is-large-meet-semilattice-Prop-Large-Locale =
has-top-element-Prop-Large-Locale
+
+Prop-Large-Meet-Semilattice : Large-Meet-Semilattice lsuc (_⊔_)
+Prop-Large-Meet-Semilattice =
+ make-Large-Meet-Semilattice
+ ( Prop-Large-Poset)
+ ( is-large-meet-semilattice-Prop-Large-Locale)
```
### Suprema in the large poset of propositions
@@ -111,6 +130,10 @@ sup-has-least-upper-bound-family-of-elements-Large-Poset
is-least-upper-bound-sup-has-least-upper-bound-family-of-elements-Large-Poset
( is-large-suplattice-Prop-Large-Locale {I = I} P) R =
inv-iff (up-exists R)
+
+Prop-Large-Suplattice : Large-Suplattice lsuc (_⊔_) lzero
+Prop-Large-Suplattice =
+ make-Large-Suplattice Prop-Large-Poset is-large-suplattice-Prop-Large-Locale
```
### The large frame of propositions
diff --git a/src/foundation/large-locale-of-subtypes.lagda.md b/src/foundation/large-locale-of-subtypes.lagda.md
index 4bec76c43e..2661d3c83d 100644
--- a/src/foundation/large-locale-of-subtypes.lagda.md
+++ b/src/foundation/large-locale-of-subtypes.lagda.md
@@ -14,21 +14,25 @@ open import foundation.universe-levels
open import foundation-core.identity-types
open import foundation-core.sets
+open import order-theory.bottom-elements-large-posets
open import order-theory.greatest-lower-bounds-large-posets
open import order-theory.large-locales
open import order-theory.large-meet-semilattices
open import order-theory.large-posets
+open import order-theory.large-preorders
open import order-theory.large-suplattices
open import order-theory.least-upper-bounds-large-posets
open import order-theory.powers-of-large-locales
+open import order-theory.top-elements-large-posets
```
## Idea
-The **large locale of subtypes** of a type `A` is the
-[power locale](order-theory.powers-of-large-locales.md) `A → Prop-Large-Locale`.
+The {{#concept "large locale of subtypes" Agda=powerset-Large-Locale}} of a type
+`A` is the [power locale](order-theory.powers-of-large-locales.md)
+`A → Prop-Large-Locale`.
## Definition
@@ -41,6 +45,11 @@ module _
Large-Locale (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3) lzero
powerset-Large-Locale = power-Large-Locale A Prop-Large-Locale
+ large-preorder-powerset-Large-Locale :
+ Large-Preorder (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3)
+ large-preorder-powerset-Large-Locale =
+ large-preorder-Large-Locale powerset-Large-Locale
+
large-poset-powerset-Large-Locale :
Large-Poset (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3)
large-poset-powerset-Large-Locale =
@@ -168,4 +177,16 @@ module _
sup-powerset-Large-Locale (λ j → meet-powerset-Large-Locale x (y j))
distributive-meet-sup-powerset-Large-Locale =
distributive-meet-sup-Large-Locale (powerset-Large-Locale A)
+
+ has-top-element-powerset-Large-Locale :
+ has-top-element-Large-Poset (large-poset-powerset-Large-Locale A)
+ has-top-element-powerset-Large-Locale =
+ has-top-element-Large-Locale (powerset-Large-Locale A)
+
+ has-bottom-element-powerset-Large-Locale :
+ has-bottom-element-Large-Poset (large-poset-powerset-Large-Locale A)
+ has-bottom-element-powerset-Large-Locale =
+ has-bottom-element-Π-Large-Poset
+ ( λ _ → Prop-Large-Poset)
+ ( λ _ → has-bottom-element-Prop-Large-Locale)
```
diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md
index 3ee5bdb8c0..a633ea32d6 100644
--- a/src/foundation/law-of-excluded-middle.lagda.md
+++ b/src/foundation/law-of-excluded-middle.lagda.md
@@ -22,8 +22,9 @@ open import univalent-combinatorics.2-element-types
## Idea
-The {{#concept "law of excluded middle" Agda=LEM}} asserts that any
-[proposition](foundation-core.propositions.md) `P` is
+The
+{{#concept "law of excluded middle" WD="principle of excluded middle" WDID=Q468422 Agda=LEM}}
+asserts that any [proposition](foundation-core.propositions.md) `P` is
[decidable](foundation.decidable-types.md).
## Definition
diff --git a/src/foundation/lesser-limited-principle-of-omniscience.lagda.md b/src/foundation/lesser-limited-principle-of-omniscience.lagda.md
index a44fa6433b..242c9b27b6 100644
--- a/src/foundation/lesser-limited-principle-of-omniscience.lagda.md
+++ b/src/foundation/lesser-limited-principle-of-omniscience.lagda.md
@@ -10,35 +10,44 @@ module foundation.lesser-limited-principle-of-omniscience where
open import elementary-number-theory.natural-numbers
open import elementary-number-theory.parity-natural-numbers
+open import foundation.booleans
+open import foundation.dependent-pair-types
open import foundation.disjunction
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.fibers-of-maps
open import foundation-core.propositions
open import foundation-core.sets
-
-open import univalent-combinatorics.standard-finite-types
```
## Statement
-The **lesser limited principle of omniscience** asserts that for any sequence
-`f : ℕ → Fin 2` containing at most one `1`, either `f n = 0` for all even `n`
-or `f n = 0` for all odd `n`.
+The {{#concept "lesser limited principle of omniscience" Agda=LLPO}} (LLPO)
+asserts that for any [sequence](foundation.sequences.md) of
+[booleans](foundation.booleans.md) `f : ℕ → bool` such that `f n` is true for
+[at most one](foundation-core.propositions.md) `n`, then either `f n` is false
+for all even `n` or `f n` is false for all odd `n`.
```agda
+prop-LLPO : Prop lzero
+prop-LLPO =
+ ∀'
+ ( ℕ → bool)
+ ( λ f →
+ function-Prop
+ ( is-prop (Σ ℕ (λ n → is-true (f n))))
+ ( disjunction-Prop
+ ( ∀' ℕ (λ n → function-Prop (is-even-ℕ n) (is-false-Prop (f n))))
+ ( ∀' ℕ (λ n → function-Prop (is-odd-ℕ n) (is-false-Prop (f n))))))
+
LLPO : UU lzero
-LLPO =
- (f : ℕ → Fin 2) → is-prop (fiber f (one-Fin 1)) →
- type-disjunction-Prop
- ( Π-Prop ℕ
- ( λ n →
- function-Prop (is-even-ℕ n) (Id-Prop (Fin-Set 2) (f n) (zero-Fin 1))))
- ( Π-Prop ℕ
- ( λ n →
- function-Prop (is-odd-ℕ n) (Id-Prop (Fin-Set 2) (f n) (zero-Fin 1))))
+LLPO = type-Prop prop-LLPO
+
+is-prop-LLPO : is-prop LLPO
+is-prop-LLPO = is-prop-type-Prop prop-LLPO
```
## See also
@@ -46,3 +55,11 @@ LLPO =
- [The principle of omniscience](foundation.principle-of-omniscience.md)
- [The limited principle of omniscience](foundation.limited-principle-of-omniscience.md)
- [The weak limited principle of omniscience](foundation.weak-limited-principle-of-omniscience.md)
+- [Markov's principle](logic.markovs-principle.md)
+
+## External links
+
+- [Taboos.LLPO](https://martinescardo.github.io/TypeTopology/Taboos.LLPO.html)
+ at TypeTopology
+- [lesser limited principle of omniscience](https://ncatlab.org/nlab/show/lesser+limited+principle+of+omniscience)
+ at $n$Lab
diff --git a/src/foundation/limited-principle-of-omniscience.lagda.md b/src/foundation/limited-principle-of-omniscience.lagda.md
index 5b6d52ed2e..e66cea9ad0 100644
--- a/src/foundation/limited-principle-of-omniscience.lagda.md
+++ b/src/foundation/limited-principle-of-omniscience.lagda.md
@@ -9,8 +9,12 @@ module foundation.limited-principle-of-omniscience where
```agda
open import elementary-number-theory.natural-numbers
+open import foundation.booleans
+open import foundation.coproduct-types
+open import foundation.dependent-pair-types
open import foundation.disjunction
open import foundation.existential-quantification
+open import foundation.negation
open import foundation.universal-quantification
open import foundation.universe-levels
@@ -27,17 +31,36 @@ open import univalent-combinatorics.standard-finite-types
The
{{#concept "limited principle of omniscience" WDID=Q6549544 WD="limited principle of omniscience" Agda=LPO}}
-(LPO) asserts that for every [sequence](foundation.sequences.md) `f : ℕ → Fin 2`
+(LPO) asserts that for every [sequence](foundation.sequences.md) `f : ℕ → bool`
there either [exists](foundation.existential-quantification.md) an `n` such that
-`f n = 1` or for all `n` we have `f n = 0`.
+`f n` is true, [or](foundation.disjunction.md) `f n` is false for all `n`.
```agda
LPO : UU lzero
LPO =
- (f : ℕ → Fin 2) →
- type-disjunction-Prop
- ( ∃ ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (one-Fin 1)))
- ( ∀' ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))
+ (f : ℕ → bool) →
+ ( exists ℕ (λ n → is-true-Prop (f n))) +
+ ( for-all ℕ (λ n → is-false-Prop (f n)))
+```
+
+## Properties
+
+### The limited principle of omniscience is a proposition
+
+```agda
+is-prop-LPO : is-prop LPO
+is-prop-LPO =
+ is-prop-Π
+ ( λ f →
+ is-prop-coproduct
+ ( elim-exists
+ ( ¬' ∀' ℕ (λ n → is-false-Prop (f n)))
+ ( λ n t h → not-is-false-is-true (f n) t (h n)))
+ ( is-prop-exists ℕ (λ n → is-true-Prop (f n)))
+ ( is-prop-for-all-Prop ℕ (λ n → is-false-Prop (f n))))
+
+prop-LPO : Prop lzero
+prop-LPO = LPO , is-prop-LPO
```
## See also
@@ -45,3 +68,11 @@ LPO =
- [The principle of omniscience](foundation.principle-of-omniscience.md)
- [The lesser limited principle of omniscience](foundation.lesser-limited-principle-of-omniscience.md)
- [The weak limited principle of omniscience](foundation.weak-limited-principle-of-omniscience.md)
+- [Markov's principle](logic.markovs-principle.md)
+
+## External links
+
+- [Taboos.LPO](https://martinescardo.github.io/TypeTopology/Taboos.LPO.html) at
+ TypeTopology
+- [limited principle of omniscience](https://ncatlab.org/nlab/show/limited+principle+of+omniscience)
+ at $n$Lab
diff --git a/src/foundation/locale-of-propositions.lagda.md b/src/foundation/locale-of-propositions.lagda.md
new file mode 100644
index 0000000000..daef5de811
--- /dev/null
+++ b/src/foundation/locale-of-propositions.lagda.md
@@ -0,0 +1,132 @@
+# The locale of propositions
+
+```agda
+module foundation.locale-of-propositions where
+```
+
+Imports
+
+```agda
+open import foundation.conjunction
+open import foundation.dependent-pair-types
+open import foundation.existential-quantification
+open import foundation.large-locale-of-propositions
+open import foundation.logical-equivalences
+open import foundation.unit-type
+open import foundation.universe-levels
+
+open import foundation-core.function-types
+
+open import order-theory.frames
+open import order-theory.greatest-lower-bounds-posets
+open import order-theory.large-posets
+open import order-theory.large-preorders
+open import order-theory.meet-semilattices
+open import order-theory.meet-suplattices
+open import order-theory.posets
+open import order-theory.preorders
+open import order-theory.suplattices
+open import order-theory.top-elements-posets
+```
+
+
+
+## Idea
+
+The [locale](order-theory.locales.md) of
+[propositions](foundation-core.propositions.md) consists of all the propositions
+of any [universe level](foundation.universe-levels.md) and is ordered by the
+implications between them. [Conjunction](foundation.conjunction.md) gives this
+[poset](order-theory.posets.md) the structure of a
+[meet-semilattice](order-theory.meet-semilattices.md), and
+[existential quantification](foundation.existential-quantification.md) gives it
+the structure of a [suplattice](order-theory.suplattices.md).
+
+## Definitions
+
+### The preorder of propositions
+
+```agda
+Prop-Preorder : (l : Level) → Preorder (lsuc l) l
+Prop-Preorder = preorder-Large-Preorder Prop-Large-Preorder
+```
+
+### The poset of propositions
+
+```agda
+Prop-Poset : (l : Level) → Poset (lsuc l) l
+Prop-Poset = poset-Large-Poset Prop-Large-Poset
+```
+
+### The largest element in the poset of propositions
+
+```agda
+has-top-element-Prop-Locale :
+ {l : Level} → has-top-element-Poset (Prop-Poset l)
+has-top-element-Prop-Locale {l} = (raise-unit-Prop l , λ _ _ → raise-star)
+```
+
+### Meets in the poset of propositions
+
+```agda
+is-meet-semilattice-Prop-Locale :
+ {l : Level} → is-meet-semilattice-Poset (Prop-Poset l)
+is-meet-semilattice-Prop-Locale P Q =
+ ( P ∧ Q , is-greatest-binary-lower-bound-conjunction-Prop P Q)
+
+Prop-Order-Theoretic-Meet-Semilattice :
+ (l : Level) → Order-Theoretic-Meet-Semilattice (lsuc l) l
+Prop-Order-Theoretic-Meet-Semilattice l =
+ Prop-Poset l , is-meet-semilattice-Prop-Locale
+
+Prop-Meet-Semilattice :
+ (l : Level) → Meet-Semilattice (lsuc l)
+Prop-Meet-Semilattice l =
+ meet-semilattice-Order-Theoretic-Meet-Semilattice
+ ( Prop-Order-Theoretic-Meet-Semilattice l)
+```
+
+### Suprema in the poset of propositions
+
+```agda
+is-suplattice-lzero-Prop-Locale :
+ {l : Level} → is-suplattice-Poset lzero (Prop-Poset l)
+is-suplattice-lzero-Prop-Locale I P = (∃ I P , inv-iff ∘ up-exists)
+
+is-suplattice-Prop-Locale :
+ {l : Level} → is-suplattice-Poset l (Prop-Poset l)
+is-suplattice-Prop-Locale I P = (∃ I P , inv-iff ∘ up-exists)
+
+Prop-Suplattice :
+ (l : Level) → Suplattice (lsuc l) l l
+Prop-Suplattice l = Prop-Poset l , is-suplattice-Prop-Locale
+```
+
+```text
+is-meet-suplattice-Prop-Locale :
+ {l : Level} → is-meet-suplattice-Meet-Semilattice l (Prop-Meet-Semilattice l)
+is-meet-suplattice-Prop-Locale I P = ?
+
+Prop-Meet-Suplattice :
+ (l : Level) → Meet-Suplattice (lsuc l) l
+Prop-Meet-Suplattice l =
+ Prop-Meet-Semilattice l , is-meet-suplattice-Prop-Locale
+```
+
+### The frame of propositions
+
+```text
+Prop-Frame : (l : Level) → Frame (lsuc l) l
+Prop-Frame l = Prop-Meet-Suplattice l , ?
+```
+
+### The locale of propositions
+
+```text
+Prop-Locale : Locale lsuc (_⊔_) lzero
+Prop-Locale = Prop-Frame
+```
+
+## See also
+
+- [Propositional resizing](foundation.propositional-resizing.md)
diff --git a/src/foundation/logical-equivalences.lagda.md b/src/foundation/logical-equivalences.lagda.md
index e6208ce084..d806e1ca3d 100644
--- a/src/foundation/logical-equivalences.lagda.md
+++ b/src/foundation/logical-equivalences.lagda.md
@@ -220,6 +220,10 @@ iff-equiv : {l1 l2 : Level} {A : UU l1} {B : UU l2} → (A ≃ B) → (A ↔ B)
pr1 (iff-equiv e) = map-equiv e
pr2 (iff-equiv e) = map-inv-equiv e
+iff-equiv' : {l1 l2 : Level} {A : UU l1} {B : UU l2} → (A ≃ B) → (B ↔ A)
+pr1 (iff-equiv' e) = map-inv-equiv e
+pr2 (iff-equiv' e) = map-equiv e
+
is-injective-iff-equiv :
{l1 l2 : Level} {A : UU l1} {B : UU l2} → is-injective (iff-equiv {A = A} {B})
is-injective-iff-equiv p = eq-htpy-equiv (pr1 (htpy-eq-iff p))
diff --git a/src/foundation/negation.lagda.md b/src/foundation/negation.lagda.md
index 3a7b5044e1..273c37d133 100644
--- a/src/foundation/negation.lagda.md
+++ b/src/foundation/negation.lagda.md
@@ -15,6 +15,7 @@ open import foundation.universe-levels
open import foundation-core.empty-types
open import foundation-core.equivalences
+open import foundation-core.identity-types
open import foundation-core.propositions
```
@@ -34,7 +35,7 @@ using propositions as types. Thus, the negation of a type `A` is the type
```agda
is-prop-neg : {l : Level} {A : UU l} → is-prop (¬ A)
-is-prop-neg {A = A} = is-prop-function-type is-prop-empty
+is-prop-neg = is-prop-function-type is-prop-empty
neg-type-Prop : {l1 : Level} → UU l1 → Prop l1
neg-type-Prop A = ¬ A , is-prop-neg
@@ -49,6 +50,9 @@ infix 25 ¬'_
¬'_ : {l1 : Level} → Prop l1 → Prop l1
¬'_ = neg-Prop
+
+eq-neg : {l : Level} {A : UU l} {p q : ¬ A} → p = q
+eq-neg = eq-is-prop is-prop-neg
```
### Reductio ad absurdum
@@ -58,17 +62,30 @@ reductio-ad-absurdum : {l1 l2 : Level} {P : UU l1} {Q : UU l2} → P → ¬ P
reductio-ad-absurdum p np = ex-falso (np p)
```
+### Logically equivalent types have logically equivalent negations
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ where
+
+ iff-neg : X ↔ Y → ¬ X ↔ ¬ Y
+ iff-neg e = map-neg (backward-implication e) , map-neg (forward-implication e)
+
+ equiv-iff-neg : X ↔ Y → ¬ X ≃ ¬ Y
+ equiv-iff-neg e =
+ equiv-iff' (neg-type-Prop X) (neg-type-Prop Y) (iff-neg e)
+```
+
### Equivalent types have equivalent negations
```agda
-equiv-neg :
- {l1 l2 : Level} {X : UU l1} {Y : UU l2} →
- (X ≃ Y) → (¬ X ≃ ¬ Y)
-equiv-neg {l1} {l2} {X} {Y} e =
- equiv-iff'
- ( neg-type-Prop X)
- ( neg-type-Prop Y)
- ( pair (map-neg (map-inv-equiv e)) (map-neg (map-equiv e)))
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ where
+
+ equiv-neg : X ≃ Y → ¬ X ≃ ¬ Y
+ equiv-neg e = equiv-iff-neg (iff-equiv e)
```
### Negation has no fixed points
diff --git a/src/foundation/path-split-type-families.lagda.md b/src/foundation/path-split-type-families.lagda.md
new file mode 100644
index 0000000000..690d6d1a74
--- /dev/null
+++ b/src/foundation/path-split-type-families.lagda.md
@@ -0,0 +1,146 @@
+# Path-split type families
+
+```agda
+module foundation.path-split-type-families where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.dependent-pair-types
+open import foundation.embeddings
+open import foundation.transport-along-identifications
+open import foundation.universe-levels
+
+open import foundation-core.contractible-types
+open import foundation-core.dependent-identifications
+open import foundation-core.equality-dependent-pair-types
+open import foundation-core.identity-types
+open import foundation-core.injective-maps
+open import foundation-core.propositions
+open import foundation-core.sections
+open import foundation-core.subtypes
+```
+
+
+
+## Idea
+
+We say a type family `P : A → 𝒰` is
+{{#concept "path-split" Disambiguation="type family" Agda=is-path-split-family}}
+if, for every [identification](foundation-core.identity-types.md) `p : x = y`
+in `A`, and every pair of elements `u : P x` and `v : P y`, there is a
+[dependent identification](foundation-core.dependent-identifications.md) of `u`
+and `v` _over_ p. This condition is
+[equivalent](foundation.logical-equivalences.md) to asking that `P` is a family
+of [propositions](foundation-core.propositions.md).
+
+This condition is a direct rephrasing of stating that the
+[action on identifications](foundation.action-on-identifications-functions.md)
+of the first projection map `Σ A P → A` has a
+[section](foundation-core.sections.md), and in this way is closely related to
+the concept of [path-split maps](foundation-core.path-split-maps.md).
+
+## Definitions
+
+### Path-split type families
+
+```agda
+is-path-split-family :
+ {l1 l2 : Level} {A : UU l1} → (A → UU l2) → UU (l1 ⊔ l2)
+is-path-split-family {A = A} P =
+ {x y : A} (p : x = y) (s : P x) (t : P y) → dependent-identification P p s t
+```
+
+## Properties
+
+### The first projection map of a path-split type family is an embedding
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {P : A → UU l2}
+ (H : is-path-split-family P)
+ where
+
+ is-injective-pr1-is-path-split-family : is-injective (pr1 {B = P})
+ is-injective-pr1-is-path-split-family {x} {y} p =
+ eq-pair-Σ p (H p (pr2 x) (pr2 y))
+
+ is-section-is-injective-pr1-is-path-split-family :
+ {x y : Σ A P} →
+ is-section (ap pr1) (is-injective-pr1-is-path-split-family {x} {y})
+ is-section-is-injective-pr1-is-path-split-family {x , u} {.x , v} refl =
+ ap-pr1-eq-pair-eq-fiber (H refl u v)
+
+ section-ap-pr1-is-path-split-family :
+ (x y : Σ A P) → section (ap pr1 {x} {y})
+ section-ap-pr1-is-path-split-family x y =
+ is-injective-pr1-is-path-split-family ,
+ is-section-is-injective-pr1-is-path-split-family
+
+ is-emb-pr1-is-path-split-family : is-emb (pr1 {B = P})
+ is-emb-pr1-is-path-split-family =
+ is-emb-section-ap pr1 section-ap-pr1-is-path-split-family
+```
+
+### The fibers of a path-split type family are propositions
+
+We give two proofs, one is direct and the other uses the previous result.
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {P : A → UU l2}
+ (H : is-path-split-family P)
+ where
+
+ all-elements-equal-is-path-split-family : {x : A} (u v : P x) → u = v
+ all-elements-equal-is-path-split-family u v = H refl u v
+
+ is-subtype-is-path-split-family : is-subtype P
+ is-subtype-is-path-split-family x =
+ is-prop-all-elements-equal all-elements-equal-is-path-split-family
+
+ is-subtype-is-path-split-family' : is-subtype P
+ is-subtype-is-path-split-family' =
+ is-subtype-is-emb-pr1 (is-emb-pr1-is-path-split-family H)
+```
+
+### Path-splittings of type families are unique
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {P : A → UU l2}
+ where
+
+ abstract
+ is-proof-irrelevant-is-path-split-family :
+ is-proof-irrelevant (is-path-split-family P)
+ is-proof-irrelevant-is-path-split-family H =
+ is-contr-implicit-Π
+ ( λ x →
+ is-contr-implicit-Π
+ ( λ y →
+ is-contr-Π
+ ( λ p →
+ is-contr-Π
+ ( λ s →
+ is-contr-Π
+ ( is-subtype-is-path-split-family H y (tr P p s))))))
+
+ is-prop-is-path-split-family : is-prop (is-path-split-family P)
+ is-prop-is-path-split-family =
+ is-prop-is-proof-irrelevant is-proof-irrelevant-is-path-split-family
+```
+
+### Subtypes are path-split
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {P : A → UU l2}
+ where
+
+ is-path-split-family-is-subtype :
+ is-subtype P → is-path-split-family P
+ is-path-split-family-is-subtype H {x} refl s t = eq-is-prop (H x)
+```
diff --git a/src/foundation/postcomposition-functions.lagda.md b/src/foundation/postcomposition-functions.lagda.md
index 3f4e9643e9..2799d76ae4 100644
--- a/src/foundation/postcomposition-functions.lagda.md
+++ b/src/foundation/postcomposition-functions.lagda.md
@@ -109,6 +109,23 @@ module _
compute-fiber-postcomp h = compute-coherence-triangle-fiber-postcomp (f ∘ h)
```
+### Computing the fiber of the postcomposition function at the identity function
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ (f : X → Y)
+ where
+
+ compute-fiber-id-postcomp :
+ ((x : Y) → fiber f x) ≃ fiber (postcomp Y f) id
+ compute-fiber-id-postcomp = compute-Π-fiber-postcomp Y f id
+
+ inv-compute-fiber-id-postcomp :
+ fiber (postcomp Y f) id ≃ ((x : Y) → fiber f x)
+ inv-compute-fiber-id-postcomp = inv-compute-Π-fiber-postcomp Y f id
+```
+
### Postcomposition and equivalences
#### A map `f` is an equivalence if and only if postcomposing by `f` is an equivalence
diff --git a/src/foundation/powersets.lagda.md b/src/foundation/powersets.lagda.md
index 33631db4f1..5610d81426 100644
--- a/src/foundation/powersets.lagda.md
+++ b/src/foundation/powersets.lagda.md
@@ -7,15 +7,34 @@ module foundation.powersets where
Imports
```agda
+open import foundation.dependent-pair-types
+open import foundation.embeddings
+open import foundation.empty-types
+open import foundation.large-locale-of-propositions
open import foundation.subtypes
+open import foundation.unit-type
open import foundation.universe-levels
open import foundation-core.sets
+open import order-theory.bottom-elements-large-posets
+open import order-theory.bottom-elements-posets
+open import order-theory.dependent-products-large-meet-semilattices
+open import order-theory.dependent-products-large-posets
+open import order-theory.dependent-products-large-preorders
+open import order-theory.dependent-products-large-suplattices
+open import order-theory.large-meet-semilattices
open import order-theory.large-posets
open import order-theory.large-preorders
+open import order-theory.large-suplattices
+open import order-theory.meet-semilattices
+open import order-theory.order-preserving-maps-large-posets
+open import order-theory.order-preserving-maps-large-preorders
open import order-theory.posets
open import order-theory.preorders
+open import order-theory.suplattices
+open import order-theory.top-elements-large-posets
+open import order-theory.top-elements-posets
```
@@ -25,8 +44,7 @@ open import order-theory.preorders
The
{{#concept "powerset" Disambiguation="of a type" Agda=powerset WD="power set" WDID=Q205170}}
of a type is the [set](foundation-core.sets.md) of all
-[subtypes](foundation-core.subtypes.md) with respect to a given
-[universe level](foundation.universe-levels.md).
+[subtypes](foundation-core.subtypes.md).
## Definition
@@ -61,10 +79,14 @@ module _
powerset-Large-Preorder :
Large-Preorder (λ l → l1 ⊔ lsuc l) (λ l2 l3 → l1 ⊔ l2 ⊔ l3)
- type-Large-Preorder powerset-Large-Preorder l = subtype l A
- leq-prop-Large-Preorder powerset-Large-Preorder = leq-prop-subtype
- refl-leq-Large-Preorder powerset-Large-Preorder = refl-leq-subtype
- transitive-leq-Large-Preorder powerset-Large-Preorder = transitive-leq-subtype
+ powerset-Large-Preorder = Π-Large-Preorder {I = A} (λ _ → Prop-Large-Preorder)
+
+module _
+ {l1 : Level} (l2 : Level) (A : UU l1)
+ where
+
+ powerset-Preorder : Preorder (l1 ⊔ lsuc l2) (l1 ⊔ l2)
+ powerset-Preorder = preorder-Large-Preorder (powerset-Large-Preorder A) l2
```
### The powerset large poset
@@ -76,33 +98,91 @@ module _
powerset-Large-Poset :
Large-Poset (λ l → l1 ⊔ lsuc l) (λ l2 l3 → l1 ⊔ l2 ⊔ l3)
- large-preorder-Large-Poset powerset-Large-Poset = powerset-Large-Preorder A
- antisymmetric-leq-Large-Poset powerset-Large-Poset P Q =
- antisymmetric-leq-subtype P Q
+ powerset-Large-Poset = Π-Large-Poset {I = A} (λ _ → Prop-Large-Poset)
+
+module _
+ {l1 : Level} (l2 : Level) (A : UU l1)
+ where
+
+ powerset-Poset : Poset (l1 ⊔ lsuc l2) (l1 ⊔ l2)
+ powerset-Poset = poset-Large-Poset (powerset-Large-Poset A) l2
```
-### The powerset preorder at a universe level
+### The powerset poset has a top element
```agda
module _
{l1 : Level} (A : UU l1)
where
- powerset-Preorder : (l2 : Level) → Preorder (l1 ⊔ lsuc l2) (l1 ⊔ l2)
- powerset-Preorder = preorder-Large-Preorder (powerset-Large-Preorder A)
+ has-top-element-powerset-Large-Poset :
+ has-top-element-Large-Poset (powerset-Large-Poset A)
+ has-top-element-powerset-Large-Poset =
+ has-top-element-Π-Large-Poset
+ ( λ _ → Prop-Large-Poset)
+ ( λ _ → has-top-element-Prop-Large-Locale)
+
+ has-top-element-powerset-Poset :
+ {l2 : Level} → has-top-element-Poset (powerset-Poset l2 A)
+ has-top-element-powerset-Poset {l2} =
+ (λ _ → raise-unit-Prop l2) , (λ _ _ _ → raise-star)
```
-### The powerset poset at a universe level
+### The powerset poset has a bottom element
```agda
module _
{l1 : Level} (A : UU l1)
where
- powerset-Poset : (l2 : Level) → Poset (l1 ⊔ lsuc l2) (l1 ⊔ l2)
- powerset-Poset = poset-Large-Poset (powerset-Large-Poset A)
+ has-bottom-element-powerset-Large-Poset :
+ has-bottom-element-Large-Poset (powerset-Large-Poset A)
+ has-bottom-element-powerset-Large-Poset =
+ has-bottom-element-Π-Large-Poset
+ ( λ _ → Prop-Large-Poset)
+ ( λ _ → has-bottom-element-Prop-Large-Locale)
+
+ has-bottom-element-powerset-Poset :
+ {l2 : Level} → has-bottom-element-Poset (powerset-Poset l2 A)
+ has-bottom-element-powerset-Poset {l2} =
+ (λ _ → raise-empty-Prop l2) , (λ _ _ → raise-ex-falso l2)
```
+### The powerset meet semilattice
+
+```agda
+module _
+ {l1 : Level} (A : UU l1)
+ where
+
+ powerset-Large-Meet-Semilattice :
+ Large-Meet-Semilattice (λ l2 → l1 ⊔ lsuc l2) (λ l2 l3 → l1 ⊔ l2 ⊔ l3)
+ powerset-Large-Meet-Semilattice =
+ Π-Large-Meet-Semilattice {I = A} (λ _ → Prop-Large-Meet-Semilattice)
+```
+
+### The powerset suplattice
+
+```agda
+module _
+ {l1 : Level} (A : UU l1)
+ where
+
+ powerset-Large-Suplattice :
+ Large-Suplattice (λ l2 → lsuc l2 ⊔ l1) (λ l2 l3 → l2 ⊔ l3 ⊔ l1) lzero
+ powerset-Large-Suplattice =
+ Π-Large-Suplattice {I = A} (λ _ → Prop-Large-Suplattice)
+
+ powerset-Suplattice :
+ (l2 l3 : Level) → Suplattice (l1 ⊔ lsuc l2 ⊔ lsuc l3) (l1 ⊔ l2 ⊔ l3) l2
+ powerset-Suplattice = suplattice-Large-Suplattice powerset-Large-Suplattice
+```
+
+## See also
+
+- [the large locale of subtypes](foundation.large-locale-of-subtypes.md)
+- [images of subtypes](foundation.images-subtypes.md)
+
## External links
- [power object](https://ncatlab.org/nlab/show/power+object) at $n$Lab
diff --git a/src/foundation/preimages-of-subtypes.lagda.md b/src/foundation/preimages-of-subtypes.lagda.md
deleted file mode 100644
index ac1a192b66..0000000000
--- a/src/foundation/preimages-of-subtypes.lagda.md
+++ /dev/null
@@ -1,29 +0,0 @@
-# Preimages of subtypes
-
-```agda
-module foundation.preimages-of-subtypes where
-```
-
-Imports
-
-```agda
-open import foundation.universe-levels
-
-open import foundation-core.subtypes
-```
-
-
-
-## Idea
-
-The preimage of a subtype `S ⊆ B` under a map `f : A → B` is the subtype of `A`
-consisting of elements `a` such that `f a ∈ S`.
-
-## Definition
-
-```agda
-preimage-subtype :
- {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (f : A → B) →
- subtype l3 B → subtype l3 A
-preimage-subtype f S a = S (f a)
-```
diff --git a/src/foundation/principle-of-omniscience.lagda.md b/src/foundation/principle-of-omniscience.lagda.md
index e3a35258cb..cf03f741b2 100644
--- a/src/foundation/principle-of-omniscience.lagda.md
+++ b/src/foundation/principle-of-omniscience.lagda.md
@@ -19,8 +19,9 @@ open import foundation-core.propositions
## Idea
-A type `X` is said to satisfy the **principle of omniscience** if every
-[decidable subtype](foundation.decidable-subtypes.md) of `X` is either
+A type `X` is said to satisfy the
+{{#concept "principle of omniscience" Disambiguation="type" Agda=is-omniscient}}
+if every [decidable subtype](foundation.decidable-subtypes.md) of `X` is either
[inhabited](foundation.inhabited-types.md) or
[empty](foundation-core.empty-types.md).
@@ -42,3 +43,4 @@ is-omniscient X = type-Prop (is-omniscient-Prop X)
- [The limited principle of omniscience](foundation.limited-principle-of-omniscience.md)
- [The lesser limited principle of omniscience](foundation.lesser-limited-principle-of-omniscience.md)
- [The weak limited principle of omniscience](foundation.weak-limited-principle-of-omniscience.md)
+- [Markov's principle](logic.markovs-principle.md)
diff --git a/src/foundation/propositional-resizing.lagda.md b/src/foundation/propositional-resizing.lagda.md
index 46d48494aa..4abe732a36 100644
--- a/src/foundation/propositional-resizing.lagda.md
+++ b/src/foundation/propositional-resizing.lagda.md
@@ -10,7 +10,9 @@ module foundation.propositional-resizing where
open import foundation.dependent-pair-types
open import foundation.universe-levels
+open import foundation-core.equivalences
open import foundation-core.propositions
+open import foundation-core.small-types
open import foundation-core.subtypes
```
@@ -19,20 +21,172 @@ open import foundation-core.subtypes
## Idea
We say that a universe `𝒱` satisfies `𝒰`-small
-{{#concept "propositional resizing"}} if there is a type `Ω` in `𝒰`
-[equipped](foundation.structure.md) with a
+{{#concept "propositional resizing" Agda=propositional-resizing-Level}} if there
+is a type `Ω` in `𝒰` [equipped](foundation.structure.md) with a
[subtype](foundation-core.subtypes.md) `Q` such that for each proposition `P` in
`𝒱` there is an element `u : Ω` such that `Q u ≃ P`. Such a type `Ω` is called
an `𝒰`-small {{#concept "classifier" Disambiguation="of small subobjects"}} of
`𝒱`-small subobjects.
-## Definition
+## Definitions
+
+### The predicate on a type of being a subtype classifier of a given universe level
+
+```agda
+is-subtype-classifier :
+ {l1 l2 : Level} (l3 : Level) → Σ (UU l1) (subtype l2) → UU (l1 ⊔ l2 ⊔ lsuc l3)
+is-subtype-classifier l3 Ω =
+ (P : Prop l3) → Σ (pr1 Ω) (λ u → type-equiv-Prop P (pr2 Ω u))
+
+module _
+ {l1 l2 l3 : Level}
+ (Ω : Σ (UU l1) (subtype l2)) (H : is-subtype-classifier l3 Ω)
+ (P : Prop l3)
+ where
+
+ object-prop-is-subtype-classifier : pr1 Ω
+ object-prop-is-subtype-classifier = pr1 (H P)
+
+ prop-is-small-prop-is-subtype-classifier : Prop l2
+ prop-is-small-prop-is-subtype-classifier =
+ pr2 Ω object-prop-is-subtype-classifier
+
+ type-is-small-prop-is-subtype-classifier : UU l2
+ type-is-small-prop-is-subtype-classifier =
+ type-Prop prop-is-small-prop-is-subtype-classifier
+
+ equiv-is-small-prop-is-subtype-classifier :
+ type-Prop P ≃ type-is-small-prop-is-subtype-classifier
+ equiv-is-small-prop-is-subtype-classifier = pr2 (H P)
+
+ is-small-prop-is-subtype-classifier : is-small l2 (type-Prop P)
+ is-small-prop-is-subtype-classifier =
+ ( type-is-small-prop-is-subtype-classifier ,
+ equiv-is-small-prop-is-subtype-classifier)
+```
+
+### Propositional resizing between specified universes
+
+```agda
+propositional-resizing-Level' :
+ (l1 l2 l3 : Level) → UU (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3)
+propositional-resizing-Level' l1 l2 l3 =
+ Σ (Σ (UU l1) (subtype l2)) (is-subtype-classifier l3)
+
+propositional-resizing-Level :
+ (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
+propositional-resizing-Level l1 l2 =
+ propositional-resizing-Level' l1 l1 l2
+```
+
+```agda
+module _
+ {l1 l2 l3 : Level} (ρ : propositional-resizing-Level' l1 l2 l3)
+ where
+
+ subtype-classifier-propositional-resizing-Level : Σ (UU l1) (subtype l2)
+ subtype-classifier-propositional-resizing-Level = pr1 ρ
+
+ type-subtype-classifier-propositional-resizing-Level : UU l1
+ type-subtype-classifier-propositional-resizing-Level =
+ pr1 subtype-classifier-propositional-resizing-Level
+
+ is-subtype-classifier-subtype-classifier-propositional-resizing-Level :
+ is-subtype-classifier l3 subtype-classifier-propositional-resizing-Level
+ is-subtype-classifier-subtype-classifier-propositional-resizing-Level = pr2 ρ
+
+module _
+ {l1 l2 l3 : Level} (ρ : propositional-resizing-Level' l1 l2 l3) (P : Prop l3)
+ where
+
+ object-prop-propositional-resizing-Level :
+ type-subtype-classifier-propositional-resizing-Level ρ
+ object-prop-propositional-resizing-Level =
+ object-prop-is-subtype-classifier
+ ( subtype-classifier-propositional-resizing-Level ρ)
+ ( is-subtype-classifier-subtype-classifier-propositional-resizing-Level ρ)
+ ( P)
+
+ prop-is-small-prop-propositional-resizing-Level : Prop l2
+ prop-is-small-prop-propositional-resizing-Level =
+ prop-is-small-prop-is-subtype-classifier
+ ( subtype-classifier-propositional-resizing-Level ρ)
+ ( is-subtype-classifier-subtype-classifier-propositional-resizing-Level ρ)
+ ( P)
+
+ type-is-small-prop-propositional-resizing-Level : UU l2
+ type-is-small-prop-propositional-resizing-Level =
+ type-is-small-prop-is-subtype-classifier
+ ( subtype-classifier-propositional-resizing-Level ρ)
+ ( is-subtype-classifier-subtype-classifier-propositional-resizing-Level ρ)
+ ( P)
+
+ equiv-is-small-prop-propositional-resizing-Level :
+ type-Prop P ≃ type-is-small-prop-propositional-resizing-Level
+ equiv-is-small-prop-propositional-resizing-Level =
+ equiv-is-small-prop-is-subtype-classifier
+ ( subtype-classifier-propositional-resizing-Level ρ)
+ ( is-subtype-classifier-subtype-classifier-propositional-resizing-Level ρ)
+ ( P)
+
+ is-small-prop-propositional-resizing-Level : is-small l2 (type-Prop P)
+ is-small-prop-propositional-resizing-Level =
+ is-small-prop-is-subtype-classifier
+ ( subtype-classifier-propositional-resizing-Level ρ)
+ ( is-subtype-classifier-subtype-classifier-propositional-resizing-Level ρ)
+ ( P)
+```
+
+### The propositional resizing axiom
+
+```agda
+propositional-resizing : UUω
+propositional-resizing = (l1 l2 : Level) → propositional-resizing-Level l1 l2
+```
```agda
-propositional-resizing : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
-propositional-resizing l1 l2 =
- Σ ( Σ (UU l1) (subtype l1))
- ( λ Ω → (P : Prop l2) → Σ (pr1 Ω) (λ u → type-equiv-Prop (pr2 Ω u) P))
+module _
+ (ρ : propositional-resizing) (l1 l2 : Level)
+ where
+
+ subtype-classifier-propositional-resizing : Σ (UU l1) (subtype l1)
+ subtype-classifier-propositional-resizing = pr1 (ρ l1 l2)
+
+ type-subtype-classifier-propositional-resizing : UU l1
+ type-subtype-classifier-propositional-resizing =
+ type-subtype-classifier-propositional-resizing-Level (ρ l1 l2)
+
+ is-subtype-classifier-subtype-classifier-propositional-resizing :
+ is-subtype-classifier l2 subtype-classifier-propositional-resizing
+ is-subtype-classifier-subtype-classifier-propositional-resizing =
+ is-subtype-classifier-subtype-classifier-propositional-resizing-Level
+ ( ρ l1 l2)
+
+module _
+ (ρ : propositional-resizing) (l1 : Level) {l2 : Level} (P : Prop l2)
+ where
+
+ object-prop-propositional-resizing :
+ type-subtype-classifier-propositional-resizing ρ l1 l2
+ object-prop-propositional-resizing =
+ object-prop-propositional-resizing-Level (ρ l1 l2) P
+
+ prop-is-small-prop-propositional-resizing : Prop l1
+ prop-is-small-prop-propositional-resizing =
+ prop-is-small-prop-propositional-resizing-Level (ρ l1 l2) P
+
+ type-is-small-prop-propositional-resizing : UU l1
+ type-is-small-prop-propositional-resizing =
+ type-is-small-prop-propositional-resizing-Level (ρ l1 l2) P
+
+ equiv-is-small-prop-propositional-resizing :
+ type-Prop P ≃ type-is-small-prop-propositional-resizing
+ equiv-is-small-prop-propositional-resizing =
+ equiv-is-small-prop-propositional-resizing-Level (ρ l1 l2) P
+
+ is-small-prop-propositional-resizing : is-small l1 (type-Prop P)
+ is-small-prop-propositional-resizing =
+ is-small-prop-propositional-resizing-Level (ρ l1 l2) P
```
## See also
diff --git a/src/foundation/propositions.lagda.md b/src/foundation/propositions.lagda.md
index cc36431139..e4bf1eb773 100644
--- a/src/foundation/propositions.lagda.md
+++ b/src/foundation/propositions.lagda.md
@@ -11,12 +11,15 @@ open import foundation-core.propositions public
```agda
open import foundation.contractible-types
open import foundation.dependent-pair-types
+open import foundation.fibers-of-maps
open import foundation.logical-equivalences
open import foundation.retracts-of-types
+open import foundation.unit-type
open import foundation.universe-levels
open import foundation-core.embeddings
open import foundation-core.equivalences
+open import foundation-core.propositional-maps
open import foundation-core.truncated-types
open import foundation-core.truncation-levels
```
@@ -64,6 +67,24 @@ abstract
is-prop-emb = is-trunc-emb neg-two-𝕋
```
+### A type is a proposition if and only if it embeds into the unit type
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ abstract
+ is-prop-is-emb-terminal-map : is-emb (terminal-map A) → is-prop A
+ is-prop-is-emb-terminal-map H =
+ is-prop-is-emb (terminal-map A) H is-prop-unit
+
+ abstract
+ is-emb-terminal-map-is-prop : is-prop A → is-emb (terminal-map A)
+ is-emb-terminal-map-is-prop H =
+ is-emb-is-prop-map (λ y → is-prop-equiv (equiv-fiber-terminal-map y) H)
+```
+
### Two equivalent types are equivalently propositions
```agda
diff --git a/src/foundation/retracts-of-types.lagda.md b/src/foundation/retracts-of-types.lagda.md
index ce0fa62ff3..56c7476a61 100644
--- a/src/foundation/retracts-of-types.lagda.md
+++ b/src/foundation/retracts-of-types.lagda.md
@@ -15,12 +15,14 @@ open import foundation.fundamental-theorem-of-identity-types
open import foundation.homotopies
open import foundation.homotopy-algebra
open import foundation.homotopy-induction
+open import foundation.logical-equivalences
open import foundation.structure-identity-principle
open import foundation.univalence
open import foundation.universe-levels
open import foundation.whiskering-homotopies-composition
open import foundation-core.contractible-types
+open import foundation-core.function-types
open import foundation-core.functoriality-dependent-pair-types
open import foundation-core.identity-types
open import foundation-core.torsorial-type-families
@@ -166,6 +168,18 @@ module _
eq-equiv-retracts R S = map-inv-is-equiv (is-equiv-equiv-eq-retracts R S)
```
+### The underlying logical equivalence associated to a retract
+
+```agda
+iff-retract :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} → A retract-of B → A ↔ B
+iff-retract R = inclusion-retract R , map-retraction-retract R
+
+iff-retract' :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} → A retract-of B → B ↔ A
+iff-retract' = inv-iff ∘ iff-retract
+```
+
## See also
- [Retracts of maps](foundation.retracts-of-maps.md)
diff --git a/src/foundation/subsingleton-induction.lagda.md b/src/foundation/subsingleton-induction.lagda.md
index 7007417be8..1497a0c6d6 100644
--- a/src/foundation/subsingleton-induction.lagda.md
+++ b/src/foundation/subsingleton-induction.lagda.md
@@ -59,12 +59,18 @@ compute-ind-is-subsingleton is-subsingleton-A a = pr2 (is-subsingleton-A a)
### Propositions satisfy subsingleton induction
```agda
+abstract
+ ind-subsingleton' :
+ {l1 l2 : Level} {A : UU l1} → is-proof-irrelevant A →
+ {B : A → UU l2} (a : A) → B a → (x : A) → B x
+ ind-subsingleton' H {B} a = ind-singleton a (H a) B
+
abstract
ind-subsingleton :
- {l1 l2 : Level} {A : UU l1} (is-prop-A : is-prop A)
+ {l1 l2 : Level} {A : UU l1} → is-prop A →
{B : A → UU l2} (a : A) → B a → (x : A) → B x
- ind-subsingleton is-prop-A {B} a =
- ind-singleton a (is-proof-irrelevant-is-prop is-prop-A a) B
+ ind-subsingleton is-prop-A =
+ ind-subsingleton' (is-proof-irrelevant-is-prop is-prop-A)
abstract
compute-ind-subsingleton :
@@ -78,18 +84,30 @@ abstract
### A type satisfies subsingleton induction if and only if it is a proposition
```agda
+is-subsingleton-is-proof-irrelevant :
+ {l1 l2 : Level} {A : UU l1} → is-proof-irrelevant A → is-subsingleton l2 A
+is-subsingleton-is-proof-irrelevant H {B} a =
+ is-singleton-is-contr a (H a) B
+
is-subsingleton-is-prop :
{l1 l2 : Level} {A : UU l1} → is-prop A → is-subsingleton l2 A
-is-subsingleton-is-prop is-prop-A {B} a =
- is-singleton-is-contr a (is-proof-irrelevant-is-prop is-prop-A a) B
+is-subsingleton-is-prop is-prop-A =
+ is-subsingleton-is-proof-irrelevant (is-proof-irrelevant-is-prop is-prop-A)
+
+abstract
+ is-proof-irrelevant-ind-subsingleton :
+ {l1 : Level} (A : UU l1) →
+ ({l2 : Level} {B : A → UU l2} (a : A) → B a → (x : A) → B x) →
+ is-proof-irrelevant A
+ is-proof-irrelevant-ind-subsingleton A S a =
+ is-contr-ind-singleton A a (λ B → S {B = B} a)
abstract
is-prop-ind-subsingleton :
{l1 : Level} (A : UU l1) →
({l2 : Level} {B : A → UU l2} (a : A) → B a → (x : A) → B x) → is-prop A
is-prop-ind-subsingleton A S =
- is-prop-is-proof-irrelevant
- ( λ a → is-contr-ind-singleton A a (λ B → S {B = B} a))
+ is-prop-is-proof-irrelevant (is-proof-irrelevant-ind-subsingleton A S)
abstract
is-prop-is-subsingleton :
diff --git a/src/foundation/surjective-maps.lagda.md b/src/foundation/surjective-maps.lagda.md
index 59d60318b5..42ab12dd78 100644
--- a/src/foundation/surjective-maps.lagda.md
+++ b/src/foundation/surjective-maps.lagda.md
@@ -470,6 +470,9 @@ module _
is-surjective g → is-surjective h → is-surjective (g ∘ h)
is-surjective-comp {g} {h} =
is-surjective-left-map-triangle (g ∘ h) g h refl-htpy
+
+ comp-surjection : B ↠ X → A ↠ B → A ↠ X
+ comp-surjection (g , G) (h , H) = g ∘ h , is-surjective-comp G H
```
### Functoriality of products preserves being surjective
diff --git a/src/foundation/transport-along-homotopies.lagda.md b/src/foundation/transport-along-homotopies.lagda.md
index f4dd21d96f..577c47c512 100644
--- a/src/foundation/transport-along-homotopies.lagda.md
+++ b/src/foundation/transport-along-homotopies.lagda.md
@@ -69,7 +69,7 @@ module _
abstract
compute-tr-htpy :
{f g : (x : A) → B x} (H : f ~ g) → statement-compute-tr-htpy H
- compute-tr-htpy {f} {g} =
+ compute-tr-htpy {f} =
ind-htpy f
( λ _ → statement-compute-tr-htpy)
( base-case-compute-tr-htpy)
diff --git a/src/foundation/transport-along-identifications.lagda.md b/src/foundation/transport-along-identifications.lagda.md
index 10987aecc0..d137762977 100644
--- a/src/foundation/transport-along-identifications.lagda.md
+++ b/src/foundation/transport-along-identifications.lagda.md
@@ -78,4 +78,9 @@ tr-loop :
{l1 : Level} {A : UU l1} {a0 a1 : A} (p : a0 = a1) (q : a0 = a0) →
tr (λ y → y = y) p q = (inv p ∙ q) ∙ p
tr-loop refl q = inv right-unit
+
+tr-loop-self :
+ {l1 : Level} {A : UU l1} {a : A} (p : a = a) →
+ tr (λ y → y = y) p p = p
+tr-loop-self p = tr-loop p p ∙ ap (_∙ p) (left-inv p)
```
diff --git a/src/foundation/trivial-relaxed-sigma-decompositions.lagda.md b/src/foundation/trivial-relaxed-sigma-decompositions.lagda.md
index 6578bdec99..8ab55c81ac 100644
--- a/src/foundation/trivial-relaxed-sigma-decompositions.lagda.md
+++ b/src/foundation/trivial-relaxed-sigma-decompositions.lagda.md
@@ -83,13 +83,7 @@ module _
equiv-trivial-is-trivial-Relaxed-Σ-Decomposition :
equiv-Relaxed-Σ-Decomposition D (trivial-Relaxed-Σ-Decomposition l4 A)
pr1 equiv-trivial-is-trivial-Relaxed-Σ-Decomposition =
- ( map-equiv (compute-raise-unit l4) ∘
- terminal-map (indexing-type-Relaxed-Σ-Decomposition D) ,
- is-equiv-comp
- ( map-equiv (compute-raise-unit l4))
- ( terminal-map (indexing-type-Relaxed-Σ-Decomposition D))
- ( is-equiv-terminal-map-is-contr is-trivial)
- ( is-equiv-map-equiv ( compute-raise-unit l4)))
+ equiv-raise-unit-is-contr is-trivial
pr1 (pr2 equiv-trivial-is-trivial-Relaxed-Σ-Decomposition) x =
( inv-equiv (matching-correspondence-Relaxed-Σ-Decomposition D)) ∘e
( inv-left-unit-law-Σ-is-contr is-trivial x)
diff --git a/src/foundation/trivial-sigma-decompositions.lagda.md b/src/foundation/trivial-sigma-decompositions.lagda.md
index 4e0fbf766b..e50fa85044 100644
--- a/src/foundation/trivial-sigma-decompositions.lagda.md
+++ b/src/foundation/trivial-sigma-decompositions.lagda.md
@@ -101,13 +101,7 @@ module _
( A)
( is-inhabited-base-type-is-trivial-Σ-Decomposition))
pr1 equiv-trivial-is-trivial-Σ-Decomposition =
- ( map-equiv (compute-raise-unit l4) ∘
- terminal-map (indexing-type-Σ-Decomposition D) ,
- is-equiv-comp
- ( map-equiv (compute-raise-unit l4))
- ( terminal-map (indexing-type-Σ-Decomposition D))
- ( is-equiv-terminal-map-is-contr is-trivial)
- ( is-equiv-map-equiv ( compute-raise-unit l4)))
+ equiv-raise-unit-is-contr is-trivial
pr1 (pr2 equiv-trivial-is-trivial-Σ-Decomposition) =
( λ x →
( ( inv-equiv (matching-correspondence-Σ-Decomposition D)) ∘e
diff --git a/src/foundation/truncation-equivalences.lagda.md b/src/foundation/truncation-equivalences.lagda.md
index 8d8ecf5786..5462a0bf32 100644
--- a/src/foundation/truncation-equivalences.lagda.md
+++ b/src/foundation/truncation-equivalences.lagda.md
@@ -40,8 +40,10 @@ open import foundation-core.truncation-levels
## Idea
-A map `f : A → B` is said to be a `k`-equivalence if the map
-`map-trunc k f : trunc k A → trunc k B` is an equivalence.
+A map `f : A → B` is said to be a
+{{#concept "`k`-equivalence" Disambiguation="truncations of types" Agda=truncation-equivalence}}
+if the map `map-trunc k f : trunc k A → trunc k B` is an
+[equivalence](foundation-core.equivalences.md).
## Definition
@@ -300,11 +302,11 @@ 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 |
+ → Σ 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/truncations.lagda.md b/src/foundation/truncations.lagda.md
index 8b01d5b970..9523e78045 100644
--- a/src/foundation/truncations.lagda.md
+++ b/src/foundation/truncations.lagda.md
@@ -144,11 +144,8 @@ module _
unique-dependent-function-trunc B f =
is-contr-equiv'
( fiber (precomp-Π-Truncated-Type unit-trunc B) f)
- ( equiv-tot
- ( λ h → equiv-funext))
- ( is-contr-map-is-equiv
- ( dependent-universal-property-trunc B)
- ( f))
+ ( equiv-tot (λ h → equiv-funext))
+ ( is-contr-map-is-equiv (dependent-universal-property-trunc B) f)
apply-dependent-universal-property-trunc :
{l2 : Level} (B : type-trunc k A → Truncated-Type l2 k) →
diff --git a/src/foundation/type-arithmetic-dependent-function-types.lagda.md b/src/foundation/type-arithmetic-dependent-function-types.lagda.md
index dcf9aa2c94..27533716f7 100644
--- a/src/foundation/type-arithmetic-dependent-function-types.lagda.md
+++ b/src/foundation/type-arithmetic-dependent-function-types.lagda.md
@@ -36,7 +36,7 @@ module _
( left-unit-law-Π ( λ _ → B a)) ∘e
( equiv-Π
( λ _ → B a)
- ( terminal-map A , is-equiv-terminal-map-is-contr C)
+ ( equiv-unit-is-contr C)
( λ a → equiv-eq (ap B ( eq-is-contr C))))
```
diff --git a/src/foundation/uniformly-decidable-type-families.lagda.md b/src/foundation/uniformly-decidable-type-families.lagda.md
new file mode 100644
index 0000000000..01904ea182
--- /dev/null
+++ b/src/foundation/uniformly-decidable-type-families.lagda.md
@@ -0,0 +1,175 @@
+# Uniformly decidable type families
+
+```agda
+module foundation.uniformly-decidable-type-families where
+```
+
+Imports
+
+```agda
+open import foundation.contractible-types
+open import foundation.coproduct-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.equality-coproduct-types
+open import foundation.inhabited-types
+open import foundation.negation
+open import foundation.propositional-truncations
+open import foundation.propositions
+open import foundation.truncated-types
+open import foundation.truncation-levels
+open import foundation.type-arithmetic-empty-type
+open import foundation.universe-levels
+
+open import foundation-core.cartesian-product-types
+open import foundation-core.contractible-maps
+open import foundation-core.empty-types
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.homotopies
+open import foundation-core.identity-types
+```
+
+
+
+## Idea
+
+A type family `B : A → 𝒰` is
+{{#concept "uniformly decidable" Agda=is-uniformly-decidable-family}} if there
+either is an element of every fiber `B x`, or every fiber is
+[empty](foundation-core.empty-types.md).
+
+## Definitions
+
+### The predicate of being uniformly decidable
+
+```agda
+is-uniformly-decidable-family :
+ {l1 l2 : Level} {A : UU l1} → (A → UU l2) → UU (l1 ⊔ l2)
+is-uniformly-decidable-family {A = A} B =
+ ((x : A) → B x) + ((x : A) → ¬ (B x))
+```
+
+## Properties
+
+### The fibers of a uniformly decidable type family are decidable
+
+```agda
+is-decidable-is-uniformly-decidable-family :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ is-uniformly-decidable-family B →
+ (x : A) → is-decidable (B x)
+is-decidable-is-uniformly-decidable-family (inl f) x = inl (f x)
+is-decidable-is-uniformly-decidable-family (inr g) x = inr (g x)
+```
+
+### The uniform decidability predicate on a family of truncated types
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ is-contr-is-uniformly-decidable-family-is-inhabited-base'' :
+ is-contr ((x : A) → (B x)) →
+ A →
+ is-contr (is-uniformly-decidable-family B)
+ is-contr-is-uniformly-decidable-family-is-inhabited-base'' H a =
+ is-contr-equiv
+ ( (x : A) → B x)
+ ( right-unit-law-coproduct-is-empty
+ ( (x : A) → B x)
+ ( (x : A) → ¬ B x)
+ ( λ f → f a (center H a)))
+ ( H)
+
+ is-contr-is-uniformly-decidable-family-is-inhabited-base' :
+ ((x : A) → is-contr (B x)) →
+ A →
+ is-contr (is-uniformly-decidable-family B)
+ is-contr-is-uniformly-decidable-family-is-inhabited-base' H =
+ is-contr-is-uniformly-decidable-family-is-inhabited-base'' (is-contr-Π H)
+
+ is-contr-is-uniformly-decidable-family-is-inhabited-base :
+ ((x : A) → is-contr (B x)) →
+ is-inhabited A →
+ is-contr (is-uniformly-decidable-family B)
+ is-contr-is-uniformly-decidable-family-is-inhabited-base H =
+ rec-trunc-Prop
+ ( is-contr-Prop (is-uniformly-decidable-family B))
+ ( is-contr-is-uniformly-decidable-family-is-inhabited-base' H)
+```
+
+### The uniform decidability predicate on a family of propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ is-prop-is-uniformly-decidable-family-is-inhabited-base'' :
+ is-prop ((x : A) → (B x)) →
+ A →
+ is-prop (is-uniformly-decidable-family B)
+ is-prop-is-uniformly-decidable-family-is-inhabited-base'' H a =
+ is-prop-coproduct
+ ( λ f nf → nf a (f a))
+ ( H)
+ ( is-prop-Π (λ x → is-prop-neg))
+
+ is-prop-is-uniformly-decidable-family-is-inhabited-base' :
+ ((x : A) → is-prop (B x)) →
+ A →
+ is-prop (is-uniformly-decidable-family B)
+ is-prop-is-uniformly-decidable-family-is-inhabited-base' H =
+ is-prop-is-uniformly-decidable-family-is-inhabited-base'' (is-prop-Π H)
+
+ is-prop-is-uniformly-decidable-family-is-inhabited-base :
+ ((x : A) → is-prop (B x)) →
+ is-inhabited A →
+ is-prop (is-uniformly-decidable-family B)
+ is-prop-is-uniformly-decidable-family-is-inhabited-base H =
+ rec-trunc-Prop
+ ( is-prop-Prop (is-uniformly-decidable-family B))
+ ( is-prop-is-uniformly-decidable-family-is-inhabited-base' H)
+```
+
+### The uniform decidability predicate on a family of truncated types
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ is-trunc-succ-succ-is-uniformly-decidable-family' :
+ (k : 𝕋) →
+ is-trunc (succ-𝕋 (succ-𝕋 k)) ((x : A) → (B x)) →
+ is-trunc (succ-𝕋 (succ-𝕋 k)) (is-uniformly-decidable-family B)
+ is-trunc-succ-succ-is-uniformly-decidable-family' k H =
+ is-trunc-coproduct k
+ ( H)
+ ( is-trunc-is-prop (succ-𝕋 k) (is-prop-Π (λ x → is-prop-neg)))
+
+ is-trunc-succ-succ-is-uniformly-decidable-family :
+ (k : 𝕋) →
+ ((x : A) → is-trunc (succ-𝕋 (succ-𝕋 k)) (B x)) →
+ is-trunc (succ-𝕋 (succ-𝕋 k)) (is-uniformly-decidable-family B)
+ is-trunc-succ-succ-is-uniformly-decidable-family k H =
+ is-trunc-succ-succ-is-uniformly-decidable-family' k
+ ( is-trunc-Π (succ-𝕋 (succ-𝕋 k)) H)
+
+ is-trunc-is-uniformly-decidable-family-is-inhabited-base :
+ (k : 𝕋) →
+ ((x : A) → is-trunc k (B x)) →
+ is-inhabited A →
+ is-trunc k (is-uniformly-decidable-family B)
+ is-trunc-is-uniformly-decidable-family-is-inhabited-base
+ neg-two-𝕋 =
+ is-contr-is-uniformly-decidable-family-is-inhabited-base
+ is-trunc-is-uniformly-decidable-family-is-inhabited-base
+ ( succ-𝕋 neg-two-𝕋) =
+ is-prop-is-uniformly-decidable-family-is-inhabited-base
+ is-trunc-is-uniformly-decidable-family-is-inhabited-base
+ ( succ-𝕋 (succ-𝕋 k)) H _ =
+ is-trunc-succ-succ-is-uniformly-decidable-family k H
+```
diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md
index f1bd0aee75..a3ae84790a 100644
--- a/src/foundation/unions-subtypes.lagda.md
+++ b/src/foundation/unions-subtypes.lagda.md
@@ -17,6 +17,10 @@ open import foundation.universe-levels
open import foundation-core.subtypes
+open import logic.de-morgan-propositions
+open import logic.de-morgan-subtypes
+open import logic.double-negation-stable-subtypes
+
open import order-theory.least-upper-bounds-large-posets
```
@@ -49,6 +53,15 @@ module _
union-decidable-subtype P Q x = disjunction-Decidable-Prop (P x) (Q x)
```
+### Unions of De Morgan subtypes
+
+```agda
+ union-de-morgan-subtype :
+ de-morgan-subtype l1 X → de-morgan-subtype l2 X →
+ de-morgan-subtype (l1 ⊔ l2) X
+ union-de-morgan-subtype P Q x = disjunction-De-Morgan-Prop (P x) (Q x)
+```
+
### Unions of families of subtypes
```agda
diff --git a/src/foundation/unit-type.lagda.md b/src/foundation/unit-type.lagda.md
index ca159a3636..cc388a234b 100644
--- a/src/foundation/unit-type.lagda.md
+++ b/src/foundation/unit-type.lagda.md
@@ -7,6 +7,7 @@ module foundation.unit-type where
Imports
```agda
+open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.diagonal-maps-of-types
open import foundation.raising-universe-levels
@@ -18,6 +19,7 @@ open import foundation-core.equivalences
open import foundation-core.identity-types
open import foundation-core.injective-maps
open import foundation-core.propositions
+open import foundation-core.retractions
open import foundation-core.sets
open import foundation-core.truncated-types
open import foundation-core.truncation-levels
@@ -96,7 +98,12 @@ inv-compute-raise-unit l = inv-compute-raise l unit
abstract
is-contr-unit : is-contr unit
pr1 is-contr-unit = star
- pr2 is-contr-unit star = refl
+ pr2 is-contr-unit _ = refl
+
+abstract
+ is-contr-raise-unit : {l1 : Level} → is-contr (raise-unit l1)
+ is-contr-raise-unit {l1} =
+ is-contr-equiv' unit (compute-raise l1 unit) is-contr-unit
```
### Any contractible type is equivalent to the unit type
@@ -106,22 +113,58 @@ module _
{l : Level} {A : UU l}
where
- abstract
- is-equiv-terminal-map-is-contr :
- is-contr A → is-equiv (terminal-map A)
- pr1 (pr1 (is-equiv-terminal-map-is-contr H)) = ind-unit (center H)
- pr2 (pr1 (is-equiv-terminal-map-is-contr H)) = ind-unit refl
- pr1 (pr2 (is-equiv-terminal-map-is-contr H)) x = center H
- pr2 (pr2 (is-equiv-terminal-map-is-contr H)) = contraction H
+ is-equiv-terminal-map-is-contr : is-contr A → is-equiv (terminal-map A)
+ is-equiv-terminal-map-is-contr H =
+ is-equiv-is-invertible (λ _ → center H) (λ _ → refl) (contraction H)
equiv-unit-is-contr : is-contr A → A ≃ unit
- pr1 (equiv-unit-is-contr H) = terminal-map A
- pr2 (equiv-unit-is-contr H) = is-equiv-terminal-map-is-contr H
+ equiv-unit-is-contr H = terminal-map A , is-equiv-terminal-map-is-contr H
+
+ is-contr-retraction-terminal-map : retraction (terminal-map A) → is-contr A
+ is-contr-retraction-terminal-map (h , H) = h star , H
+
+ is-contr-is-equiv-terminal-map : is-equiv (terminal-map A) → is-contr A
+ is-contr-is-equiv-terminal-map H =
+ is-contr-retraction-terminal-map (retraction-is-equiv H)
+
+ is-contr-equiv-unit : A ≃ unit → is-contr A
+ is-contr-equiv-unit e = map-inv-equiv e star , is-retraction-map-inv-equiv e
+```
+
+### Any contractible type is equivalent to the raised unit type
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1}
+ where
- abstract
- is-contr-is-equiv-terminal-map : is-equiv (terminal-map A) → is-contr A
- pr1 (is-contr-is-equiv-terminal-map ((g , G) , (h , H))) = h star
- pr2 (is-contr-is-equiv-terminal-map ((g , G) , (h , H))) = H
+ is-equiv-raise-terminal-map-is-contr :
+ is-contr A → is-equiv (raise-terminal-map {l2 = l2} A)
+ is-equiv-raise-terminal-map-is-contr H =
+ is-equiv-is-invertible
+ ( λ _ → center H)
+ ( λ where (map-raise x) → refl)
+ ( contraction H)
+
+ equiv-raise-unit-is-contr : is-contr A → A ≃ raise-unit l2
+ equiv-raise-unit-is-contr H =
+ raise-terminal-map A , is-equiv-raise-terminal-map-is-contr H
+
+ is-contr-retraction-raise-terminal-map :
+ retraction (raise-terminal-map {l2 = l2} A) → is-contr A
+ is-contr-retraction-raise-terminal-map (h , H) = h raise-star , H
+
+ is-contr-is-equiv-raise-terminal-map :
+ is-equiv (raise-terminal-map {l2 = l2} A) → is-contr A
+ is-contr-is-equiv-raise-terminal-map H =
+ is-contr-retraction-raise-terminal-map (retraction-is-equiv H)
+
+ is-contr-equiv-raise-unit : A ≃ raise-unit l2 → is-contr A
+ is-contr-equiv-raise-unit e =
+ ( map-inv-equiv e raise-star) ,
+ ( λ x →
+ ap (map-inv-equiv e) (eq-is-contr is-contr-raise-unit) ∙
+ is-retraction-map-inv-equiv e x)
```
### The unit type is a proposition
@@ -132,8 +175,14 @@ abstract
is-prop-unit = is-prop-is-contr is-contr-unit
unit-Prop : Prop lzero
-pr1 unit-Prop = unit
-pr2 unit-Prop = is-prop-unit
+unit-Prop = unit , is-prop-unit
+
+abstract
+ is-prop-raise-unit : {l1 : Level} → is-prop (raise-unit l1)
+ is-prop-raise-unit {l1} = is-prop-equiv' (compute-raise l1 unit) is-prop-unit
+
+raise-unit-Prop : (l1 : Level) → Prop l1
+raise-unit-Prop l1 = raise-unit l1 , is-prop-raise-unit
```
### The unit type is a set
@@ -144,36 +193,14 @@ abstract
is-set-unit = is-trunc-succ-is-trunc neg-one-𝕋 is-prop-unit
unit-Set : Set lzero
-pr1 unit-Set = unit
-pr2 unit-Set = is-set-unit
-```
-
-```agda
-abstract
- is-contr-raise-unit :
- {l1 : Level} → is-contr (raise-unit l1)
- is-contr-raise-unit {l1} =
- is-contr-equiv' unit (compute-raise l1 unit) is-contr-unit
-
-abstract
- is-prop-raise-unit :
- {l1 : Level} → is-prop (raise-unit l1)
- is-prop-raise-unit {l1} =
- is-prop-equiv' (compute-raise l1 unit) is-prop-unit
-
-raise-unit-Prop :
- (l1 : Level) → Prop l1
-pr1 (raise-unit-Prop l1) = raise-unit l1
-pr2 (raise-unit-Prop l1) = is-prop-raise-unit
+unit-Set = unit , is-set-unit
abstract
- is-set-raise-unit :
- {l1 : Level} → is-set (raise-unit l1)
+ is-set-raise-unit : {l1 : Level} → is-set (raise-unit l1)
is-set-raise-unit = is-trunc-succ-is-trunc neg-one-𝕋 is-prop-raise-unit
raise-unit-Set : (l1 : Level) → Set l1
-pr1 (raise-unit-Set l1) = raise-unit l1
-pr2 (raise-unit-Set l1) = is-set-raise-unit
+raise-unit-Set l1 = raise-unit l1 , is-set-raise-unit
```
### All parallel maps into `unit` are equal
diff --git a/src/foundation/univalence-implies-function-extensionality.lagda.md b/src/foundation/univalence-implies-function-extensionality.lagda.md
index 8a0b581a0e..8845653a49 100644
--- a/src/foundation/univalence-implies-function-extensionality.lagda.md
+++ b/src/foundation/univalence-implies-function-extensionality.lagda.md
@@ -21,6 +21,7 @@ open import foundation-core.fibers-of-maps
open import foundation-core.function-types
open import foundation-core.homotopies
open import foundation-core.identity-types
+open import foundation-core.retracts-of-types
open import foundation-core.transport-along-identifications
```
@@ -34,14 +35,20 @@ The [univalence axiom](foundation-core.univalence.md) implies
## Theorem
```agda
+retract-compute-fiber-id-postcomp-pr1 :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ ((a : A) → B a) retract-of (fiber (postcomp A (pr1 {B = B})) id)
+retract-compute-fiber-id-postcomp-pr1 {B = B} =
+ ( λ f → ((λ x → (x , f x)) , refl)) ,
+ ( λ h x → tr B (htpy-eq (pr2 h) x) (pr2 (pr1 h x))) ,
+ ( refl-htpy)
+
abstract
weak-funext-univalence : {l : Level} → weak-function-extensionality-Level l l
weak-funext-univalence A B is-contr-B =
is-contr-retract-of
( fiber (postcomp A pr1) id)
- ( ( λ f → ((λ x → (x , f x)) , refl)) ,
- ( λ h x → tr B (htpy-eq (pr2 h) x) (pr2 (pr1 h x))) ,
- ( refl-htpy))
+ ( retract-compute-fiber-id-postcomp-pr1)
( is-contr-map-is-equiv
( is-equiv-postcomp-univalence A (equiv-pr1 is-contr-B))
( id))
diff --git a/src/foundation/weak-limited-principle-of-omniscience.lagda.md b/src/foundation/weak-limited-principle-of-omniscience.lagda.md
index b4a805bc24..bc9b213e06 100644
--- a/src/foundation/weak-limited-principle-of-omniscience.lagda.md
+++ b/src/foundation/weak-limited-principle-of-omniscience.lagda.md
@@ -9,38 +9,37 @@ module foundation.weak-limited-principle-of-omniscience where
```agda
open import elementary-number-theory.natural-numbers
+open import foundation.booleans
open import foundation.disjunction
open import foundation.negation
open import foundation.universal-quantification
open import foundation.universe-levels
+open import foundation-core.decidable-propositions
open import foundation-core.propositions
open import foundation-core.sets
-
-open import univalent-combinatorics.standard-finite-types
```
## Statement
-The {{#concept "Weak Limited Principle of Omniscience"}} asserts that for any
-[sequence](foundation.sequences.md) `f : ℕ → Fin 2` either `f n = 0` for all
-`n : ℕ` or not. In particular, it is a restricted form of the
+The {{#concept "weak limited principle of omniscience"}} (WLPO) asserts that for
+any [sequence](foundation.sequences.md) `f : ℕ → bool` either `f n` is true for
+all `n : ℕ` or it is [not](foundation-core.negation.md). In particular, it is a
+restricted form of the
[law of excluded middle](foundation.law-of-excluded-middle.md).
```agda
-WLPO-Prop : Prop lzero
-WLPO-Prop =
- ∀'
- ( ℕ → Fin 2)
- ( λ f →
- disjunction-Prop
- ( ∀' ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))
- ( ¬' (∀' ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))))
+prop-WLPO : Prop lzero
+prop-WLPO =
+ ∀' (ℕ → bool) (λ f → is-decidable-Prop (∀' ℕ (λ n → is-true-Prop (f n))))
WLPO : UU lzero
-WLPO = type-Prop WLPO-Prop
+WLPO = type-Prop prop-WLPO
+
+is-prop-WLPO : is-prop WLPO
+is-prop-WLPO = is-prop-type-Prop prop-WLPO
```
## See also
@@ -48,3 +47,11 @@ WLPO = type-Prop WLPO-Prop
- [The principle of omniscience](foundation.principle-of-omniscience.md)
- [The limited principle of omniscience](foundation.limited-principle-of-omniscience.md)
- [The lesser limited principle of omniscience](foundation.lesser-limited-principle-of-omniscience.md)
+- [Markov's principle](logic.markovs-principle.md)
+
+## External links
+
+- [Taboos.WLPO](https://martinescardo.github.io/TypeTopology/Taboos.WLPO.html)
+ at TypeTopology
+- [weak limited principle of omniscience](https://ncatlab.org/nlab/show/weak+limited+principle+of+omniscience)
+ at $n$Lab
diff --git a/src/logic.lagda.md b/src/logic.lagda.md
new file mode 100644
index 0000000000..4f76bcf78b
--- /dev/null
+++ b/src/logic.lagda.md
@@ -0,0 +1,23 @@
+# Logic
+
+## Modules in the logic namespace
+
+```agda
+module logic where
+
+open import logic.complements-de-morgan-subtypes public
+open import logic.complements-decidable-subtypes public
+open import logic.complements-double-negation-stable-subtypes public
+open import logic.de-morgan-embeddings public
+open import logic.de-morgan-maps public
+open import logic.de-morgan-propositions public
+open import logic.de-morgan-subtypes public
+open import logic.de-morgan-types public
+open import logic.de-morgans-law public
+open import logic.double-negation-eliminating-maps public
+open import logic.double-negation-elimination public
+open import logic.double-negation-stable-embeddings public
+open import logic.double-negation-stable-subtypes public
+open import logic.markovian-types public
+open import logic.markovs-principle public
+```
diff --git a/src/logic/complements-de-morgan-subtypes.lagda.md b/src/logic/complements-de-morgan-subtypes.lagda.md
new file mode 100644
index 0000000000..fbfe0b0e4a
--- /dev/null
+++ b/src/logic/complements-de-morgan-subtypes.lagda.md
@@ -0,0 +1,101 @@
+# Complements of De Morgan subtypes
+
+```agda
+module logic.complements-de-morgan-subtypes where
+```
+
+Imports
+
+```agda
+open import foundation.complements-subtypes
+open import foundation.decidable-subtypes
+open import foundation.dependent-pair-types
+open import foundation.double-negation
+open import foundation.full-subtypes
+open import foundation.involutions
+open import foundation.negation
+open import foundation.postcomposition-functions
+open import foundation.powersets
+open import foundation.propositional-truncations
+open import foundation.subtypes
+open import foundation.unions-subtypes
+open import foundation.universe-levels
+
+open import foundation-core.function-types
+
+open import logic.complements-decidable-subtypes
+open import logic.de-morgan-propositions
+open import logic.de-morgan-subtypes
+```
+
+
+
+## Idea
+
+The
+{{#concept "complement" Disambiguation="of a De Morgan subtype" Agda=complement-de-morgan-subtype}}
+of a [De Morgan subtype](logic.de-morgan-subtypes.md) `B ⊆ A` consists of the
+elements that are not in `B`.
+
+## Definition
+
+### Complements of De Morgan subtypes
+
+```agda
+complement-de-morgan-subtype :
+ {l1 l2 : Level} {A : UU l1} → de-morgan-subtype l2 A → de-morgan-subtype l2 A
+complement-de-morgan-subtype P x = neg-De-Morgan-Prop (P x)
+```
+
+## Properties
+
+### Complement of De Morgan subtypes are decidable
+
+```agda
+is-decidable-complement-de-morgan-subtype :
+ {l1 l2 : Level} {A : UU l1} (P : de-morgan-subtype l2 A) →
+ is-decidable-subtype
+ ( subtype-de-morgan-subtype (complement-de-morgan-subtype P))
+is-decidable-complement-de-morgan-subtype P = is-de-morgan-de-morgan-subtype P
+```
+
+### The union of the complement of a subtype `P` with its double complement is the full subtype if and only if `P` is De Morgan
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1}
+ where
+
+ is-full-union-complement-subtype-double-complement-subtype :
+ (P : subtype l2 A) → is-de-morgan-subtype P →
+ is-full-subtype
+ ( union-subtype
+ ( complement-subtype P)
+ ( complement-subtype (complement-subtype P)))
+ is-full-union-complement-subtype-double-complement-subtype P =
+ is-full-union-subtype-complement-subtype (complement-subtype P)
+
+ is-de-morgan-subtype-is-full-union-complement-subtype-double-complement-subtype :
+ (P : subtype l2 A) →
+ is-full-subtype
+ ( union-subtype
+ ( complement-subtype P)
+ ( complement-subtype (complement-subtype P))) →
+ is-de-morgan-subtype P
+ is-de-morgan-subtype-is-full-union-complement-subtype-double-complement-subtype
+ P =
+ is-decidable-subtype-is-full-union-subtype-complement-subtype
+ ( complement-subtype P)
+
+ is-full-union-subtype-complement-de-morgan-subtype :
+ (P : de-morgan-subtype l2 A) →
+ is-full-subtype
+ ( union-subtype
+ ( complement-subtype (subtype-de-morgan-subtype P))
+ ( complement-subtype
+ ( complement-subtype (subtype-de-morgan-subtype P))))
+ is-full-union-subtype-complement-de-morgan-subtype P =
+ is-full-union-complement-subtype-double-complement-subtype
+ ( subtype-de-morgan-subtype P)
+ ( is-de-morgan-de-morgan-subtype P)
+```
diff --git a/src/logic/complements-decidable-subtypes.lagda.md b/src/logic/complements-decidable-subtypes.lagda.md
new file mode 100644
index 0000000000..ae2a775f33
--- /dev/null
+++ b/src/logic/complements-decidable-subtypes.lagda.md
@@ -0,0 +1,100 @@
+# Complements of decidable subtypes
+
+```agda
+module logic.complements-decidable-subtypes where
+```
+
+Imports
+
+```agda
+open import foundation.complements-subtypes
+open import foundation.coproduct-types
+open import foundation.decidable-propositions
+open import foundation.decidable-subtypes
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation-stable-propositions
+open import foundation.evaluation-functions
+open import foundation.full-subtypes
+open import foundation.involutions
+open import foundation.negation
+open import foundation.postcomposition-functions
+open import foundation.powersets
+open import foundation.propositional-truncations
+open import foundation.unions-subtypes
+open import foundation.universe-levels
+
+open import foundation-core.function-types
+open import foundation-core.subtypes
+
+open import logic.double-negation-stable-subtypes
+```
+
+
+
+## Idea
+
+The
+{{#concept "complement" Disambiguation="of a decidable subtype" Agda=complement-decidable-subtype}}
+of a [decidable subtype](foundation.decidable-subtypes.md) `B ⊆ A` consists of
+the elements that are not in `B`.
+
+## Definition
+
+### Complements of decidable subtypes
+
+```agda
+complement-decidable-subtype :
+ {l1 l2 : Level} {A : UU l1} → decidable-subtype l2 A → decidable-subtype l2 A
+complement-decidable-subtype P x = neg-Decidable-Prop (P x)
+```
+
+## Properties
+
+### Taking complements is an involution on decidable subtypes
+
+```agda
+is-involution-complement-decidable-subtype :
+ {l1 l2 : Level} {A : UU l1} →
+ is-involution (complement-decidable-subtype {l1} {l2} {A})
+is-involution-complement-decidable-subtype P =
+ eq-has-same-elements-decidable-subtype
+ ( complement-decidable-subtype (complement-decidable-subtype P))
+ ( P)
+ ( λ x →
+ double-negation-elim-is-decidable (is-decidable-decidable-subtype P x) ,
+ ev)
+```
+
+### The union of a subtype `P` with its complement is the full subtype if and only if `P` is a decidable subtype
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1}
+ where
+
+ is-full-union-subtype-complement-subtype :
+ (P : subtype l2 A) → is-decidable-subtype P →
+ is-full-subtype (union-subtype P (complement-subtype P))
+ is-full-union-subtype-complement-subtype P d x =
+ unit-trunc-Prop (d x)
+
+ is-decidable-subtype-is-full-union-subtype-complement-subtype :
+ (P : subtype l2 A) →
+ is-full-subtype (union-subtype P (complement-subtype P)) →
+ is-decidable-subtype P
+ is-decidable-subtype-is-full-union-subtype-complement-subtype P H x =
+ apply-universal-property-trunc-Prop
+ ( H x)
+ ( is-decidable-Prop (P x))
+ ( id)
+
+ is-full-union-subtype-complement-decidable-subtype :
+ (P : decidable-subtype l2 A) →
+ is-full-decidable-subtype
+ ( union-decidable-subtype P (complement-decidable-subtype P))
+ is-full-union-subtype-complement-decidable-subtype P =
+ is-full-union-subtype-complement-subtype
+ ( subtype-decidable-subtype P)
+ ( is-decidable-decidable-subtype P)
+```
diff --git a/src/logic/complements-double-negation-stable-subtypes.lagda.md b/src/logic/complements-double-negation-stable-subtypes.lagda.md
new file mode 100644
index 0000000000..e4ba546190
--- /dev/null
+++ b/src/logic/complements-double-negation-stable-subtypes.lagda.md
@@ -0,0 +1,66 @@
+# Complements of double negation stable subtypes
+
+```agda
+module logic.complements-double-negation-stable-subtypes where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.double-negation
+open import foundation.double-negation-stable-propositions
+open import foundation.full-subtypes
+open import foundation.involutions
+open import foundation.negation
+open import foundation.postcomposition-functions
+open import foundation.powersets
+open import foundation.propositional-truncations
+open import foundation.subtypes
+open import foundation.unions-subtypes
+open import foundation.universe-levels
+
+open import foundation-core.function-types
+
+open import logic.double-negation-stable-subtypes
+```
+
+
+
+## Idea
+
+The
+{{#concept "complement" Disambiguation="of a double negation stable subtype" Agda=complement-double-negation-stable-subtype}}
+of a [double negation stable subtype](logic.double-negation-stable-subtypes.md)
+`B ⊆ A` consists of the elements that are not in `B`.
+
+## Definition
+
+### Complements of double negation stable subtypes
+
+```agda
+complement-double-negation-stable-subtype :
+ {l1 l2 : Level} {A : UU l1} →
+ double-negation-stable-subtype l2 A →
+ double-negation-stable-subtype l2 A
+complement-double-negation-stable-subtype P x =
+ neg-Double-Negation-Stable-Prop (P x)
+```
+
+## Properties
+
+### Taking complements is an involution on double negation stable subtypes
+
+```agda
+is-involution-complement-double-negation-stable-subtype :
+ {l1 l2 : Level} {A : UU l1} →
+ is-involution (complement-double-negation-stable-subtype {l1} {l2} {A})
+is-involution-complement-double-negation-stable-subtype P =
+ eq-has-same-elements-double-negation-stable-subtype
+ ( complement-double-negation-stable-subtype
+ ( complement-double-negation-stable-subtype P))
+ ( P)
+ ( λ x →
+ ( is-double-negation-stable-double-negation-stable-subtype P x ,
+ intro-double-negation))
+```
diff --git a/src/logic/de-morgan-embeddings.lagda.md b/src/logic/de-morgan-embeddings.lagda.md
new file mode 100644
index 0000000000..3391dea43b
--- /dev/null
+++ b/src/logic/de-morgan-embeddings.lagda.md
@@ -0,0 +1,676 @@
+# De Morgan embeddings
+
+```agda
+module logic.de-morgan-embeddings where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.cartesian-morphisms-arrows
+open import foundation.decidable-embeddings
+open import foundation.decidable-maps
+open import foundation.decidable-propositions
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.embeddings
+open import foundation.fibers-of-maps
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-coproduct-types
+open import foundation.fundamental-theorem-of-identity-types
+open import foundation.homotopy-induction
+open import foundation.identity-types
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.propositional-maps
+open import foundation.propositions
+open import foundation.retracts-of-maps
+open import foundation.subtype-identity-principle
+open import foundation.type-arithmetic-dependent-pair-types
+open import foundation.unit-type
+open import foundation.universal-property-equivalences
+open import foundation.universe-levels
+
+open import foundation-core.cartesian-product-types
+open import foundation-core.coproduct-types
+open import foundation-core.empty-types
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.functoriality-dependent-pair-types
+open import foundation-core.homotopies
+open import foundation-core.injective-maps
+open import foundation-core.torsorial-type-families
+
+open import logic.de-morgan-maps
+open import logic.de-morgan-propositions
+open import logic.de-morgan-types
+open import logic.double-negation-eliminating-maps
+open import logic.double-negation-elimination
+```
+
+
+
+## Idea
+
+A [map](foundation-core.function-types.md) is said to be a
+{{#concept "De Morgan embedding" Disambiguation="of types" Agda=is-de-morgan-emb}}
+if it is an [embedding](foundation-core.embeddings.md) and the
+[negations](foundation-core.negation.md) of its
+[fibers](foundation-core.fibers-of-maps.md) are
+[decidable](foundation.decidable-maps.md).
+
+Equivalently, a De Morgan embedding is a map whose fibers are
+[De Morgan propositions](logic.de-morgan-propositions.md). We refer to this
+condition as being a
+{{#concept "De Morgan propositional map" Disambiguation="of types" Agda=is-de-morgan-prop-map}}.
+
+## Definitions
+
+### The condition on a map of being a De Morgan embedding
+
+```agda
+is-de-morgan-emb :
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} → (X → Y) → UU (l1 ⊔ l2)
+is-de-morgan-emb f = is-emb f × is-de-morgan-map f
+
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y}
+ (F : is-de-morgan-emb f)
+ where
+
+ is-emb-is-de-morgan-emb : is-emb f
+ is-emb-is-de-morgan-emb = pr1 F
+
+ is-de-morgan-map-is-de-morgan-emb :
+ is-de-morgan-map f
+ is-de-morgan-map-is-de-morgan-emb = pr2 F
+
+ is-prop-map-is-de-morgan-emb : is-prop-map f
+ is-prop-map-is-de-morgan-emb =
+ is-prop-map-is-emb is-emb-is-de-morgan-emb
+
+ is-injective-is-de-morgan-emb : is-injective f
+ is-injective-is-de-morgan-emb =
+ is-injective-is-emb is-emb-is-de-morgan-emb
+```
+
+### De Morgan propositional maps
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ where
+
+ is-de-morgan-prop-map : (X → Y) → UU (l1 ⊔ l2)
+ is-de-morgan-prop-map f =
+ (y : Y) → is-de-morgan-prop (fiber f y)
+
+ is-prop-is-de-morgan-prop-map :
+ (f : X → Y) → is-prop (is-de-morgan-prop-map f)
+ is-prop-is-de-morgan-prop-map f =
+ is-prop-Π (λ y → is-prop-is-de-morgan-prop (fiber f y))
+
+ is-de-morgan-prop-map-Prop : (X → Y) → Prop (l1 ⊔ l2)
+ is-de-morgan-prop-map-Prop f =
+ ( is-de-morgan-prop-map f , is-prop-is-de-morgan-prop-map f)
+
+ is-prop-map-is-de-morgan-prop-map :
+ {f : X → Y} → is-de-morgan-prop-map f → is-prop-map f
+ is-prop-map-is-de-morgan-prop-map H y =
+ is-prop-type-is-de-morgan-prop (H y)
+
+ is-de-morgan-map-is-de-morgan-prop-map :
+ {f : X → Y} → is-de-morgan-prop-map f → is-de-morgan-map f
+ is-de-morgan-map-is-de-morgan-prop-map H y =
+ is-de-morgan-type-is-de-morgan-prop (H y)
+```
+
+### The type of De Morgan embeddings
+
+```agda
+infix 5 _↪ᵈᵐ_
+_↪ᵈᵐ_ : {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+X ↪ᵈᵐ Y = Σ (X → Y) (is-de-morgan-emb)
+
+de-morgan-emb : {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+de-morgan-emb = _↪ᵈᵐ_
+
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪ᵈᵐ Y)
+ where
+
+ 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-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-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
+ emb-de-morgan-emb = map-de-morgan-emb , is-emb-map-de-morgan-emb
+```
+
+## Properties
+
+### Any map of which the fibers are De Morgan propositions is a De Morgan embedding
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y}
+ where
+
+ abstract
+ is-de-morgan-emb-is-de-morgan-prop-map :
+ is-de-morgan-prop-map f → is-de-morgan-emb f
+ pr1 (is-de-morgan-emb-is-de-morgan-prop-map H) =
+ is-emb-is-prop-map (is-prop-map-is-de-morgan-prop-map H)
+ pr2 (is-de-morgan-emb-is-de-morgan-prop-map H) =
+ is-de-morgan-map-is-de-morgan-prop-map H
+
+ abstract
+ is-de-morgan-prop-map-is-de-morgan-emb :
+ is-de-morgan-emb f → is-de-morgan-prop-map f
+ pr1 (is-de-morgan-prop-map-is-de-morgan-emb H y) =
+ is-prop-map-is-de-morgan-emb H y
+ pr2 (is-de-morgan-prop-map-is-de-morgan-emb H y) =
+ is-de-morgan-map-is-de-morgan-emb H y
+```
+
+### The first projection map of a dependent sum of De Morgan propositions is a De Morgan embedding
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (Q : A → De-Morgan-Prop l2)
+ where
+
+ is-de-morgan-prop-map-pr1 :
+ is-de-morgan-prop-map
+ ( pr1 {B = type-De-Morgan-Prop ∘ Q})
+ is-de-morgan-prop-map-pr1 y =
+ is-de-morgan-prop-equiv
+ ( equiv-fiber-pr1 (type-De-Morgan-Prop ∘ Q) y)
+ ( is-de-morgan-prop-type-De-Morgan-Prop (Q y))
+
+ is-de-morgan-emb-pr1 :
+ is-de-morgan-emb
+ ( pr1 {B = type-De-Morgan-Prop ∘ Q})
+ is-de-morgan-emb-pr1 =
+ is-de-morgan-emb-is-de-morgan-prop-map
+ ( is-de-morgan-prop-map-pr1)
+```
+
+### Decidable embeddings are De Morgan
+
+```agda
+is-de-morgan-map-is-decidable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-emb f → is-de-morgan-map f
+is-de-morgan-map-is-decidable-emb H =
+ is-de-morgan-map-is-decidable-map
+ ( is-decidable-map-is-decidable-emb H)
+
+abstract
+ is-de-morgan-emb-is-decidable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-emb f → is-de-morgan-emb f
+ is-de-morgan-emb-is-decidable-emb H =
+ ( is-emb-is-decidable-emb H ,
+ is-de-morgan-map-is-decidable-emb H)
+```
+
+### Equivalences are De Morgan embeddings
+
+```agda
+abstract
+ is-de-morgan-emb-is-equiv :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-equiv f → is-de-morgan-emb f
+ is-de-morgan-emb-is-equiv H =
+ ( is-emb-is-equiv H , is-de-morgan-map-is-equiv H)
+```
+
+### Identity maps are De Morgan embeddings
+
+```agda
+is-de-morgan-emb-id :
+ {l : Level} {A : UU l} → is-de-morgan-emb (id {A = A})
+is-de-morgan-emb-id =
+ ( is-emb-id , is-de-morgan-map-id)
+
+de-morgan-emb-id : {l : Level} {A : UU l} → A ↪ᵈᵐ A
+de-morgan-emb-id = (id , is-de-morgan-emb-id)
+
+is-de-morgan-prop-map-id :
+ {l : Level} {A : UU l} → is-de-morgan-prop-map (id {A = A})
+is-de-morgan-prop-map-id y =
+ is-de-morgan-prop-is-contr (is-torsorial-Id' y)
+```
+
+### Being a De Morgan embedding is a property
+
+```agda
+abstract
+ is-prop-is-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) →
+ is-prop (is-de-morgan-emb f)
+ is-prop-is-de-morgan-emb f =
+ is-prop-product (is-property-is-emb f) is-prop-is-de-morgan-map
+```
+
+### De Morgan embeddings are closed under homotopies
+
+```agda
+abstract
+ is-de-morgan-emb-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
+ f ~ g → is-de-morgan-emb g → is-de-morgan-emb f
+ is-de-morgan-emb-htpy {f = f} {g} H K =
+ ( is-emb-htpy H (is-emb-is-de-morgan-emb K) ,
+ is-de-morgan-map-htpy H
+ ( is-de-morgan-map-is-de-morgan-emb K))
+```
+
+### De Morgan embeddings are closed under composition
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {g : B → C} {f : A → B}
+ where
+
+ is-de-morgan-map-comp-is-decidable-emb :
+ is-decidable-emb g →
+ is-de-morgan-map f →
+ is-de-morgan-map (g ∘ f)
+ is-de-morgan-map-comp-is-decidable-emb G F =
+ is-de-morgan-map-comp-is-decidable-map
+ ( is-injective-is-decidable-emb G)
+ ( is-decidable-map-is-decidable-emb G)
+ ( F)
+
+ is-de-morgan-prop-map-comp-is-decidable-prop-map :
+ is-decidable-prop-map g →
+ is-de-morgan-prop-map f →
+ is-de-morgan-prop-map (g ∘ f)
+ is-de-morgan-prop-map-comp-is-decidable-prop-map K H z =
+ is-de-morgan-prop-equiv
+ ( compute-fiber-comp g f z)
+ ( is-de-morgan-prop-Σ (K z) (H ∘ pr1))
+
+ is-de-morgan-emb-comp-is-decidable-emb :
+ is-decidable-emb g →
+ is-de-morgan-emb f →
+ is-de-morgan-emb (g ∘ f)
+ is-de-morgan-emb-comp-is-decidable-emb K H =
+ is-de-morgan-emb-is-de-morgan-prop-map
+ ( is-de-morgan-prop-map-comp-is-decidable-prop-map
+ ( is-decidable-prop-map-is-decidable-emb K)
+ ( is-de-morgan-prop-map-is-de-morgan-emb H))
+
+comp-de-morgan-emb :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ B ↪ᵈ C → A ↪ᵈᵐ B → A ↪ᵈᵐ C
+comp-de-morgan-emb (g , G) (f , F) =
+ ( g ∘ f , is-de-morgan-emb-comp-is-decidable-emb G F)
+```
+
+### Left cancellation for De Morgan embeddings
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {f : A → B} {g : B → C}
+ where
+
+ is-de-morgan-emb-right-factor' :
+ is-de-morgan-emb (g ∘ f) →
+ is-emb g →
+ is-de-morgan-emb f
+ is-de-morgan-emb-right-factor' GF G =
+ ( is-emb-right-factor g f G (is-emb-is-de-morgan-emb GF) ,
+ is-de-morgan-map-right-factor'
+ ( is-injective-is-emb G)
+ ( is-de-morgan-map-is-de-morgan-emb GF))
+
+ is-de-morgan-emb-right-factor :
+ is-de-morgan-emb (g ∘ f) →
+ is-de-morgan-emb g →
+ is-de-morgan-emb f
+ is-de-morgan-emb-right-factor GF G =
+ is-de-morgan-emb-right-factor'
+ ( GF)
+ ( is-emb-is-de-morgan-emb G)
+```
+
+### In a commuting triangle of maps, if the top and right maps are De Morgan embeddings so is the left map
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {top : A → B} {left : A → C} {right : B → C}
+ (H : left ~ right ∘ top)
+ where
+
+ is-de-morgan-emb-left-map-triangle-is-decidable-emb-top :
+ is-de-morgan-emb top →
+ is-decidable-emb right →
+ is-de-morgan-emb left
+ is-de-morgan-emb-left-map-triangle-is-decidable-emb-top T R =
+ is-de-morgan-emb-htpy H (is-de-morgan-emb-comp-is-decidable-emb R T)
+```
+
+### In a commuting triangle of maps, if the left and right maps are De Morgan embeddings so is the top map
+
+In fact, the right map need only be an embedding.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {top : A → B} {left : A → C} {right : B → C}
+ (H : left ~ right ∘ top)
+ where
+
+ is-de-morgan-emb-top-map-triangle' :
+ is-emb right →
+ is-de-morgan-emb left →
+ is-de-morgan-emb top
+ is-de-morgan-emb-top-map-triangle' R' L =
+ is-de-morgan-emb-right-factor'
+ ( is-de-morgan-emb-htpy (inv-htpy H) L)
+ ( R')
+
+ is-de-morgan-emb-top-map-triangle :
+ is-de-morgan-emb right →
+ is-de-morgan-emb left →
+ is-de-morgan-emb top
+ is-de-morgan-emb-top-map-triangle R L =
+ is-de-morgan-emb-right-factor
+ ( is-de-morgan-emb-htpy (inv-htpy H) L)
+ ( R)
+```
+
+### Characterizing equality in the type of De Morgan embeddings
+
+```agda
+htpy-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪ᵈᵐ B) → UU (l1 ⊔ l2)
+htpy-de-morgan-emb f g =
+ map-de-morgan-emb f ~ map-de-morgan-emb g
+
+refl-htpy-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A ↪ᵈᵐ B) →
+ htpy-de-morgan-emb f f
+refl-htpy-de-morgan-emb f = refl-htpy
+
+htpy-eq-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪ᵈᵐ B) →
+ f = g → htpy-de-morgan-emb f g
+htpy-eq-de-morgan-emb f .f refl =
+ refl-htpy-de-morgan-emb f
+
+abstract
+ is-torsorial-htpy-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A ↪ᵈᵐ B) →
+ is-torsorial (htpy-de-morgan-emb f)
+ is-torsorial-htpy-de-morgan-emb f =
+ is-torsorial-Eq-subtype
+ ( is-torsorial-htpy (map-de-morgan-emb f))
+ ( is-prop-is-de-morgan-emb)
+ ( map-de-morgan-emb f)
+ ( refl-htpy)
+ ( is-de-morgan-emb-map-de-morgan-emb f)
+
+abstract
+ is-equiv-htpy-eq-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪ᵈᵐ B) →
+ is-equiv (htpy-eq-de-morgan-emb f g)
+ is-equiv-htpy-eq-de-morgan-emb f =
+ fundamental-theorem-id
+ ( is-torsorial-htpy-de-morgan-emb f)
+ ( htpy-eq-de-morgan-emb f)
+
+eq-htpy-de-morgan-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪ᵈᵐ B) →
+ htpy-de-morgan-emb f g → f = g
+eq-htpy-de-morgan-emb f g =
+ map-inv-is-equiv (is-equiv-htpy-eq-de-morgan-emb f g)
+```
+
+### Any map out of the empty type is a De Morgan embedding
+
+```agda
+abstract
+ is-de-morgan-emb-ex-falso :
+ {l : Level} {X : UU l} → is-de-morgan-emb (ex-falso {l} {X})
+ is-de-morgan-emb-ex-falso =
+ ( is-emb-ex-falso , is-de-morgan-map-ex-falso)
+
+de-morgan-emb-ex-falso :
+ {l : Level} {X : UU l} → empty ↪ᵈᵐ X
+de-morgan-emb-ex-falso =
+ ( ex-falso , is-de-morgan-emb-ex-falso)
+```
+
+### The map on total spaces induced by a family of De Morgan embeddings is a De Morgan embedding
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3}
+ where
+
+ is-de-morgan-emb-tot :
+ {f : (x : A) → B x → C x} →
+ ((x : A) → is-de-morgan-emb (f x)) →
+ is-de-morgan-emb (tot f)
+ is-de-morgan-emb-tot H =
+ ( is-emb-tot (is-emb-is-de-morgan-emb ∘ H) ,
+ is-de-morgan-map-tot
+ ( is-de-morgan-map-is-de-morgan-emb ∘ H))
+
+ de-morgan-emb-tot : ((x : A) → B x ↪ᵈᵐ C x) → Σ A B ↪ᵈᵐ Σ A C
+ 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))
+```
+
+### The map on total spaces induced by a De Morgan embedding on the base is a De Morgan embedding
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3)
+ where
+
+ is-de-morgan-emb-map-Σ-map-base :
+ {f : A → B} →
+ is-de-morgan-emb f →
+ is-de-morgan-emb (map-Σ-map-base f C)
+ is-de-morgan-emb-map-Σ-map-base H =
+ ( is-emb-map-Σ-map-base C (is-emb-is-de-morgan-emb H) ,
+ is-de-morgan-map-Σ-map-base C
+ ( is-de-morgan-map-is-de-morgan-emb H))
+
+ de-morgan-emb-map-Σ-map-base :
+ (f : A ↪ᵈᵐ B) → Σ A (C ∘ map-de-morgan-emb f) ↪ᵈᵐ Σ B C
+ 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))
+```
+
+### The functoriality of dependent pair types on De Morgan embeddings
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} (D : B → UU l4)
+ where
+
+ is-de-morgan-emb-map-Σ :
+ {f : A → B} {g : (x : A) → C x → D (f x)} →
+ is-decidable-emb f →
+ ((x : A) → is-de-morgan-emb (g x)) →
+ is-de-morgan-emb (map-Σ D f g)
+ is-de-morgan-emb-map-Σ {f} {g} F G =
+ is-de-morgan-emb-left-map-triangle-is-decidable-emb-top
+ ( triangle-map-Σ D f g)
+ ( is-de-morgan-emb-tot G)
+ ( is-decidable-emb-map-Σ-map-base D F)
+
+ de-morgan-emb-Σ :
+ (f : A ↪ᵈ B) →
+ ((x : A) → C x ↪ᵈᵐ D (map-decidable-emb f x)) →
+ Σ A C ↪ᵈᵐ Σ B D
+ 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)))
+```
+
+### Products of De Morgan embeddings are De Morgan embeddings
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-de-morgan-emb-map-product :
+ {f : A → B} {g : C → D} →
+ is-de-morgan-emb f →
+ is-de-morgan-emb g →
+ is-de-morgan-emb (map-product f g)
+ is-de-morgan-emb-map-product (eF , dF) (eG , dG) =
+ ( is-emb-map-product eF eG ,
+ is-de-morgan-map-product dF dG)
+
+ de-morgan-emb-product :
+ A ↪ᵈᵐ B → C ↪ᵈᵐ D → A × C ↪ᵈᵐ B × D
+ de-morgan-emb-product (f , F) (g , G) =
+ ( map-product f g , is-de-morgan-emb-map-product F G)
+```
+
+### Coproducts of De Morgan embeddings are De Morgan embeddings
+
+```agda
+module _
+ {l1 l2 l1' l2' : Level} {A : UU l1} {B : UU l2} {A' : UU l1'} {B' : UU l2'}
+ where
+
+abstract
+ is-de-morgan-emb-map-coproduct :
+ {l1 l2 l3 l4 : Level}
+ {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y} →
+ is-de-morgan-emb f →
+ is-de-morgan-emb g →
+ is-de-morgan-emb (map-coproduct f g)
+ is-de-morgan-emb-map-coproduct (eF , dF) (eG , dG) =
+ ( is-emb-map-coproduct eF eG ,
+ is-de-morgan-map-coproduct dF dG)
+```
+
+### De Morgan embeddings are closed under base change
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ {f : A → B} {g : C → D}
+ where
+
+ is-de-morgan-prop-map-base-change :
+ cartesian-hom-arrow g f →
+ is-de-morgan-prop-map f →
+ is-de-morgan-prop-map g
+ is-de-morgan-prop-map-base-change α F d =
+ is-de-morgan-prop-equiv
+ ( equiv-fibers-cartesian-hom-arrow g f α d)
+ ( F (map-codomain-cartesian-hom-arrow g f α d))
+
+ is-de-morgan-emb-base-change :
+ cartesian-hom-arrow g f →
+ is-de-morgan-emb f →
+ is-de-morgan-emb g
+ is-de-morgan-emb-base-change α F =
+ is-de-morgan-emb-is-de-morgan-prop-map
+ ( is-de-morgan-prop-map-base-change α
+ ( is-de-morgan-prop-map-is-de-morgan-emb F))
+```
+
+### De Morgan embeddings are closed under retracts of maps
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y}
+ where
+
+ is-de-morgan-prop-map-retract-map :
+ f retract-of-map g →
+ is-de-morgan-prop-map g →
+ is-de-morgan-prop-map f
+ is-de-morgan-prop-map-retract-map R G x =
+ is-de-morgan-prop-retract-of
+ ( retract-fiber-retract-map f g R x)
+ ( G (map-codomain-inclusion-retract-map f g R x))
+
+ is-de-morgan-emb-retract-map :
+ f retract-of-map g →
+ is-de-morgan-emb g →
+ is-de-morgan-emb f
+ is-de-morgan-emb-retract-map R G =
+ is-de-morgan-emb-is-de-morgan-prop-map
+ ( is-de-morgan-prop-map-retract-map R
+ ( is-de-morgan-prop-map-is-de-morgan-emb G))
+```
+
+### A type is a De Morgan proposition if and only if its terminal map is a De Morgan embedding
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ is-de-morgan-prop-is-de-morgan-emb-terminal-map :
+ is-de-morgan-emb (terminal-map A) →
+ is-de-morgan-prop A
+ is-de-morgan-prop-is-de-morgan-emb-terminal-map H =
+ is-de-morgan-prop-equiv'
+ ( equiv-fiber-terminal-map star)
+ ( is-de-morgan-prop-map-is-de-morgan-emb H star)
+
+ is-de-morgan-emb-terminal-map-is-de-morgan-prop :
+ is-de-morgan-prop A →
+ is-de-morgan-emb (terminal-map A)
+ is-de-morgan-emb-terminal-map-is-de-morgan-prop H =
+ is-de-morgan-emb-is-de-morgan-prop-map
+ ( λ y →
+ is-de-morgan-prop-equiv (equiv-fiber-terminal-map y) H)
+```
+
+### If a dependent sum of propositions over a proposition is De Morgan, then the family is a family of De Morgan propositions
+
+```agda
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : type-Prop P → Prop l2)
+ where
+
+ is-de-morgan-prop-family-is-de-morgan-Σ :
+ is-de-morgan (Σ (type-Prop P) (type-Prop ∘ Q)) →
+ (p : type-Prop P) → is-de-morgan (type-Prop (Q p))
+ is-de-morgan-prop-family-is-de-morgan-Σ H p =
+ is-de-morgan-equiv'
+ ( equiv-fiber-pr1 (type-Prop ∘ Q) p)
+ ( is-de-morgan-map-is-de-morgan-emb
+ ( is-de-morgan-emb-right-factor'
+ ( is-de-morgan-emb-terminal-map-is-de-morgan-prop
+ ( is-prop-Σ (is-prop-type-Prop P) (is-prop-type-Prop ∘ Q) , H))
+ ( is-emb-terminal-map-is-prop (is-prop-type-Prop P)))
+ ( p))
+```
diff --git a/src/logic/de-morgan-maps.lagda.md b/src/logic/de-morgan-maps.lagda.md
new file mode 100644
index 0000000000..17b5a68d6e
--- /dev/null
+++ b/src/logic/de-morgan-maps.lagda.md
@@ -0,0 +1,339 @@
+# De Morgan maps
+
+```agda
+module logic.de-morgan-maps where
+```
+
+Imports
+
+```agda
+open import elementary-number-theory.natural-numbers
+
+open import foundation.action-on-identifications-functions
+open import foundation.cartesian-morphisms-arrows
+open import foundation.coproduct-types
+open import foundation.decidable-equality
+open import foundation.decidable-maps
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation
+open import foundation.embeddings
+open import foundation.empty-types
+open import foundation.existential-quantification
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-coproduct-types
+open import foundation.identity-types
+open import foundation.injective-maps
+open import foundation.negation
+open import foundation.propositions
+open import foundation.retractions
+open import foundation.retracts-of-maps
+open import foundation.retracts-of-types
+open import foundation.transport-along-identifications
+open import foundation.unit-type
+open import foundation.universal-property-equivalences
+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 logic.de-morgan-types
+open import logic.de-morgans-law
+open import logic.double-negation-eliminating-maps
+open import logic.double-negation-elimination
+```
+
+
+
+## Idea
+
+A [map](foundation-core.function-types.md) is said to be
+{{#concept "De Morgan" Disambiguation="map of types" Agda=is-de-morgan-map}} if
+the [negation](foundation-core.negation.md) of its
+[fibers](foundation-core.fibers-of-maps.md) are
+[decidable](foundation.decidable-types.md). I.e., the map `f : A → B` is De
+Morgan if for every `y : B`, the fiber `fiber f y` is either
+[empty](foundation.empty-types.md) or not empty. This is equivalent to asking
+that the fibers satisfy [De Morgan's law](logic.de-morgans-law.md), but is a
+[small](foundation.small-types.md) condition.
+
+## Definintion
+
+### De Morgan maps
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-map : (A → B) → UU (l1 ⊔ l2)
+ is-de-morgan-map f = (y : B) → is-de-morgan (fiber f y)
+
+ is-prop-is-de-morgan-map : {f : A → B} → is-prop (is-de-morgan-map f)
+ is-prop-is-de-morgan-map {f} =
+ is-prop-Π (λ y → is-prop-is-de-morgan (fiber f y))
+
+ is-de-morgan-map-Prop : (A → B) → Prop (l1 ⊔ l2)
+ is-de-morgan-map-Prop f = is-de-morgan-map f , is-prop-is-de-morgan-map
+```
+
+### The type of De Morgan maps
+
+```agda
+infix 5 _→ᵈᵐ_
+
+_→ᵈᵐ_ : {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+A →ᵈᵐ B = Σ (A → B) (is-de-morgan-map)
+
+de-morgan-map : {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+de-morgan-map = _→ᵈᵐ_
+
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A →ᵈᵐ B)
+ where
+
+ map-de-morgan-map : A → B
+ map-de-morgan-map = pr1 f
+
+ is-de-morgan-de-morgan-map : is-de-morgan-map map-de-morgan-map
+ is-de-morgan-de-morgan-map = pr2 f
+```
+
+### Self-De-Morgan maps
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-self-de-morgan-map : (A → B) → UU (l1 ⊔ l2)
+ is-self-de-morgan-map f =
+ (y z : B) → de-morgans-law' (fiber f y) (fiber f z)
+```
+
+## Properties
+
+### De Morgan maps are closed under homotopy
+
+```agda
+abstract
+ is-de-morgan-map-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
+ f ~ g →
+ is-de-morgan-map g →
+ is-de-morgan-map f
+ is-de-morgan-map-htpy H K b =
+ is-decidable-equiv'
+ ( equiv-precomp (equiv-tot (λ a → equiv-concat (inv (H a)) b)) empty)
+ ( K b)
+```
+
+### Decidable maps are De Morgan
+
+```agda
+is-de-morgan-map-is-decidable-map :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-map f → is-de-morgan-map f
+is-de-morgan-map-is-decidable-map H y = is-de-morgan-is-decidable (H y)
+```
+
+### Double negation eliminating De Morgan maps are decidable
+
+```agda
+is-decidable-map-is-double-negation-eliminating-de-morgan-map :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-de-morgan-map f → is-double-negation-eliminating-map f → is-decidable-map f
+is-decidable-map-is-double-negation-eliminating-de-morgan-map H K y =
+ is-decidable-is-decidable-neg-has-double-negation-elim (K y) (H y)
+```
+
+### Left cancellation for De Morgan maps
+
+If a composite `g ∘ f` is De Morgan and the left factor `g` is injective, then
+the right factor `f` is De Morgan.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {f : A → B} {g : B → C}
+ where
+
+ is-de-morgan-map-right-factor' :
+ is-injective g →
+ is-de-morgan-map (g ∘ f) →
+ is-de-morgan-map f
+ is-de-morgan-map-right-factor' H GF y =
+ rec-coproduct
+ ( λ ngfy → inl (λ p → ngfy (pr1 p , ap g (pr2 p))))
+ ( λ nngfy → inr (λ nq → nngfy (λ p → nq (pr1 p , H (pr2 p)))))
+ ( GF (g y))
+```
+
+### Composition of De Morgan maps with decidable maps
+
+If `g` is a decidable injection and `f` is a De Morgan map, then `g ∘ f` is De
+Morgan.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {f : A → B} {g : B → C}
+ where
+
+ is-de-morgan-map-comp-is-decidable-map :
+ is-injective g →
+ is-decidable-map g →
+ is-de-morgan-map f →
+ is-de-morgan-map (g ∘ f)
+ is-de-morgan-map-comp-is-decidable-map H G F y =
+ rec-coproduct
+ ( λ u →
+ is-de-morgan-iff
+ ( λ v → (pr1 v) , ap g (pr2 v) ∙ pr2 u)
+ ( λ w → pr1 w , H (pr2 w ∙ inv (pr2 u)))
+ ( F (pr1 u)))
+ ( λ ng → inl (λ u → ng (f (pr1 u) , pr2 u)))
+ ( G y)
+```
+
+### Any map out of the empty type is De Morgan
+
+```agda
+abstract
+ is-de-morgan-map-ex-falso :
+ {l : Level} {X : UU l} → is-de-morgan-map (ex-falso {l} {X})
+ is-de-morgan-map-ex-falso =
+ is-de-morgan-map-is-decidable-map is-decidable-map-ex-falso
+```
+
+### The identity map is De Morgan
+
+```agda
+abstract
+ is-de-morgan-map-id :
+ {l : Level} {X : UU l} → is-de-morgan-map (id {l} {X})
+ is-de-morgan-map-id =
+ is-de-morgan-map-is-decidable-map is-decidable-map-id
+```
+
+### Equivalences are De Morgan maps
+
+```agda
+abstract
+ is-de-morgan-map-is-equiv :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-equiv f → is-de-morgan-map f
+ is-de-morgan-map-is-equiv H =
+ is-de-morgan-map-is-decidable-map (is-decidable-map-is-equiv H)
+```
+
+### The map on total spaces induced by a family of De Morgan maps is De Morgan
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3}
+ where
+
+ is-de-morgan-map-tot :
+ {f : (x : A) → B x → C x} →
+ ((x : A) → is-de-morgan-map (f x)) →
+ is-de-morgan-map (tot f)
+ is-de-morgan-map-tot {f} H x =
+ is-decidable-equiv (equiv-neg (compute-fiber-tot f x)) (H (pr1 x) (pr2 x))
+```
+
+### The map on total spaces induced by a De Morgan map on the base is De Morgan
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3)
+ where
+
+ is-de-morgan-map-Σ-map-base :
+ {f : A → B} →
+ is-de-morgan-map f →
+ is-de-morgan-map (map-Σ-map-base f C)
+ is-de-morgan-map-Σ-map-base {f} H x =
+ is-decidable-equiv'
+ ( equiv-neg (compute-fiber-map-Σ-map-base f C x))
+ ( H (pr1 x))
+```
+
+### Products of De Morgan maps are De Morgan
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-de-morgan-map-product :
+ {f : A → B} {g : C → D} →
+ is-de-morgan-map f →
+ is-de-morgan-map g →
+ is-de-morgan-map (map-product f g)
+ is-de-morgan-map-product {f} {g} F G y =
+ is-de-morgan-equiv
+ ( compute-fiber-map-product f g y)
+ ( is-de-morgan-product (F (pr1 y)) (G (pr2 y)))
+```
+
+### Coproducts of De Morgan maps are De Morgan
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-de-morgan-map-coproduct :
+ {f : A → B} {g : C → D} →
+ is-de-morgan-map f →
+ is-de-morgan-map g →
+ is-de-morgan-map (map-coproduct f g)
+ is-de-morgan-map-coproduct {f} {g} F G (inl x) =
+ is-decidable-equiv'
+ ( equiv-neg (compute-fiber-inl-map-coproduct f g x))
+ ( F x)
+ is-de-morgan-map-coproduct {f} {g} F G (inr x) =
+ is-decidable-equiv'
+ ( equiv-neg (compute-fiber-inr-map-coproduct f g x))
+ ( G x)
+```
+
+### De Morgan maps are closed under base change
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ {f : A → B} {g : C → D}
+ where
+
+ is-de-morgan-map-base-change :
+ cartesian-hom-arrow g f →
+ is-de-morgan-map f →
+ is-de-morgan-map g
+ is-de-morgan-map-base-change α F d =
+ is-decidable-equiv
+ ( equiv-neg (equiv-fibers-cartesian-hom-arrow g f α d))
+ ( F (map-codomain-cartesian-hom-arrow g f α d))
+```
+
+### De Morgan maps are closed under retracts of maps
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y}
+ where
+
+ is-de-morgan-retract-map :
+ f retract-of-map g →
+ is-de-morgan-map g →
+ is-de-morgan-map f
+ is-de-morgan-retract-map R G x =
+ is-decidable-iff
+ ( map-neg (inclusion-retract (retract-fiber-retract-map f g R x)))
+ ( map-neg (map-retraction-retract (retract-fiber-retract-map f g R x)))
+ ( G (map-codomain-inclusion-retract-map f g R x))
+```
diff --git a/src/logic/de-morgan-propositions.lagda.md b/src/logic/de-morgan-propositions.lagda.md
new file mode 100644
index 0000000000..5ca81ba117
--- /dev/null
+++ b/src/logic/de-morgan-propositions.lagda.md
@@ -0,0 +1,351 @@
+# De Morgan propositions
+
+```agda
+module logic.de-morgan-propositions where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.conjunction
+open import foundation.contractible-types
+open import foundation.coproduct-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.disjunction
+open import foundation.double-negation
+open import foundation.embeddings
+open import foundation.empty-types
+open import foundation.equivalences
+open import foundation.evaluation-functions
+open import foundation.function-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.identity-types
+open import foundation.irrefutable-propositions
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.propositional-extensionality
+open import foundation.propositional-truncations
+open import foundation.propositions
+open import foundation.retracts-of-types
+open import foundation.sets
+open import foundation.subtypes
+open import foundation.transport-along-identifications
+open import foundation.universe-levels
+
+open import foundation-core.decidable-propositions
+
+open import logic.de-morgan-types
+open import logic.de-morgans-law
+```
+
+
+
+## Idea
+
+In classical logic, i.e., logic where we assume
+[the law of excluded middle](foundation.law-of-excluded-middle.md), the _De
+Morgan laws_ refers to the pair of logical equivalences
+
+```text
+ ¬ (P ∨ Q) ⇔ (¬ P) ∧ (¬ Q)
+ ¬ (P ∧ Q) ⇔ (¬ P) ∨ (¬ Q).
+```
+
+Out of these in total four logical implications, all but one are validated in
+constructive mathematics. The odd one out is
+
+```text
+ ¬ (P ∧ Q) ⇒ (¬ P) ∨ (¬ Q).
+```
+
+Indeed, this would state that we could constructively deduce from a proof that
+not both of `P` and `Q` are true, which of `P` and `Q` that is false. This
+logical law is what we refer to as [De Morgan's Law](logic.de-morgans-law.md).
+If a proposition `P` is such that for every other proposition `Q`, the De Morgan
+implication
+
+```text
+ ¬ (P ∧ Q) ⇒ (¬ P) ∨ (¬ Q)
+```
+
+holds, we say `P` is {{#concept "De Morgan" Disambiguation="proposition"}}.
+
+Equivalently, a proposition is De Morgan if its negation is decidable. Since
+this is a [small](foundation.small-types.md) condition, it is frequently more
+convenient to use.
+
+## Definition
+
+### The predicate on propositions of being De Morgan
+
+```agda
+module _
+ {l : Level} (P : UU l)
+ where
+
+ is-de-morgan-prop : UU l
+ is-de-morgan-prop = is-prop P × is-de-morgan P
+
+ is-prop-is-de-morgan-prop : is-prop is-de-morgan-prop
+ is-prop-is-de-morgan-prop =
+ is-prop-product (is-prop-is-prop P) (is-prop-is-de-morgan P)
+
+ is-de-morgan-prop-Prop : Prop l
+ is-de-morgan-prop-Prop = is-de-morgan-prop , is-prop-is-de-morgan-prop
+
+module _
+ {l : Level} {P : UU l} (H : is-de-morgan-prop P)
+ where
+
+ is-prop-type-is-de-morgan-prop : is-prop P
+ is-prop-type-is-de-morgan-prop = pr1 H
+
+ is-de-morgan-type-is-de-morgan-prop : is-de-morgan P
+ is-de-morgan-type-is-de-morgan-prop = pr2 H
+```
+
+### The subuniverse of De Morgan propositions
+
+```agda
+De-Morgan-Prop : (l : Level) → UU (lsuc l)
+De-Morgan-Prop l = Σ (UU l) (is-de-morgan-prop)
+
+module _
+ {l : Level} (A : De-Morgan-Prop l)
+ where
+
+ type-De-Morgan-Prop : UU l
+ type-De-Morgan-Prop = pr1 A
+
+ is-de-morgan-prop-type-De-Morgan-Prop : is-de-morgan-prop type-De-Morgan-Prop
+ is-de-morgan-prop-type-De-Morgan-Prop = pr2 A
+
+ is-prop-type-De-Morgan-Prop : is-prop type-De-Morgan-Prop
+ is-prop-type-De-Morgan-Prop =
+ is-prop-type-is-de-morgan-prop is-de-morgan-prop-type-De-Morgan-Prop
+
+ is-de-morgan-type-De-Morgan-Prop : is-de-morgan type-De-Morgan-Prop
+ is-de-morgan-type-De-Morgan-Prop =
+ is-de-morgan-type-is-de-morgan-prop is-de-morgan-prop-type-De-Morgan-Prop
+
+ prop-De-Morgan-Prop : Prop l
+ prop-De-Morgan-Prop = type-De-Morgan-Prop , is-prop-type-De-Morgan-Prop
+
+ de-morgan-type-De-Morgan-Prop : De-Morgan-Type l
+ de-morgan-type-De-Morgan-Prop =
+ type-De-Morgan-Prop , is-de-morgan-type-De-Morgan-Prop
+```
+
+## Properties
+
+### The forgetful map from De Morgan propositions to propositions is an embedding
+
+```agda
+is-emb-prop-De-Morgan-Prop :
+ {l : Level} → is-emb (prop-De-Morgan-Prop {l})
+is-emb-prop-De-Morgan-Prop =
+ is-emb-tot
+ ( λ X →
+ is-emb-inclusion-subtype (λ _ → is-de-morgan X , is-prop-is-de-morgan X))
+
+emb-prop-De-Morgan-Prop :
+ {l : Level} → De-Morgan-Prop l ↪ Prop l
+emb-prop-De-Morgan-Prop =
+ ( prop-De-Morgan-Prop , is-emb-prop-De-Morgan-Prop)
+```
+
+### The subuniverse of De Morgan propositions is a set
+
+```agda
+is-set-De-Morgan-Prop : {l : Level} → is-set (De-Morgan-Prop l)
+is-set-De-Morgan-Prop {l} =
+ is-set-emb emb-prop-De-Morgan-Prop is-set-type-Prop
+
+set-De-Morgan-Prop : (l : Level) → Set (lsuc l)
+set-De-Morgan-Prop l = (De-Morgan-Prop l , is-set-De-Morgan-Prop)
+```
+
+### Extensionality of De Morgan propositions
+
+```agda
+module _
+ {l : Level} (P Q : De-Morgan-Prop l)
+ where
+
+ extensionality-De-Morgan-Prop :
+ (P = Q) ≃ (type-De-Morgan-Prop P ↔ type-De-Morgan-Prop Q)
+ extensionality-De-Morgan-Prop =
+ ( propositional-extensionality
+ ( prop-De-Morgan-Prop P)
+ ( prop-De-Morgan-Prop Q)) ∘e
+ ( equiv-ap-emb emb-prop-De-Morgan-Prop)
+
+ iff-eq-De-Morgan-Prop : P = Q → type-De-Morgan-Prop P ↔ type-De-Morgan-Prop Q
+ iff-eq-De-Morgan-Prop = map-equiv extensionality-De-Morgan-Prop
+
+ eq-iff-De-Morgan-Prop' : type-De-Morgan-Prop P ↔ type-De-Morgan-Prop Q → P = Q
+ eq-iff-De-Morgan-Prop' = map-inv-equiv extensionality-De-Morgan-Prop
+
+ eq-iff-De-Morgan-Prop :
+ (type-De-Morgan-Prop P → type-De-Morgan-Prop Q) →
+ (type-De-Morgan-Prop Q → type-De-Morgan-Prop P) →
+ P = Q
+ eq-iff-De-Morgan-Prop f g = eq-iff-De-Morgan-Prop' (f , g)
+```
+
+### De Morgan propositions are closed under retracts
+
+```agda
+is-de-morgan-prop-retract-of :
+ {l1 l2 : Level} {P : UU l1} {Q : UU l2} →
+ P retract-of Q → is-de-morgan-prop Q → is-de-morgan-prop P
+is-de-morgan-prop-retract-of R H =
+ is-prop-retract-of R (is-prop-type-is-de-morgan-prop H) ,
+ is-de-morgan-iff' (iff-retract' R) (is-de-morgan-type-is-de-morgan-prop H)
+```
+
+### De Morgan propositions are closed under equivalences
+
+```agda
+is-de-morgan-prop-equiv :
+ {l1 l2 : Level} {P : UU l1} {Q : UU l2} →
+ P ≃ Q → is-de-morgan-prop Q → is-de-morgan-prop P
+is-de-morgan-prop-equiv e = is-de-morgan-prop-retract-of (retract-equiv e)
+
+is-de-morgan-prop-equiv' :
+ {l1 l2 : Level} {P : UU l1} {Q : UU l2} →
+ Q ≃ P → is-de-morgan-prop Q → is-de-morgan-prop P
+is-de-morgan-prop-equiv' e = is-de-morgan-prop-retract-of (retract-inv-equiv e)
+```
+
+### Irrefutable propositions are De Morgan
+
+```agda
+is-de-morgan-prop-is-irrefutable-prop :
+ {l : Level} {P : UU l} → is-irrefutable-prop P → is-de-morgan-prop P
+is-de-morgan-prop-is-irrefutable-prop = tot (λ _ → is-de-morgan-is-irrefutable)
+```
+
+### Contractible types are De Morgan propositions
+
+```agda
+is-de-morgan-prop-is-contr :
+ {l : Level} {P : UU l} → is-contr P → is-de-morgan-prop P
+is-de-morgan-prop-is-contr H =
+ is-de-morgan-prop-is-irrefutable-prop (is-irrefutable-prop-is-contr H)
+```
+
+### Empty types are De Morgan propositions
+
+```agda
+is-de-morgan-prop-is-empty :
+ {l : Level} {P : UU l} → is-empty P → is-de-morgan-prop P
+is-de-morgan-prop-is-empty H = is-prop-is-empty H , is-de-morgan-is-empty H
+```
+
+### Dependent sums of De Morgan propositions over decidable propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ is-de-morgan-prop-Σ' :
+ is-decidable-prop A → ((x : A) → is-de-morgan (B x)) → is-de-morgan (Σ A B)
+ is-de-morgan-prop-Σ' (is-prop-A , inl a) b =
+ rec-coproduct
+ ( λ nb → inl λ ab → nb (tr B (eq-is-prop is-prop-A) (pr2 ab)))
+ ( λ x → inr (λ z → x (λ b → z (a , b))))
+ ( b a)
+ is-de-morgan-prop-Σ' (is-prop-A , inr na) b = inl (λ ab → na (pr1 ab))
+
+ is-de-morgan-prop-Σ :
+ is-decidable-prop A →
+ ((x : A) → is-de-morgan-prop (B x)) →
+ is-de-morgan-prop (Σ A B)
+ is-de-morgan-prop-Σ a b =
+ ( is-prop-Σ
+ ( is-prop-type-is-decidable-prop a)
+ ( is-prop-type-is-de-morgan-prop ∘ b)) ,
+ ( is-de-morgan-prop-Σ' a (is-de-morgan-type-is-de-morgan-prop ∘ b))
+```
+
+### The negation operation on decidable propositions
+
+```agda
+is-de-morgan-prop-neg :
+ {l1 : Level} {A : UU l1} → is-de-morgan A → is-de-morgan-prop (¬ A)
+is-de-morgan-prop-neg is-de-morgan-A =
+ ( is-prop-neg , is-de-morgan-neg is-de-morgan-A)
+
+neg-type-De-Morgan-Prop :
+ {l1 : Level} (A : UU l1) → is-de-morgan A → De-Morgan-Prop l1
+neg-type-De-Morgan-Prop A is-de-morgan-A =
+ ( ¬ A , is-de-morgan-prop-neg is-de-morgan-A)
+
+neg-De-Morgan-Prop :
+ {l1 : Level} → De-Morgan-Prop l1 → De-Morgan-Prop l1
+neg-De-Morgan-Prop P =
+ neg-type-De-Morgan-Prop
+ ( type-De-Morgan-Prop P)
+ ( is-de-morgan-type-De-Morgan-Prop P)
+
+type-neg-De-Morgan-Prop :
+ {l1 : Level} → De-Morgan-Prop l1 → UU l1
+type-neg-De-Morgan-Prop P = type-De-Morgan-Prop (neg-De-Morgan-Prop P)
+```
+
+### Propositional truncations of De Morgan types are De Morgan propositions
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ is-de-morgan-prop-trunc-Prop :
+ is-de-morgan A → is-de-morgan-prop (type-trunc-Prop A)
+ is-de-morgan-prop-trunc-Prop a =
+ ( is-prop-type-trunc-Prop , is-de-morgan-trunc a)
+```
+
+### Disjunctions of De Morgan types are De Morgan propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-prop-disjunction :
+ is-de-morgan A → is-de-morgan B → is-de-morgan-prop (disjunction-type A B)
+ is-de-morgan-prop-disjunction a b =
+ is-de-morgan-prop-trunc-Prop (is-de-morgan-coproduct a b)
+
+module _
+ {l1 l2 : Level} (P : De-Morgan-Prop l1) (Q : De-Morgan-Prop l2)
+ where
+
+ disjunction-De-Morgan-Prop : De-Morgan-Prop (l1 ⊔ l2)
+ disjunction-De-Morgan-Prop =
+ disjunction-type (type-De-Morgan-Prop P) (type-De-Morgan-Prop Q) ,
+ is-de-morgan-prop-disjunction
+ ( is-de-morgan-type-De-Morgan-Prop P)
+ ( is-de-morgan-type-De-Morgan-Prop Q)
+```
+
+### Negation has no fixed points on decidable propositions
+
+```agda
+abstract
+ no-fixed-points-neg-De-Morgan-Prop :
+ {l : Level} (P : De-Morgan-Prop l) →
+ ¬ (type-De-Morgan-Prop P ↔ ¬ (type-De-Morgan-Prop P))
+ no-fixed-points-neg-De-Morgan-Prop P =
+ no-fixed-points-neg (type-De-Morgan-Prop P)
+```
+
+## External links
+
+- [De Morgan laws, in constructive mathematics](https://ncatlab.org/nlab/show/De+Morgan+laws#in_constructive_mathematics)
+ at $n$Lab
diff --git a/src/logic/de-morgan-subtypes.lagda.md b/src/logic/de-morgan-subtypes.lagda.md
new file mode 100644
index 0000000000..8981009f85
--- /dev/null
+++ b/src/logic/de-morgan-subtypes.lagda.md
@@ -0,0 +1,334 @@
+# De Morgan subtypes
+
+```agda
+module logic.de-morgan-subtypes where
+```
+
+Imports
+
+```agda
+open import foundation.1-types
+open import foundation.coproduct-types
+open import foundation.dependent-pair-types
+open import foundation.equality-dependent-function-types
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.logical-equivalences
+open import foundation.propositional-maps
+open import foundation.sets
+open import foundation.structured-type-duality
+open import foundation.subtypes
+open import foundation.type-theoretic-principle-of-choice
+open import foundation.universe-levels
+
+open import foundation-core.embeddings
+open import foundation-core.equivalences
+open import foundation-core.fibers-of-maps
+open import foundation-core.function-types
+open import foundation-core.identity-types
+open import foundation-core.injective-maps
+open import foundation-core.propositions
+open import foundation-core.truncated-types
+open import foundation-core.truncation-levels
+
+open import logic.de-morgan-embeddings
+open import logic.de-morgan-maps
+open import logic.de-morgan-propositions
+open import logic.de-morgan-types
+```
+
+
+
+## Idea
+
+A
+{{#concept "De Morgan subtype" Disambiguation="of a type" Agda=is-de-morgan-subtype Agda=de-morgan-subtype}}
+of a type consists of a family of
+[De Morgan propositions](logic.de-morgan-propositions.md) over it.
+
+## Definitions
+
+### De Morgan subtypes
+
+```agda
+is-de-morgan-subtype-Prop :
+ {l1 l2 : Level} {A : UU l1} → subtype l2 A → Prop (l1 ⊔ l2)
+is-de-morgan-subtype-Prop {A = A} P =
+ Π-Prop A (λ a → is-de-morgan-Prop (type-Prop (P a)))
+
+is-de-morgan-subtype :
+ {l1 l2 : Level} {A : UU l1} → subtype l2 A → UU (l1 ⊔ l2)
+is-de-morgan-subtype P =
+ type-Prop (is-de-morgan-subtype-Prop P)
+
+is-prop-is-de-morgan-subtype :
+ {l1 l2 : Level} {A : UU l1} (P : subtype l2 A) →
+ is-prop (is-de-morgan-subtype P)
+is-prop-is-de-morgan-subtype P =
+ is-prop-type-Prop (is-de-morgan-subtype-Prop P)
+
+de-morgan-subtype :
+ {l1 : Level} (l : Level) (X : UU l1) → UU (l1 ⊔ lsuc l)
+de-morgan-subtype l X = X → De-Morgan-Prop l
+```
+
+### The underlying subtype of a De Morgan subtype
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : de-morgan-subtype l2 A)
+ where
+
+ subtype-de-morgan-subtype : subtype l2 A
+ subtype-de-morgan-subtype a =
+ prop-De-Morgan-Prop (P a)
+
+ is-de-morgan-de-morgan-subtype :
+ is-de-morgan-subtype subtype-de-morgan-subtype
+ is-de-morgan-de-morgan-subtype a =
+ is-de-morgan-type-De-Morgan-Prop (P a)
+
+ is-in-de-morgan-subtype : A → UU l2
+ is-in-de-morgan-subtype =
+ is-in-subtype subtype-de-morgan-subtype
+
+ is-prop-is-in-de-morgan-subtype :
+ (a : A) → is-prop (is-in-de-morgan-subtype a)
+ is-prop-is-in-de-morgan-subtype =
+ is-prop-is-in-subtype subtype-de-morgan-subtype
+```
+
+### The underlying type of a De Morgan subtype
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : de-morgan-subtype l2 A)
+ where
+
+ type-de-morgan-subtype : UU (l1 ⊔ l2)
+ type-de-morgan-subtype =
+ type-subtype (subtype-de-morgan-subtype P)
+
+ inclusion-de-morgan-subtype :
+ type-de-morgan-subtype → A
+ inclusion-de-morgan-subtype =
+ inclusion-subtype (subtype-de-morgan-subtype P)
+
+ is-emb-inclusion-de-morgan-subtype :
+ is-emb inclusion-de-morgan-subtype
+ is-emb-inclusion-de-morgan-subtype =
+ is-emb-inclusion-subtype (subtype-de-morgan-subtype P)
+
+ is-de-morgan-map-inclusion-de-morgan-subtype :
+ is-de-morgan-map inclusion-de-morgan-subtype
+ is-de-morgan-map-inclusion-de-morgan-subtype x =
+ is-de-morgan-equiv
+ ( equiv-fiber-pr1 (type-De-Morgan-Prop ∘ P) x)
+ ( is-de-morgan-type-De-Morgan-Prop (P x))
+
+ is-injective-inclusion-de-morgan-subtype :
+ is-injective inclusion-de-morgan-subtype
+ is-injective-inclusion-de-morgan-subtype =
+ is-injective-inclusion-subtype (subtype-de-morgan-subtype P)
+
+ emb-de-morgan-subtype : type-de-morgan-subtype ↪ A
+ emb-de-morgan-subtype =
+ emb-subtype (subtype-de-morgan-subtype P)
+
+ is-de-morgan-emb-inclusion-de-morgan-subtype :
+ is-de-morgan-emb inclusion-de-morgan-subtype
+ is-de-morgan-emb-inclusion-de-morgan-subtype =
+ ( is-emb-inclusion-de-morgan-subtype ,
+ is-de-morgan-map-inclusion-de-morgan-subtype)
+
+ de-morgan-emb-de-morgan-subtype :
+ type-de-morgan-subtype ↪ᵈᵐ A
+ de-morgan-emb-de-morgan-subtype =
+ ( inclusion-de-morgan-subtype ,
+ is-de-morgan-emb-inclusion-de-morgan-subtype)
+```
+
+### The De Morgan subtype associated to a De Morgan embedding
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X ↪ᵈᵐ Y)
+ where
+
+ de-morgan-subtype-de-morgan-emb :
+ de-morgan-subtype (l1 ⊔ l2) Y
+ pr1 (de-morgan-subtype-de-morgan-emb y) =
+ 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)
+ ( y)
+
+ compute-type-de-morgan-type-de-morgan-emb :
+ type-de-morgan-subtype
+ de-morgan-subtype-de-morgan-emb ≃
+ X
+ compute-type-de-morgan-type-de-morgan-emb =
+ equiv-total-fiber (map-de-morgan-emb f)
+
+ inv-compute-type-de-morgan-type-de-morgan-emb :
+ X ≃
+ type-de-morgan-subtype
+ de-morgan-subtype-de-morgan-emb
+ inv-compute-type-de-morgan-type-de-morgan-emb =
+ inv-equiv-total-fiber (map-de-morgan-emb f)
+```
+
+## Examples
+
+### The De Morgan subtypes of left and right elements in a coproduct type
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-is-left :
+ (x : A + B) → is-de-morgan (is-left x)
+ is-de-morgan-is-left (inl x) = is-de-morgan-unit
+ is-de-morgan-is-left (inr x) = is-de-morgan-empty
+
+ is-left-De-Morgan-Prop :
+ A + B → De-Morgan-Prop lzero
+ pr1 (is-left-De-Morgan-Prop x) = is-left x
+ pr1 (pr2 (is-left-De-Morgan-Prop x)) = is-prop-is-left x
+ pr2 (pr2 (is-left-De-Morgan-Prop x)) =
+ is-de-morgan-is-left x
+
+ is-de-morgan-is-right :
+ (x : A + B) → is-de-morgan (is-right x)
+ is-de-morgan-is-right (inl x) = is-de-morgan-empty
+ is-de-morgan-is-right (inr x) = is-de-morgan-unit
+
+ is-right-De-Morgan-Prop :
+ A + B → De-Morgan-Prop lzero
+ pr1 (is-right-De-Morgan-Prop x) = is-right x
+ pr1 (pr2 (is-right-De-Morgan-Prop x)) = is-prop-is-right x
+ pr2 (pr2 (is-right-De-Morgan-Prop x)) =
+ is-de-morgan-is-right x
+```
+
+## Properties
+
+### A De Morgan subtype of a `k+1`-truncated type is `k+1`-truncated
+
+```agda
+module _
+ {l1 l2 : Level} (k : 𝕋) {A : UU l1} (P : de-morgan-subtype l2 A)
+ where
+
+ abstract
+ is-trunc-type-de-morgan-subtype :
+ is-trunc (succ-𝕋 k) A → is-trunc (succ-𝕋 k)
+ (type-de-morgan-subtype P)
+ is-trunc-type-de-morgan-subtype =
+ is-trunc-type-subtype k (subtype-de-morgan-subtype P)
+
+module _
+ {l1 l2 : Level} {A : UU l1} (P : de-morgan-subtype l2 A)
+ where
+
+ abstract
+ is-prop-type-de-morgan-subtype :
+ is-prop A → is-prop (type-de-morgan-subtype P)
+ is-prop-type-de-morgan-subtype =
+ is-prop-type-subtype (subtype-de-morgan-subtype P)
+
+ abstract
+ is-set-type-de-morgan-subtype :
+ is-set A → is-set (type-de-morgan-subtype P)
+ is-set-type-de-morgan-subtype =
+ is-set-type-subtype (subtype-de-morgan-subtype P)
+
+ abstract
+ is-1-type-type-de-morgan-subtype :
+ is-1-type A → is-1-type (type-de-morgan-subtype P)
+ is-1-type-type-de-morgan-subtype =
+ is-1-type-type-subtype (subtype-de-morgan-subtype P)
+
+prop-de-morgan-subprop :
+ {l1 l2 : Level} (A : Prop l1)
+ (P : de-morgan-subtype l2 (type-Prop A)) →
+ Prop (l1 ⊔ l2)
+prop-de-morgan-subprop A P =
+ prop-subprop A (subtype-de-morgan-subtype P)
+
+set-de-morgan-subset :
+ {l1 l2 : Level} (A : Set l1)
+ (P : de-morgan-subtype l2 (type-Set A)) →
+ Set (l1 ⊔ l2)
+set-de-morgan-subset A P =
+ set-subset A (subtype-de-morgan-subtype P)
+```
+
+### The type of De Morgan subtypes of a type is a set
+
+```agda
+is-set-de-morgan-subtype :
+ {l1 l2 : Level} {X : UU l1} → is-set (de-morgan-subtype l2 X)
+is-set-de-morgan-subtype =
+ is-set-function-type is-set-De-Morgan-Prop
+```
+
+### Extensionality of the type of De Morgan subtypes
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : de-morgan-subtype l2 A)
+ where
+
+ has-same-elements-de-morgan-subtype :
+ {l3 : Level} → de-morgan-subtype l3 A → UU (l1 ⊔ l2 ⊔ l3)
+ has-same-elements-de-morgan-subtype Q =
+ has-same-elements-subtype
+ ( subtype-de-morgan-subtype P)
+ ( subtype-de-morgan-subtype Q)
+
+ extensionality-de-morgan-subtype :
+ (Q : de-morgan-subtype l2 A) →
+ (P = Q) ≃ has-same-elements-de-morgan-subtype Q
+ extensionality-de-morgan-subtype =
+ extensionality-Π P
+ ( λ x Q →
+ ( type-De-Morgan-Prop (P x)) ↔
+ ( type-De-Morgan-Prop Q))
+ ( λ x Q → extensionality-De-Morgan-Prop (P x) Q)
+
+ has-same-elements-eq-de-morgan-subtype :
+ (Q : de-morgan-subtype l2 A) →
+ (P = Q) → has-same-elements-de-morgan-subtype Q
+ has-same-elements-eq-de-morgan-subtype Q =
+ map-equiv (extensionality-de-morgan-subtype Q)
+
+ eq-has-same-elements-de-morgan-subtype :
+ (Q : de-morgan-subtype l2 A) →
+ has-same-elements-de-morgan-subtype Q → P = Q
+ eq-has-same-elements-de-morgan-subtype Q =
+ map-inv-equiv (extensionality-de-morgan-subtype Q)
+
+ refl-extensionality-de-morgan-subtype :
+ map-equiv (extensionality-de-morgan-subtype P) refl =
+ (λ x → id , id)
+ refl-extensionality-de-morgan-subtype = refl
+```
+
+### The type of De Morgan subtypes of `A` is equivalent to the type of all De Morgan embeddings into a type `A`
+
+```agda
+equiv-Fiber-De-Morgan-Prop :
+ (l : Level) {l1 : Level} (A : UU l1) →
+ Σ (UU (l1 ⊔ l)) (λ X → X ↪ᵈᵐ A) ≃ (de-morgan-subtype (l1 ⊔ l) A)
+equiv-Fiber-De-Morgan-Prop l A =
+ ( equiv-Fiber-structure l is-de-morgan-prop A) ∘e
+ ( equiv-tot
+ ( λ X →
+ equiv-tot
+ ( λ f →
+ ( inv-distributive-Π-Σ) ∘e
+ ( equiv-product-left (equiv-is-prop-map-is-emb f)))))
+```
diff --git a/src/logic/de-morgan-types.lagda.md b/src/logic/de-morgan-types.lagda.md
new file mode 100644
index 0000000000..a87cd18593
--- /dev/null
+++ b/src/logic/de-morgan-types.lagda.md
@@ -0,0 +1,466 @@
+# De Morgan types
+
+```agda
+module logic.de-morgan-types where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.conjunction
+open import foundation.contractible-types
+open import foundation.coproduct-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.disjunction
+open import foundation.double-negation
+open import foundation.empty-types
+open import foundation.evaluation-functions
+open import foundation.function-types
+open import foundation.identity-types
+open import foundation.irrefutable-propositions
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.precomposition-functions
+open import foundation.propositional-truncations
+open import foundation.retracts-of-types
+open import foundation.truncation-levels
+open import foundation.truncations
+open import foundation.unit-type
+open import foundation.universe-levels
+
+open import foundation-core.decidable-propositions
+open import foundation-core.equivalences
+open import foundation-core.propositions
+
+open import logic.de-morgans-law
+```
+
+
+
+## Idea
+
+In classical logic, i.e., logic where we assume
+[the law of excluded middle](foundation.law-of-excluded-middle.md), the _De
+Morgan laws_ refers to the pair of logical equivalences
+
+```text
+ ¬ (P ∨ Q) ⇔ (¬ P) ∧ (¬ Q)
+ ¬ (P ∧ Q) ⇔ (¬ P) ∨ (¬ Q).
+```
+
+Out of these in total four logical implications, all but one are validated in
+constructive mathematics. The odd one out is
+
+```text
+ ¬ (P ∧ Q) ⇒ (¬ P) ∨ (¬ Q).
+```
+
+Indeed, this would state that we could constructively deduce from a proof that
+not both of `P` and `Q` are true, which of `P` and `Q` that is false. This
+logical law is what we refer to as [De Morgan's Law](logic.de-morgans-law.md).
+If a type `P` is such that for every other type `Q`, _the De Morgan implication_
+
+```text
+ ¬ (P ∧ Q) ⇒ (¬ P) ∨ (¬ Q)
+```
+
+holds, we say `P` is
+{{#concept "De Morgan" Disambiguation="type" Agda=is-de-morgan Agda=De-Morgan-Type Agda=satisfies-de-morgans-law-type}}.
+
+Equivalently, a type is De Morgan [iff](foundation.logical-equivalences.md) its
+[negation](foundation-core.negation.md) is
+[decidable](foundation.decidable-types.md). Since this is a
+[small](foundation.small-types.md) condition, it is frequently more convenient
+to use and is what we take as the main definition.
+
+## Definition
+
+### The small condition of being a De Morgan type
+
+I.e., types whose negation is decidable.
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ is-de-morgan : UU l
+ is-de-morgan = is-decidable (¬ A)
+
+ is-prop-is-de-morgan : is-prop is-de-morgan
+ is-prop-is-de-morgan = is-prop-is-decidable is-prop-neg
+
+ is-de-morgan-Prop : Prop l
+ is-de-morgan-Prop = is-decidable-Prop (neg-type-Prop A)
+```
+
+### The subuniverse of De Morgan types
+
+We use the decidability of the negation condition to define the subuniverse of
+De Morgan types.
+
+```agda
+De-Morgan-Type : (l : Level) → UU (lsuc l)
+De-Morgan-Type l = Σ (UU l) (is-de-morgan)
+
+module _
+ {l : Level} (A : De-Morgan-Type l)
+ where
+
+ type-De-Morgan-Type : UU l
+ type-De-Morgan-Type = pr1 A
+
+ is-de-morgan-type-De-Morgan-Type : is-de-morgan type-De-Morgan-Type
+ is-de-morgan-type-De-Morgan-Type = pr2 A
+```
+
+### Types that satisfy De Morgan's law
+
+```agda
+satisfies-de-morgans-law-type-Level :
+ {l1 : Level} (l2 : Level) (A : UU l1) → UU (l1 ⊔ lsuc l2)
+satisfies-de-morgans-law-type-Level l2 A =
+ (B : UU l2) → ¬ (A × B) → disjunction-type (¬ A) (¬ B)
+
+satisfies-de-morgans-law-type : {l1 : Level} (A : UU l1) → UUω
+satisfies-de-morgans-law-type A =
+ {l2 : Level} (B : UU l2) → ¬ (A × B) → disjunction-type (¬ A) (¬ B)
+
+is-prop-satisfies-de-morgans-law-type-Level :
+ {l1 l2 : Level} {A : UU l1} →
+ is-prop (satisfies-de-morgans-law-type-Level l2 A)
+is-prop-satisfies-de-morgans-law-type-Level {A = A} =
+ is-prop-Π (λ B → is-prop-Π (λ p → is-prop-disjunction-type (¬ A) (¬ B)))
+
+satisfies-de-morgans-law-type-Prop :
+ {l1 : Level} (l2 : Level) (A : UU l1) → Prop (l1 ⊔ lsuc l2)
+satisfies-de-morgans-law-type-Prop l2 A =
+ ( satisfies-de-morgans-law-type-Level l2 A ,
+ is-prop-satisfies-de-morgans-law-type-Level)
+```
+
+```agda
+satisfies-de-morgans-law-type-Level' :
+ {l1 : Level} (l2 : Level) (A : UU l1) → UU (l1 ⊔ lsuc l2)
+satisfies-de-morgans-law-type-Level' l2 A =
+ (B : UU l2) → ¬ (B × A) → disjunction-type (¬ B) (¬ A)
+
+satisfies-de-morgans-law-type' : {l1 : Level} (A : UU l1) → UUω
+satisfies-de-morgans-law-type' A =
+ {l2 : Level} (B : UU l2) → ¬ (B × A) → disjunction-type (¬ B) (¬ A)
+
+is-prop-satisfies-de-morgans-law-type-Level' :
+ {l1 l2 : Level} {A : UU l1} →
+ is-prop (satisfies-de-morgans-law-type-Level' l2 A)
+is-prop-satisfies-de-morgans-law-type-Level' {A = A} =
+ is-prop-Π (λ B → is-prop-Π (λ p → is-prop-disjunction-type (¬ B) (¬ A)))
+
+satisfies-de-morgans-law-type-Prop' :
+ {l1 : Level} (l2 : Level) (A : UU l1) → Prop (l1 ⊔ lsuc l2)
+satisfies-de-morgans-law-type-Prop' l2 A =
+ ( satisfies-de-morgans-law-type-Level' l2 A ,
+ is-prop-satisfies-de-morgans-law-type-Level')
+```
+
+## Properties
+
+### If a type satisfies De Morgan's law then its negation is decidable
+
+Indeed, one need only check that `A` and `¬ A` satisfy De Morgan's law, as then
+the hypothesis of the implication
+
+```text
+ ¬ (A ∧ ¬ A) ⇒ ¬ A ∨ ¬¬ A
+```
+
+is true.
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ is-de-morgan-satisfies-de-morgans-law' :
+ ({l' : Level} (B : UU l') → ¬ (A × B) → ¬ A + ¬ B) → is-de-morgan A
+ is-de-morgan-satisfies-de-morgans-law' H = H (¬ A) (λ f → pr2 f (pr1 f))
+
+ is-merely-decidable-neg-satisfies-de-morgans-law :
+ satisfies-de-morgans-law-type A → is-merely-decidable (¬ A)
+ is-merely-decidable-neg-satisfies-de-morgans-law H =
+ H (¬ A) (λ f → pr2 f (pr1 f))
+
+ is-de-morgan-satisfies-de-morgans-law :
+ satisfies-de-morgans-law-type A → is-de-morgan A
+ is-de-morgan-satisfies-de-morgans-law H =
+ rec-trunc-Prop
+ ( is-decidable-Prop (neg-type-Prop A))
+ ( id)
+ ( H (¬ A) (λ f → pr2 f (pr1 f)))
+```
+
+### If the negation of a type is decidable then it satisfies De Morgan's law
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ satisfies-de-morgans-law-is-de-morgan-left :
+ is-de-morgan A → ¬ (A × B) → ¬ A + ¬ B
+ satisfies-de-morgans-law-is-de-morgan-left (inl na) f =
+ inl na
+ satisfies-de-morgans-law-is-de-morgan-left (inr nna) f =
+ inr (λ y → nna (λ x → f (x , y)))
+
+ satisfies-de-morgans-law-is-de-morgan-right :
+ is-de-morgan B → ¬ (A × B) → ¬ A + ¬ B
+ satisfies-de-morgans-law-is-de-morgan-right (inl nb) f =
+ inr nb
+ satisfies-de-morgans-law-is-de-morgan-right (inr nnb) f =
+ inl (λ x → nnb (λ y → f (x , y)))
+
+satisfies-de-morgans-law-is-de-morgan :
+ {l : Level} {A : UU l} → is-de-morgan A → satisfies-de-morgans-law-type A
+satisfies-de-morgans-law-is-de-morgan x B q =
+ unit-trunc-Prop (satisfies-de-morgans-law-is-de-morgan-left x q)
+
+satisfies-de-morgans-law-is-de-morgan' :
+ {l : Level} {A : UU l} → is-de-morgan A → satisfies-de-morgans-law-type' A
+satisfies-de-morgans-law-is-de-morgan' x B q =
+ unit-trunc-Prop (satisfies-de-morgans-law-is-de-morgan-right x q)
+
+module _
+ {l : Level} (A : De-Morgan-Type l)
+ where
+
+ satisfies-de-morgans-law-type-De-Morgan-Type :
+ satisfies-de-morgans-law-type (type-De-Morgan-Type A)
+ satisfies-de-morgans-law-type-De-Morgan-Type =
+ satisfies-de-morgans-law-is-de-morgan (is-de-morgan-type-De-Morgan-Type A)
+
+ satisfies-de-morgans-law-type-De-Morgan-Type' :
+ satisfies-de-morgans-law-type' (type-De-Morgan-Type A)
+ satisfies-de-morgans-law-type-De-Morgan-Type' =
+ satisfies-de-morgans-law-is-de-morgan' (is-de-morgan-type-De-Morgan-Type A)
+```
+
+### It is irrefutable that a type is De Morgan
+
+```agda
+is-irrefutable-is-de-morgan : {l : Level} {A : UU l} → ¬¬ (is-de-morgan A)
+is-irrefutable-is-de-morgan = is-irrefutable-is-decidable
+```
+
+### Decidable types are De Morgan
+
+```agda
+is-de-morgan-is-decidable :
+ {l : Level} {A : UU l} → is-decidable A → is-de-morgan A
+is-de-morgan-is-decidable (inl x) = inr (intro-double-negation x)
+is-de-morgan-is-decidable (inr x) = inl x
+
+satisfies-de-morgans-law-is-decidable :
+ {l : Level} {A : UU l} → is-decidable A → satisfies-de-morgans-law-type A
+satisfies-de-morgans-law-is-decidable H =
+ satisfies-de-morgans-law-is-de-morgan (is-de-morgan-is-decidable H)
+
+satisfies-de-morgans-law-is-decidable' :
+ {l : Level} {A : UU l} → is-decidable A → satisfies-de-morgans-law-type' A
+satisfies-de-morgans-law-is-decidable' H =
+ satisfies-de-morgans-law-is-de-morgan' (is-de-morgan-is-decidable H)
+```
+
+### Irrefutable types are De Morgan
+
+```agda
+is-de-morgan-is-irrefutable :
+ {l : Level} {A : UU l} → ¬¬ A → is-de-morgan A
+is-de-morgan-is-irrefutable = inr
+```
+
+### Contractible types are De Morgan
+
+```agda
+is-de-morgan-is-contr :
+ {l : Level} {A : UU l} → is-contr A → is-de-morgan A
+is-de-morgan-is-contr H =
+ is-de-morgan-is-irrefutable (intro-double-negation (center H))
+
+is-de-morgan-unit : is-de-morgan unit
+is-de-morgan-unit = is-de-morgan-is-contr is-contr-unit
+```
+
+### Empty types are De Morgan
+
+```agda
+is-de-morgan-is-empty : {l : Level} {A : UU l} → is-empty A → is-de-morgan A
+is-de-morgan-is-empty = inl
+
+is-de-morgan-empty : is-de-morgan empty
+is-de-morgan-empty = is-de-morgan-is-empty id
+```
+
+### De Morgan types are closed under logical equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-iff :
+ (A → B) → (B → A) → is-de-morgan A → is-de-morgan B
+ is-de-morgan-iff f g = is-decidable-iff (map-neg g) (map-neg f)
+
+ is-de-morgan-iff' :
+ (A ↔ B) → is-de-morgan A → is-de-morgan B
+ is-de-morgan-iff' (f , g) = is-de-morgan-iff f g
+
+ satisfies-de-morgans-law-iff :
+ (A → B) → (B → A) →
+ satisfies-de-morgans-law-type A → satisfies-de-morgans-law-type B
+ satisfies-de-morgans-law-iff f g a =
+ satisfies-de-morgans-law-is-de-morgan
+ ( is-de-morgan-iff f g (is-de-morgan-satisfies-de-morgans-law A a))
+
+ satisfies-de-morgans-law-iff' :
+ (A ↔ B) → satisfies-de-morgans-law-type A → satisfies-de-morgans-law-type B
+ satisfies-de-morgans-law-iff' (f , g) = satisfies-de-morgans-law-iff f g
+```
+
+### De Morgan types are closed under retracts
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (R : B retract-of A)
+ where
+
+ is-de-morgan-retract-of : is-de-morgan A → is-de-morgan B
+ is-de-morgan-retract-of = is-de-morgan-iff' (iff-retract' R)
+
+ is-de-morgan-retract-of' : is-de-morgan B → is-de-morgan A
+ is-de-morgan-retract-of' = is-de-morgan-iff' (iff-retract R)
+```
+
+### De Morgan types are closed under equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (e : A ≃ B)
+ where
+
+ is-de-morgan-equiv' : is-de-morgan A → is-de-morgan B
+ is-de-morgan-equiv' = is-de-morgan-iff' (iff-equiv e)
+
+ is-de-morgan-equiv : is-de-morgan B → is-de-morgan A
+ is-de-morgan-equiv = is-de-morgan-iff' (iff-equiv' e)
+```
+
+### Equivalent types have equivalent De Morgan predicates
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ iff-is-de-morgan : A ↔ B → is-de-morgan A ↔ is-de-morgan B
+ iff-is-de-morgan e = iff-is-decidable (iff-neg e)
+
+ equiv-iff-is-de-morgan : A ↔ B → is-de-morgan A ≃ is-de-morgan B
+ equiv-iff-is-de-morgan e = equiv-is-decidable (equiv-iff-neg e)
+
+ equiv-is-de-morgan : A ≃ B → is-de-morgan A ≃ is-de-morgan B
+ equiv-is-de-morgan e = equiv-iff-is-de-morgan (iff-equiv e)
+```
+
+### The truncation of a De Morgan type is De Morgan
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ is-de-morgan-trunc : {k : 𝕋} → is-de-morgan A → is-de-morgan (type-trunc k A)
+ is-de-morgan-trunc {neg-two-𝕋} a =
+ is-de-morgan-is-contr is-trunc-type-trunc
+ is-de-morgan-trunc {succ-𝕋 k} (inl na) =
+ inl (map-universal-property-trunc (empty-Truncated-Type k) na)
+ is-de-morgan-trunc {succ-𝕋 k} (inr nna) =
+ inr (λ nn|a| → nna (λ a → nn|a| (unit-trunc a)))
+```
+
+### If the truncation of a type is De Morgan then the type is De Morgan
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ equiv-is-de-morgan-trunc :
+ {k : 𝕋} → is-de-morgan (type-trunc (succ-𝕋 k) A) ≃ is-de-morgan A
+ equiv-is-de-morgan-trunc {k} =
+ equiv-is-decidable
+ ( map-neg unit-trunc , is-truncation-trunc (empty-Truncated-Type k))
+
+ is-de-morgan-is-de-morgan-trunc :
+ {k : 𝕋} → is-de-morgan (type-trunc (succ-𝕋 k) A) → is-de-morgan A
+ is-de-morgan-is-de-morgan-trunc = map-equiv equiv-is-de-morgan-trunc
+```
+
+### Products of De Morgan types are De Morgan
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-product : is-de-morgan A → is-de-morgan B → is-de-morgan (A × B)
+ is-de-morgan-product (inl na) b =
+ inl (λ ab → na (pr1 ab))
+ is-de-morgan-product (inr nna) (inl nb) =
+ inl (λ ab → nb (pr2 ab))
+ is-de-morgan-product (inr nna) (inr nnb) =
+ inr (is-irrefutable-product nna nnb)
+```
+
+### Coproducts of De Morgan types are De Morgan
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-de-morgan-coproduct :
+ is-de-morgan A → is-de-morgan B → is-de-morgan (A + B)
+ is-de-morgan-coproduct (inl na) (inl nb) =
+ inl (rec-coproduct na nb)
+ is-de-morgan-coproduct (inl na) (inr nnb) =
+ inr (λ nab → nnb (λ nb → nab (inr nb)))
+ is-de-morgan-coproduct (inr nna) _ =
+ inr (λ nab → nna (λ na → nab (inl na)))
+
+ is-de-morgan-disjunction :
+ is-de-morgan A → is-de-morgan B → is-de-morgan (disjunction-type A B)
+ is-de-morgan-disjunction a b = is-de-morgan-trunc (is-de-morgan-coproduct a b)
+```
+
+### The negation of a De Morgan type is De Morgan
+
+```agda
+is-de-morgan-neg : {l : Level} {A : UU l} → is-de-morgan A → is-de-morgan (¬ A)
+is-de-morgan-neg = is-decidable-neg
+```
+
+### The identity types of De Morgan types are not generally De Morgan
+
+Consider any type `A`, then its suspension `ΣA` is De Morgan since it is
+inhabited. However, its identity type `N = S` is equivalent to `A`, so cannot
+be De Morgan unless `A` is.
+
+> This remains to be formalized.
+
+## External links
+
+- [De Morgan laws, in constructive mathematics](https://ncatlab.org/nlab/show/De+Morgan+laws#in_constructive_mathematics)
+ at $n$Lab
diff --git a/src/logic/de-morgans-law.lagda.md b/src/logic/de-morgans-law.lagda.md
new file mode 100644
index 0000000000..b680aef549
--- /dev/null
+++ b/src/logic/de-morgans-law.lagda.md
@@ -0,0 +1,144 @@
+# De Morgan's law
+
+```agda
+module logic.de-morgans-law where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.conjunction
+open import foundation.coproduct-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.disjunction
+open import foundation.double-negation
+open import foundation.empty-types
+open import foundation.evaluation-functions
+open import foundation.function-types
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.universe-levels
+
+open import foundation-core.decidable-propositions
+open import foundation-core.propositions
+
+open import univalent-combinatorics.2-element-types
+```
+
+
+
+## Idea
+
+In classical logic, i.e., logic where we assume
+[the law of excluded middle](foundation.law-of-excluded-middle.md), the _De
+Morgan laws_ refers to the pair of logical equivalences
+
+```text
+ ¬ (P ∨ Q) ⇔ (¬ P) ∧ (¬ Q)
+ ¬ (P ∧ Q) ⇔ (¬ P) ∨ (¬ Q).
+```
+
+Out of these in total four logical implications, all but one are validated in
+constructive mathematics. The odd one out is
+
+```text
+ ¬ (P ∧ Q) ⇒ (¬ P) ∨ (¬ Q).
+```
+
+Indeed, this would state that we could constructively deduce from a proof that
+not both of `P` and `Q` are true, which of `P` and `Q` that is false. This
+logical law is what we refer to as
+{{#concept "De Morgan's Law" Agda=De-Morgans-Law}}.
+
+## Definition
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ de-morgans-law-Prop' : Prop (l1 ⊔ l2)
+ de-morgans-law-Prop' =
+ neg-type-Prop (A × B) ⇒ (neg-type-Prop A) ∨ (neg-type-Prop B)
+
+ de-morgans-law' : UU (l1 ⊔ l2)
+ de-morgans-law' = ¬ (A × B) → disjunction-type (¬ A) (¬ B)
+
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ de-morgans-law-Prop : Prop (l1 ⊔ l2)
+ de-morgans-law-Prop = ¬' (P ∧ Q) ⇒ (¬' P) ∨ (¬' Q)
+
+ de-morgans-law : UU (l1 ⊔ l2)
+ de-morgans-law = type-Prop de-morgans-law-Prop
+
+De-Morgans-Law-Level : (l1 l2 : Level) → UU (lsuc l1 ⊔ lsuc l2)
+De-Morgans-Law-Level l1 l2 =
+ (P : Prop l1) (Q : Prop l2) → de-morgans-law P Q
+
+prop-De-Morgans-Law-Level : (l1 l2 : Level) → Prop (lsuc l1 ⊔ lsuc l2)
+prop-De-Morgans-Law-Level l1 l2 =
+ Π-Prop
+ ( Prop l1)
+ ( λ P → Π-Prop (Prop l2) (λ Q → de-morgans-law-Prop P Q))
+
+De-Morgans-Law : UUω
+De-Morgans-Law = {l1 l2 : Level} → De-Morgans-Law-Level l1 l2
+```
+
+## Properties
+
+### The constructively valid De Morgan's laws
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ forward-implication-constructive-de-morgan : ¬ (A + B) → (¬ A) × (¬ B)
+ forward-implication-constructive-de-morgan z = (z ∘ inl) , (z ∘ inr)
+
+ backward-implication-constructive-de-morgan : (¬ A) × (¬ B) → ¬ (A + B)
+ backward-implication-constructive-de-morgan (na , nb) (inl x) = na x
+ backward-implication-constructive-de-morgan (na , nb) (inr y) = nb y
+
+ constructive-de-morgan : ¬ (A + B) ↔ (¬ A) × (¬ B)
+ constructive-de-morgan =
+ ( forward-implication-constructive-de-morgan ,
+ backward-implication-constructive-de-morgan)
+
+ backward-implication-de-morgan : ¬ A + ¬ B → ¬ (A × B)
+ backward-implication-de-morgan (inl na) (x , y) = na x
+ backward-implication-de-morgan (inr nb) (x , y) = nb y
+```
+
+### Given the hypothesis of De Morgan's law, the conclusion is irrefutable
+
+```agda
+double-negation-de-morgan :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} → ¬ (A × B) → ¬¬ (¬ A + ¬ B)
+double-negation-de-morgan f v =
+ v (inl (λ x → v (inr (λ y → f (x , y)))))
+```
+
+### De Morgan's law is irrefutable
+
+```agda
+is-irrefutable-de-morgans-law :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} → ¬¬ (¬ (A × B) → ¬ A + ¬ B)
+is-irrefutable-de-morgans-law u =
+ u (λ _ → inl (λ x → u (λ f → inr (λ y → f (x , y)))))
+```
+
+## See also
+
+- [De Morgan types](logic.de-morgan-types.md)
+
+## External links
+
+- [De Morgan laws, in constructive mathematics](https://ncatlab.org/nlab/show/De+Morgan+laws#in_constructive_mathematics)
+ at $n$Lab
diff --git a/src/logic/double-negation-eliminating-maps.lagda.md b/src/logic/double-negation-eliminating-maps.lagda.md
new file mode 100644
index 0000000000..836e3191e1
--- /dev/null
+++ b/src/logic/double-negation-eliminating-maps.lagda.md
@@ -0,0 +1,335 @@
+# Double negation eliminating maps
+
+```agda
+module logic.double-negation-eliminating-maps where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.cartesian-morphisms-arrows
+open import foundation.coproduct-types
+open import foundation.decidable-equality
+open import foundation.decidable-maps
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation
+open import foundation.empty-types
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-coproduct-types
+open import foundation.identity-types
+open import foundation.injective-maps
+open import foundation.retractions
+open import foundation.retracts-of-maps
+open import foundation.retracts-of-types
+open import foundation.transport-along-identifications
+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 logic.double-negation-elimination
+```
+
+
+
+## Idea
+
+A [map](foundation-core.function-types.md) is said to be
+{{#concept "double negation eliminating" Disambiguation="map of types" Agda=is-double-negation-eliminating-map}}
+if its [fibers](foundation-core.fibers-of-maps.md) satisfy
+[untruncated double negation elimination](logic.double-negation-elimination.md).
+I.e., for every `y : B`, if `fiber f y` is
+[irrefutable](foundation.irrefutable-propositions.md), then we do in fact have
+an element of the fiber `p : fiber f y`. In other words, double negation
+eliminating maps come [equipped](foundation.structure.md) with a map
+
+```text
+ (y : B) → ¬¬ (fiber f y) → fiber f y.
+```
+
+## Definintion
+
+### Double negation elimination structure on a map
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-double-negation-eliminating-map : (A → B) → UU (l1 ⊔ l2)
+ is-double-negation-eliminating-map f =
+ (y : B) → has-double-negation-elim (fiber f y)
+```
+
+### The type of double negation eliminating maps
+
+```agda
+infix 5 _→¬¬_
+
+_→¬¬_ : {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+A →¬¬ B = Σ (A → B) (is-double-negation-eliminating-map)
+
+double-negation-eliminating-map :
+ {l1 l2 : Level} (A : UU l1) (B : UU l2) → UU (l1 ⊔ l2)
+double-negation-eliminating-map = _→¬¬_
+
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A →¬¬ B)
+ where
+
+ 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 map-double-negation-eliminating-map
+ is-double-negation-eliminating-double-negation-eliminating-map = pr2 f
+```
+
+## Properties
+
+### Double negation eliminating maps are closed under homotopy
+
+```agda
+abstract
+ is-double-negation-eliminating-map-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
+ f ~ g →
+ is-double-negation-eliminating-map g →
+ is-double-negation-eliminating-map f
+ is-double-negation-eliminating-map-htpy H K b =
+ has-double-negation-elim-equiv
+ ( equiv-tot (λ a → equiv-concat (inv (H a)) b))
+ ( K b)
+```
+
+### Decidable maps are double negation eliminating
+
+```agda
+is-double-negation-eliminating-map-is-decidable-map :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-map f → is-double-negation-eliminating-map f
+is-double-negation-eliminating-map-is-decidable-map H y =
+ double-negation-elim-is-decidable (H y)
+```
+
+### Composition of double negation eliminating maps
+
+Given a composition `g ∘ f` of double negation eliminating maps where the left
+factor `g` is injective, then the composition is double negation eliminating.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {g : B → C} {f : A → B}
+ where
+
+ fiber-left-is-double-negation-eliminating-map-left :
+ is-double-negation-eliminating-map g →
+ (z : C) → ¬¬ (fiber (g ∘ f) z) → fiber g z
+ fiber-left-is-double-negation-eliminating-map-left G z nngfz =
+ G z (λ x → nngfz (λ w → x (f (pr1 w) , pr2 w)))
+
+ fiber-right-is-double-negation-eliminating-map-comp :
+ is-injective g →
+ (G : is-double-negation-eliminating-map g) →
+ is-double-negation-eliminating-map f →
+ (z : C) (nngfz : ¬¬ (fiber (g ∘ f) z)) →
+ fiber f (pr1 (fiber-left-is-double-negation-eliminating-map-left G z nngfz))
+ fiber-right-is-double-negation-eliminating-map-comp H G F z nngfz =
+ F ( pr1
+ ( fiber-left-is-double-negation-eliminating-map-left G z nngfz))
+ ( λ x →
+ nngfz
+ ( λ w →
+ x ( pr1 w ,
+ H ( pr2 w ∙
+ inv
+ ( pr2
+ ( fiber-left-is-double-negation-eliminating-map-left
+ G z nngfz))))))
+
+ is-double-negation-eliminating-map-comp :
+ is-injective g →
+ is-double-negation-eliminating-map g →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map (g ∘ f)
+ is-double-negation-eliminating-map-comp H G F z nngfz =
+ map-inv-compute-fiber-comp g f z
+ ( ( fiber-left-is-double-negation-eliminating-map-left G z nngfz) ,
+ ( fiber-right-is-double-negation-eliminating-map-comp H G F z nngfz))
+```
+
+### Left cancellation for double negation eliminating maps
+
+If a composite `g ∘ f` is double negation eliminating and the left factor `g` is
+injective, then the right factor `f` is double negation eliminating.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {f : A → B} {g : B → C}
+ (GF : is-double-negation-eliminating-map (g ∘ f))
+ where
+
+ fiber-comp-is-double-negation-eliminating-map-right-factor' :
+ (y : B) → ¬¬ (fiber f y) → Σ (fiber g (g y)) (λ t → fiber f (pr1 t))
+ fiber-comp-is-double-negation-eliminating-map-right-factor' y nnfy =
+ map-compute-fiber-comp g f (g y)
+ ( GF (g y) (λ ngfgy → nnfy λ x → ngfgy ((pr1 x) , ap g (pr2 x))))
+
+ is-double-negation-eliminating-map-right-factor' :
+ is-injective g → is-double-negation-eliminating-map f
+ is-double-negation-eliminating-map-right-factor' G y nnfy =
+ tr
+ ( fiber f)
+ ( G ( pr2
+ ( pr1
+ ( fiber-comp-is-double-negation-eliminating-map-right-factor'
+ ( y)
+ ( nnfy)))))
+ ( pr2
+ ( fiber-comp-is-double-negation-eliminating-map-right-factor' y nnfy))
+```
+
+### Any map out of the empty type is double negation eliminating
+
+```agda
+abstract
+ is-double-negation-eliminating-map-ex-falso :
+ {l : Level} {X : UU l} →
+ is-double-negation-eliminating-map (ex-falso {l} {X})
+ is-double-negation-eliminating-map-ex-falso x f = ex-falso (f λ ())
+```
+
+### The identity map is double negation eliminating
+
+```agda
+abstract
+ is-double-negation-eliminating-map-id :
+ {l : Level} {X : UU l} → is-double-negation-eliminating-map (id {l} {X})
+ is-double-negation-eliminating-map-id x y = (x , refl)
+```
+
+### Equivalences are double negation eliminating maps
+
+```agda
+abstract
+ is-double-negation-eliminating-map-is-equiv :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-equiv f → is-double-negation-eliminating-map f
+ is-double-negation-eliminating-map-is-equiv H x =
+ double-negation-elim-is-contr (is-contr-map-is-equiv H x)
+```
+
+### The map on total spaces induced by a family of double negation eliminating maps is double negation eliminating
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3}
+ where
+
+ is-double-negation-eliminating-map-tot :
+ {f : (x : A) → B x → C x} →
+ ((x : A) → is-double-negation-eliminating-map (f x)) →
+ is-double-negation-eliminating-map (tot f)
+ is-double-negation-eliminating-map-tot {f} H x =
+ has-double-negation-elim-equiv (compute-fiber-tot f x) (H (pr1 x) (pr2 x))
+```
+
+### The map on total spaces induced by a double negation eliminating map on the base is double negation eliminating
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3)
+ where
+
+ is-double-negation-eliminating-map-Σ-map-base :
+ {f : A → B} →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map (map-Σ-map-base f C)
+ is-double-negation-eliminating-map-Σ-map-base {f} H x =
+ has-double-negation-elim-equiv'
+ ( compute-fiber-map-Σ-map-base f C x)
+ ( H (pr1 x))
+```
+
+### Products of double negation eliminating maps are double negation eliminating
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-double-negation-eliminating-map-product :
+ {f : A → B} {g : C → D} →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map g →
+ is-double-negation-eliminating-map (map-product f g)
+ is-double-negation-eliminating-map-product {f} {g} F G x =
+ has-double-negation-elim-equiv
+ ( compute-fiber-map-product f g x)
+ ( double-negation-elim-product (F (pr1 x)) (G (pr2 x)))
+```
+
+### Coproducts of double negation eliminating maps are double negation eliminating
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-double-negation-eliminating-map-coproduct :
+ {f : A → B} {g : C → D} →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map g →
+ is-double-negation-eliminating-map (map-coproduct f g)
+ is-double-negation-eliminating-map-coproduct {f} {g} F G (inl x) =
+ has-double-negation-elim-equiv'
+ ( compute-fiber-inl-map-coproduct f g x)
+ ( F x)
+ is-double-negation-eliminating-map-coproduct {f} {g} F G (inr y) =
+ has-double-negation-elim-equiv'
+ ( compute-fiber-inr-map-coproduct f g y)
+ ( G y)
+```
+
+### Double negation eliminating maps are closed under base change
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ {f : A → B} {g : C → D}
+ where
+
+ is-double-negation-eliminating-map-base-change :
+ cartesian-hom-arrow g f →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map g
+ is-double-negation-eliminating-map-base-change α F d =
+ has-double-negation-elim-equiv
+ ( equiv-fibers-cartesian-hom-arrow g f α d)
+ ( F (map-codomain-cartesian-hom-arrow g f α d))
+```
+
+### Double negation eliminating maps are closed under retracts of maps
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y}
+ where
+
+ is-double-negation-eliminating-retract-map :
+ f retract-of-map g →
+ is-double-negation-eliminating-map g →
+ is-double-negation-eliminating-map f
+ is-double-negation-eliminating-retract-map R G x =
+ has-double-negation-elim-retract
+ ( retract-fiber-retract-map f g R x)
+ ( G (map-codomain-inclusion-retract-map f g R x))
+```
diff --git a/src/logic/double-negation-elimination.lagda.md b/src/logic/double-negation-elimination.lagda.md
new file mode 100644
index 0000000000..6fd70fe6e7
--- /dev/null
+++ b/src/logic/double-negation-elimination.lagda.md
@@ -0,0 +1,270 @@
+# Double negation elimination
+
+```agda
+module logic.double-negation-elimination where
+```
+
+Imports
+
+```agda
+open import foundation.cartesian-product-types
+open import foundation.coproduct-types
+open import foundation.decidable-propositions
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation
+open import foundation.empty-types
+open import foundation.evaluation-functions
+open import foundation.hilberts-epsilon-operators
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.retracts-of-types
+open import foundation.transport-along-identifications
+open import foundation.unit-type
+open import foundation.universe-levels
+
+open import foundation-core.contractible-types
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.propositions
+```
+
+
+
+## Idea
+
+We say a type `A` satisfies
+{{#concept "untruncated double negation elimination" Disambiguation="on a type" Agda=has-double-negation-elim}}
+if there is a map
+
+```text
+ ¬¬A → A.
+```
+
+## Definitions
+
+### Untruncated double negation elimination
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ has-double-negation-elim : UU l
+ has-double-negation-elim = ¬¬ A → A
+```
+
+## Properties
+
+### Having double negation elimination is a property for propositions
+
+```agda
+is-prop-has-double-negation-elim :
+ {l : Level} {A : UU l} → is-prop A → is-prop (has-double-negation-elim A)
+is-prop-has-double-negation-elim = is-prop-function-type
+```
+
+### Double negation elimination is preserved by logical equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ has-double-negation-elim-iff :
+ A ↔ B → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-iff e H =
+ backward-implication e ∘ H ∘ map-double-negation (forward-implication e)
+
+ has-double-negation-elim-iff' :
+ B ↔ A → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-iff' e = has-double-negation-elim-iff (inv-iff e)
+```
+
+### Double negation elimination is preserved by equivalences
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ has-double-negation-elim-equiv :
+ A ≃ B → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-equiv e =
+ has-double-negation-elim-iff (iff-equiv e)
+
+ has-double-negation-elim-equiv' :
+ B ≃ A → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-equiv' e =
+ has-double-negation-elim-iff' (iff-equiv e)
+```
+
+### Double negation elimination is preserved by retracts
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ has-double-negation-elim-retract :
+ A retract-of B → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-retract e =
+ has-double-negation-elim-iff (iff-retract e)
+
+ has-double-negation-elim-retract' :
+ B retract-of A → has-double-negation-elim B → has-double-negation-elim A
+ has-double-negation-elim-retract' e =
+ has-double-negation-elim-iff' (iff-retract e)
+```
+
+### If the negation of a type with double negation elimination is decidable, then the type is decidable
+
+```agda
+module _
+ {l1 : Level} {A : UU l1}
+ where
+
+ is-decidable-is-decidable-neg-has-double-negation-elim :
+ has-double-negation-elim A → is-decidable (¬ A) → is-decidable A
+ is-decidable-is-decidable-neg-has-double-negation-elim f (inl nx) =
+ inr nx
+ is-decidable-is-decidable-neg-has-double-negation-elim f (inr nnx) =
+ inl (f nnx)
+```
+
+**Remark.** It is an established fact that both the property of satisfying
+double negation elimination, and the property of having decidable negation, are
+strictly weaker conditions than being decidable. Therefore, this result
+demonstrates that they are independent too.
+
+### Double negation elimination for empty types
+
+```agda
+double-negation-elim-empty : has-double-negation-elim empty
+double-negation-elim-empty e = e id
+
+double-negation-elim-is-empty :
+ {l : Level} {A : UU l} → is-empty A → has-double-negation-elim A
+double-negation-elim-is-empty H q = ex-falso (q H)
+```
+
+### Double negation elimination for the unit type
+
+```agda
+double-negation-elim-unit : has-double-negation-elim unit
+double-negation-elim-unit _ = star
+```
+
+### Double negation elimination for contractible types
+
+```agda
+double-negation-elim-is-contr :
+ {l : Level} {A : UU l} → is-contr A → has-double-negation-elim A
+double-negation-elim-is-contr H _ = center H
+```
+
+### Double negation elimination for negations of types
+
+```agda
+double-negation-elim-neg :
+ {l : Level} (A : UU l) → has-double-negation-elim (¬ A)
+double-negation-elim-neg A f p = f (ev p)
+```
+
+### Double negation elimination for universal quantification over double negations
+
+```agda
+module _
+ {l1 l2 : Level} {P : UU l1} {Q : P → UU l2}
+ where
+
+ double-negation-elim-for-all-neg-neg :
+ has-double-negation-elim ((p : P) → ¬¬ (Q p))
+ double-negation-elim-for-all-neg-neg f p =
+ double-negation-elim-neg
+ ( ¬ (Q p))
+ ( map-double-negation (λ (g : (u : P) → ¬¬ (Q u)) → g p) f)
+
+ double-negation-elim-for-all :
+ ((p : P) → has-double-negation-elim (Q p)) →
+ has-double-negation-elim ((p : P) → Q p)
+ double-negation-elim-for-all H f p = H p (λ nq → f (λ g → nq (g p)))
+```
+
+### Double negation elimination for function types into double negations
+
+```agda
+module _
+ {l1 l2 : Level} {P : UU l1} {Q : UU l2}
+ where
+
+ double-negation-elim-exp-neg-neg :
+ has-double-negation-elim (P → ¬¬ Q)
+ double-negation-elim-exp-neg-neg =
+ double-negation-elim-for-all-neg-neg
+
+ double-negation-elim-exp :
+ has-double-negation-elim Q →
+ has-double-negation-elim (P → Q)
+ double-negation-elim-exp q = double-negation-elim-for-all (λ _ → q)
+```
+
+### Double negation elimination for decidable propositions
+
+```text
+double-negation-elim-is-decidable :
+ {l : Level} {A : UU l} → is-decidable A → has-double-negation-elim A
+double-negation-elim-is-decidable = double-negation-elim-is-decidable
+```
+
+### Double negation elimination for dependent sums of types with double negation elimination over a double negation stable proposition
+
+```agda
+double-negation-elim-Σ-is-prop-base :
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2} →
+ is-prop A →
+ has-double-negation-elim A →
+ ((a : A) → has-double-negation-elim (B a)) →
+ has-double-negation-elim (Σ A B)
+double-negation-elim-Σ-is-prop-base {B = B} is-prop-A f g h =
+ ( f (λ na → h (na ∘ pr1))) ,
+ ( g ( f (λ na → h (na ∘ pr1)))
+ ( λ nb → h (λ y → nb (tr B (eq-is-prop is-prop-A) (pr2 y)))))
+
+double-negation-elim-Σ-is-decidable-prop-base :
+ {l1 l2 : Level} {P : UU l1} {B : P → UU l2} →
+ is-decidable-prop P →
+ ((x : P) → has-double-negation-elim (B x)) →
+ has-double-negation-elim (Σ P B)
+double-negation-elim-Σ-is-decidable-prop-base (H , d) =
+ double-negation-elim-Σ-is-prop-base H (double-negation-elim-is-decidable d)
+```
+
+### Double negation elimination for products of types with double negation elimination
+
+```agda
+double-negation-elim-product :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} →
+ has-double-negation-elim A →
+ has-double-negation-elim B →
+ has-double-negation-elim (A × B)
+double-negation-elim-product f g h =
+ f (λ na → h (na ∘ pr1)) , g (λ nb → h (nb ∘ pr2))
+```
+
+### If a type satisfies untruncated double negation elimination then it has a Hilbert ε-operator
+
+```agda
+ε-operator-Hilbert-has-double-negation-elim :
+ {l1 : Level} {A : UU l1} →
+ has-double-negation-elim A →
+ ε-operator-Hilbert A
+ε-operator-Hilbert-has-double-negation-elim {A = A} H =
+ H ∘ double-negation-double-negation-type-trunc-Prop A ∘ intro-double-negation
+```
+
+## See also
+
+- [The double negation modality](foundation.double-negation-modality.md)
+- [Irrefutable propositions](foundation.irrefutable-propositions.md) are double
+ negation connected types.
diff --git a/src/logic/double-negation-stable-embeddings.lagda.md b/src/logic/double-negation-stable-embeddings.lagda.md
new file mode 100644
index 0000000000..4e6a207085
--- /dev/null
+++ b/src/logic/double-negation-stable-embeddings.lagda.md
@@ -0,0 +1,731 @@
+# Double negation stable embeddings
+
+```agda
+module logic.double-negation-stable-embeddings where
+```
+
+Imports
+
+```agda
+open import foundation.action-on-identifications-functions
+open import foundation.cartesian-morphisms-arrows
+open import foundation.decidable-embeddings
+open import foundation.decidable-maps
+open import foundation.decidable-propositions
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation-stable-propositions
+open import foundation.embeddings
+open import foundation.fibers-of-maps
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-coproduct-types
+open import foundation.fundamental-theorem-of-identity-types
+open import foundation.homotopy-induction
+open import foundation.identity-types
+open import foundation.logical-equivalences
+open import foundation.negation
+open import foundation.propositional-maps
+open import foundation.propositions
+open import foundation.retracts-of-maps
+open import foundation.subtype-identity-principle
+open import foundation.type-arithmetic-dependent-pair-types
+open import foundation.unit-type
+open import foundation.universal-property-equivalences
+open import foundation.universe-levels
+
+open import foundation-core.cartesian-product-types
+open import foundation-core.coproduct-types
+open import foundation-core.empty-types
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.functoriality-dependent-pair-types
+open import foundation-core.homotopies
+open import foundation-core.injective-maps
+open import foundation-core.torsorial-type-families
+
+open import logic.double-negation-eliminating-maps
+open import logic.double-negation-elimination
+```
+
+
+
+## Idea
+
+A [map](foundation-core.function-types.md) is said to be a
+{{#concept "double negation stable embedding" Disambiguation="of types" Agda=is-double-negation-stable-emb}}
+if it is an [embedding](foundation-core.embeddings.md) and its
+[fibers](foundation-core.fibers-of-maps.md) satisfy
+[double negation elimination](logic.double-negation-elimination.md).
+
+Equivalently, a double negation stable embedding is a map whose fibers are
+[double negation stable propositions](foundation.double-negation-stable-propositions.md).
+We refer to this condition as being a
+{{#concept "double negation stable propositional map" Disambiguation="of types" Agda=is-double-negation-stable-prop-map}}.
+
+## Definitions
+
+### The condition on a map of being a double negation stable embedding
+
+```agda
+is-double-negation-stable-emb :
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} → (X → Y) → UU (l1 ⊔ l2)
+is-double-negation-stable-emb f =
+ is-emb f × is-double-negation-eliminating-map f
+
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y}
+ (F : is-double-negation-stable-emb f)
+ where
+
+ is-emb-is-double-negation-stable-emb : is-emb f
+ is-emb-is-double-negation-stable-emb = pr1 F
+
+ is-double-negation-eliminating-map-is-double-negation-stable-emb :
+ is-double-negation-eliminating-map f
+ is-double-negation-eliminating-map-is-double-negation-stable-emb = pr2 F
+
+ is-prop-map-is-double-negation-stable-emb : is-prop-map f
+ is-prop-map-is-double-negation-stable-emb =
+ is-prop-map-is-emb is-emb-is-double-negation-stable-emb
+
+ is-injective-is-double-negation-stable-emb : is-injective f
+ is-injective-is-double-negation-stable-emb =
+ is-injective-is-emb is-emb-is-double-negation-stable-emb
+```
+
+### Double negation stable propositional maps
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2}
+ where
+
+ is-double-negation-stable-prop-map : (X → Y) → UU (l1 ⊔ l2)
+ is-double-negation-stable-prop-map f =
+ (y : Y) → is-double-negation-stable-prop (fiber f y)
+
+ is-prop-is-double-negation-stable-prop-map :
+ (f : X → Y) → is-prop (is-double-negation-stable-prop-map f)
+ is-prop-is-double-negation-stable-prop-map f =
+ is-prop-Π (λ y → is-prop-is-double-negation-stable-prop (fiber f y))
+
+ is-double-negation-stable-prop-map-Prop : (X → Y) → Prop (l1 ⊔ l2)
+ is-double-negation-stable-prop-map-Prop f =
+ ( is-double-negation-stable-prop-map f ,
+ is-prop-is-double-negation-stable-prop-map f)
+
+ is-prop-map-is-double-negation-stable-prop-map :
+ {f : X → Y} → is-double-negation-stable-prop-map f → is-prop-map f
+ is-prop-map-is-double-negation-stable-prop-map H y = pr1 (H y)
+
+ is-double-negation-eliminating-map-is-double-negation-stable-prop-map :
+ {f : X → Y} →
+ is-double-negation-stable-prop-map f →
+ is-double-negation-eliminating-map f
+ is-double-negation-eliminating-map-is-double-negation-stable-prop-map H y =
+ pr2 (H y)
+```
+
+### The type of double negation stable embeddings
+
+```agda
+infix 5 _↪¬¬_
+_↪¬¬_ :
+ {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+X ↪¬¬ Y = Σ (X → Y) is-double-negation-stable-emb
+
+double-negation-stable-emb :
+ {l1 l2 : Level} (X : UU l1) (Y : UU l2) → UU (l1 ⊔ l2)
+double-negation-stable-emb = _↪¬¬_
+
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (e : X ↪¬¬ Y)
+ where
+
+ 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 map-double-negation-stable-emb
+ is-double-negation-stable-emb-map-double-negation-stable-emb = pr2 e
+
+ is-emb-map-double-negation-stable-emb :
+ is-emb map-double-negation-stable-emb
+ is-emb-map-double-negation-stable-emb =
+ is-emb-is-double-negation-stable-emb
+ is-double-negation-stable-emb-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-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
+
+ emb-double-negation-stable-emb : X ↪ Y
+ emb-double-negation-stable-emb =
+ map-double-negation-stable-emb , is-emb-map-double-negation-stable-emb
+```
+
+## Properties
+
+### Any map of which the fibers are double negation stable propositions is a double negation stable embedding
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} {f : X → Y}
+ where
+
+ abstract
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map :
+ is-double-negation-stable-prop-map f → is-double-negation-stable-emb f
+ pr1 (is-double-negation-stable-emb-is-double-negation-stable-prop-map H) =
+ is-emb-is-prop-map (is-prop-map-is-double-negation-stable-prop-map H)
+ pr2 (is-double-negation-stable-emb-is-double-negation-stable-prop-map H) =
+ is-double-negation-eliminating-map-is-double-negation-stable-prop-map H
+
+ abstract
+ is-double-negation-stable-prop-map-is-double-negation-stable-emb :
+ is-double-negation-stable-emb f → is-double-negation-stable-prop-map f
+ pr1 (is-double-negation-stable-prop-map-is-double-negation-stable-emb H y) =
+ is-prop-map-is-double-negation-stable-emb H y
+ pr2 (is-double-negation-stable-prop-map-is-double-negation-stable-emb H y) =
+ is-double-negation-eliminating-map-is-double-negation-stable-emb H y
+```
+
+### The first projection map of a dependent sum of double negation stable propositions is a double negation stable embedding
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (Q : A → Double-Negation-Stable-Prop l2)
+ where
+
+ is-double-negation-stable-prop-map-pr1 :
+ is-double-negation-stable-prop-map
+ ( pr1 {B = type-Double-Negation-Stable-Prop ∘ Q})
+ is-double-negation-stable-prop-map-pr1 y =
+ is-double-negation-stable-prop-equiv
+ ( equiv-fiber-pr1 (type-Double-Negation-Stable-Prop ∘ Q) y)
+ ( is-double-negation-stable-prop-type-Double-Negation-Stable-Prop (Q y))
+
+ is-double-negation-stable-emb-pr1 :
+ is-double-negation-stable-emb
+ ( pr1 {B = type-Double-Negation-Stable-Prop ∘ Q})
+ is-double-negation-stable-emb-pr1 =
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map
+ ( is-double-negation-stable-prop-map-pr1)
+```
+
+### Decidable embeddings are double negation stable
+
+```agda
+is-double-negation-eliminating-map-is-decidable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-emb f → is-double-negation-eliminating-map f
+is-double-negation-eliminating-map-is-decidable-emb H =
+ is-double-negation-eliminating-map-is-decidable-map
+ ( is-decidable-map-is-decidable-emb H)
+
+abstract
+ is-double-negation-stable-emb-is-decidable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-decidable-emb f → is-double-negation-stable-emb f
+ is-double-negation-stable-emb-is-decidable-emb H =
+ ( is-emb-is-decidable-emb H ,
+ is-double-negation-eliminating-map-is-decidable-emb H)
+```
+
+### Equivalences are double negation stable embeddings
+
+```agda
+abstract
+ is-double-negation-stable-emb-is-equiv :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
+ is-equiv f → is-double-negation-stable-emb f
+ is-double-negation-stable-emb-is-equiv H =
+ ( is-emb-is-equiv H , is-double-negation-eliminating-map-is-equiv H)
+```
+
+### Identity maps are double negation stable embeddings
+
+```agda
+is-double-negation-stable-emb-id :
+ {l : Level} {A : UU l} → is-double-negation-stable-emb (id {A = A})
+is-double-negation-stable-emb-id =
+ ( is-emb-id , is-double-negation-eliminating-map-id)
+
+double-negation-stable-emb-id : {l : Level} {A : UU l} → A ↪¬¬ A
+double-negation-stable-emb-id = (id , is-double-negation-stable-emb-id)
+
+is-double-negation-stable-prop-map-id :
+ {l : Level} {A : UU l} → is-double-negation-stable-prop-map (id {A = A})
+is-double-negation-stable-prop-map-id y =
+ is-double-negation-stable-prop-is-contr (is-torsorial-Id' y)
+```
+
+### Being a double negation stable embedding is a property
+
+```agda
+abstract
+ is-prop-is-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) →
+ is-prop (is-double-negation-stable-emb f)
+ is-prop-is-double-negation-stable-emb f =
+ is-prop-has-element
+ ( λ H →
+ is-prop-product
+ ( is-property-is-emb f)
+ ( is-prop-Π
+ ( is-prop-has-double-negation-elim ∘ is-prop-map-is-emb (pr1 H))))
+```
+
+### Double negation stable embeddings are closed under homotopies
+
+```agda
+abstract
+ is-double-negation-stable-emb-htpy :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f g : A → B} →
+ f ~ g → is-double-negation-stable-emb g → is-double-negation-stable-emb f
+ is-double-negation-stable-emb-htpy {f = f} {g} H K =
+ ( is-emb-htpy H (is-emb-is-double-negation-stable-emb K) ,
+ is-double-negation-eliminating-map-htpy H
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb K))
+```
+
+### Double negation stable embeddings are closed under composition
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {g : B → C} {f : A → B}
+ where
+
+ is-double-negation-eliminating-map-comp-is-double-negation-stable-emb :
+ is-double-negation-stable-emb g →
+ is-double-negation-eliminating-map f →
+ is-double-negation-eliminating-map (g ∘ f)
+ is-double-negation-eliminating-map-comp-is-double-negation-stable-emb G F =
+ is-double-negation-eliminating-map-comp
+ ( is-injective-is-double-negation-stable-emb G)
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb G)
+ ( F)
+
+ is-double-negation-stable-prop-map-comp :
+ is-double-negation-stable-prop-map g →
+ is-double-negation-stable-prop-map f →
+ is-double-negation-stable-prop-map (g ∘ f)
+ is-double-negation-stable-prop-map-comp K H z =
+ is-double-negation-stable-prop-equiv
+ ( compute-fiber-comp g f z)
+ ( is-double-negation-stable-prop-Σ (K z) (H ∘ pr1))
+
+ is-double-negation-stable-emb-comp :
+ is-double-negation-stable-emb g →
+ is-double-negation-stable-emb f →
+ is-double-negation-stable-emb (g ∘ f)
+ is-double-negation-stable-emb-comp K H =
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map
+ ( is-double-negation-stable-prop-map-comp
+ ( is-double-negation-stable-prop-map-is-double-negation-stable-emb K)
+ ( is-double-negation-stable-prop-map-is-double-negation-stable-emb H))
+
+comp-double-negation-stable-emb :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ B ↪¬¬ C → A ↪¬¬ B → A ↪¬¬ C
+comp-double-negation-stable-emb (g , G) (f , F) =
+ ( g ∘ f , is-double-negation-stable-emb-comp G F)
+
+infixr 15 _∘¬¬_
+_∘¬¬_ :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ B ↪¬¬ C → A ↪¬¬ B → A ↪¬¬ C
+_∘¬¬_ = comp-double-negation-stable-emb
+```
+
+### Left cancellation for double negation stable embeddings
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {f : A → B} {g : B → C}
+ where
+
+ is-double-negation-stable-emb-right-factor' :
+ is-double-negation-stable-emb (g ∘ f) →
+ is-emb g →
+ is-double-negation-stable-emb f
+ is-double-negation-stable-emb-right-factor' GH G =
+ ( is-emb-right-factor g f G (is-emb-is-double-negation-stable-emb GH) ,
+ is-double-negation-eliminating-map-right-factor'
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb GH)
+ ( is-injective-is-emb G))
+
+ is-double-negation-stable-emb-right-factor :
+ is-double-negation-stable-emb (g ∘ f) →
+ is-double-negation-stable-emb g →
+ is-double-negation-stable-emb f
+ is-double-negation-stable-emb-right-factor GH G =
+ is-double-negation-stable-emb-right-factor'
+ ( GH)
+ ( is-emb-is-double-negation-stable-emb G)
+```
+
+### In a commuting triangle of maps, if the top and right maps are double negation stable embeddings so is the left map
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {top : A → B} {left : A → C} {right : B → C}
+ (H : left ~ right ∘ top)
+ where
+
+ is-double-negation-stable-emb-left-map-triangle :
+ is-double-negation-stable-emb top →
+ is-double-negation-stable-emb right →
+ is-double-negation-stable-emb left
+ is-double-negation-stable-emb-left-map-triangle T R =
+ is-double-negation-stable-emb-htpy H
+ ( is-double-negation-stable-emb-comp R T)
+```
+
+### In a commuting triangle of maps, if the left and right maps are double negation stable embeddings so is the top map
+
+In fact, the right map need only be an embedding.
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ {top : A → B} {left : A → C} {right : B → C}
+ (H : left ~ right ∘ top)
+ where
+
+ is-double-negation-stable-emb-top-map-triangle' :
+ is-emb right →
+ is-double-negation-stable-emb left →
+ is-double-negation-stable-emb top
+ is-double-negation-stable-emb-top-map-triangle' R' L =
+ is-double-negation-stable-emb-right-factor'
+ ( is-double-negation-stable-emb-htpy (inv-htpy H) L)
+ ( R')
+
+ is-double-negation-stable-emb-top-map-triangle :
+ is-double-negation-stable-emb right →
+ is-double-negation-stable-emb left →
+ is-double-negation-stable-emb top
+ is-double-negation-stable-emb-top-map-triangle R L =
+ is-double-negation-stable-emb-right-factor
+ ( is-double-negation-stable-emb-htpy (inv-htpy H) L)
+ ( R)
+```
+
+### Characterizing equality in the type of double negation stable embeddings
+
+```agda
+htpy-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪¬¬ B) → UU (l1 ⊔ l2)
+htpy-double-negation-stable-emb f g =
+ map-double-negation-stable-emb f ~ map-double-negation-stable-emb g
+
+refl-htpy-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A ↪¬¬ B) →
+ htpy-double-negation-stable-emb f f
+refl-htpy-double-negation-stable-emb f = refl-htpy
+
+htpy-eq-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪¬¬ B) →
+ f = g → htpy-double-negation-stable-emb f g
+htpy-eq-double-negation-stable-emb f .f refl =
+ refl-htpy-double-negation-stable-emb f
+
+abstract
+ is-torsorial-htpy-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A ↪¬¬ B) →
+ is-torsorial (htpy-double-negation-stable-emb f)
+ is-torsorial-htpy-double-negation-stable-emb f =
+ is-torsorial-Eq-subtype
+ ( is-torsorial-htpy (map-double-negation-stable-emb f))
+ ( 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)
+
+abstract
+ is-equiv-htpy-eq-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪¬¬ B) →
+ is-equiv (htpy-eq-double-negation-stable-emb f g)
+ is-equiv-htpy-eq-double-negation-stable-emb f =
+ fundamental-theorem-id
+ ( is-torsorial-htpy-double-negation-stable-emb f)
+ ( htpy-eq-double-negation-stable-emb f)
+
+eq-htpy-double-negation-stable-emb :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} (f g : A ↪¬¬ B) →
+ htpy-double-negation-stable-emb f g → f = g
+eq-htpy-double-negation-stable-emb f g =
+ map-inv-is-equiv (is-equiv-htpy-eq-double-negation-stable-emb f g)
+```
+
+### Precomposing double negation stable embeddings with equivalences
+
+```agda
+equiv-precomp-double-negation-stable-emb-equiv :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (e : A ≃ B) →
+ (C : UU l3) → (B ↪¬¬ C) ≃ (A ↪¬¬ C)
+equiv-precomp-double-negation-stable-emb-equiv e C =
+ equiv-Σ
+ ( is-double-negation-stable-emb)
+ ( equiv-precomp e C)
+ ( λ g →
+ equiv-iff-is-prop
+ ( is-prop-is-double-negation-stable-emb g)
+ ( is-prop-is-double-negation-stable-emb (g ∘ map-equiv e))
+ ( λ H →
+ is-double-negation-stable-emb-comp H
+ ( is-double-negation-stable-emb-is-equiv (pr2 e)))
+ ( λ d →
+ is-double-negation-stable-emb-htpy
+ ( λ b → ap g (inv (is-section-map-inv-equiv e b)))
+ ( is-double-negation-stable-emb-comp
+ ( d)
+ ( is-double-negation-stable-emb-is-equiv
+ ( is-equiv-map-inv-equiv e)))))
+```
+
+### Any map out of the empty type is a double negation stable embedding
+
+```agda
+abstract
+ is-double-negation-stable-emb-ex-falso :
+ {l : Level} {X : UU l} → is-double-negation-stable-emb (ex-falso {l} {X})
+ is-double-negation-stable-emb-ex-falso =
+ ( is-emb-ex-falso , is-double-negation-eliminating-map-ex-falso)
+
+double-negation-stable-emb-ex-falso :
+ {l : Level} {X : UU l} → empty ↪¬¬ X
+double-negation-stable-emb-ex-falso =
+ ( ex-falso , is-double-negation-stable-emb-ex-falso)
+
+double-negation-stable-emb-is-empty :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} → is-empty A → A ↪¬¬ B
+double-negation-stable-emb-is-empty {A = A} f =
+ map-equiv
+ ( equiv-precomp-double-negation-stable-emb-equiv (equiv-is-empty f id) _)
+ ( double-negation-stable-emb-ex-falso)
+```
+
+### The map on total spaces induced by a family of double negation stable embeddings is a double negation stable embedding
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} {C : A → UU l3}
+ where
+
+ is-double-negation-stable-emb-tot :
+ {f : (x : A) → B x → C x} →
+ ((x : A) → is-double-negation-stable-emb (f x)) →
+ is-double-negation-stable-emb (tot f)
+ is-double-negation-stable-emb-tot H =
+ ( is-emb-tot (is-emb-is-double-negation-stable-emb ∘ H) ,
+ is-double-negation-eliminating-map-tot
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb ∘ H))
+
+ double-negation-stable-emb-tot : ((x : A) → B x ↪¬¬ C x) → Σ A B ↪¬¬ Σ A C
+ 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))
+```
+
+### The map on total spaces induced by a double negation stable embedding on the base is a double negation stable embedding
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (C : B → UU l3)
+ where
+
+ is-double-negation-stable-emb-map-Σ-map-base :
+ {f : A → B} →
+ is-double-negation-stable-emb f →
+ is-double-negation-stable-emb (map-Σ-map-base f C)
+ is-double-negation-stable-emb-map-Σ-map-base H =
+ ( is-emb-map-Σ-map-base C (is-emb-is-double-negation-stable-emb H) ,
+ is-double-negation-eliminating-map-Σ-map-base C
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb H))
+
+ double-negation-stable-emb-map-Σ-map-base :
+ (f : A ↪¬¬ B) → Σ A (C ∘ map-double-negation-stable-emb f) ↪¬¬ Σ B C
+ 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))
+```
+
+### The functoriality of dependent pair types preserves double negation stable embeddings
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : A → UU l3} (D : B → UU l4)
+ where
+
+ is-double-negation-stable-emb-map-Σ :
+ {f : A → B} {g : (x : A) → C x → D (f x)} →
+ is-double-negation-stable-emb f →
+ ((x : A) → is-double-negation-stable-emb (g x)) →
+ is-double-negation-stable-emb (map-Σ D f g)
+ is-double-negation-stable-emb-map-Σ {f} {g} F G =
+ is-double-negation-stable-emb-left-map-triangle
+ ( triangle-map-Σ D f g)
+ ( is-double-negation-stable-emb-tot G)
+ ( is-double-negation-stable-emb-map-Σ-map-base D F)
+
+ double-negation-stable-emb-Σ :
+ (f : A ↪¬¬ B) →
+ ((x : A) → C x ↪¬¬ D (map-double-negation-stable-emb f x)) →
+ Σ A C ↪¬¬ Σ B D
+ double-negation-stable-emb-Σ f g =
+ ( ( map-Σ D
+ ( 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)))
+```
+
+### Products of double negation stable embeddings are double negation stable embeddings
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ where
+
+ is-double-negation-stable-emb-map-product :
+ {f : A → B} {g : C → D} →
+ is-double-negation-stable-emb f →
+ is-double-negation-stable-emb g →
+ is-double-negation-stable-emb (map-product f g)
+ is-double-negation-stable-emb-map-product (eF , dF) (eG , dG) =
+ ( is-emb-map-product eF eG ,
+ is-double-negation-eliminating-map-product dF dG)
+
+ double-negation-stable-emb-product :
+ A ↪¬¬ B → C ↪¬¬ D → A × C ↪¬¬ B × D
+ double-negation-stable-emb-product (f , F) (g , G) =
+ ( map-product f g , is-double-negation-stable-emb-map-product F G)
+```
+
+### Coproducts of double negation stable embeddings are double negation stable embeddings
+
+```agda
+module _
+ {l1 l2 l1' l2' : Level} {A : UU l1} {B : UU l2} {A' : UU l1'} {B' : UU l2'}
+ where
+
+abstract
+ is-double-negation-stable-emb-map-coproduct :
+ {l1 l2 l3 l4 : Level}
+ {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y} →
+ is-double-negation-stable-emb f →
+ is-double-negation-stable-emb g →
+ is-double-negation-stable-emb (map-coproduct f g)
+ is-double-negation-stable-emb-map-coproduct (eF , dF) (eG , dG) =
+ ( is-emb-map-coproduct eF eG ,
+ is-double-negation-eliminating-map-coproduct dF dG)
+```
+
+### Double negation stable embeddings are closed under base change
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {C : UU l3} {D : UU l4}
+ {f : A → B} {g : C → D}
+ where
+
+ is-double-negation-stable-prop-map-base-change :
+ cartesian-hom-arrow g f →
+ is-double-negation-stable-prop-map f →
+ is-double-negation-stable-prop-map g
+ is-double-negation-stable-prop-map-base-change α F d =
+ is-double-negation-stable-prop-equiv
+ ( equiv-fibers-cartesian-hom-arrow g f α d)
+ ( F (map-codomain-cartesian-hom-arrow g f α d))
+
+ is-double-negation-stable-emb-base-change :
+ cartesian-hom-arrow g f →
+ is-double-negation-stable-emb f →
+ is-double-negation-stable-emb g
+ is-double-negation-stable-emb-base-change α F =
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map
+ ( is-double-negation-stable-prop-map-base-change α
+ ( is-double-negation-stable-prop-map-is-double-negation-stable-emb F))
+```
+
+### Double negation stable embeddings are closed under retracts of maps
+
+```agda
+module _
+ {l1 l2 l3 l4 : Level} {A : UU l1} {B : UU l2} {X : UU l3} {Y : UU l4}
+ {f : A → B} {g : X → Y}
+ where
+
+ is-double-negation-stable-prop-map-retract-map :
+ f retract-of-map g →
+ is-double-negation-stable-prop-map g →
+ is-double-negation-stable-prop-map f
+ is-double-negation-stable-prop-map-retract-map R G x =
+ is-double-negation-stable-prop-retract
+ ( retract-fiber-retract-map f g R x)
+ ( G (map-codomain-inclusion-retract-map f g R x))
+
+ is-double-negation-stable-emb-retract-map :
+ f retract-of-map g →
+ is-double-negation-stable-emb g →
+ is-double-negation-stable-emb f
+ is-double-negation-stable-emb-retract-map R G =
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map
+ ( is-double-negation-stable-prop-map-retract-map R
+ ( is-double-negation-stable-prop-map-is-double-negation-stable-emb G))
+```
+
+### A type is a double negation stable proposition if and only if its terminal map is a double negation stable embedding
+
+```agda
+module _
+ {l : Level} {A : UU l}
+ where
+
+ is-double-negation-stable-prop-is-double-negation-stable-emb-terminal-map :
+ is-double-negation-stable-emb (terminal-map A) →
+ is-double-negation-stable-prop A
+ is-double-negation-stable-prop-is-double-negation-stable-emb-terminal-map H =
+ is-double-negation-stable-prop-equiv'
+ ( equiv-fiber-terminal-map star)
+ ( is-double-negation-stable-prop-map-is-double-negation-stable-emb H star)
+
+ is-double-negation-stable-emb-terminal-map-is-double-negation-stable-prop :
+ is-double-negation-stable-prop A →
+ is-double-negation-stable-emb (terminal-map A)
+ is-double-negation-stable-emb-terminal-map-is-double-negation-stable-prop H =
+ is-double-negation-stable-emb-is-double-negation-stable-prop-map
+ ( λ y →
+ is-double-negation-stable-prop-equiv (equiv-fiber-terminal-map y) H)
+```
+
+### If a dependent sum of propositions over a proposition is double negation stable, then the family is a family of double negation stable propositions
+
+```agda
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : type-Prop P → Prop l2)
+ where
+
+ is-double-negation-stable-prop-family-has-double-negation-elim-Σ :
+ has-double-negation-elim (Σ (type-Prop P) (type-Prop ∘ Q)) →
+ (p : type-Prop P) → has-double-negation-elim (type-Prop (Q p))
+ is-double-negation-stable-prop-family-has-double-negation-elim-Σ H p =
+ has-double-negation-elim-equiv'
+ ( equiv-fiber-pr1 (type-Prop ∘ Q) p)
+ ( is-double-negation-eliminating-map-is-double-negation-stable-emb
+ ( is-double-negation-stable-emb-right-factor'
+ ( is-double-negation-stable-emb-terminal-map-is-double-negation-stable-prop
+ ( is-prop-Σ (is-prop-type-Prop P) (is-prop-type-Prop ∘ Q) , H))
+ ( is-emb-terminal-map-is-prop (is-prop-type-Prop P)))
+ ( p))
+```
diff --git a/src/logic/double-negation-stable-subtypes.lagda.md b/src/logic/double-negation-stable-subtypes.lagda.md
new file mode 100644
index 0000000000..6614510a1b
--- /dev/null
+++ b/src/logic/double-negation-stable-subtypes.lagda.md
@@ -0,0 +1,338 @@
+# Double negation stable subtypes
+
+```agda
+module logic.double-negation-stable-subtypes where
+```
+
+Imports
+
+```agda
+open import foundation.1-types
+open import foundation.coproduct-types
+open import foundation.dependent-pair-types
+open import foundation.double-negation-stable-propositions
+open import foundation.equality-dependent-function-types
+open import foundation.functoriality-cartesian-product-types
+open import foundation.functoriality-dependent-function-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.logical-equivalences
+open import foundation.propositional-maps
+open import foundation.sets
+open import foundation.structured-type-duality
+open import foundation.subtypes
+open import foundation.type-theoretic-principle-of-choice
+open import foundation.universe-levels
+
+open import foundation-core.embeddings
+open import foundation-core.equivalences
+open import foundation-core.fibers-of-maps
+open import foundation-core.function-types
+open import foundation-core.identity-types
+open import foundation-core.injective-maps
+open import foundation-core.propositions
+open import foundation-core.transport-along-identifications
+open import foundation-core.truncated-types
+open import foundation-core.truncation-levels
+
+open import logic.double-negation-eliminating-maps
+open import logic.double-negation-elimination
+open import logic.double-negation-stable-embeddings
+```
+
+
+
+## Idea
+
+A
+{{#concept "double negation stable subtype" Disambiguation="of a type" Agda=is-double-negation-stable-subtype Agda=double-negation-stable-subtype}}
+of a type consists of a family of
+[double negation stable propositions](foundation.double-negation-stable-propositions.md)
+over it.
+
+## Definitions
+
+### Decidable subtypes
+
+```agda
+is-double-negation-stable-subtype-Prop :
+ {l1 l2 : Level} {A : UU l1} → subtype l2 A → Prop (l1 ⊔ l2)
+is-double-negation-stable-subtype-Prop {A = A} P =
+ Π-Prop A (λ a → is-double-negation-stable-Prop (P a))
+
+is-double-negation-stable-subtype :
+ {l1 l2 : Level} {A : UU l1} → subtype l2 A → UU (l1 ⊔ l2)
+is-double-negation-stable-subtype P =
+ type-Prop (is-double-negation-stable-subtype-Prop P)
+
+is-prop-is-double-negation-stable-subtype :
+ {l1 l2 : Level} {A : UU l1} (P : subtype l2 A) →
+ is-prop (is-double-negation-stable-subtype P)
+is-prop-is-double-negation-stable-subtype P =
+ is-prop-type-Prop (is-double-negation-stable-subtype-Prop P)
+
+double-negation-stable-subtype :
+ {l1 : Level} (l : Level) (X : UU l1) → UU (l1 ⊔ lsuc l)
+double-negation-stable-subtype l X = X → Double-Negation-Stable-Prop l
+```
+
+### The underlying subtype of a double negation stable subtype
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : double-negation-stable-subtype l2 A)
+ where
+
+ subtype-double-negation-stable-subtype : subtype l2 A
+ subtype-double-negation-stable-subtype a =
+ prop-Double-Negation-Stable-Prop (P a)
+
+ is-double-negation-stable-double-negation-stable-subtype :
+ is-double-negation-stable-subtype subtype-double-negation-stable-subtype
+ is-double-negation-stable-double-negation-stable-subtype a =
+ has-double-negation-elim-type-Double-Negation-Stable-Prop (P a)
+
+ is-in-double-negation-stable-subtype : A → UU l2
+ is-in-double-negation-stable-subtype =
+ is-in-subtype subtype-double-negation-stable-subtype
+
+ is-prop-is-in-double-negation-stable-subtype :
+ (a : A) → is-prop (is-in-double-negation-stable-subtype a)
+ is-prop-is-in-double-negation-stable-subtype =
+ is-prop-is-in-subtype subtype-double-negation-stable-subtype
+```
+
+### The underlying type of a double negation stable subtype
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : double-negation-stable-subtype l2 A)
+ where
+
+ type-double-negation-stable-subtype : UU (l1 ⊔ l2)
+ type-double-negation-stable-subtype =
+ type-subtype (subtype-double-negation-stable-subtype P)
+
+ inclusion-double-negation-stable-subtype :
+ type-double-negation-stable-subtype → A
+ inclusion-double-negation-stable-subtype =
+ inclusion-subtype (subtype-double-negation-stable-subtype P)
+
+ is-emb-inclusion-double-negation-stable-subtype :
+ is-emb inclusion-double-negation-stable-subtype
+ is-emb-inclusion-double-negation-stable-subtype =
+ is-emb-inclusion-subtype (subtype-double-negation-stable-subtype P)
+
+ is-double-negation-eliminating-map-inclusion-double-negation-stable-subtype :
+ is-double-negation-eliminating-map inclusion-double-negation-stable-subtype
+ is-double-negation-eliminating-map-inclusion-double-negation-stable-subtype
+ x =
+ has-double-negation-elim-equiv
+ ( equiv-fiber-pr1 (type-Double-Negation-Stable-Prop ∘ P) x)
+ ( has-double-negation-elim-type-Double-Negation-Stable-Prop (P x))
+
+ is-injective-inclusion-double-negation-stable-subtype :
+ is-injective inclusion-double-negation-stable-subtype
+ is-injective-inclusion-double-negation-stable-subtype =
+ is-injective-inclusion-subtype (subtype-double-negation-stable-subtype P)
+
+ emb-double-negation-stable-subtype : type-double-negation-stable-subtype ↪ A
+ emb-double-negation-stable-subtype =
+ emb-subtype (subtype-double-negation-stable-subtype P)
+
+ is-double-negation-stable-emb-inclusion-double-negation-stable-subtype :
+ is-double-negation-stable-emb inclusion-double-negation-stable-subtype
+ is-double-negation-stable-emb-inclusion-double-negation-stable-subtype =
+ ( is-emb-inclusion-double-negation-stable-subtype ,
+ is-double-negation-eliminating-map-inclusion-double-negation-stable-subtype)
+
+ double-negation-stable-emb-double-negation-stable-subtype :
+ type-double-negation-stable-subtype ↪¬¬ A
+ double-negation-stable-emb-double-negation-stable-subtype =
+ ( inclusion-double-negation-stable-subtype ,
+ is-double-negation-stable-emb-inclusion-double-negation-stable-subtype)
+```
+
+### The double negation stable subtype associated to a double negation stable embedding
+
+```agda
+module _
+ {l1 l2 : Level} {X : UU l1} {Y : UU l2} (f : X ↪¬¬ Y)
+ where
+
+ double-negation-stable-subtype-double-negation-stable-emb :
+ double-negation-stable-subtype (l1 ⊔ l2) Y
+ pr1 (double-negation-stable-subtype-double-negation-stable-emb y) =
+ 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)
+ ( y)
+
+ compute-type-double-negation-stable-type-double-negation-stable-emb :
+ type-double-negation-stable-subtype
+ double-negation-stable-subtype-double-negation-stable-emb ≃
+ X
+ compute-type-double-negation-stable-type-double-negation-stable-emb =
+ equiv-total-fiber (map-double-negation-stable-emb f)
+
+ inv-compute-type-double-negation-stable-type-double-negation-stable-emb :
+ X ≃
+ type-double-negation-stable-subtype
+ double-negation-stable-subtype-double-negation-stable-emb
+ inv-compute-type-double-negation-stable-type-double-negation-stable-emb =
+ inv-equiv-total-fiber (map-double-negation-stable-emb f)
+```
+
+## Examples
+
+### The double negation stable subtypes of left and right elements in a coproduct type
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-double-negation-stable-is-left :
+ (x : A + B) → has-double-negation-elim (is-left x)
+ is-double-negation-stable-is-left (inl x) = double-negation-elim-unit
+ is-double-negation-stable-is-left (inr x) = double-negation-elim-empty
+
+ is-left-Double-Negation-Stable-Prop :
+ A + B → Double-Negation-Stable-Prop lzero
+ pr1 (is-left-Double-Negation-Stable-Prop x) = is-left x
+ pr1 (pr2 (is-left-Double-Negation-Stable-Prop x)) = is-prop-is-left x
+ pr2 (pr2 (is-left-Double-Negation-Stable-Prop x)) =
+ is-double-negation-stable-is-left x
+
+ is-double-negation-stable-is-right :
+ (x : A + B) → has-double-negation-elim (is-right x)
+ is-double-negation-stable-is-right (inl x) = double-negation-elim-empty
+ is-double-negation-stable-is-right (inr x) = double-negation-elim-unit
+
+ is-right-Double-Negation-Stable-Prop :
+ A + B → Double-Negation-Stable-Prop lzero
+ pr1 (is-right-Double-Negation-Stable-Prop x) = is-right x
+ pr1 (pr2 (is-right-Double-Negation-Stable-Prop x)) = is-prop-is-right x
+ pr2 (pr2 (is-right-Double-Negation-Stable-Prop x)) =
+ is-double-negation-stable-is-right x
+```
+
+## Properties
+
+### A double negation stable subtype of a `k+1`-truncated type is `k+1`-truncated
+
+```agda
+module _
+ {l1 l2 : Level} (k : 𝕋) {A : UU l1} (P : double-negation-stable-subtype l2 A)
+ where
+
+ abstract
+ is-trunc-type-double-negation-stable-subtype :
+ is-trunc (succ-𝕋 k) A → is-trunc (succ-𝕋 k)
+ (type-double-negation-stable-subtype P)
+ is-trunc-type-double-negation-stable-subtype =
+ is-trunc-type-subtype k (subtype-double-negation-stable-subtype P)
+
+module _
+ {l1 l2 : Level} {A : UU l1} (P : double-negation-stable-subtype l2 A)
+ where
+
+ abstract
+ is-prop-type-double-negation-stable-subtype :
+ is-prop A → is-prop (type-double-negation-stable-subtype P)
+ is-prop-type-double-negation-stable-subtype =
+ is-prop-type-subtype (subtype-double-negation-stable-subtype P)
+
+ abstract
+ is-set-type-double-negation-stable-subtype :
+ is-set A → is-set (type-double-negation-stable-subtype P)
+ is-set-type-double-negation-stable-subtype =
+ is-set-type-subtype (subtype-double-negation-stable-subtype P)
+
+ abstract
+ is-1-type-type-double-negation-stable-subtype :
+ is-1-type A → is-1-type (type-double-negation-stable-subtype P)
+ is-1-type-type-double-negation-stable-subtype =
+ is-1-type-type-subtype (subtype-double-negation-stable-subtype P)
+
+prop-double-negation-stable-subprop :
+ {l1 l2 : Level} (A : Prop l1)
+ (P : double-negation-stable-subtype l2 (type-Prop A)) →
+ Prop (l1 ⊔ l2)
+prop-double-negation-stable-subprop A P =
+ prop-subprop A (subtype-double-negation-stable-subtype P)
+
+set-double-negation-stable-subset :
+ {l1 l2 : Level} (A : Set l1)
+ (P : double-negation-stable-subtype l2 (type-Set A)) →
+ Set (l1 ⊔ l2)
+set-double-negation-stable-subset A P =
+ set-subset A (subtype-double-negation-stable-subtype P)
+```
+
+### The type of double negation stable subtypes of a type is a set
+
+```agda
+is-set-double-negation-stable-subtype :
+ {l1 l2 : Level} {X : UU l1} → is-set (double-negation-stable-subtype l2 X)
+is-set-double-negation-stable-subtype =
+ is-set-function-type is-set-Double-Negation-Stable-Prop
+```
+
+### Extensionality of the type of double negation stable subtypes
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} (P : double-negation-stable-subtype l2 A)
+ where
+
+ has-same-elements-double-negation-stable-subtype :
+ {l3 : Level} → double-negation-stable-subtype l3 A → UU (l1 ⊔ l2 ⊔ l3)
+ has-same-elements-double-negation-stable-subtype Q =
+ has-same-elements-subtype
+ ( subtype-double-negation-stable-subtype P)
+ ( subtype-double-negation-stable-subtype Q)
+
+ extensionality-double-negation-stable-subtype :
+ (Q : double-negation-stable-subtype l2 A) →
+ (P = Q) ≃ has-same-elements-double-negation-stable-subtype Q
+ extensionality-double-negation-stable-subtype =
+ extensionality-Π P
+ ( λ x Q →
+ ( type-Double-Negation-Stable-Prop (P x)) ↔
+ ( type-Double-Negation-Stable-Prop Q))
+ ( λ x Q → extensionality-Double-Negation-Stable-Prop (P x) Q)
+
+ has-same-elements-eq-double-negation-stable-subtype :
+ (Q : double-negation-stable-subtype l2 A) →
+ (P = Q) → has-same-elements-double-negation-stable-subtype Q
+ has-same-elements-eq-double-negation-stable-subtype Q =
+ map-equiv (extensionality-double-negation-stable-subtype Q)
+
+ eq-has-same-elements-double-negation-stable-subtype :
+ (Q : double-negation-stable-subtype l2 A) →
+ has-same-elements-double-negation-stable-subtype Q → P = Q
+ eq-has-same-elements-double-negation-stable-subtype Q =
+ map-inv-equiv (extensionality-double-negation-stable-subtype Q)
+
+ refl-extensionality-double-negation-stable-subtype :
+ map-equiv (extensionality-double-negation-stable-subtype P) refl =
+ (λ x → id , id)
+ refl-extensionality-double-negation-stable-subtype = refl
+```
+
+### The type of double negation stable subtypes of `A` is equivalent to the type of all double negation stable embeddings into a type `A`
+
+```agda
+equiv-Fiber-Double-Negation-Stable-Prop :
+ (l : Level) {l1 : Level} (A : UU l1) →
+ Σ (UU (l1 ⊔ l)) (λ X → X ↪¬¬ A) ≃ (double-negation-stable-subtype (l1 ⊔ l) A)
+equiv-Fiber-Double-Negation-Stable-Prop l A =
+ ( equiv-Fiber-structure l is-double-negation-stable-prop A) ∘e
+ ( equiv-tot
+ ( λ X →
+ equiv-tot
+ ( λ f →
+ ( inv-distributive-Π-Σ) ∘e
+ ( equiv-product-left (equiv-is-prop-map-is-emb f)))))
+```
diff --git a/src/logic/markovian-types.lagda.md b/src/logic/markovian-types.lagda.md
new file mode 100644
index 0000000000..54b779cddd
--- /dev/null
+++ b/src/logic/markovian-types.lagda.md
@@ -0,0 +1,101 @@
+# Markovian types
+
+```agda
+module logic.markovian-types where
+```
+
+Imports
+
+```agda
+open import elementary-number-theory.natural-numbers
+
+open import foundation.booleans
+open import foundation.decidable-subtypes
+open import foundation.dependent-pair-types
+open import foundation.disjunction
+open import foundation.existential-quantification
+open import foundation.function-types
+open import foundation.inhabited-types
+open import foundation.negation
+open import foundation.universal-quantification
+open import foundation.universe-levels
+
+open import foundation-core.identity-types
+open import foundation-core.propositions
+open import foundation-core.sets
+
+open import univalent-combinatorics.standard-finite-types
+```
+
+
+
+## Idea
+
+A type `A` is {{#concept "Markovian" Disambiguation="type" Agda=is-markovian}}
+if, for every [decidable subtype](foundation.decidable-subtypes.md) `𝒫` of `A`,
+if `𝒫` is [not](foundation-core.negation.md) [full](foundation.full-subtypes.md)
+then there is an element of `A` that is
+[not in](foundation.complements-subtypes.md) `𝒫`.
+
+## Definitions
+
+### The predicate of being Markovian
+
+We phrase the condition using the [type of booleans](foundation.booleans.md) so
+that the predicate is small.
+
+```agda
+is-markovian : {l : Level} → UU l → UU l
+is-markovian A =
+ (𝒫 : A → bool) →
+ ¬ ((x : A) → is-true (𝒫 x)) →
+ exists A (λ x → is-false-Prop (𝒫 x))
+
+is-prop-is-markovian : {l : Level} (A : UU l) → is-prop (is-markovian A)
+is-prop-is-markovian A =
+ is-prop-Π (λ 𝒫 → is-prop-function-type (is-prop-exists A (is-false-Prop ∘ 𝒫)))
+```
+
+### The predicate of being Markovian at a universe level
+
+```agda
+module _
+ {l1 : Level} (l2 : Level) (A : UU l1)
+ where
+
+ is-markovian-prop-Level : Prop (l1 ⊔ lsuc l2)
+ is-markovian-prop-Level =
+ Π-Prop
+ ( decidable-subtype l2 A)
+ ( λ P →
+ ¬' (∀' A (subtype-decidable-subtype P)) ⇒
+ ∃ A (¬'_ ∘ subtype-decidable-subtype P))
+
+ is-markovian-Level : UU (l1 ⊔ lsuc l2)
+ is-markovian-Level =
+ (P : decidable-subtype l2 A) →
+ ¬ ((x : A) → is-in-decidable-subtype P x) →
+ exists A (¬'_ ∘ subtype-decidable-subtype P)
+
+ is-prop-is-markovian-Level : is-prop is-markovian-Level
+ is-prop-is-markovian-Level = is-prop-type-Prop is-markovian-prop-Level
+```
+
+## Properties
+
+### A type is Markovian if and only if it is Markovian at any universe level
+
+> This remains to be formalized.
+
+### A type is Markovian if and only if it is Markovian at all universe levels
+
+> This remains to be formalized.
+
+## See also
+
+- [Markov's principle](logic.markovs-principle.md)
+
+## External links
+
+- [limited principle of omniscience](https://ncatlab.org/nlab/show/limited+principle+of+omniscience)
+ at $n$Lab
diff --git a/src/logic/markovs-principle.lagda.md b/src/logic/markovs-principle.lagda.md
new file mode 100644
index 0000000000..370789da8e
--- /dev/null
+++ b/src/logic/markovs-principle.lagda.md
@@ -0,0 +1,86 @@
+# Markov's principle
+
+```agda
+module logic.markovs-principle where
+```
+
+Imports
+
+```agda
+open import elementary-number-theory.natural-numbers
+
+open import foundation.booleans
+open import foundation.decidable-subtypes
+open import foundation.dependent-pair-types
+open import foundation.disjunction
+open import foundation.existential-quantification
+open import foundation.function-types
+open import foundation.inhabited-types
+open import foundation.negation
+open import foundation.universal-quantification
+open import foundation.universe-levels
+
+open import foundation-core.identity-types
+open import foundation-core.propositions
+open import foundation-core.sets
+
+open import logic.markovian-types
+
+open import univalent-combinatorics.standard-finite-types
+```
+
+
+
+## Idea
+
+{{#concept "Markov's principle" WDID=Q3922074 WD="Markov's principle" Agda=Markov's-Principle}}
+asserts that if a [decidable subtype](foundation.decidable-subtypes.md) `𝒫` of
+the [natural numbers](elementary-number-theory.natural-numbers.md) `ℕ` is not
+[full](foundation.full-subtypes.md), then
+[there is](foundation.existential-quantification.md) a natural number `n` that
+is not in `𝒫`.
+
+Markov's principle is an example of a _constructive taboo_. It is a consequence
+of the [law of excluded middle](foundation.law-of-excluded-middle.md) that is
+not provable generally in constructive mathematics.
+
+## Definitions
+
+### Markov's principle
+
+```agda
+Markov's-Principle : UU lzero
+Markov's-Principle = is-markovian ℕ
+```
+
+## Properties
+
+### Markov's principle is constructively valid for ascending chains of decidable propositions
+
+**Proof.** Assume given an ascending chain of decidable propositions `Pᵢ ⇒ Pᵢ₊₁`
+indexed by the natural numbers `ℕ`. This gives a decidable subtype `𝒫` of `ℕ`
+given by `i ∈ 𝒫` iff `Pᵢ` is true. Observe that if `i ∈ 𝒫` then every `j ≥ i` is
+also in `𝒫`, and there must exist a least such `i ∈ 𝒫`. Therefore,
+`𝒫 = Σ (m ∈ ℕ) (m ≥ k)` for some `k`. So, if `¬ (∀ᵢ Pᵢ)` it is necessarily the
+case that `¬ P₀`.
+
+```agda
+markovs-principle-ascending-chains :
+ {l : Level} (P : ℕ → UU l)
+ (H : (n : ℕ) → P n → P (succ-ℕ n)) → ¬ ((n : ℕ) → P n) → Σ ℕ (¬_ ∘ P)
+markovs-principle-ascending-chains P H q = (0 , λ x → q (ind-ℕ x H))
+```
+
+## See also
+
+- [The principle of omniscience](foundation.principle-of-omniscience.md)
+- [The limited principle of omniscience](foundation.limited-principle-of-omniscience.md)
+- [The lesser limited principle of omniscience](foundation.lesser-limited-principle-of-omniscience.md)
+- [The weak limited principle of omniscience](foundation.weak-limited-principle-of-omniscience.md)
+
+## External links
+
+- [Taboos.MarkovsPrinciple](https://martinescardo.github.io/TypeTopology/Taboos.MarkovsPrinciple.html)
+ at TypeTopology
+- [limited principle of omniscience](https://ncatlab.org/nlab/show/limited+principle+of+omniscience)
+ at $n$Lab
diff --git a/src/modal-type-theory/transport-along-crisp-identifications.lagda.md b/src/modal-type-theory/transport-along-crisp-identifications.lagda.md
index 5ca52cc131..3405430500 100644
--- a/src/modal-type-theory/transport-along-crisp-identifications.lagda.md
+++ b/src/modal-type-theory/transport-along-crisp-identifications.lagda.md
@@ -10,6 +10,7 @@ module modal-type-theory.transport-along-crisp-identifications where
```agda
open import foundation.action-on-identifications-functions
+open import foundation.function-types
open import foundation.identity-types
open import foundation.universe-levels
@@ -119,6 +120,17 @@ crisp-tr-ap {A = A} {B} {C} {D} f g {x} p z =
( p)
```
+### Substitution law for crisp transport
+
+```agda
+substitution-law-crisp-tr :
+ {@♭ l1 l2 l3 : Level} {@♭ X : UU l1} {@♭ A : UU l2} (@♭ B : @♭ A → UU l3)
+ (@♭ f : X → A)
+ {@♭ x y : X} (@♭ p : x = y) {@♭ x' : B (f x)} →
+ crisp-tr B (ap f p) x' = crisp-tr (λ (@♭ x : X) → B (f x)) p x'
+substitution-law-crisp-tr B f p {x'} = crisp-tr-ap f (λ _ → id) p x'
+```
+
### Computing maps out of crisp identity types as crisp transports
```agda
@@ -152,18 +164,3 @@ crisp-tr-Id-right :
crisp-tr-Id-right {a = a} q p =
crisp-based-ind-Id (λ y q → crisp-tr (a =_) q p = (p ∙ q)) (inv right-unit) q
```
-
-### Substitution law for crisp transport
-
-```agda
-substitution-law-crisp-tr :
- {@♭ l1 l2 l3 : Level} {@♭ X : UU l1} {@♭ A : UU l2} (@♭ B : @♭ A → UU l3)
- (@♭ f : X → A)
- {@♭ x y : X} (@♭ p : x = y) {@♭ x' : B (f x)} →
- crisp-tr B (ap f p) x' = crisp-tr (λ (@♭ x : X) → B (f x)) p x'
-substitution-law-crisp-tr {X = X} B f p {x'} =
- crisp-based-ind-Id
- ( λ y p → crisp-tr B (ap f p) x' = crisp-tr (λ (@♭ x : X) → B (f x)) p x')
- ( refl)
- ( p)
-```
diff --git a/src/species/small-cauchy-composition-species-of-types-in-subuniverses.lagda.md b/src/species/small-cauchy-composition-species-of-types-in-subuniverses.lagda.md
index b535b78e80..78eac4cbc1 100644
--- a/src/species/small-cauchy-composition-species-of-types-in-subuniverses.lagda.md
+++ b/src/species/small-cauchy-composition-species-of-types-in-subuniverses.lagda.md
@@ -193,12 +193,9 @@ module _
pr2 (equiv-Σ-extension-small-cauchy-composition-unit-subuniverse X) =
is-equiv-is-invertible
( λ S →
- ( tr
+ ( inv-tr
( is-in-subuniverse P)
- ( eq-equiv
- ( ( inv-equiv
- ( terminal-map X , is-equiv-terminal-map-is-contr S)) ∘e
- ( inv-equiv (compute-raise-unit l1))))
+ ( eq-equiv (equiv-raise-unit-is-contr S))
( C4) ,
map-equiv-is-small (C5 X) S))
( λ x → eq-is-prop is-property-is-contr)
diff --git a/src/synthetic-homotopy-theory.lagda.md b/src/synthetic-homotopy-theory.lagda.md
index 454cf9d9d3..ab877c2eea 100644
--- a/src/synthetic-homotopy-theory.lagda.md
+++ b/src/synthetic-homotopy-theory.lagda.md
@@ -120,6 +120,7 @@ open import synthetic-homotopy-theory.spheres public
open import synthetic-homotopy-theory.suspension-prespectra public
open import synthetic-homotopy-theory.suspension-structures public
open import synthetic-homotopy-theory.suspensions-of-pointed-types public
+open import synthetic-homotopy-theory.suspensions-of-propositions public
open import synthetic-homotopy-theory.suspensions-of-types public
open import synthetic-homotopy-theory.tangent-spheres public
open import synthetic-homotopy-theory.total-cocones-families-sequential-diagrams public
diff --git a/src/synthetic-homotopy-theory/codiagonals-of-maps.lagda.md b/src/synthetic-homotopy-theory/codiagonals-of-maps.lagda.md
index 20311a86ad..b483d3e9a8 100644
--- a/src/synthetic-homotopy-theory/codiagonals-of-maps.lagda.md
+++ b/src/synthetic-homotopy-theory/codiagonals-of-maps.lagda.md
@@ -103,12 +103,8 @@ module _
( fiber f b)
( unit)
( unit)
- ( inv-equiv
- ( terminal-map (fiber id b) ,
- ( is-equiv-terminal-map-is-contr (is-torsorial-Id' b))))
- ( inv-equiv
- ( terminal-map (fiber id b) ,
- ( is-equiv-terminal-map-is-contr (is-torsorial-Id' b))))
+ ( inv-equiv (equiv-unit-is-contr (is-torsorial-Id' b)))
+ ( inv-equiv (equiv-unit-is-contr (is-torsorial-Id' b)))
( id-equiv)
( terminal-map (fiber f b))
( terminal-map (fiber f b))
diff --git a/src/synthetic-homotopy-theory/suspension-structures.lagda.md b/src/synthetic-homotopy-theory/suspension-structures.lagda.md
index 761190971f..7d579d350f 100644
--- a/src/synthetic-homotopy-theory/suspension-structures.lagda.md
+++ b/src/synthetic-homotopy-theory/suspension-structures.lagda.md
@@ -86,7 +86,7 @@ module _
where
suspension-structure : UU (l1 ⊔ l2)
- suspension-structure = Σ Y (λ N → Σ Y (λ S → (x : X) → N = S))
+ suspension-structure = Σ Y (λ N → Σ Y (λ S → X → N = S))
module _
{l1 l2 : Level} {X : UU l1} {Y : UU l2}
diff --git a/src/synthetic-homotopy-theory/suspensions-of-propositions.lagda.md b/src/synthetic-homotopy-theory/suspensions-of-propositions.lagda.md
new file mode 100644
index 0000000000..0ad717c882
--- /dev/null
+++ b/src/synthetic-homotopy-theory/suspensions-of-propositions.lagda.md
@@ -0,0 +1,503 @@
+# Suspensions of propositions
+
+```agda
+module synthetic-homotopy-theory.suspensions-of-propositions where
+```
+
+Imports
+
+```agda
+open import foundation.booleans
+open import foundation.contractible-types
+open import foundation.dependent-pair-types
+open import foundation.equivalences
+open import foundation.existential-quantification
+open import foundation.function-extensionality
+open import foundation.function-types
+open import foundation.fundamental-theorem-of-identity-types
+open import foundation.homotopies
+open import foundation.identity-types
+open import foundation.propositions
+open import foundation.sets
+open import foundation.subsingleton-induction
+open import foundation.surjective-maps
+open import foundation.torsorial-type-families
+open import foundation.transport-along-identifications
+open import foundation.unit-type
+open import foundation.univalence
+open import foundation.universe-levels
+
+open import synthetic-homotopy-theory.dependent-suspension-structures
+open import synthetic-homotopy-theory.suspension-structures
+open import synthetic-homotopy-theory.suspensions-of-types
+
+open import univalent-combinatorics.kuratowski-finite-sets
+```
+
+
+
+## Idea
+
+The
+{{#concept "suspension" Disambiguation="of a proposition" Agda=suspension-Prop}}
+of a [proposition](foundation-core.propositions.md) `P` is the
+[set](foundation.sets.md) satisfying the
+[universal property of the pushout](synthetic-homotopy-theory.universal-property-pushouts.md)
+of the [span](foundation.spans.md)
+
+```text
+ 1 <--- P ---> 1.
+```
+
+## Definitions
+
+### The suspension of a proposition
+
+```agda
+module _
+ {l : Level} (P : Prop l)
+ where
+
+ suspension-Prop : UU l
+ suspension-Prop = suspension (type-Prop P)
+```
+
+### Suspensions of propositions are sets
+
+We characterize equality in the suspension of a proposition `P` by the following
+four equations
+
+```text
+ (N = N) ≃ 1
+ (N = S) ≃ P
+ (S = N) ≃ P
+ (S = S) ≃ 1.
+```
+
+Since these are all propositions, `ΣP` forms a set. This is Lemma 10.1.13 in
+{{#cite UF13}}.
+
+First we define the observational equality family on `ΣP`.
+
+```agda
+module _
+ {l : Level} (P : Prop l)
+ where
+
+ suspension-structure-Eq-north-suspension-Prop :
+ suspension-structure (type-Prop P) (UU l)
+ suspension-structure-Eq-north-suspension-Prop =
+ ( raise-unit l) ,
+ ( type-Prop P) ,
+ ( λ x →
+ inv
+ ( eq-equiv
+ ( equiv-raise-unit-is-contr (is-proof-irrelevant-type-Prop P x))))
+
+ Eq-north-suspension-Prop :
+ suspension (type-Prop P) → UU l
+ Eq-north-suspension-Prop =
+ cogap-suspension suspension-structure-Eq-north-suspension-Prop
+
+ suspension-structure-Eq-south-suspension-Prop :
+ suspension-structure (type-Prop P) (UU l)
+ suspension-structure-Eq-south-suspension-Prop =
+ ( type-Prop P) ,
+ ( raise-unit l) ,
+ ( λ x →
+ eq-equiv (equiv-raise-unit-is-contr (is-proof-irrelevant-type-Prop P x)))
+
+ Eq-south-suspension-Prop :
+ suspension (type-Prop P) → UU l
+ Eq-south-suspension-Prop =
+ cogap-suspension suspension-structure-Eq-south-suspension-Prop
+
+ abstract
+ coh-Eq-suspension-Prop :
+ type-Prop P → Eq-north-suspension-Prop ~ Eq-south-suspension-Prop
+ coh-Eq-suspension-Prop x =
+ dependent-cogap-suspension
+ ( eq-value
+ ( Eq-north-suspension-Prop)
+ ( Eq-south-suspension-Prop))
+ ( ( ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop)) ∙
+ ( inv
+ ( eq-equiv
+ ( equiv-raise-unit-is-contr
+ ( is-proof-irrelevant-type-Prop P x)))) ∙
+ ( inv
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop)))) ,
+ ( ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop)) ∙
+ ( eq-equiv
+ ( equiv-raise-unit-is-contr
+ ( is-proof-irrelevant-type-Prop P x))) ∙
+ ( inv
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop)))) ,
+ ind-subsingleton
+ ( is-prop-type-Prop P)
+ ( x)
+ ( eq-is-contr
+ ( is-contr-equiv
+ ( type-Prop P ≃ raise-unit l)
+ ( equiv-univalence ∘e
+ equiv-binary-concat
+ ( inv
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop)))
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop)))
+ ( is-contr-equiv-is-contr
+ ( is-proof-irrelevant-type-Prop P x)
+ ( is-contr-raise-unit)))))
+
+ suspension-structure-Eq-suspension-Prop :
+ suspension-structure (type-Prop P) (suspension-Prop P → UU l)
+ suspension-structure-Eq-suspension-Prop =
+ ( Eq-north-suspension-Prop ,
+ Eq-south-suspension-Prop ,
+ eq-htpy ∘ coh-Eq-suspension-Prop)
+
+ Eq-suspension-Prop :
+ suspension-Prop P → suspension-Prop P → UU l
+ Eq-suspension-Prop =
+ cogap-suspension suspension-structure-Eq-suspension-Prop
+```
+
+We compute the observational equality of the suspension of a proposition at the
+poles.
+
+```agda
+ eq-compute-Eq-north-north-suspension-Prop :
+ Eq-suspension-Prop north-suspension north-suspension = raise-unit l
+ eq-compute-Eq-north-north-suspension-Prop =
+ htpy-eq
+ ( compute-north-cogap-suspension suspension-structure-Eq-suspension-Prop)
+ ( north-suspension) ∙
+ compute-north-cogap-suspension suspension-structure-Eq-north-suspension-Prop
+
+ compute-Eq-north-north-suspension-Prop :
+ Eq-suspension-Prop north-suspension north-suspension ≃ unit
+ compute-Eq-north-north-suspension-Prop =
+ inv-compute-raise-unit l ∘e
+ equiv-eq eq-compute-Eq-north-north-suspension-Prop
+
+ is-contr-Eq-north-north-suspension-Prop :
+ is-contr (Eq-suspension-Prop north-suspension north-suspension)
+ is-contr-Eq-north-north-suspension-Prop =
+ is-contr-equiv-raise-unit
+ ( equiv-eq eq-compute-Eq-north-north-suspension-Prop)
+
+ eq-compute-Eq-south-south-suspension-Prop :
+ Eq-suspension-Prop south-suspension south-suspension = raise-unit l
+ eq-compute-Eq-south-south-suspension-Prop =
+ htpy-eq
+ ( compute-south-cogap-suspension suspension-structure-Eq-suspension-Prop)
+ ( south-suspension) ∙
+ compute-south-cogap-suspension suspension-structure-Eq-south-suspension-Prop
+
+ compute-Eq-south-south-suspension-Prop :
+ Eq-suspension-Prop south-suspension south-suspension ≃ unit
+ compute-Eq-south-south-suspension-Prop =
+ inv-compute-raise-unit l ∘e
+ equiv-eq eq-compute-Eq-south-south-suspension-Prop
+
+ is-contr-Eq-south-south-suspension-Prop :
+ is-contr (Eq-suspension-Prop south-suspension south-suspension)
+ is-contr-Eq-south-south-suspension-Prop =
+ is-contr-equiv-raise-unit
+ ( equiv-eq eq-compute-Eq-south-south-suspension-Prop)
+
+ eq-compute-Eq-north-south-suspension-Prop :
+ Eq-suspension-Prop north-suspension south-suspension = type-Prop P
+ eq-compute-Eq-north-south-suspension-Prop =
+ htpy-eq
+ ( compute-north-cogap-suspension suspension-structure-Eq-suspension-Prop)
+ ( south-suspension) ∙
+ compute-south-cogap-suspension suspension-structure-Eq-north-suspension-Prop
+
+ eq-compute-Eq-south-north-suspension-Prop :
+ Eq-suspension-Prop south-suspension north-suspension = type-Prop P
+ eq-compute-Eq-south-north-suspension-Prop =
+ htpy-eq
+ ( compute-south-cogap-suspension suspension-structure-Eq-suspension-Prop)
+ ( north-suspension) ∙
+ compute-north-cogap-suspension suspension-structure-Eq-south-suspension-Prop
+```
+
+The observational equality on the suspension of a proposition is propositional.
+
+```agda
+ abstract
+ is-prop-Eq-north-suspension-Prop :
+ (y : suspension-Prop P) → is-prop (Eq-north-suspension-Prop y)
+ is-prop-Eq-north-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ x → is-prop (Eq-north-suspension-Prop x))
+ ( ( inv-tr
+ ( is-prop)
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop))
+ ( is-prop-raise-unit)) ,
+ ( inv-tr
+ ( is-prop)
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop))
+ ( is-prop-type-Prop P)) ,
+ ( λ _ →
+ eq-is-prop
+ ( is-prop-is-prop (Eq-north-suspension-Prop south-suspension))))
+
+ abstract
+ is-prop-Eq-south-suspension-Prop :
+ (y : suspension-Prop P) → is-prop (Eq-south-suspension-Prop y)
+ is-prop-Eq-south-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ x → is-prop (Eq-south-suspension-Prop x))
+ ( ( inv-tr
+ ( is-prop)
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop))
+ ( is-prop-type-Prop P)) ,
+ ( inv-tr
+ ( is-prop)
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop))
+ ( is-prop-raise-unit)) ,
+ ( λ _ →
+ eq-is-prop
+ ( is-prop-is-prop (Eq-south-suspension-Prop south-suspension))))
+
+ abstract
+ is-prop-Eq-suspension-Prop :
+ (x y : suspension-Prop P) → is-prop (Eq-suspension-Prop x y)
+ is-prop-Eq-suspension-Prop x y =
+ dependent-cogap-suspension
+ ( λ x → is-prop (Eq-suspension-Prop x y))
+ ( ( inv-tr
+ ( is-prop)
+ ( htpy-eq
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-suspension-Prop))
+ ( y))
+ ( is-prop-Eq-north-suspension-Prop y)) ,
+ ( inv-tr
+ ( is-prop)
+ ( htpy-eq
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-suspension-Prop))
+ ( y))
+ ( is-prop-Eq-south-suspension-Prop y)) ,
+ ( λ _ →
+ eq-is-prop
+ ( is-prop-is-prop (Eq-suspension-Prop south-suspension y))))
+ ( x)
+```
+
+The observational equality characterizes equality in `ΣP`.
+
+```agda
+ refl-Eq-suspension-Prop : (x : suspension-Prop P) → Eq-suspension-Prop x x
+ refl-Eq-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ x → Eq-suspension-Prop x x)
+ ( ( map-inv-eq
+ ( ( htpy-eq
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-suspension-Prop))
+ ( north-suspension)) ∙
+ ( compute-north-cogap-suspension
+ ( suspension-structure-Eq-north-suspension-Prop)))
+ ( raise-star)) ,
+ ( map-inv-eq
+ ( ( htpy-eq
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-suspension-Prop))
+ ( south-suspension)) ∙
+ ( compute-south-cogap-suspension
+ ( suspension-structure-Eq-south-suspension-Prop)))
+ ( raise-star)) ,
+ ( λ _ →
+ eq-is-contr
+ ( inv-tr
+ ( is-contr)
+ ( eq-compute-Eq-south-south-suspension-Prop)
+ ( is-contr-raise-unit))))
+
+ Eq-eq-suspension-Prop :
+ (x y : suspension-Prop P) → x = y → Eq-suspension-Prop x y
+ Eq-eq-suspension-Prop x .x refl = refl-Eq-suspension-Prop x
+
+ dependent-suspension-structure-eq-north-Eq-north-suspension-Prop :
+ dependent-suspension-structure
+ ( λ y → Eq-suspension-Prop north-suspension y → north-suspension = y)
+ ( suspension-structure-suspension (type-Prop P))
+ dependent-suspension-structure-eq-north-Eq-north-suspension-Prop =
+ ( ( λ _ → refl) ,
+ ( λ u →
+ meridian-suspension
+ ( map-eq eq-compute-Eq-north-south-suspension-Prop u)) ,
+ ( λ p →
+ eq-is-contr
+ ( is-contr-function-type
+ ( is-prop-is-contr
+ ( is-contr-suspension-is-contr
+ ( is-proof-irrelevant-type-Prop P p))
+ ( north-suspension)
+ ( south-suspension)))))
+
+ eq-north-Eq-north-suspension-Prop :
+ (y : suspension-Prop P) →
+ Eq-suspension-Prop north-suspension y → north-suspension = y
+ eq-north-Eq-north-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ y → Eq-suspension-Prop north-suspension y → north-suspension = y)
+ ( dependent-suspension-structure-eq-north-Eq-north-suspension-Prop)
+
+ dependent-suspension-structure-eq-south-Eq-south-suspension-Prop :
+ dependent-suspension-structure
+ ( λ y → Eq-suspension-Prop south-suspension y → south-suspension = y)
+ ( suspension-structure-suspension (type-Prop P))
+ dependent-suspension-structure-eq-south-Eq-south-suspension-Prop =
+ ( ( λ u →
+ inv
+ ( meridian-suspension
+ ( map-eq eq-compute-Eq-south-north-suspension-Prop u))) ,
+ ( λ _ → refl) ,
+ ( λ p →
+ eq-is-contr
+ ( is-contr-function-type
+ ( is-prop-is-contr
+ ( is-contr-suspension-is-contr
+ ( is-proof-irrelevant-type-Prop P p))
+ ( south-suspension)
+ ( south-suspension)))))
+
+ eq-south-Eq-south-suspension-Prop :
+ (y : suspension-Prop P) →
+ Eq-suspension-Prop south-suspension y → south-suspension = y
+ eq-south-Eq-south-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ y → Eq-suspension-Prop south-suspension y → south-suspension = y)
+ ( dependent-suspension-structure-eq-south-Eq-south-suspension-Prop)
+
+ eq-Eq-suspension-Prop :
+ (x y : suspension-Prop P) → Eq-suspension-Prop x y → x = y
+ eq-Eq-suspension-Prop =
+ dependent-cogap-suspension
+ ( λ x → (y : suspension-Prop P) → Eq-suspension-Prop x y → x = y)
+ ( ( eq-north-Eq-north-suspension-Prop) ,
+ ( eq-south-Eq-south-suspension-Prop) ,
+ ( λ p →
+ eq-is-contr
+ ( is-contr-Π
+ ( λ y →
+ is-contr-function-type
+ ( is-prop-is-contr
+ ( is-contr-suspension-is-contr
+ ( is-proof-irrelevant-type-Prop P p))
+ ( south-suspension)
+ ( y))))))
+
+ abstract
+ is-equiv-Eq-eq-suspension-Prop :
+ (x y : suspension-Prop P) → is-equiv (Eq-eq-suspension-Prop x y)
+ is-equiv-Eq-eq-suspension-Prop x =
+ fundamental-theorem-id-section
+ ( x)
+ ( Eq-eq-suspension-Prop x)
+ ( λ y →
+ ( eq-Eq-suspension-Prop x y) ,
+ ( λ _ → eq-is-prop (is-prop-Eq-suspension-Prop x y)))
+
+ abstract
+ is-equiv-eq-Eq-suspension-Prop :
+ (x y : suspension-Prop P) → is-equiv (eq-Eq-suspension-Prop x y)
+ is-equiv-eq-Eq-suspension-Prop x =
+ fundamental-theorem-id-retraction
+ ( x)
+ ( eq-Eq-suspension-Prop x)
+ ( λ y →
+ ( Eq-eq-suspension-Prop x y) ,
+ ( λ _ → eq-is-prop (is-prop-Eq-suspension-Prop x y)))
+
+ equiv-Eq-eq-suspension-Prop :
+ (x y : suspension-Prop P) → (x = y) ≃ Eq-suspension-Prop x y
+ equiv-Eq-eq-suspension-Prop x y =
+ ( Eq-eq-suspension-Prop x y , is-equiv-Eq-eq-suspension-Prop x y)
+
+ equiv-eq-Eq-suspension-Prop :
+ (x y : suspension-Prop P) → Eq-suspension-Prop x y ≃ (x = y)
+ equiv-eq-Eq-suspension-Prop x y =
+ ( eq-Eq-suspension-Prop x y , is-equiv-eq-Eq-suspension-Prop x y)
+
+ abstract
+ is-torsorial-Eq-suspension-Prop :
+ (x : suspension-Prop P) → is-torsorial (Eq-suspension-Prop x)
+ is-torsorial-Eq-suspension-Prop x =
+ fundamental-theorem-id'
+ ( Eq-eq-suspension-Prop x)
+ ( is-equiv-Eq-eq-suspension-Prop x)
+```
+
+Piecing it all together, we conclude that suspensions of propositions are sets.
+
+```agda
+ abstract
+ is-set-suspension-Prop : is-set (suspension-Prop P)
+ is-set-suspension-Prop x y =
+ is-prop-equiv
+ ( equiv-Eq-eq-suspension-Prop x y)
+ ( is-prop-Eq-suspension-Prop x y)
+
+ suspension-set-Prop : Set l
+ suspension-set-Prop = suspension-Prop P , is-set-suspension-Prop
+```
+
+```agda
+ compute-eq-north-south-suspension-Prop :
+ (north-suspension {X = type-Prop P} = south-suspension) ≃ type-Prop P
+ compute-eq-north-south-suspension-Prop =
+ equiv-eq eq-compute-Eq-north-south-suspension-Prop ∘e
+ equiv-Eq-eq-suspension-Prop north-suspension south-suspension
+
+ compute-eq-south-north-suspension-Prop :
+ (south-suspension {X = type-Prop P} = north-suspension) ≃ type-Prop P
+ compute-eq-south-north-suspension-Prop =
+ equiv-eq eq-compute-Eq-south-north-suspension-Prop ∘e
+ equiv-Eq-eq-suspension-Prop south-suspension north-suspension
+```
+
+### Suspensions of propositions are Kuratowski finite sets
+
+```agda
+module _
+ {l : Level} (P : Prop l)
+ where
+
+ is-kuratowski-finite-set-suspension-Prop :
+ is-kuratowski-finite-set (suspension-set-Prop P)
+ is-kuratowski-finite-set-suspension-Prop =
+ intro-exists 2
+ ( comp-surjection
+ ( surjection-bool-suspension)
+ ( surjection-equiv equiv-bool-Fin-two-ℕ))
+```
+
+## See also
+
+- Suspensions of propositions are used in the proof of
+ [Diaconescu's theorem](foundation.diaconescus-theorem.md), which states that
+ the axiom of choice implies the law of excluded middle.
+
+## References
+
+{{#bibliography}}
+
+## External links
+
+- [Truncatedness, properties of suspensions](https://1lab.dev/Homotopy.Space.Suspension.Properties.html#truncatedness)
+ at 1lab
diff --git a/src/synthetic-homotopy-theory/suspensions-of-types.lagda.md b/src/synthetic-homotopy-theory/suspensions-of-types.lagda.md
index b66b91c00a..b0d6084cbc 100644
--- a/src/synthetic-homotopy-theory/suspensions-of-types.lagda.md
+++ b/src/synthetic-homotopy-theory/suspensions-of-types.lagda.md
@@ -9,6 +9,7 @@ module synthetic-homotopy-theory.suspensions-of-types where
```agda
open import foundation.action-on-identifications-dependent-functions
open import foundation.action-on-identifications-functions
+open import foundation.booleans
open import foundation.commuting-squares-of-identifications
open import foundation.connected-types
open import foundation.constant-maps
@@ -18,6 +19,7 @@ open import foundation.dependent-pair-types
open import foundation.diagonal-maps-of-types
open import foundation.equivalence-extensionality
open import foundation.equivalences
+open import foundation.fibers-of-maps
open import foundation.function-extensionality
open import foundation.function-types
open import foundation.functoriality-dependent-function-types
@@ -25,8 +27,11 @@ open import foundation.functoriality-dependent-pair-types
open import foundation.homotopies
open import foundation.identity-types
open import foundation.path-algebra
+open import foundation.propositional-truncations
+open import foundation.propositions
open import foundation.retractions
open import foundation.sections
+open import foundation.surjective-maps
open import foundation.torsorial-type-families
open import foundation.transport-along-identifications
open import foundation.truncated-types
@@ -50,12 +55,12 @@ open import synthetic-homotopy-theory.universal-property-suspensions
## Idea
-The **suspension** of a type `X` is the
-[pushout](synthetic-homotopy-theory.pushouts.md) of the
+The {{#concept "suspension" WD="suspension" WDID=Q1307987 Agda=suspension}} of a
+type `X` is the [pushout](synthetic-homotopy-theory.pushouts.md) of the
[span](foundation.spans.md)
```text
-unit <-- X --> unit
+ 1 <--- X ---> 1
```
Suspensions play an important role in synthetic homotopy theory. For example,
@@ -488,9 +493,9 @@ module _
( c)))
is-section-htpy-htpy-function-out-of-suspension :
- ( ( htpy-function-out-of-suspension-htpy) ∘
- ( htpy-htpy-function-out-of-suspension)) ~
- ( id)
+ is-section
+ htpy-function-out-of-suspension-htpy
+ htpy-htpy-function-out-of-suspension
is-section-htpy-htpy-function-out-of-suspension c =
( ap
( htpy-function-out-of-suspension-htpy)
@@ -516,22 +521,54 @@ module _
( is-section-htpy-htpy-function-out-of-suspension c))
```
+### The booleans surject onto suspensions
+
+```agda
+module _
+ {l : Level} {X : UU l}
+ where
+
+ map-surjection-bool-suspension : bool → suspension X
+ map-surjection-bool-suspension true = north-suspension
+ map-surjection-bool-suspension false = south-suspension
+
+ is-surjective-map-surjection-bool-suspension :
+ is-surjective map-surjection-bool-suspension
+ is-surjective-map-surjection-bool-suspension =
+ dependent-cogap-suspension
+ ( λ x → ║ fiber map-surjection-bool-suspension x ║₋₁)
+ ( unit-trunc-Prop (true , refl) ,
+ unit-trunc-Prop (false , refl) ,
+ ( λ _ →
+ eq-type-Prop
+ ( trunc-Prop
+ ( fiber
+ ( map-surjection-bool-suspension)
+ ( south-suspension)))))
+
+ surjection-bool-suspension : bool ↠ suspension X
+ surjection-bool-suspension =
+ map-surjection-bool-suspension ,
+ is-surjective-map-surjection-bool-suspension
+```
+
### The suspension of a contractible type is contractible
```agda
-is-contr-suspension-is-contr :
- {l : Level} {X : UU l} → is-contr X → is-contr (suspension X)
-is-contr-suspension-is-contr {l} {X} is-contr-X =
- is-contr-is-equiv'
- ( unit)
- ( pr1 (pr2 (cocone-suspension X)))
- ( is-equiv-universal-property-pushout
- ( terminal-map X)
- ( terminal-map X)
- ( cocone-suspension X)
- ( is-equiv-is-contr (terminal-map X) is-contr-X is-contr-unit)
- ( up-suspension' X))
- ( is-contr-unit)
+abstract
+ is-contr-suspension-is-contr :
+ {l : Level} {X : UU l} → is-contr X → is-contr (suspension X)
+ is-contr-suspension-is-contr {X = X} is-contr-X =
+ is-contr-is-equiv'
+ ( unit)
+ ( pr1 (pr2 (cocone-suspension X)))
+ ( is-equiv-universal-property-pushout
+ ( terminal-map X)
+ ( terminal-map X)
+ ( cocone-suspension X)
+ ( is-equiv-is-contr (terminal-map X) is-contr-X is-contr-unit)
+ ( up-suspension' X))
+ ( is-contr-unit)
```
### Suspensions increase connectedness
diff --git a/src/univalent-combinatorics.lagda.md b/src/univalent-combinatorics.lagda.md
index d0099b7719..afd6a44b8c 100644
--- a/src/univalent-combinatorics.lagda.md
+++ b/src/univalent-combinatorics.lagda.md
@@ -43,6 +43,7 @@ open import univalent-combinatorics.cubes public
open import univalent-combinatorics.cycle-partitions public
open import univalent-combinatorics.cycle-prime-decomposition-natural-numbers public
open import univalent-combinatorics.cyclic-finite-types public
+open import univalent-combinatorics.de-morgans-law public
open import univalent-combinatorics.decidable-dependent-function-types public
open import univalent-combinatorics.decidable-dependent-pair-types public
open import univalent-combinatorics.decidable-equivalence-relations public
diff --git a/src/univalent-combinatorics/de-morgans-law.lagda.md b/src/univalent-combinatorics/de-morgans-law.lagda.md
new file mode 100644
index 0000000000..bca666f108
--- /dev/null
+++ b/src/univalent-combinatorics/de-morgans-law.lagda.md
@@ -0,0 +1,98 @@
+# De Morgan's law for finite families of propositions
+
+```agda
+module univalent-combinatorics.de-morgans-law where
+```
+
+Imports
+
+```agda
+open import elementary-number-theory.natural-numbers
+
+open import foundation.coproduct-types
+open import foundation.decidable-dependent-pair-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.empty-types
+open import foundation.equivalences
+open import foundation.existential-quantification
+open import foundation.function-types
+open import foundation.functoriality-dependent-pair-types
+open import foundation.negation
+open import foundation.unit-type
+open import foundation.universe-levels
+
+open import logic.de-morgan-propositions
+open import logic.de-morgan-types
+
+open import univalent-combinatorics.counting
+open import univalent-combinatorics.standard-finite-types
+```
+
+
+
+## Idea
+
+Given a [finite](univalent-combinatorics.finite-types.md) family of
+[De Morgan types](logic.de-morgan-types.md) `P : Fin k → De-Morgan-Type`, then
+the _finitary De Morgan's law_
+
+```text
+ ¬ (∀ i, P i) ⇒ ∃ i, ¬ (P i)
+```
+
+holds.
+
+## Result
+
+```agda
+cases-decide-de-morgans-law-family-Fin-is-de-morgan-family :
+ {l : Level} (k : ℕ) {P : Fin (succ-ℕ k) → UU l} →
+ ((x : Fin (succ-ℕ k)) → is-de-morgan (P x)) →
+ ¬ ((x : Fin (succ-ℕ k)) → P x) →
+ ¬ P (inr star) + ¬ ((x : Fin k) → P (inl x))
+cases-decide-de-morgans-law-family-Fin-is-de-morgan-family k {P} H q =
+ rec-coproduct
+ ( inl)
+ ( λ z → inr (λ y → z (λ y' → q (ind-coproduct P y (λ _ → y')))))
+ ( H (inr star))
+
+decide-de-morgans-law-family-Fin-is-de-morgan-family :
+ {l : Level} (k : ℕ) {P : Fin k → UU l} →
+ ((x : Fin k) → is-de-morgan (P x)) →
+ ¬ ((x : Fin k) → P x) →
+ Σ (Fin k) (¬_ ∘ P)
+decide-de-morgans-law-family-Fin-is-de-morgan-family zero-ℕ {P} H q =
+ q (λ ()) , (λ x → q (λ ()))
+decide-de-morgans-law-family-Fin-is-de-morgan-family (succ-ℕ k) {P} H q =
+ rec-coproduct
+ ( inr star ,_)
+ ( λ q' →
+ map-Σ-map-base inl
+ ( ¬_ ∘ P)
+ ( decide-de-morgans-law-family-Fin-is-de-morgan-family k (H ∘ inl) q'))
+ ( cases-decide-de-morgans-law-family-Fin-is-de-morgan-family k H q)
+```
+
+```agda
+module _
+ {l : Level} {k : ℕ} (P : Fin k → De-Morgan-Type l)
+ where
+
+ decide-de-morgans-law-family-family-Fin :
+ ¬ ((x : Fin k) → type-De-Morgan-Type (P x)) →
+ Σ (Fin k) (¬_ ∘ type-De-Morgan-Type ∘ P)
+ decide-de-morgans-law-family-family-Fin =
+ decide-de-morgans-law-family-Fin-is-de-morgan-family k
+ ( is-de-morgan-type-De-Morgan-Type ∘ P)
+```
+
+## Comment
+
+It is likely that this finitary De Morgan's law can be generalized to families
+of De Morgan types indexed by _searchable types_ in the sense of Escardó
+{{#cite Esc08}}.
+
+## References
+
+{{#bibliography}}
diff --git a/src/univalent-combinatorics/kuratowski-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowski-finite-sets.lagda.md
index 8255b3ad10..1917de6c09 100644
--- a/src/univalent-combinatorics/kuratowski-finite-sets.lagda.md
+++ b/src/univalent-combinatorics/kuratowski-finite-sets.lagda.md
@@ -89,6 +89,10 @@ has-decidable-equality-is-finite-type-𝔽-Kuratowski X =
has-decidable-equality-is-finite
```
+### Kuratowski finite sets are Markovian
+
+> This remains to be formalized. ([Markovian types](logic.markovian-types.md))
+
## See also
- [Finite types](univalent-combinatorics.finite-types.md)