diff --git a/MIXFIX-OPERATORS.md b/MIXFIX-OPERATORS.md
index 4ec18bd8cc..c0e723b26e 100644
--- a/MIXFIX-OPERATORS.md
+++ b/MIXFIX-OPERATORS.md
@@ -163,22 +163,22 @@ Below, we outline a list of general rules when assigning associativities.
## Full table of assigned precedences
-| Precedence level | Operators |
-| ---------------- | ------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------- |
-| 50 | Unary nonparametric operators. This class is currently empty |
-| 45 | Exponentiative arithmetic operators |
-| 40 | Multiplicative arithmetic operators |
-| 36 | Subtractive arithmetic operators |
-| 35 | Additive arithmetic operators |
-| 30 | Relational arithmetic operators like`_≤-ℕ_` and `_<-ℕ_` |
-| 25 | Parametric unary operators. This class is currently empty |
-| 20 | Parametric exponentiative operators. This class is currently empty |
-| 17 | Left homotopy whiskering `_·l_` |
-| 16 | Right homotopy whiskering `_·r_` |
-| 15 | Parametric multiplicative operators like `_×_`,`_×∗_`, `_∧_`, `_∧∗_`, function composition operators like `_∘_`,`_∘∗_`, `_∘e_`, and `_∘iff_`, concatenation operators like `_∙_` and `_∙h_` |
-| 10 | Parametric additive operators like `_+_`, `_∨_`, `_∨∗_`, list concatenation, monadic bind operators for the type checking monad |
-| 6 | Parametric relational operators like `_=_`, `_~_`, `_≃_`, and `_↔_`, elementhood relations, subtype relations |
-| 5 | Directed function type-like formation operators, e.g. `_⇒_`, `_→∗_`, `_↠_`, `_↪_`, `_↪ᵈ_`, and `_⊆_` |
-| 3 | The pairing operators `_,_` and `_,ω_` |
-| 0-1 | Reasoning syntaxes |
-| -∞ | Function type formation `_→_` |
+| Precedence level | Operators |
+| ---------------- | -------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------------- |
+| 50 | Unary nonparametric operators. This class is currently empty |
+| 45 | Exponentiative arithmetic operators |
+| 40 | Multiplicative arithmetic operators |
+| 36 | Subtractive arithmetic operators |
+| 35 | Additive arithmetic operators |
+| 30 | Relational arithmetic operators like`_≤-ℕ_` and `_<-ℕ_` |
+| 25 | Parametric unary operators like `¬_` and `¬¬_` |
+| 20 | Parametric exponentiative operators. This class is currently empty |
+| 17 | Left homotopy whiskering `_·l_` |
+| 16 | Right homotopy whiskering `_·r_` |
+| 15 | Parametric multiplicative operators like `_×_`,`_×∗_`, `_∧_`, `_∧∗_`, `_*_`, function composition operators like `_∘_`,`_∘∗_`, `_∘e_`, and `_∘iff_`, concatenation operators like `_∙_` and `_∙h_` |
+| 10 | Parametric additive operators like `_+_`, `_∨_`, `_∨∗_`, list concatenation, monadic bind operators for the type checking monad |
+| 6 | Parametric relational operators like `_=_`, `_~_`, `_≃_`, `_⇔_`, and `_↔_`, elementhood relations, subtype relations |
+| 5 | Directed function type-like formation operators, e.g. `_⇒_`, `_↪_`, `_→∗_`, `_↠_`, `_↪ᵈ_`, and `_⊆_` |
+| 3 | The pairing operators `_,_` and `_,ω_` |
+| 0-1 | Reasoning syntaxes |
+| -∞ | Function type formation `_→_` |
diff --git a/src/category-theory/categories.lagda.md b/src/category-theory/categories.lagda.md
index e015106c69..810689aba3 100644
--- a/src/category-theory/categories.lagda.md
+++ b/src/category-theory/categories.lagda.md
@@ -18,6 +18,7 @@ open import foundation.cartesian-product-types
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.strictly-involutive-identity-types
@@ -312,7 +313,7 @@ module _
is-equiv-is-category-is-surjective-iso-eq-Preunivalent-Category :
is-equiv is-category-is-surjective-iso-eq-Preunivalent-Category
is-equiv-is-category-is-surjective-iso-eq-Preunivalent-Category =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-surjective-iso-eq-Precategory
( precategory-Preunivalent-Category C))
( is-prop-is-category-Precategory (precategory-Preunivalent-Category C))
@@ -321,7 +322,7 @@ module _
is-equiv-is-surjective-iso-eq-is-category-Preunivalent-Category :
is-equiv is-surjective-iso-eq-is-category-Preunivalent-Category
is-equiv-is-surjective-iso-eq-is-category-Preunivalent-Category =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-category-Precategory (precategory-Preunivalent-Category C))
( is-prop-is-surjective-iso-eq-Precategory
( precategory-Preunivalent-Category C))
diff --git a/src/category-theory/coproducts-in-precategories.lagda.md b/src/category-theory/coproducts-in-precategories.lagda.md
index 1ef104ed59..cc72f57907 100644
--- a/src/category-theory/coproducts-in-precategories.lagda.md
+++ b/src/category-theory/coproducts-in-precategories.lagda.md
@@ -16,7 +16,7 @@ open import foundation.dependent-pair-types
open import foundation.identity-types
open import foundation.iterated-dependent-product-types
open import foundation.propositions
-open import foundation.unique-existence
+open import foundation.uniqueness-quantification
open import foundation.universe-levels
```
@@ -39,7 +39,7 @@ module _
(z : obj-Precategory C)
(f : hom-Precategory C x z) →
(g : hom-Precategory C y z) →
- ∃!
+ uniquely-exists-structure
( hom-Precategory C p z)
( λ h →
( comp-hom-Precategory C h l = f) ×
diff --git a/src/category-theory/dependent-products-of-large-precategories.lagda.md b/src/category-theory/dependent-products-of-large-precategories.lagda.md
index 83d18e5945..eacb31a0b5 100644
--- a/src/category-theory/dependent-products-of-large-precategories.lagda.md
+++ b/src/category-theory/dependent-products-of-large-precategories.lagda.md
@@ -14,6 +14,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.strictly-involutive-identity-types
@@ -215,7 +216,7 @@ module _
(f : hom-Π-Large-Precategory I C x y) →
is-equiv (is-fiberwise-iso-is-iso-Π-Large-Precategory f)
is-equiv-is-fiberwise-iso-is-iso-Π-Large-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Large-Precategory (Π-Large-Precategory I C) f)
( is-prop-Π (λ i → is-prop-is-iso-Large-Precategory (C i) (f i)))
( is-iso-is-fiberwise-iso-Π-Large-Precategory f)
@@ -233,7 +234,7 @@ module _
(f : hom-Π-Large-Precategory I C x y) →
is-equiv (is-iso-is-fiberwise-iso-Π-Large-Precategory f)
is-equiv-is-iso-is-fiberwise-iso-Π-Large-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-Π (λ i → is-prop-is-iso-Large-Precategory (C i) (f i)))
( is-prop-is-iso-Large-Precategory (Π-Large-Precategory I C) f)
( is-fiberwise-iso-is-iso-Π-Large-Precategory f)
diff --git a/src/category-theory/dependent-products-of-precategories.lagda.md b/src/category-theory/dependent-products-of-precategories.lagda.md
index 6e7448dc99..4ee2a1858a 100644
--- a/src/category-theory/dependent-products-of-precategories.lagda.md
+++ b/src/category-theory/dependent-products-of-precategories.lagda.md
@@ -15,6 +15,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.strictly-involutive-identity-types
@@ -183,7 +184,7 @@ module _
(f : hom-Π-Precategory I C x y) →
is-equiv (is-fiberwise-iso-is-iso-Π-Precategory f)
is-equiv-is-fiberwise-iso-is-iso-Π-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Precategory (Π-Precategory I C) f)
( is-prop-Π (λ i → is-prop-is-iso-Precategory (C i) (f i)))
( is-iso-is-fiberwise-iso-Π-Precategory f)
@@ -201,7 +202,7 @@ module _
(f : hom-Π-Precategory I C x y) →
is-equiv (is-iso-is-fiberwise-iso-Π-Precategory f)
is-equiv-is-iso-is-fiberwise-iso-Π-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-Π (λ i → is-prop-is-iso-Precategory (C i) (f i)))
( is-prop-is-iso-Precategory (Π-Precategory I C) f)
( is-fiberwise-iso-is-iso-Π-Precategory f)
diff --git a/src/category-theory/essentially-surjective-functors-precategories.lagda.md b/src/category-theory/essentially-surjective-functors-precategories.lagda.md
index ef49c3c337..b33efd03a0 100644
--- a/src/category-theory/essentially-surjective-functors-precategories.lagda.md
+++ b/src/category-theory/essentially-surjective-functors-precategories.lagda.md
@@ -43,14 +43,15 @@ module _
Π-Prop
( obj-Precategory D)
( λ y →
- ∃-Prop
+ exists-structure-Prop
( obj-Precategory C)
( λ x → iso-Precategory D (obj-functor-Precategory C D F x) y))
is-essentially-surjective-functor-Precategory : UU (l1 ⊔ l3 ⊔ l4)
is-essentially-surjective-functor-Precategory =
( y : obj-Precategory D) →
- ∃ ( obj-Precategory C)
+ exists-structure
+ ( obj-Precategory C)
( λ x → iso-Precategory D (obj-functor-Precategory C D F x) y)
is-prop-is-essentially-surjective-functor-Precategory :
diff --git a/src/category-theory/exponential-objects-precategories.lagda.md b/src/category-theory/exponential-objects-precategories.lagda.md
index 941b98f1d7..e2fad05451 100644
--- a/src/category-theory/exponential-objects-precategories.lagda.md
+++ b/src/category-theory/exponential-objects-precategories.lagda.md
@@ -13,7 +13,7 @@ open import category-theory.products-in-precategories
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.identity-types
-open import foundation.unique-existence
+open import foundation.uniqueness-quantification
open import foundation.universe-levels
```
@@ -47,7 +47,7 @@ module _
is-exponential-obj-Precategory x y e ev =
(z : obj-Precategory C)
(f : hom-Precategory C (object-product-obj-Precategory C p z x) y) →
- ∃!
+ uniquely-exists-structure
( hom-Precategory C z e)
( λ g →
comp-hom-Precategory C ev
diff --git a/src/category-theory/faithful-maps-precategories.lagda.md b/src/category-theory/faithful-maps-precategories.lagda.md
index 40b26f802d..69cc0cd483 100644
--- a/src/category-theory/faithful-maps-precategories.lagda.md
+++ b/src/category-theory/faithful-maps-precategories.lagda.md
@@ -16,7 +16,9 @@ open import foundation.equivalences
open import foundation.function-types
open import foundation.injective-maps
open import foundation.iterated-dependent-product-types
+open import foundation.logical-equivalences
open import foundation.propositions
+open import foundation.sets
open import foundation.universe-levels
```
@@ -156,7 +158,7 @@ module _
is-equiv-is-injective-hom-is-faithful-map-Precategory :
is-equiv is-injective-hom-is-faithful-map-Precategory
is-equiv-is-injective-hom-is-faithful-map-Precategory =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-faithful-map-Precategory C D F)
( is-prop-is-injective-hom-map-Precategory C D F)
( is-faithful-is-injective-hom-map-Precategory)
@@ -164,7 +166,7 @@ module _
is-equiv-is-faithful-is-injective-hom-map-Precategory :
is-equiv is-faithful-is-injective-hom-map-Precategory
is-equiv-is-faithful-is-injective-hom-map-Precategory =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-injective-hom-map-Precategory C D F)
( is-prop-is-faithful-map-Precategory C D F)
( is-injective-hom-is-faithful-map-Precategory)
diff --git a/src/category-theory/isomorphisms-in-categories.lagda.md b/src/category-theory/isomorphisms-in-categories.lagda.md
index 5efbcdb841..a0f12490f9 100644
--- a/src/category-theory/isomorphisms-in-categories.lagda.md
+++ b/src/category-theory/isomorphisms-in-categories.lagda.md
@@ -212,7 +212,7 @@ module _
### The type of isomorphisms form a set
The type of isomorphisms between objects `x y : A` is a
-[subtype](foundation-core.subtypes.md) of the [foundation-core.sets.md)
+[subtype](foundation-core.subtypes.md) of the [set](foundation-core.sets.md)
`hom x y` since being an isomorphism is a proposition.
```agda
diff --git a/src/category-theory/isomorphisms-in-precategories.lagda.md b/src/category-theory/isomorphisms-in-precategories.lagda.md
index 4e7be2a906..31a63e5d4f 100644
--- a/src/category-theory/isomorphisms-in-precategories.lagda.md
+++ b/src/category-theory/isomorphisms-in-precategories.lagda.md
@@ -248,7 +248,7 @@ module _
### The type of isomorphisms form a set
The type of isomorphisms between objects `x y : A` is a
-[subtype](foundation-core.subtypes.md) of the [foundation-core.sets.md)
+[subtype](foundation-core.subtypes.md) of the [set](foundation-core.sets.md)
`hom x y` since being an isomorphism is a proposition.
```agda
diff --git a/src/category-theory/natural-numbers-object-precategories.lagda.md b/src/category-theory/natural-numbers-object-precategories.lagda.md
index b55eb9b842..33cad6e351 100644
--- a/src/category-theory/natural-numbers-object-precategories.lagda.md
+++ b/src/category-theory/natural-numbers-object-precategories.lagda.md
@@ -14,7 +14,7 @@ open import foundation.action-on-identifications-functions
open import foundation.cartesian-product-types
open import foundation.dependent-pair-types
open import foundation.identity-types
-open import foundation.unique-existence
+open import foundation.uniqueness-quantification
open import foundation.universe-levels
```
@@ -43,7 +43,7 @@ module _
(x : obj-Precategory C)
(q : hom-Precategory C t x)
(f : hom-Precategory C x x) →
- ∃!
+ uniquely-exists-structure
( hom-Precategory C n x)
( λ u →
( comp-hom-Precategory C u z = q) ×
diff --git a/src/category-theory/precategory-of-functors.lagda.md b/src/category-theory/precategory-of-functors.lagda.md
index 9130ba3fe0..d0dd6873e4 100644
--- a/src/category-theory/precategory-of-functors.lagda.md
+++ b/src/category-theory/precategory-of-functors.lagda.md
@@ -20,6 +20,7 @@ open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.functoriality-dependent-pair-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.strictly-involutive-identity-types
open import foundation.universe-levels
@@ -199,7 +200,7 @@ module _
(f : natural-transformation-Precategory C D F G) →
is-equiv (is-iso-functor-is-natural-isomorphism-Precategory f)
is-equiv-is-iso-functor-is-natural-isomorphism-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-natural-isomorphism-Precategory C D F G f)
( is-prop-is-iso-Precategory
( functor-precategory-Precategory C D) {F} {G} f)
@@ -209,7 +210,7 @@ module _
(f : natural-transformation-Precategory C D F G) →
is-equiv (is-natural-isomorphism-is-iso-functor-Precategory f)
is-equiv-is-natural-isomorphism-is-iso-functor-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Precategory
( functor-precategory-Precategory C D) {F} {G} f)
( is-prop-is-natural-isomorphism-Precategory C D F G f)
diff --git a/src/category-theory/precategory-of-maps-precategories.lagda.md b/src/category-theory/precategory-of-maps-precategories.lagda.md
index bb248ce3b4..8fc3c449a5 100644
--- a/src/category-theory/precategory-of-maps-precategories.lagda.md
+++ b/src/category-theory/precategory-of-maps-precategories.lagda.md
@@ -20,6 +20,7 @@ open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.functoriality-dependent-pair-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.strictly-involutive-identity-types
open import foundation.universe-levels
@@ -199,7 +200,7 @@ module _
(f : natural-transformation-map-Precategory C D F G) →
is-equiv (is-iso-map-is-natural-isomorphism-map-Precategory f)
is-equiv-is-iso-map-is-natural-isomorphism-map-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-natural-isomorphism-map-Precategory C D F G f)
( is-prop-is-iso-Precategory
( map-precategory-Precategory C D) {F} {G} f)
@@ -209,7 +210,7 @@ module _
(f : natural-transformation-map-Precategory C D F G) →
is-equiv (is-natural-isomorphism-map-is-iso-map-Precategory f)
is-equiv-is-natural-isomorphism-map-is-iso-map-Precategory f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Precategory
( map-precategory-Precategory C D) {F} {G} f)
( is-prop-is-natural-isomorphism-map-Precategory C D F G f)
diff --git a/src/category-theory/products-in-precategories.lagda.md b/src/category-theory/products-in-precategories.lagda.md
index bb27e473b9..78e0668af4 100644
--- a/src/category-theory/products-in-precategories.lagda.md
+++ b/src/category-theory/products-in-precategories.lagda.md
@@ -16,7 +16,7 @@ open import foundation.dependent-pair-types
open import foundation.identity-types
open import foundation.iterated-dependent-product-types
open import foundation.propositions
-open import foundation.unique-existence
+open import foundation.uniqueness-quantification
open import foundation.universe-levels
```
@@ -52,8 +52,11 @@ module _
(z : obj-Precategory C)
(f : hom-Precategory C z x) →
(g : hom-Precategory C z y) →
- (∃! (hom-Precategory C z p) λ h →
- (comp-hom-Precategory C l h = f) × (comp-hom-Precategory C r h = g))
+ ( uniquely-exists-structure
+ ( hom-Precategory C z p)
+ ( λ h →
+ ( comp-hom-Precategory C l h = f) ×
+ ( comp-hom-Precategory C r h = g)))
product-obj-Precategory : obj-Precategory C → obj-Precategory C → UU (l1 ⊔ l2)
product-obj-Precategory x y =
diff --git a/src/category-theory/pullbacks-in-precategories.lagda.md b/src/category-theory/pullbacks-in-precategories.lagda.md
index 1d18cc1147..aaefb61ec6 100644
--- a/src/category-theory/pullbacks-in-precategories.lagda.md
+++ b/src/category-theory/pullbacks-in-precategories.lagda.md
@@ -16,7 +16,7 @@ open import foundation.dependent-pair-types
open import foundation.identity-types
open import foundation.iterated-dependent-product-types
open import foundation.propositions
-open import foundation.unique-existence
+open import foundation.uniqueness-quantification
open import foundation.universe-levels
```
@@ -62,7 +62,7 @@ module _
(p₁' : hom-Precategory C w' y) →
(p₂' : hom-Precategory C w' z) →
comp-hom-Precategory C f p₁' = comp-hom-Precategory C g p₂' →
- ∃!
+ uniquely-exists-structure
( hom-Precategory C w' w)
( λ h →
( comp-hom-Precategory C p₁ h = p₁') ×
diff --git a/src/category-theory/replete-subprecategories.lagda.md b/src/category-theory/replete-subprecategories.lagda.md
index 2f89bd6296..50d39bd9b3 100644
--- a/src/category-theory/replete-subprecategories.lagda.md
+++ b/src/category-theory/replete-subprecategories.lagda.md
@@ -18,6 +18,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.functoriality-dependent-pair-types
open import foundation.iterated-dependent-product-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.subsingleton-induction
open import foundation.subtypes
@@ -216,7 +217,7 @@ module _
is-equiv-is-iso-is-iso-base-is-replete-Subprecategory :
is-equiv is-iso-is-iso-base-is-replete-Subprecategory
is-equiv-is-iso-is-iso-base-is-replete-Subprecategory =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Precategory C (inclusion-hom-Subprecategory C P x y f))
( is-prop-is-iso-Subprecategory C P f)
( is-iso-base-is-iso-Subprecategory C P)
diff --git a/src/category-theory/representing-arrow-category.lagda.md b/src/category-theory/representing-arrow-category.lagda.md
index ff5fd7b484..dbdfec5962 100644
--- a/src/category-theory/representing-arrow-category.lagda.md
+++ b/src/category-theory/representing-arrow-category.lagda.md
@@ -15,6 +15,7 @@ open import foundation.booleans
open import foundation.dependent-pair-types
open import foundation.empty-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.subtypes
@@ -118,7 +119,7 @@ representing-arrow-Precategory =
is-category-representing-arrow-Category :
is-category-Precategory representing-arrow-Precategory
is-category-representing-arrow-Category true true =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-bool true true)
( is-prop-type-subtype
( is-iso-prop-Precategory representing-arrow-Precategory {true} {true})
@@ -133,7 +134,7 @@ is-category-representing-arrow-Category false true =
( iso-eq-Precategory representing-arrow-Precategory false true)
( hom-inv-iso-Precategory representing-arrow-Precategory)
is-category-representing-arrow-Category false false =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-bool false false)
( is-prop-type-subtype
( is-iso-prop-Precategory representing-arrow-Precategory {false} {false})
diff --git a/src/category-theory/slice-precategories.lagda.md b/src/category-theory/slice-precategories.lagda.md
index 00eaee9ffd..ccef207457 100644
--- a/src/category-theory/slice-precategories.lagda.md
+++ b/src/category-theory/slice-precategories.lagda.md
@@ -22,6 +22,7 @@ open import foundation.function-types
open import foundation.functoriality-dependent-pair-types
open import foundation.homotopies
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.strictly-involutive-identity-types
@@ -263,7 +264,11 @@ module _
eq-hom-Slice-Precategory C A _ _ (pr2 (pr2 (pr1 (ϕ Z h₁ h₂ β₂))))
q :
- ∀ k →
+ (k :
+ hom-Precategory
+ ( Slice-Precategory C A)
+ ( Z , comp-hom-Precategory C f h₁)
+ ( W , p)) →
is-prop
( ( comp-hom-Precategory
(Slice-Precategory C A) (p₁ , α₁) k = (h₁ , refl)) ×
@@ -275,7 +280,11 @@ module _
( is-set-hom-Slice-Precategory C A _ _ _ _)
σ :
- ∀ k →
+ (k :
+ hom-Precategory
+ ( Slice-Precategory C A)
+ ( Z , comp-hom-Precategory C f h₁)
+ ( W , p)) →
( ( comp-hom-Precategory
( Slice-Precategory C A)
( p₁ , α₁)
@@ -331,7 +340,7 @@ module _
( p₂' , α')))))
q :
- ∀ k' →
+ (k' : hom-Precategory C W' W) →
is-prop
(( comp-hom-Precategory C p₁ k' = p₁') ×
( comp-hom-Precategory C p₂ k' = p₂'))
@@ -363,7 +372,7 @@ module _
is-product-obj-Precategory
(Slice-Precategory C A) (X , f) (Y , g) (W , p) (p₁ , α₁) (p₂ , α₂)
equiv-is-pullback-is-product-Slice-Precategory =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-pullback-obj-Precategory C A X Y f g W p₁ p₂ α)
( is-prop-is-product-obj-Precategory
(Slice-Precategory C A) (X , f) (Y , g) (W , p) (p₁ , α₁) (p₂ , α₂))
diff --git a/src/commutative-algebra/intersections-radical-ideals-commutative-rings.lagda.md b/src/commutative-algebra/intersections-radical-ideals-commutative-rings.lagda.md
index 850554b720..1bf031cb1c 100644
--- a/src/commutative-algebra/intersections-radical-ideals-commutative-rings.lagda.md
+++ b/src/commutative-algebra/intersections-radical-ideals-commutative-rings.lagda.md
@@ -177,7 +177,7 @@ module _
( intersection-ideal-Commutative-Ring A I J)
( x))
( λ (m , K') →
- intro-∃
+ intro-exists
( add-ℕ n m)
( ( is-closed-under-eq-ideal-Commutative-Ring A I
( is-closed-under-right-multiplication-ideal-Commutative-Ring
@@ -209,7 +209,8 @@ module _
( ideal-radical-of-ideal-Commutative-Ring A I)
( ideal-radical-of-ideal-Commutative-Ring A J)
( x))
- ( λ (n , H' , K') → (intro-∃ n H' , intro-∃ n K'))
+ ( λ (n , H' , K') →
+ ( intro-exists n H' , intro-exists n K'))
preserves-intersection-radical-of-ideal-Commutative-Ring :
( intersection-ideal-Commutative-Ring A
diff --git a/src/commutative-algebra/nilradical-commutative-rings.lagda.md b/src/commutative-algebra/nilradical-commutative-rings.lagda.md
index 8acddc3744..e8cae03fe9 100644
--- a/src/commutative-algebra/nilradical-commutative-rings.lagda.md
+++ b/src/commutative-algebra/nilradical-commutative-rings.lagda.md
@@ -49,7 +49,7 @@ contains-zero-nilradical-Commutative-Ring :
{l : Level} (A : Commutative-Ring l) →
contains-zero-subset-Commutative-Ring A
( subset-nilradical-Commutative-Ring A)
-contains-zero-nilradical-Commutative-Ring A = intro-∃ 1 refl
+contains-zero-nilradical-Commutative-Ring A = intro-exists 1 refl
```
### The nilradical is closed under addition
diff --git a/src/commutative-algebra/nilradicals-commutative-semirings.lagda.md b/src/commutative-algebra/nilradicals-commutative-semirings.lagda.md
index f60df9ffde..994c83bc29 100644
--- a/src/commutative-algebra/nilradicals-commutative-semirings.lagda.md
+++ b/src/commutative-algebra/nilradicals-commutative-semirings.lagda.md
@@ -42,7 +42,7 @@ contains-zero-nilradical-Commutative-Semiring :
{l : Level} (A : Commutative-Semiring l) →
contains-zero-subset-Commutative-Semiring A
( subset-nilradical-Commutative-Semiring A)
-contains-zero-nilradical-Commutative-Semiring A = intro-∃ 1 refl
+contains-zero-nilradical-Commutative-Semiring A = intro-exists 1 refl
```
### The nilradical is closed under addition
diff --git a/src/commutative-algebra/prime-ideals-commutative-rings.lagda.md b/src/commutative-algebra/prime-ideals-commutative-rings.lagda.md
index 317ca4828a..de9a327c13 100644
--- a/src/commutative-algebra/prime-ideals-commutative-rings.lagda.md
+++ b/src/commutative-algebra/prime-ideals-commutative-rings.lagda.md
@@ -164,11 +164,10 @@ is-radical-prime-ideal-Commutative-Ring R P x zero-ℕ p =
( p)
( x)
is-radical-prime-ideal-Commutative-Ring R P x (succ-ℕ n) p =
- elim-disjunction-Prop
- ( subset-prime-ideal-Commutative-Ring R P (power-Commutative-Ring R n x))
+ elim-disjunction
( subset-prime-ideal-Commutative-Ring R P x)
- ( subset-prime-ideal-Commutative-Ring R P x)
- ( is-radical-prime-ideal-Commutative-Ring R P x n , id)
+ ( is-radical-prime-ideal-Commutative-Ring R P x n)
+ ( id)
( is-prime-ideal-ideal-prime-ideal-Commutative-Ring R P
( power-Commutative-Ring R n x)
( x)
diff --git a/src/commutative-algebra/products-radical-ideals-commutative-rings.lagda.md b/src/commutative-algebra/products-radical-ideals-commutative-rings.lagda.md
index 6177f5690e..f112058f6d 100644
--- a/src/commutative-algebra/products-radical-ideals-commutative-rings.lagda.md
+++ b/src/commutative-algebra/products-radical-ideals-commutative-rings.lagda.md
@@ -182,7 +182,7 @@ module _
( J))
( mul-Commutative-Ring A x y))
( λ (n , p) →
- intro-∃
+ intro-exists
( succ-ℕ n)
( is-closed-under-eq-ideal-Commutative-Ring' A
( product-ideal-Commutative-Ring A
diff --git a/src/commutative-algebra/radicals-of-ideals-commutative-rings.lagda.md b/src/commutative-algebra/radicals-of-ideals-commutative-rings.lagda.md
index 40c4fe1491..13a2f12ae2 100644
--- a/src/commutative-algebra/radicals-of-ideals-commutative-rings.lagda.md
+++ b/src/commutative-algebra/radicals-of-ideals-commutative-rings.lagda.md
@@ -82,8 +82,7 @@ module _
subset-radical-of-ideal-Commutative-Ring : type-Commutative-Ring A → Prop l2
subset-radical-of-ideal-Commutative-Ring f =
- ∃-Prop ℕ
- ( λ n → is-in-ideal-Commutative-Ring A I (power-Commutative-Ring A n f))
+ ∃ ℕ (λ n → subset-ideal-Commutative-Ring A I (power-Commutative-Ring A n f))
is-in-radical-of-ideal-Commutative-Ring : type-Commutative-Ring A → UU l2
is-in-radical-of-ideal-Commutative-Ring =
@@ -93,7 +92,7 @@ module _
(f : type-Commutative-Ring A) →
is-in-ideal-Commutative-Ring A I f →
is-in-radical-of-ideal-Commutative-Ring f
- contains-ideal-radical-of-ideal-Commutative-Ring f H = intro-∃ 1 H
+ contains-ideal-radical-of-ideal-Commutative-Ring f H = intro-exists 1 H
contains-zero-radical-of-ideal-Commutative-Ring :
is-in-radical-of-ideal-Commutative-Ring (zero-Commutative-Ring A)
@@ -113,7 +112,7 @@ module _
( subset-radical-of-ideal-Commutative-Ring
( add-Commutative-Ring A x y))
( λ (m , q) →
- intro-∃
+ intro-exists
( n +ℕ m)
( is-closed-under-eq-ideal-Commutative-Ring' A I
( is-closed-under-addition-ideal-Commutative-Ring A I
@@ -138,7 +137,7 @@ module _
apply-universal-property-trunc-Prop H
( subset-radical-of-ideal-Commutative-Ring (neg-Commutative-Ring A x))
( λ (n , p) →
- intro-∃ n
+ intro-exists n
( is-closed-under-eq-ideal-Commutative-Ring' A I
( is-closed-under-left-multiplication-ideal-Commutative-Ring A I _
( power-Commutative-Ring A n x)
@@ -152,7 +151,7 @@ module _
apply-universal-property-trunc-Prop H
( subset-radical-of-ideal-Commutative-Ring (mul-Commutative-Ring A x y))
( λ (n , p) →
- intro-∃ n
+ intro-exists n
( is-closed-under-eq-ideal-Commutative-Ring' A I
( is-closed-under-right-multiplication-ideal-Commutative-Ring A I
( power-Commutative-Ring A n x)
@@ -167,7 +166,7 @@ module _
apply-universal-property-trunc-Prop H
( subset-radical-of-ideal-Commutative-Ring (mul-Commutative-Ring A x y))
( λ (n , p) →
- intro-∃ n
+ intro-exists n
( is-closed-under-eq-ideal-Commutative-Ring' A I
( is-closed-under-left-multiplication-ideal-Commutative-Ring A I
( power-Commutative-Ring A n x)
@@ -190,7 +189,7 @@ module _
apply-universal-property-trunc-Prop H
( subset-radical-of-ideal-Commutative-Ring x)
( λ (m , K) →
- intro-∃
+ intro-exists
( mul-ℕ n m)
( is-closed-under-eq-ideal-Commutative-Ring' A I K
( power-mul-Commutative-Ring A n m)))
diff --git a/src/commutative-algebra/zariski-topology.lagda.md b/src/commutative-algebra/zariski-topology.lagda.md
index 3b325ea92e..b3bbbec25c 100644
--- a/src/commutative-algebra/zariski-topology.lagda.md
+++ b/src/commutative-algebra/zariski-topology.lagda.md
@@ -42,10 +42,9 @@ is-closed-subset-zariski-topology-Commutative-Ring :
(U : subtype (l1 ⊔ l2 ⊔ l3) (prime-ideal-Commutative-Ring l2 A)) →
Prop (lsuc l1 ⊔ lsuc l2 ⊔ lsuc l3)
is-closed-subset-zariski-topology-Commutative-Ring {l1} {l2} {l3} A U =
- ∃-Prop
+ exists-structure-Prop
( subtype l3 (type-Commutative-Ring A))
- ( λ V →
- standard-closed-subset-zariski-topology-Commutative-Ring A V = U)
+ ( λ V → standard-closed-subset-zariski-topology-Commutative-Ring A V = U)
closed-subset-zariski-topology-Commutative-Ring :
{l1 l2 : Level} (l3 : Level) (A : Commutative-Ring l1) →
diff --git a/src/elementary-number-theory/addition-integers.lagda.md b/src/elementary-number-theory/addition-integers.lagda.md
index 51c6422199..0074401613 100644
--- a/src/elementary-number-theory/addition-integers.lagda.md
+++ b/src/elementary-number-theory/addition-integers.lagda.md
@@ -26,6 +26,7 @@ open import foundation.function-types
open import foundation.identity-types
open import foundation.injective-maps
open import foundation.interchange-law
+open import foundation.sets
open import foundation.unit-type
```
diff --git a/src/elementary-number-theory/addition-natural-numbers.lagda.md b/src/elementary-number-theory/addition-natural-numbers.lagda.md
index 403385d149..47633f70b0 100644
--- a/src/elementary-number-theory/addition-natural-numbers.lagda.md
+++ b/src/elementary-number-theory/addition-natural-numbers.lagda.md
@@ -21,6 +21,7 @@ open import foundation.identity-types
open import foundation.injective-maps
open import foundation.interchange-law
open import foundation.negated-equality
+open import foundation.sets
```
diff --git a/src/elementary-number-theory/finitary-natural-numbers.lagda.md b/src/elementary-number-theory/finitary-natural-numbers.lagda.md
index f97fed792d..57acf9d9c0 100644
--- a/src/elementary-number-theory/finitary-natural-numbers.lagda.md
+++ b/src/elementary-number-theory/finitary-natural-numbers.lagda.md
@@ -21,6 +21,7 @@ open import foundation.empty-types
open import foundation.function-types
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.universe-levels
open import univalent-combinatorics.standard-finite-types
diff --git a/src/elementary-number-theory/inequality-rational-numbers.lagda.md b/src/elementary-number-theory/inequality-rational-numbers.lagda.md
index c4b8e44389..3b96f951fd 100644
--- a/src/elementary-number-theory/inequality-rational-numbers.lagda.md
+++ b/src/elementary-number-theory/inequality-rational-numbers.lagda.md
@@ -21,6 +21,7 @@ open import elementary-number-theory.positive-integers
open import elementary-number-theory.rational-numbers
open import foundation.cartesian-product-types
+open import foundation.conjunction
open import foundation.coproduct-types
open import foundation.decidable-propositions
open import foundation.dependent-pair-types
diff --git a/src/elementary-number-theory/initial-segments-natural-numbers.lagda.md b/src/elementary-number-theory/initial-segments-natural-numbers.lagda.md
index d6d10e541f..9823b2c9a1 100644
--- a/src/elementary-number-theory/initial-segments-natural-numbers.lagda.md
+++ b/src/elementary-number-theory/initial-segments-natural-numbers.lagda.md
@@ -38,7 +38,7 @@ holds for every `n : ℕ`.
```agda
is-initial-segment-subset-ℕ-Prop : {l : Level} (P : subtype l ℕ) → Prop l
is-initial-segment-subset-ℕ-Prop P =
- Π-Prop ℕ (λ n → implication-Prop (P (succ-ℕ n)) (P n))
+ Π-Prop ℕ (λ n → hom-Prop (P (succ-ℕ n)) (P n))
is-initial-segment-subset-ℕ : {l : Level} (P : subtype l ℕ) → UU l
is-initial-segment-subset-ℕ P = type-Prop (is-initial-segment-subset-ℕ-Prop P)
diff --git a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md
index 9b643eebd3..fe5615c83e 100644
--- a/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md
+++ b/src/elementary-number-theory/modular-arithmetic-standard-finite-types.lagda.md
@@ -26,6 +26,7 @@ open import foundation.equivalences
open import foundation.function-types
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.split-surjective-maps
open import foundation.surjective-maps
open import foundation.universe-levels
diff --git a/src/elementary-number-theory/multiplication-integers.lagda.md b/src/elementary-number-theory/multiplication-integers.lagda.md
index 8d9326a47b..5306775315 100644
--- a/src/elementary-number-theory/multiplication-integers.lagda.md
+++ b/src/elementary-number-theory/multiplication-integers.lagda.md
@@ -29,6 +29,7 @@ open import foundation.function-types
open import foundation.identity-types
open import foundation.injective-maps
open import foundation.interchange-law
+open import foundation.sets
open import foundation.transport-along-identifications
open import foundation.type-arithmetic-empty-type
open import foundation.unit-type
diff --git a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md
index 43f53a469c..8e9ffa25d0 100644
--- a/src/elementary-number-theory/multiplication-natural-numbers.lagda.md
+++ b/src/elementary-number-theory/multiplication-natural-numbers.lagda.md
@@ -19,6 +19,7 @@ open import foundation.identity-types
open import foundation.injective-maps
open import foundation.interchange-law
open import foundation.negated-equality
+open import foundation.sets
```
diff --git a/src/elementary-number-theory/nonzero-integers.lagda.md b/src/elementary-number-theory/nonzero-integers.lagda.md
index d5b2dcd318..2d7dd2e4a9 100644
--- a/src/elementary-number-theory/nonzero-integers.lagda.md
+++ b/src/elementary-number-theory/nonzero-integers.lagda.md
@@ -38,7 +38,7 @@ holds.
```agda
is-nonzero-prop-ℤ : ℤ → Prop lzero
-is-nonzero-prop-ℤ k = neg-Prop' (k = zero-ℤ)
+is-nonzero-prop-ℤ k = neg-type-Prop (k = zero-ℤ)
is-nonzero-ℤ : ℤ → UU lzero
is-nonzero-ℤ k = type-Prop (is-nonzero-prop-ℤ k)
diff --git a/src/elementary-number-theory/pisano-periods.lagda.md b/src/elementary-number-theory/pisano-periods.lagda.md
index d3d99b98ee..f765a59fb3 100644
--- a/src/elementary-number-theory/pisano-periods.lagda.md
+++ b/src/elementary-number-theory/pisano-periods.lagda.md
@@ -25,6 +25,7 @@ open import foundation.equivalences
open import foundation.identity-types
open import foundation.injective-maps
open import foundation.repetitions-sequences
+open import foundation.sets
open import foundation.universe-levels
open import univalent-combinatorics.cartesian-product-types
diff --git a/src/elementary-number-theory/prime-numbers.lagda.md b/src/elementary-number-theory/prime-numbers.lagda.md
index a7543fbf13..767f0d0529 100644
--- a/src/elementary-number-theory/prime-numbers.lagda.md
+++ b/src/elementary-number-theory/prime-numbers.lagda.md
@@ -157,7 +157,7 @@ has-unique-proper-divisor-is-prime-ℕ n H =
fundamental-theorem-id'
( λ x p → pr2 (H x) (inv p))
( λ x →
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-ℕ 1 x)
( is-prop-is-proper-divisor-ℕ n x)
( λ p → inv (pr1 (H x) p)))
diff --git a/src/elementary-number-theory/standard-cyclic-rings.lagda.md b/src/elementary-number-theory/standard-cyclic-rings.lagda.md
index 692df2b922..60ffeb3ce6 100644
--- a/src/elementary-number-theory/standard-cyclic-rings.lagda.md
+++ b/src/elementary-number-theory/standard-cyclic-rings.lagda.md
@@ -179,7 +179,7 @@ is-generating-element-one-ℤ-Mod n =
is-cyclic-ℤ-Mod-Group :
( n : ℕ) → is-cyclic-Group (ℤ-Mod-Group n)
is-cyclic-ℤ-Mod-Group n =
- intro-∃
+ intro-exists
( one-ℤ-Mod n)
( is-generating-element-one-ℤ-Mod n)
diff --git a/src/elementary-number-theory/strict-inequality-integer-fractions.lagda.md b/src/elementary-number-theory/strict-inequality-integer-fractions.lagda.md
index 0801560c95..966dad6794 100644
--- a/src/elementary-number-theory/strict-inequality-integer-fractions.lagda.md
+++ b/src/elementary-number-theory/strict-inequality-integer-fractions.lagda.md
@@ -26,6 +26,7 @@ open import elementary-number-theory.strict-inequality-integers
open import foundation.action-on-identifications-functions
open import foundation.cartesian-product-types
+open import foundation.conjunction
open import foundation.coproduct-types
open import foundation.decidable-propositions
open import foundation.dependent-pair-types
@@ -292,9 +293,9 @@ module _
where
dense-le-fraction-ℤ :
- ∃ fraction-ℤ (λ r → le-fraction-ℤ x r × le-fraction-ℤ r y)
+ exists fraction-ℤ (λ r → le-fraction-ℤ-Prop x r ∧ le-fraction-ℤ-Prop r y)
dense-le-fraction-ℤ =
- intro-∃
+ intro-exists
( mediant-fraction-ℤ x y)
( le-left-mediant-fraction-ℤ x y H , le-right-mediant-fraction-ℤ x y H)
```
@@ -307,7 +308,8 @@ module _
where
located-le-fraction-ℤ :
- le-fraction-ℤ y z → (le-fraction-ℤ-Prop y x) ∨ (le-fraction-ℤ-Prop x z)
+ le-fraction-ℤ y z →
+ type-disjunction-Prop (le-fraction-ℤ-Prop y x) (le-fraction-ℤ-Prop x z)
located-le-fraction-ℤ H =
unit-trunc-Prop
( map-coproduct
diff --git a/src/elementary-number-theory/strict-inequality-rational-numbers.lagda.md b/src/elementary-number-theory/strict-inequality-rational-numbers.lagda.md
index 0f342d811e..e6c91bbdf5 100644
--- a/src/elementary-number-theory/strict-inequality-rational-numbers.lagda.md
+++ b/src/elementary-number-theory/strict-inequality-rational-numbers.lagda.md
@@ -27,6 +27,7 @@ open import elementary-number-theory.strict-inequality-integers
open import foundation.binary-relations
open import foundation.cartesian-product-types
+open import foundation.conjunction
open import foundation.coproduct-types
open import foundation.decidable-propositions
open import foundation.dependent-pair-types
@@ -197,28 +198,30 @@ module _
(x : ℚ)
where
- left-∃-le-ℚ : ∃ ℚ (λ q → le-ℚ q x)
- left-∃-le-ℚ = intro-∃
- ( rational-fraction-ℤ frac)
- ( preserves-le-left-rational-fraction-ℤ x frac
- ( le-fraction-le-numerator-fraction-ℤ
- ( frac)
- ( fraction-ℚ x)
- ( refl)
- ( le-pred-ℤ (numerator-ℚ x))))
+ exists-lesser-ℚ : exists ℚ (λ q → le-ℚ-Prop q x)
+ exists-lesser-ℚ =
+ intro-exists
+ ( rational-fraction-ℤ frac)
+ ( preserves-le-left-rational-fraction-ℤ x frac
+ ( le-fraction-le-numerator-fraction-ℤ
+ ( frac)
+ ( fraction-ℚ x)
+ ( refl)
+ ( le-pred-ℤ (numerator-ℚ x))))
where
frac : fraction-ℤ
frac = (pred-ℤ (numerator-ℚ x) , positive-denominator-ℚ x)
- right-∃-le-ℚ : ∃ ℚ (λ r → le-ℚ x r)
- right-∃-le-ℚ = intro-∃
- ( rational-fraction-ℤ frac)
- ( preserves-le-right-rational-fraction-ℤ x frac
- ( le-fraction-le-numerator-fraction-ℤ
- ( fraction-ℚ x)
- ( frac)
- ( refl)
- ( le-succ-ℤ (numerator-ℚ x))))
+ exists-greater-ℚ : exists ℚ (λ r → le-ℚ-Prop x r)
+ exists-greater-ℚ =
+ intro-exists
+ ( rational-fraction-ℤ frac)
+ ( preserves-le-right-rational-fraction-ℤ x frac
+ ( le-fraction-le-numerator-fraction-ℤ
+ ( fraction-ℚ x)
+ ( frac)
+ ( refl)
+ ( le-succ-ℤ (numerator-ℚ x))))
where
frac : fraction-ℤ
frac = (succ-ℤ (numerator-ℚ x) , positive-denominator-ℚ x)
@@ -282,9 +285,9 @@ module _
(x y : ℚ) (H : le-ℚ x y)
where
- dense-le-ℚ : ∃ ℚ (λ r → le-ℚ x r × le-ℚ r y)
+ dense-le-ℚ : exists ℚ (λ r → le-ℚ-Prop x r ∧ le-ℚ-Prop r y)
dense-le-ℚ =
- intro-∃
+ intro-exists
( mediant-ℚ x y)
( le-left-mediant-ℚ x y H , le-right-mediant-ℚ x y H)
```
@@ -292,7 +295,8 @@ module _
### Strict inequality on the rational numbers is located
```agda
-located-le-ℚ : (x y z : ℚ) → le-ℚ y z → (le-ℚ-Prop y x) ∨ (le-ℚ-Prop x z)
+located-le-ℚ :
+ (x y z : ℚ) → le-ℚ y z → type-disjunction-Prop (le-ℚ-Prop y x) (le-ℚ-Prop x z)
located-le-ℚ x y z H =
unit-trunc-Prop
( map-coproduct
diff --git a/src/elementary-number-theory/well-ordering-principle-standard-finite-types.lagda.md b/src/elementary-number-theory/well-ordering-principle-standard-finite-types.lagda.md
index fd45692c37..a422c722f0 100644
--- a/src/elementary-number-theory/well-ordering-principle-standard-finite-types.lagda.md
+++ b/src/elementary-number-theory/well-ordering-principle-standard-finite-types.lagda.md
@@ -55,13 +55,13 @@ numbers.
### For any decidable family `P` over `Fin n`, if `P x` doesn't hold for all `x` then there exists an `x` for which `P x` is false
```agda
-exists-not-not-forall-Fin :
+exists-not-not-for-all-Fin :
{l : Level} (k : ℕ) {P : Fin k → UU l} → (is-decidable-fam P) →
¬ ((x : Fin k) → P x) → Σ (Fin k) (λ x → ¬ (P x))
-exists-not-not-forall-Fin {l} zero-ℕ d H = ex-falso (H ind-empty)
-exists-not-not-forall-Fin {l} (succ-ℕ k) {P} d H with d (inr star)
+exists-not-not-for-all-Fin {l} zero-ℕ d H = ex-falso (H ind-empty)
+exists-not-not-for-all-Fin {l} (succ-ℕ k) {P} d H with d (inr star)
... | inl p =
- T ( exists-not-not-forall-Fin k
+ T ( exists-not-not-for-all-Fin k
( λ x → d (inl x))
( λ f → H (ind-coproduct P f (ind-unit p))))
where
@@ -69,13 +69,13 @@ exists-not-not-forall-Fin {l} (succ-ℕ k) {P} d H with d (inr star)
T z = pair (inl (pr1 z)) (pr2 z)
... | inr f = pair (inr star) f
-exists-not-not-forall-count :
+exists-not-not-for-all-count :
{l1 l2 : Level} {X : UU l1} (P : X → UU l2) →
(is-decidable-fam P) → count X →
¬ ((x : X) → P x) → Σ X (λ x → ¬ (P x))
-exists-not-not-forall-count {l1} {l2} {X} P p e =
+exists-not-not-for-all-count {l1} {l2} {X} P p e =
( g) ∘
- ( ( exists-not-not-forall-Fin
+ ( ( exists-not-not-for-all-Fin
( number-of-elements-count e)
( p ∘ map-equiv-count e)) ∘ f)
where
@@ -180,7 +180,7 @@ well-ordering-principle-Σ-Fin (succ-ℕ k) {P} d (pair (inr star) p)
well-ordering-principle-∃-Fin :
{l : Level} (k : ℕ) (P : decidable-subtype l (Fin k)) →
- ∃ (Fin k) (is-in-decidable-subtype P) →
+ exists (Fin k) (subtype-decidable-subtype P) →
minimal-element-Fin k (is-in-decidable-subtype P)
well-ordering-principle-∃-Fin k P H =
apply-universal-property-trunc-Prop H
diff --git a/src/finite-algebra/semisimple-commutative-finite-rings.lagda.md b/src/finite-algebra/semisimple-commutative-finite-rings.lagda.md
index 2f4184550c..ddfea4c796 100644
--- a/src/finite-algebra/semisimple-commutative-finite-rings.lagda.md
+++ b/src/finite-algebra/semisimple-commutative-finite-rings.lagda.md
@@ -44,8 +44,7 @@ is-semisimple-Commutative-Ring-𝔽 l2 R =
exists
( ℕ)
( λ n →
- exists-Prop
- ( Fin n → Field-𝔽 l2)
+ ∃ ( Fin n → Field-𝔽 l2)
( λ A →
trunc-Prop
( hom-Commutative-Ring-𝔽
diff --git a/src/finite-group-theory/orbits-permutations.lagda.md b/src/finite-group-theory/orbits-permutations.lagda.md
index 525f6e0278..4b629bb9c1 100644
--- a/src/finite-group-theory/orbits-permutations.lagda.md
+++ b/src/finite-group-theory/orbits-permutations.lagda.md
@@ -799,13 +799,17 @@ module _
(pair k
( (equal-iterate-transposition-other-orbits k) ∙
( p)))))
- ( is-equiv-is-prop is-prop-type-trunc-Prop is-prop-type-trunc-Prop
- ( λ P' → apply-universal-property-trunc-Prop P'
- ( prop-equivalence-relation (same-orbits-permutation-count g) x y)
- ( λ (pair k p) → unit-trunc-Prop
- ( pair k
- ( (inv (equal-iterate-transposition-other-orbits k)) ∙
- ( p))))))
+ ( is-equiv-has-converse-is-prop
+ ( is-prop-type-trunc-Prop)
+ ( is-prop-type-trunc-Prop)
+ ( λ P' →
+ apply-universal-property-trunc-Prop P'
+ ( prop-equivalence-relation (same-orbits-permutation-count g) x y)
+ ( λ (pair k p) →
+ unit-trunc-Prop
+ ( ( k) ,
+ ( (inv (equal-iterate-transposition-other-orbits k)) ∙
+ ( p))))))
where
equal-iterate-transposition-other-orbits :
(k : ℕ) →
@@ -1124,7 +1128,7 @@ module _
apply-universal-property-trunc-Prop T
( coproduct-sim-equivalence-relation-a-b-Prop g P x)
(λ pa → lemma2 g (pair (pr1 pa) (inl (pr2 pa)))))
- ( is-equiv-is-prop is-prop-type-trunc-Prop
+ ( is-equiv-has-converse-is-prop is-prop-type-trunc-Prop
( is-prop-type-Prop
( coproduct-sim-equivalence-relation-a-b-Prop g P x))
( λ where
diff --git a/src/finite-group-theory/transpositions.lagda.md b/src/finite-group-theory/transpositions.lagda.md
index faa40ddb52..51817489c5 100644
--- a/src/finite-group-theory/transpositions.lagda.md
+++ b/src/finite-group-theory/transpositions.lagda.md
@@ -290,7 +290,7 @@ module _
element-is-not-identity-map-transposition :
Σ X (λ x → map-transposition Y x ≠ x)
element-is-not-identity-map-transposition =
- exists-not-not-forall-count
+ exists-not-not-for-all-count
( λ z → Id (map-transposition Y z) z)
( λ x → has-decidable-equality-count eX (map-transposition Y x) x)
( eX)
diff --git a/src/foundation-core/coproduct-types.lagda.md b/src/foundation-core/coproduct-types.lagda.md
index 2c6b56433a..02461897b7 100644
--- a/src/foundation-core/coproduct-types.lagda.md
+++ b/src/foundation-core/coproduct-types.lagda.md
@@ -14,8 +14,8 @@ open import foundation.universe-levels
## Idea
-The coproduct of two types `A` and `B` can be thought of as the disjoint union
-of `A` and `B`.
+The {{#concept "coproduct" Disambiguation="of types"}} of two types `A` and `B`
+can be thought of as the disjoint union of `A` and `B`.
## Definition
diff --git a/src/foundation-core/decidable-propositions.lagda.md b/src/foundation-core/decidable-propositions.lagda.md
index cd694d7194..7134f32b4b 100644
--- a/src/foundation-core/decidable-propositions.lagda.md
+++ b/src/foundation-core/decidable-propositions.lagda.md
@@ -28,8 +28,9 @@ open import foundation-core.subtypes
## Idea
-A **decidable proposition** is a [proposition](foundation-core.propositions.md)
-that has a [decidable](foundation.decidable-types.md) underlying type.
+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
@@ -52,7 +53,7 @@ pr2 (is-decidable-Prop P) = is-prop-is-decidable (is-prop-type-Prop P)
is-prop-is-decidable-prop :
{l : Level} (X : UU l) → is-prop (is-decidable-prop X)
is-prop-is-decidable-prop X =
- is-prop-is-inhabited
+ is-prop-has-element
( λ H →
is-prop-product
( is-prop-is-prop X)
@@ -121,6 +122,38 @@ pr1 unit-Decidable-Prop = unit
pr2 unit-Decidable-Prop = is-decidable-prop-unit
```
+### The product of two decidable propositions is a decidable proposition
+
+```agda
+module _
+ {l1 l2 : Level} (P : Decidable-Prop l1) (Q : Decidable-Prop l2)
+ where
+
+ type-product-Decidable-Prop : UU (l1 ⊔ l2)
+ type-product-Decidable-Prop =
+ type-product-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
+
+ is-prop-product-Decidable-Prop : is-prop type-product-Decidable-Prop
+ is-prop-product-Decidable-Prop =
+ is-prop-product-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
+
+ is-decidable-product-Decidable-Prop : is-decidable type-product-Decidable-Prop
+ is-decidable-product-Decidable-Prop =
+ is-decidable-product
+ ( is-decidable-Decidable-Prop P)
+ ( is-decidable-Decidable-Prop Q)
+
+ is-decidable-prop-product-Decidable-Prop :
+ is-decidable-prop type-product-Decidable-Prop
+ pr1 is-decidable-prop-product-Decidable-Prop = is-prop-product-Decidable-Prop
+ pr2 is-decidable-prop-product-Decidable-Prop =
+ is-decidable-product-Decidable-Prop
+
+ product-Decidable-Prop : Decidable-Prop (l1 ⊔ l2)
+ pr1 product-Decidable-Prop = type-product-Decidable-Prop
+ pr2 product-Decidable-Prop = is-decidable-prop-product-Decidable-Prop
+```
+
### Decidability of a propositional truncation
```agda
@@ -151,7 +184,7 @@ is-merely-decidable-is-decidable-trunc-Prop :
is-decidable (type-trunc-Prop A) → is-merely-decidable A
is-merely-decidable-is-decidable-trunc-Prop A (inl x) =
apply-universal-property-trunc-Prop x
- ( is-merely-Decidable-Prop A)
+ ( is-merely-decidable-Prop A)
( unit-trunc-Prop ∘ inl)
is-merely-decidable-is-decidable-trunc-Prop A (inr f) =
unit-trunc-Prop (inr (f ∘ unit-trunc-Prop))
diff --git a/src/foundation-core/injective-maps.lagda.md b/src/foundation-core/injective-maps.lagda.md
index a47703cb60..ea67f3fc16 100644
--- a/src/foundation-core/injective-maps.lagda.md
+++ b/src/foundation-core/injective-maps.lagda.md
@@ -17,11 +17,8 @@ open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.homotopies
open import foundation-core.identity-types
-open import foundation-core.propositional-maps
-open import foundation-core.propositions
open import foundation-core.retractions
open import foundation-core.sections
-open import foundation-core.sets
```
@@ -148,63 +145,6 @@ module _
is-injective-emb e {x} {y} = map-inv-is-equiv (is-emb-map-emb e x y)
```
-### Any injective map between sets is an embedding
-
-```agda
-abstract
- is-emb-is-injective' :
- {l1 l2 : Level} {A : UU l1} (is-set-A : is-set A)
- {B : UU l2} (is-set-B : is-set B) (f : A → B) →
- is-injective f → is-emb f
- is-emb-is-injective' is-set-A is-set-B f is-injective-f x y =
- is-equiv-is-prop
- ( is-set-A x y)
- ( is-set-B (f x) (f y))
- ( is-injective-f)
-
- is-set-is-injective :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
- is-set B → is-injective f → is-set A
- is-set-is-injective {f = f} H I =
- is-set-prop-in-id
- ( λ x y → f x = f y)
- ( λ x y → H (f x) (f y))
- ( λ x → refl)
- ( λ x y → I)
-
- is-emb-is-injective :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
- is-set B → is-injective f → is-emb f
- is-emb-is-injective {f = f} H I =
- is-emb-is-injective' (is-set-is-injective H I) H f I
-
- is-prop-map-is-injective :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → B} →
- is-set B → is-injective f → is-prop-map f
- is-prop-map-is-injective {f = f} H I =
- is-prop-map-is-emb (is-emb-is-injective H I)
-```
-
-### For a map between sets, being injective is a property
-
-```agda
-module _
- {l1 l2 : Level} {A : UU l1} {B : UU l2}
- where
-
- is-prop-is-injective :
- is-set A → (f : A → B) → is-prop (is-injective f)
- is-prop-is-injective H f =
- is-prop-implicit-Π
- ( λ x →
- is-prop-implicit-Π
- ( λ y → is-prop-function-type (H x y)))
-
- is-injective-Prop : is-set A → (A → B) → Prop (l1 ⊔ l2)
- pr1 (is-injective-Prop H f) = is-injective f
- pr2 (is-injective-Prop H f) = is-prop-is-injective H f
-```
-
### Any map out of a contractible type is injective
```agda
diff --git a/src/foundation-core/negation.lagda.md b/src/foundation-core/negation.lagda.md
index e565a46dfc..7021a1e464 100644
--- a/src/foundation-core/negation.lagda.md
+++ b/src/foundation-core/negation.lagda.md
@@ -16,14 +16,17 @@ open import foundation-core.empty-types
## Idea
-The Curry-Howard interpretation of negation in type theory is the interpretation
-of the proposition `P ⇒ ⊥` using propositions as types. Thus, the negation of a
-type `A` is the type `A → empty`.
+The Curry-Howard interpretation of _negation_ in type theory is the
+interpretation of the proposition `P ⇒ ⊥` using propositions as types. Thus, the
+{{#concept "negation" Disambiguation="of a type" WDID=Q190558 WD="logical negation" Agda=¬_}}
+of a type `A` is the type `A → empty`.
## Definition
```agda
-¬ : {l : Level} → UU l → UU l
+infix 25 ¬_
+
+¬_ : {l : Level} → UU l → UU l
¬ A = A → empty
map-neg :
@@ -31,3 +34,8 @@ map-neg :
(P → Q) → (¬ Q → ¬ P)
map-neg f nq p = nq (f p)
```
+
+## External links
+
+- [negation](https://ncatlab.org/nlab/show/negation) at $n$Lab
+- [Negation](https://en.wikipedia.org/wiki/Negation) at Wikipedia
diff --git a/src/foundation-core/propositions.lagda.md b/src/foundation-core/propositions.lagda.md
index 2afe3a74bd..041ce6881d 100644
--- a/src/foundation-core/propositions.lagda.md
+++ b/src/foundation-core/propositions.lagda.md
@@ -25,16 +25,24 @@ open import foundation-core.transport-along-identifications
## Idea
-A type is considered to be a proposition if its identity types are contractible.
-This condition is equivalent to the condition that it has up to identification
-at most one element.
+A type is a {{#concept "proposition" Agda=is-prop}} if its
+[identity types](foundation-core.identity-types.md) are
+[contractible](foundation-core.contractible-types.md). This condition is
+[equivalent](foundation-core.equivalences.md) to the condition that it has up to
+identification at most one element.
-## Definition
+## Definitions
+
+### The predicate of being a proposition
```agda
is-prop : {l : Level} (A : UU l) → UU l
is-prop A = (x y : A) → is-contr (x = y)
+```
+### The type of propositions
+
+```agda
Prop :
(l : Level) → UU (lsuc l)
Prop l = Σ (UU l) is-prop
@@ -56,20 +64,18 @@ module _
We prove here only that any contractible type is a proposition. The fact that
the empty type and the unit type are propositions can be found in
-```text
-foundation.empty-types
-foundation.unit-type
-```
+- [`foundation.empty-types`](foundation.empty-types.md), and
+- [`foundation.unit-type`](foundation.unit-type.md).
## Properties
-### To show that a type is a proposition, we may assume it is inhabited
+### To show that a type is a proposition we may assume it has an element
```agda
abstract
- is-prop-is-inhabited :
+ is-prop-has-element :
{l1 : Level} {X : UU l1} → (X → is-prop X) → is-prop X
- is-prop-is-inhabited f x y = f x x y
+ is-prop-has-element f x y = f x x y
```
### Equivalent characterizations of propositions
@@ -122,29 +128,6 @@ module _
eq-is-proof-irrelevant = eq-is-prop' ∘ is-prop-is-proof-irrelevant
```
-### A map between propositions is an equivalence if there is a map in the reverse direction
-
-```agda
-module _
- {l1 l2 : Level} {A : UU l1} {B : UU l2}
- where
-
- abstract
- is-equiv-is-prop :
- is-prop A → is-prop B → {f : A → B} → (B → A) → is-equiv f
- is-equiv-is-prop is-prop-A is-prop-B {f} g =
- is-equiv-is-invertible
- ( g)
- ( λ y → eq-is-prop is-prop-B)
- ( λ x → eq-is-prop is-prop-A)
-
- abstract
- equiv-prop : is-prop A → is-prop B → (A → B) → (B → A) → A ≃ B
- pr1 (equiv-prop is-prop-A is-prop-B f g) = f
- pr2 (equiv-prop is-prop-A is-prop-B f g) =
- is-equiv-is-prop is-prop-A is-prop-B g
-```
-
### Propositions are closed under equivalences
```agda
@@ -192,8 +175,7 @@ abstract
( λ p → K (pr1 y) (tr _ p (pr2 x)) (pr2 y)))
Σ-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : type-Prop P → Prop l2) →
- Prop (l1 ⊔ l2)
+ {l1 l2 : Level} (P : Prop l1) (Q : type-Prop P → Prop l2) → Prop (l1 ⊔ l2)
pr1 (Σ-Prop P Q) = Σ (type-Prop P) (λ p → type-Prop (Q p))
pr2 (Σ-Prop P Q) =
is-prop-Σ
@@ -210,10 +192,19 @@ abstract
is-prop A → is-prop B → is-prop (A × B)
is-prop-product H K = is-prop-Σ H (λ x → K)
-product-Prop : {l1 l2 : Level} → Prop l1 → Prop l2 → Prop (l1 ⊔ l2)
-pr1 (product-Prop P Q) = type-Prop P × type-Prop Q
-pr2 (product-Prop P Q) =
- is-prop-product (is-prop-type-Prop P) (is-prop-type-Prop Q)
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ type-product-Prop : UU (l1 ⊔ l2)
+ type-product-Prop = type-Prop P × type-Prop Q
+
+ is-prop-product-Prop : is-prop type-product-Prop
+ is-prop-product-Prop =
+ is-prop-product (is-prop-type-Prop P) (is-prop-type-Prop Q)
+
+ product-Prop : Prop (l1 ⊔ l2)
+ product-Prop = (type-product-Prop , is-prop-product-Prop)
```
### Products of families of propositions are propositions
@@ -234,15 +225,15 @@ module _
type-Π-Prop : UU (l1 ⊔ l2)
type-Π-Prop = (x : A) → type-Prop (P x)
- is-prop-type-Π-Prop : is-prop type-Π-Prop
- is-prop-type-Π-Prop = is-prop-Π (λ x → is-prop-type-Prop (P x))
+ is-prop-Π-Prop : is-prop type-Π-Prop
+ is-prop-Π-Prop = is-prop-Π (λ x → is-prop-type-Prop (P x))
Π-Prop : Prop (l1 ⊔ l2)
pr1 Π-Prop = type-Π-Prop
- pr2 Π-Prop = is-prop-type-Π-Prop
+ pr2 Π-Prop = is-prop-Π-Prop
```
-We repeat the above for implicit Π-types.
+We now repeat the above for implicit Π-types.
```agda
abstract
@@ -262,13 +253,13 @@ module _
type-implicit-Π-Prop : UU (l1 ⊔ l2)
type-implicit-Π-Prop = {x : A} → type-Prop (P x)
- is-prop-type-implicit-Π-Prop : is-prop type-implicit-Π-Prop
- is-prop-type-implicit-Π-Prop =
+ is-prop-implicit-Π-Prop : is-prop type-implicit-Π-Prop
+ is-prop-implicit-Π-Prop =
is-prop-implicit-Π (λ x → is-prop-type-Prop (P x))
implicit-Π-Prop : Prop (l1 ⊔ l2)
pr1 implicit-Π-Prop = type-implicit-Π-Prop
- pr2 implicit-Π-Prop = is-prop-type-implicit-Π-Prop
+ pr2 implicit-Π-Prop = is-prop-implicit-Π-Prop
```
### The type of functions into a proposition is a proposition
@@ -284,41 +275,33 @@ type-function-Prop :
{l1 l2 : Level} → UU l1 → Prop l2 → UU (l1 ⊔ l2)
type-function-Prop A P = A → type-Prop P
-is-prop-type-function-Prop :
- {l1 l2 : Level} (A : UU l1) (P : Prop l2) →
+is-prop-function-Prop :
+ {l1 l2 : Level} {A : UU l1} (P : Prop l2) →
is-prop (type-function-Prop A P)
-is-prop-type-function-Prop A P =
+is-prop-function-Prop P =
is-prop-function-type (is-prop-type-Prop P)
function-Prop :
{l1 l2 : Level} → UU l1 → Prop l2 → Prop (l1 ⊔ l2)
pr1 (function-Prop A P) = type-function-Prop A P
-pr2 (function-Prop A P) = is-prop-type-function-Prop A P
+pr2 (function-Prop A P) = is-prop-function-Prop P
type-hom-Prop :
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2) → UU (l1 ⊔ l2)
type-hom-Prop P = type-function-Prop (type-Prop P)
-is-prop-type-hom-Prop :
+is-prop-hom-Prop :
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
is-prop (type-hom-Prop P Q)
-is-prop-type-hom-Prop P = is-prop-type-function-Prop (type-Prop P)
+is-prop-hom-Prop P = is-prop-function-Prop
hom-Prop :
{l1 l2 : Level} → Prop l1 → Prop l2 → Prop (l1 ⊔ l2)
pr1 (hom-Prop P Q) = type-hom-Prop P Q
-pr2 (hom-Prop P Q) = is-prop-type-hom-Prop P Q
-
-implication-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → Prop (l1 ⊔ l2)
-implication-Prop = hom-Prop
-
-type-implication-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → UU (l1 ⊔ l2)
-type-implication-Prop = type-hom-Prop
+pr2 (hom-Prop P Q) = is-prop-hom-Prop P Q
infixr 5 _⇒_
-_⇒_ = type-implication-Prop
+_⇒_ = hom-Prop
```
### The type of equivalences between two propositions is a proposition
@@ -386,3 +369,61 @@ is-prop-Prop : {l : Level} (A : UU l) → Prop l
pr1 (is-prop-Prop A) = is-prop A
pr2 (is-prop-Prop A) = is-prop-is-prop A
```
+
+## See also
+
+### Operations on propositions
+
+There is a wide range of operations on propositions due to the rich structure of
+intuitionistic logic. Below we give a structured overview of a notable selection
+of such operations and their notation in the library.
+
+The list is split into two sections, the first consists of operations that
+generalize to arbitrary types and even sufficiently nice
+[subuniverses](foundation.subuniverses.md), such as
+$n$-[types](foundation-core.truncated-types.md).
+
+| Name | Operator on types | Operator on propositions/subtypes |
+| ----------------------------------------------------------- | ----------------- | --------------------------------- |
+| [Dependent sum](foundation.dependent-pair-types.md) | `Σ` | `Σ-Prop` |
+| [Dependent product](foundation.dependent-function-types.md) | `Π` | `Π-Prop` |
+| [Functions](foundation-core.function-types.md) | `→` | `⇒` |
+| [Logical equivalence](foundation.logical-equivalences.md) | `↔` | `⇔` |
+| [Product](foundation-core.cartesian-product-types.md) | `×` | `product-Prop` |
+| [Join](synthetic-homotopy-theory.joins-of-types.md) | `*` | `join-Prop` |
+| [Exclusive sum](foundation.exclusive-sum.md) | `exclusive-sum` | `exclusive-sum-Prop` |
+| [Coproduct](foundation-core.coproduct-types.md) | `+` | _N/A_ |
+
+Note that for many operations in the second section, there is an equivalent
+operation on propositions in the first.
+
+| Name | Operator on types | Operator on propositions/subtypes |
+| ---------------------------------------------------------------------------- | --------------------------- | ---------------------------------------- |
+| [Initial object](foundation-core.empty-types.md) | `empty` | `empty-Prop` |
+| [Terminal object](foundation.unit-type.md) | `unit` | `unit-Prop` |
+| [Existential quantification](foundation.existential-quantification.md) | `exists-structure` | `∃` |
+| [Unique existential quantification](foundation.uniqueness-quantification.md) | `uniquely-exists-structure` | `∃!` |
+| [Universal quantification](foundation.universal-quantification.md) | | `∀'` (equivalent to `Π-Prop`) |
+| [Conjunction](foundation.conjunction.md) | | `∧` (equivalent to `product-Prop`) |
+| [Disjunction](foundation.disjunction.md) | `disjunction-type` | `∨` (equivalent to `join-Prop`) |
+| [Exclusive disjunction](foundation.exclusive-disjunction.md) | `xor-type` | `⊻` (equivalent to `exclusive-sum-Prop`) |
+| [Negation](foundation.negation.md) | `¬` | `¬'` |
+| [Double negation](foundation.double-negation.md) | `¬¬` | `¬¬'` |
+
+We can also organize these operations by indexed and binary variants, giving us
+the following table:
+
+| Name | Binary | Indexed |
+| ---------------------- | ------ | ------- |
+| Product | `×` | `Π` |
+| Conjunction | `∧` | `∀'` |
+| Constructive existence | `+` | `Σ` |
+| Existence | `∨` | `∃` |
+| Unique existence | `⊻` | `∃!` |
+
+### Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation-core/sets.lagda.md b/src/foundation-core/sets.lagda.md
index 603e8413ec..27390c8d90 100644
--- a/src/foundation-core/sets.lagda.md
+++ b/src/foundation-core/sets.lagda.md
@@ -12,8 +12,10 @@ open import foundation.fundamental-theorem-of-identity-types
open import foundation.universe-levels
open import foundation-core.contractible-types
+open import foundation-core.embeddings
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.truncated-types
open import foundation-core.truncation-levels
@@ -46,8 +48,7 @@ module _
is-set-type-Set = pr2 X
Id-Prop : (x y : type-Set) → Prop l
- pr1 (Id-Prop x y) = (x = y)
- pr2 (Id-Prop x y) = is-set-type-Set x y
+ Id-Prop x y = (x = y , is-set-type-Set x y)
```
## Properties
@@ -166,3 +167,21 @@ module _
pr2 equiv-set-Set =
is-set-equiv-is-set (is-set-type-Set A) (is-set-type-Set B)
```
+
+### If a type injects into a set, then it is a set
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ abstract
+ is-set-is-injective :
+ {f : A → B} → is-set B → is-injective f → is-set A
+ is-set-is-injective {f} H I =
+ is-set-prop-in-id
+ ( λ x y → f x = f y)
+ ( λ x y → H (f x) (f y))
+ ( λ x → refl)
+ ( λ x y → I)
+```
diff --git a/src/foundation-core/subtypes.lagda.md b/src/foundation-core/subtypes.lagda.md
index a6a943e0c6..3aa30a0277 100644
--- a/src/foundation-core/subtypes.lagda.md
+++ b/src/foundation-core/subtypes.lagda.md
@@ -316,7 +316,11 @@ equiv-type-subtype :
pr1 (equiv-type-subtype is-subtype-P is-subtype-Q f g) = tot f
pr2 (equiv-type-subtype is-subtype-P is-subtype-Q f g) =
is-equiv-tot-is-fiberwise-equiv {f = f}
- ( λ x → is-equiv-is-prop (is-subtype-P x) (is-subtype-Q x) (g x))
+ ( λ x →
+ is-equiv-has-converse-is-prop
+ ( is-subtype-P x)
+ ( is-subtype-Q x)
+ ( g x))
```
### Equivalences of subtypes
@@ -342,7 +346,11 @@ abstract
is-equiv f → ((x : A) → (Q (f x)) → P x) → is-equiv (map-Σ Q f g)
is-equiv-subtype-is-equiv {Q = Q} is-subtype-P is-subtype-Q f g is-equiv-f h =
is-equiv-map-Σ Q is-equiv-f
- ( λ x → is-equiv-is-prop (is-subtype-P x) (is-subtype-Q (f x)) (h x))
+ ( λ x →
+ is-equiv-has-converse-is-prop
+ ( is-subtype-P x)
+ ( is-subtype-Q (f x))
+ ( h x))
abstract
is-equiv-subtype-is-equiv' :
@@ -356,6 +364,9 @@ abstract
is-equiv-subtype-is-equiv' {P = P} {Q}
is-subtype-P is-subtype-Q f g is-equiv-f h =
is-equiv-map-Σ Q is-equiv-f
- ( λ x → is-equiv-is-prop (is-subtype-P x) (is-subtype-Q (f x))
- ( (tr P (is-retraction-map-inv-is-equiv is-equiv-f x)) ∘ (h (f x))))
+ ( λ x →
+ is-equiv-has-converse-is-prop
+ ( is-subtype-P x)
+ ( is-subtype-Q (f x))
+ ( (tr P (is-retraction-map-inv-is-equiv is-equiv-f x)) ∘ (h (f x))))
```
diff --git a/src/foundation.lagda.md b/src/foundation.lagda.md
index b4ebc64e2e..ecda4a6b37 100644
--- a/src/foundation.lagda.md
+++ b/src/foundation.lagda.md
@@ -54,6 +54,7 @@ open import foundation.category-of-sets public
open import foundation.choice-of-representatives-equivalence-relation public
open import foundation.codiagonal-maps-of-types public
open import foundation.coherently-invertible-maps public
+open import foundation.coinhabited-pairs-of-types public
open import foundation.commuting-cubes-of-maps public
open import foundation.commuting-hexagons-of-identifications public
open import foundation.commuting-pentagons-of-identifications public
@@ -164,6 +165,7 @@ open import foundation.equivalences-spans public
open import foundation.equivalences-spans-families-of-types public
open import foundation.evaluation-functions public
open import foundation.exclusive-disjunction public
+open import foundation.exclusive-sum public
open import foundation.existential-quantification public
open import foundation.exponents-set-quotients public
open import foundation.extensions-types public
@@ -203,6 +205,7 @@ open import foundation.homotopies public
open import foundation.homotopies-morphisms-arrows public
open import foundation.homotopy-algebra public
open import foundation.homotopy-induction public
+open import foundation.homotopy-preorder-of-types public
open import foundation.identity-systems public
open import foundation.identity-truncated-types public
open import foundation.identity-types public
@@ -247,6 +250,8 @@ open import foundation.maybe public
open import foundation.mere-embeddings public
open import foundation.mere-equality public
open import foundation.mere-equivalences public
+open import foundation.mere-functions public
+open import foundation.mere-logical-equivalences public
open import foundation.monomorphisms public
open import foundation.morphisms-arrows public
open import foundation.morphisms-binary-relations public
@@ -409,8 +414,8 @@ 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.unions-subtypes public
-open import foundation.unique-existence public
open import foundation.uniqueness-image public
+open import foundation.uniqueness-quantification public
open import foundation.uniqueness-set-quotients public
open import foundation.uniqueness-set-truncations public
open import foundation.uniqueness-truncation public
@@ -441,6 +446,7 @@ open import foundation.universal-property-set-quotients public
open import foundation.universal-property-set-truncation public
open import foundation.universal-property-truncation public
open import foundation.universal-property-unit-type public
+open import foundation.universal-quantification public
open import foundation.universe-levels public
open import foundation.unordered-pairs public
open import foundation.unordered-pairs-of-types public
diff --git a/src/foundation/apartness-relations.lagda.md b/src/foundation/apartness-relations.lagda.md
index 973e7c3afe..baf80f0f6b 100644
--- a/src/foundation/apartness-relations.lagda.md
+++ b/src/foundation/apartness-relations.lagda.md
@@ -12,6 +12,7 @@ open import foundation.dependent-pair-types
open import foundation.disjunction
open import foundation.existential-quantification
open import foundation.propositional-truncations
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.cartesian-product-types
@@ -53,9 +54,12 @@ module _
is-consistent : UU (l1 ⊔ l2)
is-consistent = (a b : A) → (a = b) → ¬ (type-Prop (R a b))
+ is-cotransitive-Prop : Prop (l1 ⊔ l2)
+ is-cotransitive-Prop =
+ ∀' A (λ a → ∀' A (λ b → ∀' A (λ c → R a b ⇒ (R a c ∨ R b c))))
+
is-cotransitive : UU (l1 ⊔ l2)
- is-cotransitive =
- (a b c : A) → type-hom-Prop (R a b) (disjunction-Prop (R a c) (R b c))
+ is-cotransitive = type-Prop is-cotransitive-Prop
is-apartness-relation : UU (l1 ⊔ l2)
is-apartness-relation =
@@ -150,7 +154,7 @@ module _
rel-apart-function-into-Type-With-Apartness :
Relation-Prop (l1 ⊔ l3) (X → type-Type-With-Apartness Y)
rel-apart-function-into-Type-With-Apartness f g =
- ∃-Prop X (λ x → apart-Type-With-Apartness Y (f x) (g x))
+ ∃ X (λ x → rel-apart-Type-With-Apartness Y (f x) (g x))
apart-function-into-Type-With-Apartness :
Relation (l1 ⊔ l3) (X → type-Type-With-Apartness Y)
@@ -202,16 +206,8 @@ module _
( rel-apart-function-into-Type-With-Apartness X Y f h)
( rel-apart-function-into-Type-With-Apartness X Y g h))
( λ where
- ( inl b) →
- inl-disjunction-Prop
- ( rel-apart-function-into-Type-With-Apartness X Y f h)
- ( rel-apart-function-into-Type-With-Apartness X Y g h)
- ( unit-trunc-Prop (x , b))
- ( inr b) →
- inr-disjunction-Prop
- ( rel-apart-function-into-Type-With-Apartness X Y f h)
- ( rel-apart-function-into-Type-With-Apartness X Y g h)
- ( unit-trunc-Prop (x , b))))
+ ( inl b) → inl-disjunction (intro-exists x b)
+ ( inr b) → inr-disjunction (intro-exists x b)))
exp-Type-With-Apartness : Type-With-Apartness (l1 ⊔ l2) (l1 ⊔ l3)
pr1 exp-Type-With-Apartness = X → type-Type-With-Apartness Y
diff --git a/src/foundation/axiom-of-choice.lagda.md b/src/foundation/axiom-of-choice.lagda.md
index 1d28e42414..a7bd0fea94 100644
--- a/src/foundation/axiom-of-choice.lagda.md
+++ b/src/foundation/axiom-of-choice.lagda.md
@@ -10,6 +10,7 @@ module foundation.axiom-of-choice where
open import foundation.dependent-pair-types
open import foundation.function-extensionality
open import foundation.functoriality-propositional-truncation
+open import foundation.inhabited-types
open import foundation.postcomposition-functions
open import foundation.projective-types
open import foundation.propositional-truncations
@@ -17,7 +18,6 @@ open import foundation.sections
open import foundation.split-surjective-maps
open import foundation.surjective-maps
open import foundation.universe-levels
-open import foundation.whiskering-homotopies-composition
open import foundation-core.equivalences
open import foundation-core.fibers-of-maps
@@ -32,8 +32,10 @@ open import foundation-core.sets
## Idea
-The axiom of choice asserts that for every family of inhabited types indexed by
-a set, the type of sections of that family is inhabited.
+The {{#concept "axiom of choice" Agda=AC-0}} 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
@@ -43,8 +45,8 @@ a set, the type of sections of that family is inhabited.
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) → type-trunc-Prop (type-Set (B x))) →
- type-trunc-Prop ((x : type-Set A) → type-Set (B x))
+ ((x : type-Set A) → is-inhabited (type-Set (B x))) →
+ is-inhabited ((x : type-Set A) → type-Set (B x))
```
### The axiom of choice
@@ -53,8 +55,8 @@ AC-Set l1 l2 =
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) → type-trunc-Prop (B x)) →
- type-trunc-Prop ((x : type-Set A) → B x)
+ ((x : type-Set A) → is-inhabited (B x)) →
+ is-inhabited ((x : type-Set A) → B x)
```
## Properties
diff --git a/src/foundation/booleans.lagda.md b/src/foundation/booleans.lagda.md
index 844bfeeb29..704c90a92d 100644
--- a/src/foundation/booleans.lagda.md
+++ b/src/foundation/booleans.lagda.md
@@ -8,6 +8,7 @@ module foundation.booleans where
```agda
open import foundation.dependent-pair-types
+open import foundation.involutions
open import foundation.negated-equality
open import foundation.raising-universe-levels
open import foundation.unit-type
@@ -23,6 +24,7 @@ open import foundation-core.identity-types
open import foundation-core.injective-maps
open import foundation-core.negation
open import foundation-core.propositions
+open import foundation-core.sections
open import foundation-core.sets
open import univalent-combinatorics.finite-types
@@ -237,9 +239,9 @@ neq-neg-bool false ()
### Boolean negation is an involution
```agda
-neg-neg-bool : (neg-bool ∘ neg-bool) ~ id
-neg-neg-bool true = refl
-neg-neg-bool false = refl
+is-involution-neg-bool : is-involution neg-bool
+is-involution-neg-bool true = refl
+is-involution-neg-bool false = refl
```
### Boolean negation is an equivalence
@@ -247,10 +249,7 @@ neg-neg-bool false = refl
```agda
abstract
is-equiv-neg-bool : is-equiv neg-bool
- pr1 (pr1 is-equiv-neg-bool) = neg-bool
- pr2 (pr1 is-equiv-neg-bool) = neg-neg-bool
- pr1 (pr2 is-equiv-neg-bool) = neg-bool
- pr2 (pr2 is-equiv-neg-bool) = neg-neg-bool
+ is-equiv-neg-bool = is-equiv-is-involution is-involution-neg-bool
equiv-neg-bool : bool ≃ bool
pr1 equiv-neg-bool = neg-bool
@@ -261,8 +260,12 @@ pr2 equiv-neg-bool = is-equiv-neg-bool
```agda
abstract
- not-equiv-const :
+ no-section-const-bool : (b : bool) → ¬ (section (const bool bool b))
+ no-section-const-bool true (g , G) = neq-true-false-bool (G false)
+ no-section-const-bool false (g , G) = neq-false-true-bool (G true)
+
+abstract
+ is-not-equiv-const-bool :
(b : bool) → ¬ (is-equiv (const bool bool b))
- not-equiv-const true ((g , G) , _) = neq-true-false-bool (G false)
- not-equiv-const false ((g , G) , _) = neq-false-true-bool (G true)
+ is-not-equiv-const-bool b e = no-section-const-bool b (section-is-equiv e)
```
diff --git a/src/foundation/coherently-invertible-maps.lagda.md b/src/foundation/coherently-invertible-maps.lagda.md
index 4a590a17f1..e2e34ab431 100644
--- a/src/foundation/coherently-invertible-maps.lagda.md
+++ b/src/foundation/coherently-invertible-maps.lagda.md
@@ -13,6 +13,7 @@ open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.universe-levels
open import foundation.whiskering-higher-homotopies-composition
@@ -111,7 +112,7 @@ module _
is-equiv-is-coherently-invertible-is-equiv :
is-equiv (is-coherently-invertible-is-equiv {f = f})
is-equiv-is-coherently-invertible-is-equiv =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-property-is-equiv f)
( is-prop-is-coherently-invertible)
( is-equiv-is-coherently-invertible)
diff --git a/src/foundation/coinhabited-pairs-of-types.lagda.md b/src/foundation/coinhabited-pairs-of-types.lagda.md
new file mode 100644
index 0000000000..bb9be37fb9
--- /dev/null
+++ b/src/foundation/coinhabited-pairs-of-types.lagda.md
@@ -0,0 +1,104 @@
+# Coinhabited pairs of types
+
+```agda
+module foundation.coinhabited-pairs-of-types where
+```
+
+Imports
+
+```agda
+open import foundation.inhabited-types
+open import foundation.logical-equivalences
+open import foundation.propositional-truncations
+open import foundation.universe-levels
+
+open import foundation-core.propositions
+```
+
+
+
+## Idea
+
+Two types `A` and `B` are said to be
+{{#concept "coinhabited" Agda=is-coinhabited}} if `A` is
+[inhabited](foundation.inhabited-types.md)
+[if and only if](foundation.logical-equivalences.md) `B` is.
+
+## Definitions
+
+### The predicate of being coinhabited
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ is-coinhabited-Prop : Prop (l1 ⊔ l2)
+ is-coinhabited-Prop = iff-Prop (is-inhabited-Prop A) (is-inhabited-Prop B)
+
+ is-coinhabited : UU (l1 ⊔ l2)
+ is-coinhabited = type-Prop is-coinhabited-Prop
+```
+
+### Forward and backward implications of coinhabited types
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ forward-implication-is-coinhabited :
+ is-coinhabited A B → is-inhabited A → is-inhabited B
+ forward-implication-is-coinhabited = forward-implication
+
+ forward-implication-is-coinhabited' : is-coinhabited A B → A → is-inhabited B
+ forward-implication-is-coinhabited' e a =
+ forward-implication e (unit-trunc-Prop a)
+
+ backward-implication-is-coinhabited :
+ is-coinhabited A B → is-inhabited B → is-inhabited A
+ backward-implication-is-coinhabited = backward-implication
+
+ backward-implication-is-coinhabited' : is-coinhabited A B → B → is-inhabited A
+ backward-implication-is-coinhabited' e b =
+ backward-implication e (unit-trunc-Prop b)
+```
+
+### Every type is coinhabited with itself
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ is-reflexive-is-coinhabited : is-coinhabited A A
+ is-reflexive-is-coinhabited = id-iff
+```
+
+### Coinhabitedness is a transitive relation
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ where
+
+ is-transitive-is-coinhabited :
+ is-coinhabited B C → is-coinhabited A B → is-coinhabited A C
+ is-transitive-is-coinhabited = _∘iff_
+```
+
+### Coinhabitedness is a symmetric relation
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-symmetric-is-coinhabited : is-coinhabited A B → is-coinhabited B A
+ is-symmetric-is-coinhabited = inv-iff
+```
+
+## See also
+
+- [Mere logical equivalence of types](foundation.mere-logical-equivalences.md)
+ is a related but stronger notion than coinhabitedness.
diff --git a/src/foundation/conjunction.lagda.md b/src/foundation/conjunction.lagda.md
index 23f062e227..54c4cc00c9 100644
--- a/src/foundation/conjunction.lagda.md
+++ b/src/foundation/conjunction.lagda.md
@@ -1,4 +1,4 @@
-# Conjunction of propositions
+# Conjunction
```agda
module foundation.conjunction where
@@ -10,11 +10,14 @@ module foundation.conjunction where
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.logical-equivalences
+open import foundation.propositional-truncations
+open import foundation.universal-property-cartesian-product-types
open import foundation.universe-levels
open import foundation-core.cartesian-product-types
open import foundation-core.decidable-propositions
open import foundation-core.equivalences
+open import foundation-core.function-types
open import foundation-core.propositions
```
@@ -22,77 +25,204 @@ open import foundation-core.propositions
## Idea
-The **conjunction** of two [propositions](foundation-core.propositions.md) `P`
-and `Q` is the proposition that both `P` and `Q` hold.
+The
+{{#concept "conjunction" Disambiguation="of propositions" WDID=Q191081 WD="logical conjunction" Agda=conjunction-Prop}}
+`P ∧ Q` of two [propositions](foundation-core.propositions.md) `P` and `Q` is
+the proposition that both `P` and `Q` hold and is thus defined by the
+[cartesian product](foundation-core.cartesian-product-types.md) of their
+underlying types
-## Definition
+```text
+ P ∧ Q := P × Q
+```
+
+The conjunction of two propositions satisfies the universal property of the
+[meet](order-theory.greatest-lower-bounds-large-posets.md) in the
+[locale of propositions](foundation.large-locale-of-propositions.md). This means
+that any span of propositions over `P` and `Q` factor (uniquely) through the
+conjunction
+
+```text
+ R
+ / ∶ \
+ / ∶ \
+ ∨ ∨ ∨
+ P <--- P ∧ Q ---> Q.
+```
+
+In other words, we have a
+[logical equivalence](foundation.logical-equivalences.md)
+
+```text
+ (R → P) ∧ (R → Q) ↔ (R → P ∧ Q)
+```
+
+for every proposition `R`. In fact, `R` can be any type.
+
+The
+{{#concept "recursion principle" Disambiguation"of the conjunction of propositions" Agda=elimination-principle-conjunction-Prop}}
+of the conjunction of propositions states that to define a function `P ∧ Q → R`
+into a proposition `R`, or indeed any type, is equivalent to defining a function
+`P → Q → R`.
+
+## Definitions
+
+### The conjunction
```agda
-conjunction-Prop = product-Prop
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ conjunction-Prop : Prop (l1 ⊔ l2)
+ conjunction-Prop = product-Prop P Q
-type-conjunction-Prop : {l1 l2 : Level} → Prop l1 → Prop l2 → UU (l1 ⊔ l2)
-type-conjunction-Prop P Q = type-Prop (conjunction-Prop P Q)
+ type-conjunction-Prop : UU (l1 ⊔ l2)
+ type-conjunction-Prop = type-Prop conjunction-Prop
-abstract
- is-prop-type-conjunction-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- is-prop (type-conjunction-Prop P Q)
- is-prop-type-conjunction-Prop P Q = is-prop-type-Prop (conjunction-Prop P Q)
+ is-prop-conjunction-Prop :
+ is-prop type-conjunction-Prop
+ is-prop-conjunction-Prop = is-prop-type-Prop conjunction-Prop
-infixr 15 _∧_
-_∧_ = type-conjunction-Prop
+ infixr 15 _∧_
+ _∧_ : Prop (l1 ⊔ l2)
+ _∧_ = conjunction-Prop
```
**Note**: The symbol used for the conjunction `_∧_` is the
-[logical and](https://codepoints.net/U+2227) `∧` (agda-input: `\wedge` `\and`).
+[logical and](https://codepoints.net/U+2227) `∧` (agda-input: `\wedge` or
+`\and`).
+
+### The conjunction of decidable propositions
```agda
-conjunction-Decidable-Prop :
- {l1 l2 : Level} → Decidable-Prop l1 → Decidable-Prop l2 →
- Decidable-Prop (l1 ⊔ l2)
-pr1 (conjunction-Decidable-Prop P Q) =
- type-conjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr1 (pr2 (conjunction-Decidable-Prop P Q)) =
- is-prop-type-conjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr2 (pr2 (conjunction-Decidable-Prop P Q)) =
- is-decidable-product
- ( is-decidable-Decidable-Prop P)
- ( is-decidable-Decidable-Prop Q)
+module _
+ {l1 l2 : Level} (P : Decidable-Prop l1) (Q : Decidable-Prop l2)
+ where
+
+ is-decidable-conjunction-Decidable-Prop :
+ is-decidable
+ ( type-conjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q))
+ is-decidable-conjunction-Decidable-Prop =
+ is-decidable-product
+ ( is-decidable-Decidable-Prop P)
+ ( is-decidable-Decidable-Prop Q)
+
+ conjunction-Decidable-Prop : Decidable-Prop (l1 ⊔ l2)
+ conjunction-Decidable-Prop =
+ ( type-conjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q) ,
+ is-prop-conjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q) ,
+ is-decidable-conjunction-Decidable-Prop)
```
-## Properties
+### The introduction rule and projections for the conjunction of propositions
-### Introduction rule for conjunction
+While we define the introduction rule and projections for the conjunction below,
+we advice users to use the standard pairing and projection functions for the
+cartesian product types: `pair`, `pr1` and `pr2`.
```agda
-intro-conjunction-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- type-Prop P → type-Prop Q → type-conjunction-Prop P Q
-pr1 (intro-conjunction-Prop P Q p q) = p
-pr2 (intro-conjunction-Prop P Q p q) = q
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ intro-conjunction-Prop : type-Prop P → type-Prop Q → type-conjunction-Prop P Q
+ intro-conjunction-Prop = pair
+
+ pr1-conjunction-Prop : type-conjunction-Prop P Q → type-Prop P
+ pr1-conjunction-Prop = pr1
+
+ pr2-conjunction-Prop : type-conjunction-Prop P Q → type-Prop Q
+ pr2-conjunction-Prop = pr2
```
-### The universal property of conjunction
+### The elimination principle of the conjunction
```agda
-iff-universal-property-conjunction-Prop :
+module _
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
- {l3 : Level} (R : Prop l3) →
- ( type-hom-Prop R P × type-hom-Prop R Q) ↔
- ( type-hom-Prop R (conjunction-Prop P Q))
-pr1 (pr1 (iff-universal-property-conjunction-Prop P Q R) (f , g) r) = f r
-pr2 (pr1 (iff-universal-property-conjunction-Prop P Q R) (f , g) r) = g r
-pr1 (pr2 (iff-universal-property-conjunction-Prop P Q R) h) r = pr1 (h r)
-pr2 (pr2 (iff-universal-property-conjunction-Prop P Q R) h) r = pr2 (h r)
-
-equiv-universal-property-conjunction-Prop :
+ where
+
+ ev-conjunction-Prop :
+ {l : Level} (R : Prop l) → type-Prop (((P ∧ Q) ⇒ R) ⇒ P ⇒ Q ⇒ R)
+ ev-conjunction-Prop R = ev-pair
+
+ elimination-principle-conjunction-Prop : UUω
+ elimination-principle-conjunction-Prop =
+ {l : Level} (R : Prop l) → has-converse (ev-conjunction-Prop R)
+```
+
+## Properties
+
+### The conjunction satisfies the recursion principle of the conjunction
+
+```agda
+module _
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
- {l3 : Level} (R : Prop l3) →
- ( type-hom-Prop R P × type-hom-Prop R Q) ≃
- ( type-hom-Prop R (conjunction-Prop P Q))
-equiv-universal-property-conjunction-Prop P Q R =
- equiv-iff'
- ( conjunction-Prop (hom-Prop R P) (hom-Prop R Q))
- ( hom-Prop R (conjunction-Prop P Q))
- ( iff-universal-property-conjunction-Prop P Q R)
+ where
+
+ elim-conjunction-Prop : elimination-principle-conjunction-Prop P Q
+ elim-conjunction-Prop R f (p , q) = f p q
```
+
+### The conjunction is the meet in the locale of propositions
+
+```agda
+module _
+ {l1 l2 l3 : Level} (P : Prop l1) (Q : Prop l2) (C : UU l3)
+ where
+
+ map-distributive-conjunction-Prop :
+ type-conjunction-Prop (function-Prop C P) (function-Prop C Q) →
+ C → type-conjunction-Prop P Q
+ map-distributive-conjunction-Prop (f , g) y = (f y , g y)
+
+ map-inv-distributive-conjunction-Prop :
+ (C → type-conjunction-Prop P Q) →
+ type-conjunction-Prop (function-Prop C P) (function-Prop C Q)
+ map-inv-distributive-conjunction-Prop f = (pr1 ∘ f , pr2 ∘ f)
+
+ is-equiv-map-distributive-conjunction-Prop :
+ is-equiv map-distributive-conjunction-Prop
+ is-equiv-map-distributive-conjunction-Prop =
+ is-equiv-has-converse
+ ( conjunction-Prop (function-Prop C P) (function-Prop C Q))
+ ( function-Prop C (conjunction-Prop P Q))
+ ( map-inv-distributive-conjunction-Prop)
+
+ distributive-conjunction-Prop :
+ type-conjunction-Prop (function-Prop C P) (function-Prop C Q) ≃
+ (C → type-conjunction-Prop P Q)
+ distributive-conjunction-Prop =
+ ( map-distributive-conjunction-Prop ,
+ is-equiv-map-distributive-conjunction-Prop)
+
+module _
+ {l1 l2 l3 : Level} (P : Prop l1) (Q : Prop l2) (R : Prop l3)
+ where
+
+ is-greatest-binary-lower-bound-conjunction-Prop :
+ type-Prop (((R ⇒ P) ∧ (R ⇒ Q)) ⇔ (R ⇒ P ∧ Q))
+ is-greatest-binary-lower-bound-conjunction-Prop =
+ ( map-distributive-conjunction-Prop P Q (type-Prop R) ,
+ map-inv-distributive-conjunction-Prop P Q (type-Prop R))
+```
+
+## See also
+
+- The indexed counterpart to conjunction is
+ [universal quantification](foundation.universal-quantification.md).
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [logical conjunction](https://ncatlab.org/nlab/show/logical+conjunction) at
+ $n$Lab
+- [Logical conjunction](https://en.wikipedia.org/wiki/Logical_conjunction) at
+ Wikipedia
diff --git a/src/foundation/contractible-types.lagda.md b/src/foundation/contractible-types.lagda.md
index e489774da7..5d78e164f1 100644
--- a/src/foundation/contractible-types.lagda.md
+++ b/src/foundation/contractible-types.lagda.md
@@ -12,6 +12,7 @@ open import foundation-core.contractible-types public
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.function-extensionality
+open import foundation.logical-equivalences
open import foundation.subuniverses
open import foundation.unit-type
open import foundation.universe-levels
@@ -123,7 +124,7 @@ is-closed-under-is-contr-subuniverse P =
equiv-is-contr-equiv :
{l1 l2 : Level} {A : UU l1} {B : UU l2} → A ≃ B → is-contr A ≃ is-contr B
equiv-is-contr-equiv {A = A} {B = B} e =
- equiv-prop
+ equiv-iff-is-prop
( is-property-is-contr)
( is-property-is-contr)
( is-contr-retract-of A
diff --git a/src/foundation/coproduct-types.lagda.md b/src/foundation/coproduct-types.lagda.md
index a64a9b8d9b..14b74377c5 100644
--- a/src/foundation/coproduct-types.lagda.md
+++ b/src/foundation/coproduct-types.lagda.md
@@ -214,7 +214,8 @@ module _
coproduct-Prop :
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
(type-Prop P → ¬ (type-Prop Q)) → Prop (l1 ⊔ l2)
-pr1 (coproduct-Prop P Q H) = type-Prop P + type-Prop Q
+pr1 (coproduct-Prop P Q H) =
+ type-Prop P + type-Prop Q
pr2 (coproduct-Prop P Q H) =
is-prop-coproduct H (is-prop-type-Prop P) (is-prop-type-Prop Q)
```
diff --git a/src/foundation/decidable-embeddings.lagda.md b/src/foundation/decidable-embeddings.lagda.md
index 4bc7896cbb..d18d5f6df8 100644
--- a/src/foundation/decidable-embeddings.lagda.md
+++ b/src/foundation/decidable-embeddings.lagda.md
@@ -17,6 +17,7 @@ open import foundation.functoriality-cartesian-product-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.propositional-maps
open import foundation.structured-type-duality
open import foundation.subtype-identity-principle
@@ -219,7 +220,7 @@ abstract
{l1 l2 : Level} {A : UU l1} {B : UU l2} (f : A → B) →
is-prop (is-decidable-emb f)
is-prop-is-decidable-emb f =
- is-prop-is-inhabited
+ is-prop-has-element
( λ H →
is-prop-product
( is-property-is-emb f)
@@ -322,7 +323,7 @@ equiv-precomp-decidable-emb-equiv e C =
( is-decidable-emb)
( equiv-precomp e C)
( λ g →
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-decidable-emb g)
( is-prop-is-decidable-emb (g ∘ map-equiv e))
( is-decidable-emb-comp (is-decidable-emb-is-equiv (pr2 e)))
diff --git a/src/foundation/decidable-equality.lagda.md b/src/foundation/decidable-equality.lagda.md
index 5dbe76a61d..39b6f711fc 100644
--- a/src/foundation/decidable-equality.lagda.md
+++ b/src/foundation/decidable-equality.lagda.md
@@ -12,9 +12,11 @@ open import foundation.coproduct-types
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.double-negation
+open import foundation.injective-maps
open import foundation.negation
open import foundation.retracts-of-types
open import foundation.sections
+open import foundation.sets
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.unit-type
open import foundation.universe-levels
@@ -25,9 +27,7 @@ open import foundation-core.equality-dependent-pair-types
open import foundation-core.equivalences
open import foundation-core.fibers-of-maps
open import foundation-core.identity-types
-open import foundation-core.injective-maps
open import foundation-core.propositions
-open import foundation-core.sets
open import foundation-core.transport-along-identifications
```
@@ -223,7 +223,7 @@ abstract
is-prop-has-decidable-equality :
{l1 : Level} {X : UU l1} → is-prop (has-decidable-equality X)
is-prop-has-decidable-equality {l1} {X} =
- is-prop-is-inhabited
+ is-prop-has-element
( λ d →
is-prop-Π
( λ x →
diff --git a/src/foundation/decidable-equivalence-relations.lagda.md b/src/foundation/decidable-equivalence-relations.lagda.md
index 9addc2138d..948b84936d 100644
--- a/src/foundation/decidable-equivalence-relations.lagda.md
+++ b/src/foundation/decidable-equivalence-relations.lagda.md
@@ -128,7 +128,7 @@ module _
sim-Decidable-equivalence-relation x y ≃
sim-Decidable-equivalence-relation y x
equiv-symmetric-Decidable-equivalence-relation {x} {y} =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-sim-Decidable-equivalence-relation x y)
( is-prop-sim-Decidable-equivalence-relation y x)
( symmetric-Decidable-equivalence-relation x y)
@@ -169,7 +169,9 @@ module _
is-equivalence-class-Decidable-equivalence-relation :
decidable-subtype l2 X → UU (l1 ⊔ lsuc l2)
is-equivalence-class-Decidable-equivalence-relation P =
- ∃ X (λ x → P = decidable-relation-Decidable-equivalence-relation R x)
+ exists-structure
+ ( X)
+ ( λ x → P = decidable-relation-Decidable-equivalence-relation R x)
equivalence-class-Decidable-equivalence-relation : UU (l1 ⊔ lsuc l2)
equivalence-class-Decidable-equivalence-relation =
@@ -180,7 +182,7 @@ module _
pr1 (class-Decidable-equivalence-relation x) =
decidable-relation-Decidable-equivalence-relation R x
pr2 (class-Decidable-equivalence-relation x) =
- intro-∃ x refl
+ intro-exists x refl
emb-equivalence-class-Decidable-equivalence-relation :
equivalence-class-Decidable-equivalence-relation ↪ decidable-subtype l2 X
diff --git a/src/foundation/decidable-propositions.lagda.md b/src/foundation/decidable-propositions.lagda.md
index 02de050cce..877d8c5626 100644
--- a/src/foundation/decidable-propositions.lagda.md
+++ b/src/foundation/decidable-propositions.lagda.md
@@ -142,8 +142,8 @@ equiv-universes-Decidable-Prop l l' =
iff-universes-Decidable-Prop :
(l l' : Level) (P : Decidable-Prop l) →
- ( prop-Decidable-Prop P ⇔
- prop-Decidable-Prop (map-equiv (equiv-universes-Decidable-Prop l l') P))
+ ( type-Decidable-Prop P) ↔
+ ( type-Decidable-Prop (map-equiv (equiv-universes-Decidable-Prop l l') P))
pr1 (iff-universes-Decidable-Prop l l' P) p =
map-inv-equiv
( compute-equiv-bool-Decidable-Prop
diff --git a/src/foundation/decidable-subtypes.lagda.md b/src/foundation/decidable-subtypes.lagda.md
index 20472c465b..b8c8bce746 100644
--- a/src/foundation/decidable-subtypes.lagda.md
+++ b/src/foundation/decidable-subtypes.lagda.md
@@ -157,12 +157,12 @@ module _
iff-universes-decidable-subtype :
(l l' : Level) (S : decidable-subtype l X) →
( (x : X) →
- prop-Decidable-Prop (S x) ⇔
- prop-Decidable-Prop
+ type-Decidable-Prop (S x) ↔
+ type-Decidable-Prop
( map-equiv (equiv-universes-decidable-subtype l l') S x))
iff-universes-decidable-subtype l l' S x =
tr
- ( λ P → prop-Decidable-Prop (S x) ⇔ prop-Decidable-Prop P)
+ ( λ P → type-Decidable-Prop (S x) ↔ type-Decidable-Prop P)
( inv
( compute-map-equiv-Π
( λ _ → Decidable-Prop l')
@@ -246,7 +246,7 @@ module _
(P = Q) ≃ has-same-elements-decidable-subtype Q
extensionality-decidable-subtype =
extensionality-Π P
- ( λ x Q → prop-Decidable-Prop (P x) ⇔ prop-Decidable-Prop Q)
+ ( λ x Q → (type-Decidable-Prop (P x)) ↔ (type-Decidable-Prop Q))
( λ x Q → extensionality-Decidable-Prop (P x) Q)
has-same-elements-eq-decidable-subtype :
diff --git a/src/foundation/decidable-types.lagda.md b/src/foundation/decidable-types.lagda.md
index 4ffe535d77..8073a1ae80 100644
--- a/src/foundation/decidable-types.lagda.md
+++ b/src/foundation/decidable-types.lagda.md
@@ -61,13 +61,15 @@ is-inhabited-or-empty A = type-trunc-Prop A + is-empty A
### Merely decidable types
-A type `A` is said to be merely decidable if it comes equipped with an element
-of `type-trunc-Prop (is-decidable A)`.
+A type `A` is said to be
+{{#concept "merely decidable" Agda=is-merely-decidable}} if it comes equipped
+with an element of `║ is-decidable A ║₋₁`, or equivalently, the
+[disjunction](foundation.disjunction.md) `A ∨ ¬ A` holds.
```agda
-is-merely-Decidable-Prop :
+is-merely-decidable-Prop :
{l : Level} → UU l → Prop l
-is-merely-Decidable-Prop A = trunc-Prop (is-decidable A)
+is-merely-decidable-Prop A = trunc-Prop (is-decidable A)
is-merely-decidable : {l : Level} → UU l → UU l
is-merely-decidable A = type-trunc-Prop (is-decidable A)
@@ -208,8 +210,7 @@ double-negation-elim-is-decidable (inr x) p = ex-falso (p x)
```agda
double-negation-is-decidable : {l : Level} {P : UU l} → ¬¬ (is-decidable P)
double-negation-is-decidable {P = P} f =
- map-neg (inr {A = P} {B = ¬ P}) f
- ( map-neg (inl {A = P} {B = ¬ P}) f)
+ map-neg (inr {A = P} {B = ¬ P}) f (map-neg (inl {A = P} {B = ¬ P}) f)
```
### Decidable types have ε-operators
@@ -236,9 +237,9 @@ idempotent-is-decidable P (inr np) = inr (λ p → np (inl p))
```agda
abstract
- is-prop-is-inhabited-or-empty :
+ is-property-is-inhabited-or-empty :
{l1 : Level} (A : UU l1) → is-prop (is-inhabited-or-empty A)
- is-prop-is-inhabited-or-empty A =
+ is-property-is-inhabited-or-empty A =
is-prop-coproduct
( λ t → apply-universal-property-trunc-Prop t empty-Prop)
( is-prop-type-trunc-Prop)
@@ -246,7 +247,7 @@ abstract
is-inhabited-or-empty-Prop : {l1 : Level} → UU l1 → Prop l1
pr1 (is-inhabited-or-empty-Prop A) = is-inhabited-or-empty A
-pr2 (is-inhabited-or-empty-Prop A) = is-prop-is-inhabited-or-empty A
+pr2 (is-inhabited-or-empty-Prop A) = is-property-is-inhabited-or-empty A
```
### Any inhabited type is a fixed point for `is-decidable`
diff --git a/src/foundation/discrete-types.lagda.md b/src/foundation/discrete-types.lagda.md
index ae29cad8b8..c9c1243aa0 100644
--- a/src/foundation/discrete-types.lagda.md
+++ b/src/foundation/discrete-types.lagda.md
@@ -39,7 +39,7 @@ module _
where
rel-apart-Discrete-Type : Relation-Prop l (type-Discrete-Type X)
- rel-apart-Discrete-Type x y = neg-Prop' (x = y)
+ rel-apart-Discrete-Type x y = neg-type-Prop (x = y)
apart-Discrete-Type : (x y : type-Discrete-Type X) → UU l
apart-Discrete-Type x y = type-Prop (rel-apart-Discrete-Type x y)
diff --git a/src/foundation/disjunction.lagda.md b/src/foundation/disjunction.lagda.md
index 8cc5a61022..21a764be49 100644
--- a/src/foundation/disjunction.lagda.md
+++ b/src/foundation/disjunction.lagda.md
@@ -1,4 +1,4 @@
-# Disjunction of propositions
+# Disjunction
```agda
module foundation.disjunction where
@@ -7,12 +7,15 @@ module foundation.disjunction where
Imports
```agda
-open import foundation.conjunction
open import foundation.decidable-types
open import foundation.dependent-pair-types
+open import foundation.functoriality-coproduct-types
+open import foundation.inhabited-types
+open import foundation.logical-equivalences
open import foundation.propositional-truncations
open import foundation.universe-levels
+open import foundation-core.cartesian-product-types
open import foundation-core.coproduct-types
open import foundation-core.decidable-propositions
open import foundation-core.empty-types
@@ -25,96 +28,214 @@ open import foundation-core.propositions
## Idea
-The disjunction of two propositions `P` and `Q` is the proposition that `P`
-holds or `Q` holds.
+The
+{{#concept "disjunction" Disambiguation="of propositions" WDID=Q1651704 WD="logical disjunction" Agda=disjunction-Prop}}
+of two [propositions](foundation-core.propositions.md) `P` and `Q` is the
+proposition that `P` holds or `Q` holds, and is defined as
+[propositional truncation](foundation.propositional-truncations.md) of the
+[coproduct](foundation-core.coproduct-types.md) of their underlying types
-## Definition
+```text
+ P ∨ Q := ║ P + Q ║₋₁
+```
+
+The
+{{#concept "universal property" Disambiguation="of the disjunction" Agda=universal-property-disjunction-Prop}}
+of the disjunction states that, for every proposition `R`, the evaluation map
+
+```text
+ ev : ((P ∨ Q) → R) → ((P → R) ∧ (Q → R))
+```
+
+is a [logical equivalence](foundation.logical-equivalences.md), and thus the
+disjunction is the least upper bound of `P` and `Q` in the
+[locale of propositions](foundation.large-locale-of-propositions.md): there is a
+pair of logical implications `P → R` and `Q → R`, if and only if `P ∨ Q` implies
+`R`
+
+```text
+P ---> P ∨ Q <--- Q
+ \ ∶ /
+ \ ∶ /
+ ∨ ∨ ∨
+ R.
+```
+
+## Definitions
+
+### The disjunction of arbitrary types
+
+Given arbitrary types `A` and `B`, the truncation
+
+```text
+ ║ A + B ║₋₁
+```
+
+satisfies the universal property of
+
+```text
+ ║ A ║₋₁ ∨ ║ B ║₋₁
+```
+
+and is thus equivalent to it. Therefore, we may reasonably call this
+construction the
+{{#concept "disjunction" Disambiguation="of types" Agda=disjunction-type-Prop}}
+of types. It is important to keep in mind that this is not a generalization of
+the concept but rather a conflation, and should be read as the statement _`A` or
+`B` is (merely) [inhabited](foundation.inhabited-types.md)_.
+
+Because propositions are a special case of types, and Agda can generally infer
+types for us, we will continue to conflate the two in our formalizations for the
+benefit that we have to specify the propositions in our code less often. For
+instance, even though the introduction rules for disjunction are phrased in
+terms of disjunction of types, they equally apply to disjunction of
+propositions.
```agda
-disjunction-Prop : {l1 l2 : Level} → Prop l1 → Prop l2 → Prop (l1 ⊔ l2)
-disjunction-Prop P Q = trunc-Prop (type-Prop P + type-Prop Q)
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
-type-disjunction-Prop : {l1 l2 : Level} → Prop l1 → Prop l2 → UU (l1 ⊔ l2)
-type-disjunction-Prop P Q = type-Prop (disjunction-Prop P Q)
+ disjunction-type-Prop : Prop (l1 ⊔ l2)
+ disjunction-type-Prop = trunc-Prop (A + B)
-abstract
- is-prop-type-disjunction-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- is-prop (type-disjunction-Prop P Q)
- is-prop-type-disjunction-Prop P Q = is-prop-type-Prop (disjunction-Prop P Q)
+ disjunction-type : UU (l1 ⊔ l2)
+ disjunction-type = type-Prop disjunction-type-Prop
-infixr 10 _∨_
-_∨_ = type-disjunction-Prop
+ is-prop-disjunction-type : is-prop disjunction-type
+ is-prop-disjunction-type = is-prop-type-Prop disjunction-type-Prop
```
-**Note**: The symbol used for the disjunction `_∨_` is the
+### The disjunction
+
+```agda
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ disjunction-Prop : Prop (l1 ⊔ l2)
+ disjunction-Prop = disjunction-type-Prop (type-Prop P) (type-Prop Q)
+
+ type-disjunction-Prop : UU (l1 ⊔ l2)
+ type-disjunction-Prop = type-Prop disjunction-Prop
+
+ abstract
+ is-prop-disjunction-Prop : is-prop type-disjunction-Prop
+ is-prop-disjunction-Prop = is-prop-type-Prop disjunction-Prop
+
+ infixr 10 _∨_
+ _∨_ : Prop (l1 ⊔ l2)
+ _∨_ = disjunction-Prop
+```
+
+**Notation.** The symbol used for the disjunction `_∨_` is the
[logical or](https://codepoints.net/U+2228) `∨` (agda-input: `\vee` `\or`), and
not the [latin small letter v](https://codepoints.net/U+0076) `v`.
+### The introduction rules for the disjunction
+
```agda
-disjunction-Decidable-Prop :
- {l1 l2 : Level} →
- Decidable-Prop l1 → Decidable-Prop l2 → Decidable-Prop (l1 ⊔ l2)
-pr1 (disjunction-Decidable-Prop P Q) =
- type-disjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr1 (pr2 (disjunction-Decidable-Prop P Q)) =
- is-prop-type-disjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr2 (pr2 (disjunction-Decidable-Prop P Q)) =
- is-decidable-trunc-Prop-is-merely-decidable
- ( type-Decidable-Prop P + type-Decidable-Prop Q)
- ( unit-trunc-Prop
- ( is-decidable-coproduct
- ( is-decidable-Decidable-Prop P)
- ( is-decidable-Decidable-Prop Q)))
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ inl-disjunction : A → disjunction-type A B
+ inl-disjunction = unit-trunc-Prop ∘ inl
+
+ inr-disjunction : B → disjunction-type A B
+ inr-disjunction = unit-trunc-Prop ∘ inr
+```
+
+**Note.** Even though the introduction rules are formalized in terms of
+disjunction of types, it equally applies to disjunction of propositions. This is
+because propositions are a special case of types. The benefit of this approach
+is that Agda can infer types for us, but not generally propositions.
+
+### The universal property of disjunctions
+
+```agda
+ev-disjunction :
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3} →
+ (disjunction-type A B → C) → (A → C) × (B → C)
+pr1 (ev-disjunction h) = h ∘ inl-disjunction
+pr2 (ev-disjunction h) = h ∘ inr-disjunction
+
+universal-property-disjunction-type :
+ {l1 l2 l3 : Level} → UU l1 → UU l2 → Prop l3 → UUω
+universal-property-disjunction-type A B S =
+ {l : Level} (R : Prop l) →
+ (type-Prop S → type-Prop R) ↔ ((A → type-Prop R) × (B → type-Prop R))
+
+universal-property-disjunction-Prop :
+ {l1 l2 l3 : Level} → Prop l1 → Prop l2 → Prop l3 → UUω
+universal-property-disjunction-Prop P Q =
+ universal-property-disjunction-type (type-Prop P) (type-Prop Q)
```
## Properties
-### The introduction rules for disjunction
+### The disjunction satisfies the universal property of disjunctions
```agda
-inl-disjunction-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- type-hom-Prop P (disjunction-Prop P Q)
-inl-disjunction-Prop P Q = unit-trunc-Prop ∘ inl
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ elim-disjunction' :
+ {l : Level} (R : Prop l) →
+ (A → type-Prop R) × (B → type-Prop R) →
+ disjunction-type A B → type-Prop R
+ elim-disjunction' R (f , g) =
+ map-universal-property-trunc-Prop R (rec-coproduct f g)
-inr-disjunction-Prop :
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- type-hom-Prop Q (disjunction-Prop P Q)
-inr-disjunction-Prop P Q = unit-trunc-Prop ∘ inr
+ up-disjunction :
+ universal-property-disjunction-type A B (disjunction-type-Prop A B)
+ up-disjunction R = ev-disjunction , elim-disjunction' R
```
-### The elimination rule and universal property of disjunction
+### The elimination principle of disjunctions
```agda
-ev-disjunction-Prop :
- {l1 l2 l3 : Level} (P : Prop l1) (Q : Prop l2) (R : Prop l3) →
- type-hom-Prop
- ( hom-Prop (disjunction-Prop P Q) R)
- ( conjunction-Prop (hom-Prop P R) (hom-Prop Q R))
-pr1 (ev-disjunction-Prop P Q R h) = h ∘ (inl-disjunction-Prop P Q)
-pr2 (ev-disjunction-Prop P Q R h) = h ∘ (inr-disjunction-Prop P Q)
-
-elim-disjunction-Prop :
- {l1 l2 l3 : Level} (P : Prop l1) (Q : Prop l2) (R : Prop l3) →
- type-hom-Prop
- ( conjunction-Prop (hom-Prop P R) (hom-Prop Q R))
- ( hom-Prop (disjunction-Prop P Q) R)
-elim-disjunction-Prop P Q R (pair f g) =
- map-universal-property-trunc-Prop R (rec-coproduct f g)
-
-abstract
- is-equiv-ev-disjunction-Prop :
- {l1 l2 l3 : Level} (P : Prop l1) (Q : Prop l2) (R : Prop l3) →
- is-equiv (ev-disjunction-Prop P Q R)
- is-equiv-ev-disjunction-Prop P Q R =
- is-equiv-is-prop
- ( is-prop-type-Prop (hom-Prop (disjunction-Prop P Q) R))
- ( is-prop-type-Prop (conjunction-Prop (hom-Prop P R) (hom-Prop Q R)))
- ( elim-disjunction-Prop P Q R)
-```
-
-### The unit laws for disjunction
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (R : Prop l3)
+ where
+
+ elim-disjunction :
+ (A → type-Prop R) → (B → type-Prop R) →
+ disjunction-type A B → type-Prop R
+ elim-disjunction f g = elim-disjunction' R (f , g)
+```
+
+### Propositions that satisfy the universal property of a disjunction are equivalent to the disjunction
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} (Q : Prop l3)
+ (up-Q : universal-property-disjunction-type A B Q)
+ where
+
+ forward-implication-iff-universal-property-disjunction :
+ disjunction-type A B → type-Prop Q
+ forward-implication-iff-universal-property-disjunction =
+ elim-disjunction Q
+ ( pr1 (forward-implication (up-Q Q) id))
+ ( pr2 (forward-implication (up-Q Q) id))
+
+ backward-implication-iff-universal-property-disjunction :
+ type-Prop Q → disjunction-type A B
+ backward-implication-iff-universal-property-disjunction =
+ backward-implication
+ ( up-Q (disjunction-type-Prop A B))
+ ( inl-disjunction , inr-disjunction)
+
+ iff-universal-property-disjunction :
+ disjunction-type A B ↔ type-Prop Q
+ iff-universal-property-disjunction =
+ ( forward-implication-iff-universal-property-disjunction ,
+ backward-implication-iff-universal-property-disjunction)
+```
+
+### The unit laws for the disjunction
```agda
module _
@@ -124,10 +245,110 @@ module _
map-left-unit-law-disjunction-is-empty-Prop :
is-empty (type-Prop P) → type-disjunction-Prop P Q → type-Prop Q
map-left-unit-law-disjunction-is-empty-Prop f =
- elim-disjunction-Prop P Q Q (ex-falso ∘ f , id)
+ elim-disjunction Q (ex-falso ∘ f) id
map-right-unit-law-disjunction-is-empty-Prop :
is-empty (type-Prop Q) → type-disjunction-Prop P Q → type-Prop P
map-right-unit-law-disjunction-is-empty-Prop f =
- elim-disjunction-Prop P Q P (id , ex-falso ∘ f)
+ elim-disjunction P id (ex-falso ∘ f)
+```
+
+### The unit laws for the disjunction of types
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ map-left-unit-law-disjunction-is-empty-type :
+ is-empty A → disjunction-type A B → is-inhabited B
+ map-left-unit-law-disjunction-is-empty-type f =
+ elim-disjunction (is-inhabited-Prop B) (ex-falso ∘ f) unit-trunc-Prop
+
+ map-right-unit-law-disjunction-is-empty-type :
+ is-empty B → disjunction-type A B → is-inhabited A
+ map-right-unit-law-disjunction-is-empty-type f =
+ elim-disjunction (is-inhabited-Prop A) unit-trunc-Prop (ex-falso ∘ f)
+```
+
+### The disjunction of arbitrary types is the disjunction of inhabitedness propositions
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ universal-property-disjunction-trunc :
+ universal-property-disjunction-type A B
+ ( disjunction-Prop (trunc-Prop A) (trunc-Prop B))
+ universal-property-disjunction-trunc R =
+ ( λ f →
+ ( f ∘ inl-disjunction ∘ unit-trunc-Prop ,
+ f ∘ inr-disjunction ∘ unit-trunc-Prop)) ,
+ ( λ (f , g) →
+ rec-trunc-Prop R
+ ( rec-coproduct (rec-trunc-Prop R f) (rec-trunc-Prop R g)))
+
+ iff-compute-disjunction-trunc :
+ disjunction-type A B ↔ type-disjunction-Prop (trunc-Prop A) (trunc-Prop B)
+ iff-compute-disjunction-trunc =
+ iff-universal-property-disjunction
+ ( disjunction-Prop (trunc-Prop A) (trunc-Prop B))
+ ( universal-property-disjunction-trunc)
+```
+
+### The disjunction preserves decidability
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-decidable-disjunction :
+ is-decidable A → is-decidable B → is-decidable (disjunction-type A B)
+ is-decidable-disjunction is-decidable-A is-decidable-B =
+ is-decidable-trunc-Prop-is-merely-decidable
+ ( A + B)
+ ( unit-trunc-Prop (is-decidable-coproduct is-decidable-A is-decidable-B))
+
+module _
+ {l1 l2 : Level} (P : Decidable-Prop l1) (Q : Decidable-Prop l2)
+ where
+
+ type-disjunction-Decidable-Prop : UU (l1 ⊔ l2)
+ type-disjunction-Decidable-Prop =
+ type-disjunction-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
+
+ is-prop-disjunction-Decidable-Prop :
+ is-prop type-disjunction-Decidable-Prop
+ is-prop-disjunction-Decidable-Prop =
+ is-prop-disjunction-Prop
+ ( prop-Decidable-Prop P)
+ ( prop-Decidable-Prop Q)
+
+ disjunction-Decidable-Prop : Decidable-Prop (l1 ⊔ l2)
+ disjunction-Decidable-Prop =
+ ( type-disjunction-Decidable-Prop ,
+ is-prop-disjunction-Decidable-Prop ,
+ is-decidable-disjunction
+ ( is-decidable-Decidable-Prop P)
+ ( is-decidable-Decidable-Prop Q))
```
+
+## See also
+
+- The indexed counterpart to disjunction is
+ [existential quantification](foundation.existential-quantification.md).
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [disjunction](https://ncatlab.org/nlab/show/disjunction) at $n$Lab
+- [Logical disjunction](https://en.wikipedia.org/wiki/Logical_disjunction) at
+ Wikipedia
diff --git a/src/foundation/double-negation-modality.lagda.md b/src/foundation/double-negation-modality.lagda.md
index 1e406c460d..d313b548d3 100644
--- a/src/foundation/double-negation-modality.lagda.md
+++ b/src/foundation/double-negation-modality.lagda.md
@@ -33,7 +33,7 @@ The [double negation](foundation.double-negation.md) operation `¬¬` is a
```agda
operator-double-negation-modality :
(l : Level) → operator-modality l l
-operator-double-negation-modality _ = ¬¬
+operator-double-negation-modality _ = ¬¬_
unit-double-negation-modality :
{l : Level} → unit-modality (operator-double-negation-modality l)
@@ -57,7 +57,7 @@ is-uniquely-eliminating-modality-double-negation-modality {l} {A} P =
double-negation-extend
( λ (a : A) →
tr
- ( ¬¬ ∘ P)
+ ( ¬¬_ ∘ P)
( eq-is-prop is-prop-double-negation)
( f a))
( z))
diff --git a/src/foundation/double-negation.lagda.md b/src/foundation/double-negation.lagda.md
index 70aecd6fc0..4529283ca7 100644
--- a/src/foundation/double-negation.lagda.md
+++ b/src/foundation/double-negation.lagda.md
@@ -26,10 +26,12 @@ open import foundation-core.propositions
We define double negation and triple negation
```agda
-¬¬ : {l : Level} → UU l → UU l
+infix 25 ¬¬_ ¬¬¬_
+
+¬¬_ : {l : Level} → UU l → UU l
¬¬ P = ¬ (¬ P)
-¬¬¬ : {l : Level} → UU l → UU l
+¬¬¬_ : {l : Level} → UU l → UU l
¬¬¬ P = ¬ (¬ (¬ P))
```
@@ -41,7 +43,7 @@ intro-double-negation : {l : Level} {P : UU l} → P → ¬¬ P
intro-double-negation p f = f p
map-double-negation :
- {l1 l2 : Level} {P : UU l1} {Q : UU l2} → (P → Q) → (¬¬ P → ¬¬ Q)
+ {l1 l2 : Level} {P : UU l1} {Q : UU l2} → (P → Q) → ¬¬ P → ¬¬ Q
map-double-negation f = map-neg (map-neg f)
```
@@ -50,17 +52,22 @@ map-double-negation f = map-neg (map-neg f)
### The double negation of a type is a proposition
```agda
-double-negation-Prop' :
+double-negation-type-Prop :
{l : Level} (A : UU l) → Prop l
-double-negation-Prop' A = neg-Prop' (¬ A)
+double-negation-type-Prop A = neg-type-Prop (¬ A)
double-negation-Prop :
{l : Level} (P : Prop l) → Prop l
-double-negation-Prop P = double-negation-Prop' (type-Prop P)
+double-negation-Prop P = double-negation-type-Prop (type-Prop P)
is-prop-double-negation :
{l : Level} {A : UU l} → is-prop (¬¬ A)
is-prop-double-negation = is-prop-neg
+
+infix 25 ¬¬'_
+
+¬¬'_ : {l : Level} (P : Prop l) → Prop l
+¬¬'_ = double-negation-Prop
```
### Double negations of classical laws
@@ -109,10 +116,10 @@ double-negation-elim-exp {P = P} {Q = Q} f p =
( ¬ Q)
( map-double-negation (λ (g : P → ¬¬ Q) → g p) f)
-double-negation-elim-forall :
+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-forall {P = P} {Q = Q} f 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)
@@ -137,7 +144,7 @@ abstract
double-negation-double-negation-type-trunc-Prop A =
double-negation-extend
( map-universal-property-trunc-Prop
- ( double-negation-Prop' A)
+ ( double-negation-type-Prop A)
( intro-double-negation))
abstract
@@ -146,3 +153,10 @@ abstract
double-negation-type-trunc-Prop-double-negation =
map-double-negation unit-trunc-Prop
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/double-powersets.lagda.md b/src/foundation/double-powersets.lagda.md
index 9788978ec3..45d84375c0 100644
--- a/src/foundation/double-powersets.lagda.md
+++ b/src/foundation/double-powersets.lagda.md
@@ -81,7 +81,7 @@ union-double-powerset :
{l1 l2 l3 : Level} {A : UU l1} →
double-powerset l2 l3 A → powerset (l1 ⊔ lsuc l2 ⊔ l3) A
union-double-powerset F a =
- ∃-Prop (type-subtype F) (λ X → is-in-subtype (inclusion-subtype F X) a)
+ ∃ (type-subtype F) (λ X → inclusion-subtype F X a)
module _
{l1 l2 l3 : Level} {A : UU l1} (F : double-powerset l2 l3 A)
@@ -92,7 +92,7 @@ module _
inclusion-union-double-powerset :
(X : type-subtype F) → inclusion-subtype F X ⊆ union-double-powerset F
- inclusion-union-double-powerset X a = intro-∃ X
+ inclusion-union-double-powerset X a = intro-exists X
universal-property-union-double-powerset :
{l : Level} (P : powerset l A)
diff --git a/src/foundation/dubuc-penon-compact-types.lagda.md b/src/foundation/dubuc-penon-compact-types.lagda.md
index bfb43239fc..7a80fcc588 100644
--- a/src/foundation/dubuc-penon-compact-types.lagda.md
+++ b/src/foundation/dubuc-penon-compact-types.lagda.md
@@ -8,6 +8,7 @@ module foundation.dubuc-penon-compact-types where
```agda
open import foundation.disjunction
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.propositions
@@ -18,9 +19,12 @@ open import foundation-core.subtypes
## Idea
-A type is said to be Dubuc-Penon compact if for every proposition `P` and every
-subtype `Q` of `X` such that `P ∨ Q x` holds for all `x`, then either `P` is
-true or `Q` contains every element of `X`.
+A type is said to be
+{{#concept "Dubuc-Penon compact" Agda=is-dubuc-penon-compact}} if for every
+[proposition](foundation-core.propositions.md) `P` and every
+[subtype](foundation-core.subtypes.md) `Q` of `X` such that the
+[disjunction](foundation.disjunction.md) `P ∨ Q x` holds for all `x`, then
+either `P` is true or `Q` contains every element of `X`.
## Definition
@@ -28,15 +32,12 @@ true or `Q` contains every element of `X`.
is-dubuc-penon-compact-Prop :
{l : Level} (l1 l2 : Level) → UU l → Prop (l ⊔ lsuc l1 ⊔ lsuc l2)
is-dubuc-penon-compact-Prop l1 l2 X =
- Π-Prop
+ ∀'
( Prop l1)
( λ P →
- Π-Prop
+ ∀'
( subtype l2 X)
- ( λ Q →
- function-Prop
- ( (x : X) → type-disjunction-Prop P (Q x))
- ( disjunction-Prop P (Π-Prop X Q))))
+ ( λ Q → (∀' X (λ x → P ∨ Q x)) ⇒ (P ∨ (∀' X Q))))
is-dubuc-penon-compact :
{l : Level} (l1 l2 : Level) → UU l → UU (l ⊔ lsuc l1 ⊔ lsuc l2)
diff --git a/src/foundation/epimorphisms-with-respect-to-sets.lagda.md b/src/foundation/epimorphisms-with-respect-to-sets.lagda.md
index 01a52b2c70..f34760598e 100644
--- a/src/foundation/epimorphisms-with-respect-to-sets.lagda.md
+++ b/src/foundation/epimorphisms-with-respect-to-sets.lagda.md
@@ -12,6 +12,7 @@ open import foundation.dependent-pair-types
open import foundation.existential-quantification
open import foundation.function-extensionality
open import foundation.identity-types
+open import foundation.injective-maps
open import foundation.propositional-extensionality
open import foundation.propositional-truncations
open import foundation.sets
@@ -24,7 +25,6 @@ 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.precomposition-functions
open import foundation-core.propositional-maps
open import foundation-core.propositions
@@ -100,7 +100,7 @@ abstract
g : B → Prop (l1 ⊔ l2)
g y = raise-unit-Prop (l1 ⊔ l2)
h : B → Prop (l1 ⊔ l2)
- h y = ∃-Prop A (λ x → f x = y)
+ h y = exists-structure-Prop A (λ x → f x = y)
```
### There is at most one extension of a map into a set along a surjection
diff --git a/src/foundation/equivalence-classes.lagda.md b/src/foundation/equivalence-classes.lagda.md
index 2a8551e680..71954bae57 100644
--- a/src/foundation/equivalence-classes.lagda.md
+++ b/src/foundation/equivalence-classes.lagda.md
@@ -7,6 +7,7 @@ module foundation.equivalence-classes where
Imports
```agda
+open import foundation.conjunction
open import foundation.dependent-pair-types
open import foundation.effective-maps-equivalence-relations
open import foundation.existential-quantification
@@ -56,7 +57,7 @@ module _
is-equivalence-class-Prop : subtype l2 A → Prop (l1 ⊔ l2)
is-equivalence-class-Prop P =
- ∃-Prop A (λ x → has-same-elements-subtype P (prop-equivalence-relation R x))
+ ∃ A (λ x → has-same-elements-subtype-Prop P (prop-equivalence-relation R x))
is-equivalence-class : subtype l2 A → UU (l1 ⊔ l2)
is-equivalence-class P = type-Prop (is-equivalence-class-Prop P)
@@ -246,8 +247,9 @@ module _
share-common-element-equivalence-class-Prop :
(C D : equivalence-class R) → Prop (l1 ⊔ l2)
share-common-element-equivalence-class-Prop C D =
- ∃-Prop A
- ( λ x → is-in-equivalence-class R C x × is-in-equivalence-class R D x)
+ ∃ ( A)
+ ( λ x →
+ is-in-equivalence-class-Prop R C x ∧ is-in-equivalence-class-Prop R D x)
share-common-element-equivalence-class :
(C D : equivalence-class R) → UU (l1 ⊔ l2)
diff --git a/src/foundation/equivalences.lagda.md b/src/foundation/equivalences.lagda.md
index 7ffbfa28a8..3b62d1bde4 100644
--- a/src/foundation/equivalences.lagda.md
+++ b/src/foundation/equivalences.lagda.md
@@ -15,6 +15,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalence-extensionality
open import foundation.function-extensionality
open import foundation.functoriality-fibers-of-maps
+open import foundation.logical-equivalences
open import foundation.transposition-identifications-along-equivalences
open import foundation.truncated-maps
open import foundation.universal-property-equivalences
@@ -188,7 +189,7 @@ module _
( H : coherence-triangle-maps h (map-equiv e) f) →
is-equiv f ≃ is-equiv h
equiv-is-equiv-right-map-triangle {f} e h H =
- equiv-prop
+ equiv-iff-is-prop
( is-property-is-equiv f)
( is-property-is-equiv h)
( λ is-equiv-f →
@@ -215,7 +216,7 @@ module _
( H : coherence-triangle-maps h f (map-equiv e)) →
is-equiv f ≃ is-equiv h
equiv-is-equiv-top-map-triangle e {f} h H =
- equiv-prop
+ equiv-iff-is-prop
( is-property-is-equiv f)
( is-property-is-equiv h)
( is-equiv-left-map-triangle h f (map-equiv e) H (is-equiv-map-equiv e))
@@ -380,7 +381,7 @@ module _
{ f g : A → B} → (f ~ g) →
is-equiv f ≃ is-equiv g
equiv-is-equiv-htpy {f} {g} H =
- equiv-prop
+ equiv-iff-is-prop
( is-property-is-equiv f)
( is-property-is-equiv g)
( is-equiv-htpy f (inv-htpy H))
diff --git a/src/foundation/exclusive-disjunction.lagda.md b/src/foundation/exclusive-disjunction.lagda.md
index 78aacaf3cd..5fb3e08ee3 100644
--- a/src/foundation/exclusive-disjunction.lagda.md
+++ b/src/foundation/exclusive-disjunction.lagda.md
@@ -1,4 +1,4 @@
-# Exclusive disjunction of propositions
+# Exclusive disjunctions
```agda
module foundation.exclusive-disjunction where
@@ -7,61 +7,88 @@ module foundation.exclusive-disjunction where
Imports
```agda
-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.equality-coproduct-types
+open import foundation.exclusive-sum
open import foundation.functoriality-coproduct-types
-open import foundation.negation
-open import foundation.propositional-extensionality
-open import foundation.symmetric-operations
+open import foundation.propositional-truncations
open import foundation.type-arithmetic-cartesian-product-types
open import foundation.type-arithmetic-coproduct-types
open import foundation.universal-property-coproduct-types
open import foundation.universe-levels
-open import foundation.unordered-pairs
-open import foundation-core.cartesian-product-types
-open import foundation-core.decidable-propositions
open import foundation-core.embeddings
-open import foundation-core.empty-types
-open import foundation-core.equality-dependent-pair-types
open import foundation-core.equivalences
open import foundation-core.functoriality-dependent-function-types
open import foundation-core.functoriality-dependent-pair-types
open import foundation-core.identity-types
open import foundation-core.propositions
-open import foundation-core.transport-along-identifications
-
-open import univalent-combinatorics.2-element-types
-open import univalent-combinatorics.equality-finite-types
-open import univalent-combinatorics.finite-types
-open import univalent-combinatorics.standard-finite-types
```
## Idea
-The exclusive disjunction of two propositions `P` and `Q` is the proposition
-that `P` holds and `Q` doesn't hold or `P` doesn't hold and `Q` holds.
+The
+{{#concept "exclusive disjunction" Disambiguation="of propositions" WDID=Q498186 WD="exclusive or" Agda=xor-Prop}}
+of two [propositions](foundation-core.propositions.md) `P` and `Q` is the
+proposition that precisely one of `P` and `Q` holds, and is defined as the
+proposition that the [coproduct](foundation-core.coproduct-types.md) of their
+underlying types is [contractible](foundation-core.contractible-types.md)
+
+```text
+ P ⊻ Q := is-contr (P + Q)
+```
+
+It necessarily follows that precisely one of the two propositions hold, and the
+other does not. This is captured by the
+[exclusive sum](foundation.exclusive-sum.md).
## Definitions
-### Exclusive disjunction of types
+### The exclusive disjunction of arbitrary types
+
+The definition of exclusive sum is sometimes generalized to arbitrary types,
+which we record here for completeness.
+
+The
+{{#concept "exclusive disjunction" Disambiguation="of types" Agda=xor-type-Prop}}
+of the types `A` and `B` is the proposition that their coproduct is contractible
+
+```text
+ A ⊻ B := is-contr (A + B).
+```
+
+Note that unlike the case for [disjunction](foundation.disjunction.md) and
+[existential quantification](foundation.existential-quantification.md), but
+analogous to the case of
+[uniqueness quantification](foundation.uniqueness-quantification.md), the
+exclusive disjunction of types does _not_ coincide with the exclusive
+disjunction of the summands'
+[propositional reflections](foundation.propositional-truncations.md):
+
+```text
+ A ⊻ B ≠ ║ A ║₋₁ ⊻ ║ B ║₋₁.
+```
```agda
module _
{l1 l2 : Level} (A : UU l1) (B : UU l2)
where
- xor : UU (l1 ⊔ l2)
- xor = (A × ¬ B) + (B × ¬ A)
+ xor-type-Prop : Prop (l1 ⊔ l2)
+ xor-type-Prop = is-contr-Prop (A + B)
+
+ xor-type : UU (l1 ⊔ l2)
+ xor-type = type-Prop xor-type-Prop
+
+ is-prop-xor-type : is-prop xor-type
+ is-prop-xor-type = is-prop-type-Prop xor-type-Prop
```
-### Exclusive disjunction of propositions
+### The exclusive disjunction
```agda
module _
@@ -69,119 +96,52 @@ module _
where
xor-Prop : Prop (l1 ⊔ l2)
- xor-Prop =
- coproduct-Prop
- ( conjunction-Prop P (neg-Prop Q))
- ( conjunction-Prop Q (neg-Prop P))
- ( λ p q → pr2 q (pr1 p))
+ xor-Prop = xor-type-Prop (type-Prop P) (type-Prop Q)
type-xor-Prop : UU (l1 ⊔ l2)
type-xor-Prop = type-Prop xor-Prop
- abstract
- is-prop-type-xor-Prop : is-prop type-xor-Prop
- is-prop-type-xor-Prop = is-prop-type-Prop xor-Prop
-```
+ is-prop-xor-Prop : is-prop type-xor-Prop
+ is-prop-xor-Prop = is-prop-type-Prop xor-Prop
-### The symmetric operation of exclusive disjunction
-
-```agda
-predicate-symmetric-xor :
- {l : Level} (p : unordered-pair (UU l)) → type-unordered-pair p → UU l
-predicate-symmetric-xor p x =
- ( element-unordered-pair p x) × (¬ (other-element-unordered-pair p x))
-
-symmetric-xor : {l : Level} → symmetric-operation (UU l) (UU l)
-symmetric-xor p = Σ (type-unordered-pair p) (predicate-symmetric-xor p)
+ infixr 10 _⊻_
+ _⊻_ : Prop (l1 ⊔ l2)
+ _⊻_ = xor-Prop
```
-### The symmetric operation of exclusive disjunction of propositions
+**Notation.** The
+[symbol used for exclusive disjunction](https://codepoints.net/U+22BB?lang=en)
+`⊻` can be written with the escape sequence `\veebar`.
-```agda
-predicate-symmetric-xor-Prop :
- {l : Level} (p : unordered-pair (Prop l)) →
- type-unordered-pair p → UU l
-predicate-symmetric-xor-Prop p =
- predicate-symmetric-xor (map-unordered-pair type-Prop p)
-
-type-symmetric-xor-Prop :
- {l : Level} → symmetric-operation (Prop l) (UU l)
-type-symmetric-xor-Prop p = symmetric-xor (map-unordered-pair type-Prop p)
-
-all-elements-equal-type-symmetric-xor-Prop :
- {l : Level} (p : unordered-pair (Prop l)) →
- all-elements-equal (type-symmetric-xor-Prop p)
-all-elements-equal-type-symmetric-xor-Prop (pair X P) x y =
- cases-is-prop-type-symmetric-xor-Prop
- ( has-decidable-equality-is-finite
- ( is-finite-type-UU-Fin 2 X)
- ( pr1 x)
- ( pr1 y))
- where
- cases-is-prop-type-symmetric-xor-Prop :
- is-decidable (pr1 x = pr1 y) → x = y
- cases-is-prop-type-symmetric-xor-Prop (inl p) =
- eq-pair-Σ
- ( p)
- ( eq-is-prop
- ( is-prop-product (is-prop-type-Prop (P (pr1 y))) is-prop-neg))
- cases-is-prop-type-symmetric-xor-Prop (inr np) =
- ex-falso
- ( tr
- ( λ z → ¬ (type-Prop (P z)))
- ( compute-swap-2-Element-Type X (pr1 x) (pr1 y) np)
- ( pr2 (pr2 x))
- ( pr1 (pr2 y)))
-
-is-prop-type-symmetric-xor-Prop :
- {l : Level} (p : unordered-pair (Prop l)) →
- is-prop (type-symmetric-xor-Prop p)
-is-prop-type-symmetric-xor-Prop p =
- is-prop-all-elements-equal
- ( all-elements-equal-type-symmetric-xor-Prop p)
-
-symmetric-xor-Prop :
- {l : Level} → symmetric-operation (Prop l) (Prop l)
-pr1 (symmetric-xor-Prop E) = type-symmetric-xor-Prop E
-pr2 (symmetric-xor-Prop E) = is-prop-type-symmetric-xor-Prop E
-```
+## Properties
-### Second definition of exclusive disjunction
+### The canonical map from the exclusive disjunction into the exclusive sum
```agda
module _
- {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
where
- xor-Prop' : Prop (l1 ⊔ l2)
- xor-Prop' = is-contr-Prop (type-Prop P + type-Prop Q)
-
- type-xor-Prop' : UU (l1 ⊔ l2)
- type-xor-Prop' = type-Prop xor-Prop'
+ map-exclusive-sum-xor : xor-type A B → exclusive-sum A B
+ map-exclusive-sum-xor (inl a , H) =
+ inl (a , (λ b → is-empty-eq-coproduct-inl-inr a b (H (inr b))))
+ map-exclusive-sum-xor (inr b , H) =
+ inr (b , (λ a → is-empty-eq-coproduct-inr-inl b a (H (inl a))))
```
-## Properties
-
-### The definitions `xor-Prop` and `xor-Prop'` are equivalent
+### The exclusive disjunction of two propositions is equivalent to their exclusive sum
```agda
module _
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
where
- map-equiv-xor-Prop : type-xor-Prop' P Q → type-xor-Prop P Q
- map-equiv-xor-Prop (pair (inl p) H) =
- inl (pair p (λ q → is-empty-eq-coproduct-inl-inr p q (H (inr q))))
- map-equiv-xor-Prop (pair (inr q) H) =
- inr (pair q (λ p → is-empty-eq-coproduct-inr-inl q p (H (inl p))))
-
- equiv-xor-Prop : type-xor-Prop' P Q ≃ type-xor-Prop P Q
- equiv-xor-Prop =
+ equiv-exclusive-sum-xor-Prop : type-xor-Prop P Q ≃ type-exclusive-sum-Prop P Q
+ equiv-exclusive-sum-xor-Prop =
( equiv-coproduct
( equiv-tot
( λ p →
- ( ( equiv-Π-equiv-family
- ( λ q → compute-eq-coproduct-inl-inr p q)) ∘e
+ ( ( equiv-Π-equiv-family (compute-eq-coproduct-inl-inr p)) ∘e
( left-unit-law-product-is-contr
( is-contr-Π
( λ p' →
@@ -189,11 +149,10 @@ module _
( p = p')
( equiv-ap-emb (emb-inl (type-Prop P) (type-Prop Q)))
( is-prop-type-Prop P p p'))))) ∘e
- ( equiv-dependent-universal-property-coproduct (λ x → inl p = x))))
+ ( equiv-dependent-universal-property-coproduct (inl p =_))))
( equiv-tot
( λ q →
- ( ( equiv-Π-equiv-family
- ( λ p → compute-eq-coproduct-inr-inl q p)) ∘e
+ ( ( equiv-Π-equiv-family (compute-eq-coproduct-inr-inl q)) ∘e
( right-unit-law-product-is-contr
( is-contr-Π
( λ q' →
@@ -201,211 +160,28 @@ module _
( q = q')
( equiv-ap-emb (emb-inr (type-Prop P) (type-Prop Q)))
( is-prop-type-Prop Q q q'))))) ∘e
- ( equiv-dependent-universal-property-coproduct
- ( λ x → inr q = x))))) ∘e
+ ( equiv-dependent-universal-property-coproduct (inr q =_))))) ∘e
( right-distributive-Σ-coproduct
( type-Prop P)
( type-Prop Q)
( λ x → (y : type-Prop P + type-Prop Q) → x = y))
```
-### The symmetric exclusive disjunction at a standard unordered pair
+## See also
-```agda
-module _
- {l : Level} {A B : UU l}
- where
+- The indexed counterpart to exclusive disjunction is
+ [unique existence](foundation.uniqueness-quantification.md).
- xor-symmetric-xor :
- symmetric-xor (standard-unordered-pair A B) → xor A B
- xor-symmetric-xor (pair (inl (inr _)) (pair p nq)) =
- inl
- ( pair p
- ( tr
- ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
- ( compute-swap-Fin-two-ℕ (zero-Fin 1))
- ( nq)))
- xor-symmetric-xor (pair (inr _) (pair q np)) =
- inr
- ( pair
- ( q)
- ( tr
- ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
- ( compute-swap-Fin-two-ℕ (one-Fin 1))
- ( np)))
-
- symmetric-xor-xor :
- xor A B → symmetric-xor (standard-unordered-pair A B)
- pr1 (symmetric-xor-xor (inl (pair a nb))) = (zero-Fin 1)
- pr1 (pr2 (symmetric-xor-xor (inl (pair a nb)))) = a
- pr2 (pr2 (symmetric-xor-xor (inl (pair a nb)))) =
- tr
- ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
- ( inv (compute-swap-Fin-two-ℕ (zero-Fin 1)))
- ( nb)
- pr1 (symmetric-xor-xor (inr (pair na b))) = (one-Fin 1)
- pr1 (pr2 (symmetric-xor-xor (inr (pair b na)))) = b
- pr2 (pr2 (symmetric-xor-xor (inr (pair b na)))) =
- tr
- ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
- ( inv (compute-swap-Fin-two-ℕ (one-Fin 1)))
- ( na)
-
-{-
- eq-equiv-Prop
- ( ( ( equiv-coproduct
- ( ( ( left-unit-law-coproduct (type-Prop (conjunction-Prop P (neg-Prop Q)))) ∘e
- ( equiv-coproduct
- ( left-absorption-Σ
- ( λ x →
- ( type-Prop
- ( pr2 (standard-unordered-pair P Q) (inl (inl x)))) ×
- ( ¬ ( type-Prop
- ( pr2
- ( standard-unordered-pair P Q)
- ( map-swap-2-Element-Type
- ( pr1 (standard-unordered-pair P Q))
- ( inl (inl x))))))))
- ( ( equiv-product
- ( id-equiv)
- ( tr
- ( λ b →
- ( ¬ (type-Prop (pr2 (standard-unordered-pair P Q) b))) ≃
- ( ¬ (type-Prop Q)))
- ( inv (compute-swap-Fin-two-ℕ (zero-Fin ?)))
- ( id-equiv))) ∘e
- ( left-unit-law-Σ
- ( λ y →
- ( type-Prop
- ( pr2 (standard-unordered-pair P Q) (inl (inr y)))) ×
- ( ¬ ( type-Prop
- ( pr2
- ( standard-unordered-pair P Q)
- ( map-swap-2-Element-Type
- ( pr1 (standard-unordered-pair P Q))
- ( inl (inr y))))))))))) ∘e
- ( ( right-distributive-Σ-coproduct
- ( Fin 0)
- ( unit)
- ( λ x →
- ( type-Prop (pr2 (standard-unordered-pair P Q) (inl x))) ×
- ( ¬ ( type-Prop
- ( pr2
- ( standard-unordered-pair P Q)
- ( map-swap-2-Element-Type
- ( pr1 (standard-unordered-pair P Q)) (inl x)))))))))
- ( ( equiv-product
- ( tr
- ( λ b →
- ¬ (type-Prop (pr2 (standard-unordered-pair P Q) b)) ≃
- ¬ (type-Prop P))
- ( inv (compute-swap-Fin-two-ℕ (one-Fin ?)))
- ( id-equiv))
- ( id-equiv)) ∘e
- ( ( commutative-product) ∘e
- ( left-unit-law-Σ
- ( λ y →
- ( type-Prop (pr2 (standard-unordered-pair P Q) (inr y))) ×
- ( ¬ ( type-Prop
- ( pr2
- ( standard-unordered-pair P Q)
- ( map-swap-2-Element-Type
- ( pr1 (standard-unordered-pair P Q))
- ( inr y)))))))))) ∘e
- ( right-distributive-Σ-coproduct
- ( Fin 1)
- ( unit)
- ( λ x →
- ( type-Prop (pr2 (standard-unordered-pair P Q) x)) ×
- ( ¬ ( type-Prop
- ( pr2
- ( standard-unordered-pair P Q)
- ( map-swap-2-Element-Type
- ( pr1 (standard-unordered-pair P Q))
- ( x)))))))))
--}
-```
+## Table of files about propositional logic
-```agda
-module _
- {l : Level} (P Q : Prop l)
- where
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
- xor-symmetric-xor-Prop :
- type-hom-Prop
- ( symmetric-xor-Prop (standard-unordered-pair P Q))
- ( xor-Prop P Q)
- xor-symmetric-xor-Prop (pair (inl (inr _)) (pair p nq)) =
- inl
- ( pair p
- ( tr
- ( λ t →
- ¬ ( type-Prop
- ( element-unordered-pair (standard-unordered-pair P Q) t)))
- ( compute-swap-Fin-two-ℕ (zero-Fin 1))
- ( nq)))
- xor-symmetric-xor-Prop (pair (inr _) (pair q np)) =
- inr
- ( pair q
- ( tr
- ( λ t →
- ¬ ( type-Prop
- ( element-unordered-pair (standard-unordered-pair P Q) t)))
- ( compute-swap-Fin-two-ℕ (one-Fin 1))
- ( np)))
-
- symmetric-xor-xor-Prop :
- type-hom-Prop
- ( xor-Prop P Q)
- ( symmetric-xor-Prop (standard-unordered-pair P Q))
- pr1 (symmetric-xor-xor-Prop (inl (pair p nq))) = (zero-Fin 1)
- pr1 (pr2 (symmetric-xor-xor-Prop (inl (pair p nq)))) = p
- pr2 (pr2 (symmetric-xor-xor-Prop (inl (pair p nq)))) =
- tr
- ( λ t →
- ¬ (type-Prop (element-unordered-pair (standard-unordered-pair P Q) t)))
- ( inv (compute-swap-Fin-two-ℕ (zero-Fin 1)))
- ( nq)
- pr1 (symmetric-xor-xor-Prop (inr (pair q np))) = (one-Fin 1)
- pr1 (pr2 (symmetric-xor-xor-Prop (inr (pair q np)))) = q
- pr2 (pr2 (symmetric-xor-xor-Prop (inr (pair q np)))) =
- tr
- ( λ t →
- ¬ (type-Prop (element-unordered-pair (standard-unordered-pair P Q) t)))
- ( inv (compute-swap-Fin-two-ℕ (one-Fin 1)))
- ( np)
-
-eq-commmutative-xor-xor :
- {l : Level} (P Q : Prop l) →
- symmetric-xor-Prop (standard-unordered-pair P Q) = xor-Prop P Q
-eq-commmutative-xor-xor P Q =
- eq-iff (xor-symmetric-xor-Prop P Q) (symmetric-xor-xor-Prop P Q)
-```
+{{#include tables/propositional-logic.md}}
-### Exclusive disjunction of decidable propositions
+## External links
-```agda
-is-decidable-xor :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} →
- is-decidable A → is-decidable B → is-decidable (xor A B)
-is-decidable-xor d e =
- is-decidable-coproduct
- ( is-decidable-product d (is-decidable-neg e))
- ( is-decidable-product e (is-decidable-neg d))
-
-xor-Decidable-Prop :
- {l1 l2 : Level} → Decidable-Prop l1 → Decidable-Prop l2 →
- Decidable-Prop (l1 ⊔ l2)
-pr1 (xor-Decidable-Prop P Q) =
- type-xor-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr1 (pr2 (xor-Decidable-Prop P Q)) =
- is-prop-type-xor-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
-pr2 (pr2 (xor-Decidable-Prop P Q)) =
- is-decidable-coproduct
- ( is-decidable-product
- ( is-decidable-Decidable-Prop P)
- ( is-decidable-neg (is-decidable-Decidable-Prop Q)))
- ( is-decidable-product
- ( is-decidable-Decidable-Prop Q)
- ( is-decidable-neg (is-decidable-Decidable-Prop P)))
-```
+- [exclusive disjunction](https://ncatlab.org/nlab/show/exclusive+disjunction)
+ at $n$Lab
+- [Exclusive disjunction](https://simple.wikipedia.org/wiki/Exclusive_disjunction)
+ at Wikipedia
diff --git a/src/foundation/exclusive-sum.lagda.md b/src/foundation/exclusive-sum.lagda.md
new file mode 100644
index 0000000000..fc12a94489
--- /dev/null
+++ b/src/foundation/exclusive-sum.lagda.md
@@ -0,0 +1,395 @@
+# Exclusive sums
+
+```agda
+module foundation.exclusive-sum where
+```
+
+Imports
+
+```agda
+open import foundation.conjunction
+open import foundation.coproduct-types
+open import foundation.decidable-types
+open import foundation.dependent-pair-types
+open import foundation.negation
+open import foundation.propositional-extensionality
+open import foundation.symmetric-operations
+open import foundation.universal-quantification
+open import foundation.universe-levels
+open import foundation.unordered-pairs
+
+open import foundation-core.cartesian-product-types
+open import foundation-core.decidable-propositions
+open import foundation-core.embeddings
+open import foundation-core.empty-types
+open import foundation-core.equality-dependent-pair-types
+open import foundation-core.identity-types
+open import foundation-core.propositions
+open import foundation-core.transport-along-identifications
+
+open import univalent-combinatorics.2-element-types
+open import univalent-combinatorics.equality-finite-types
+open import univalent-combinatorics.finite-types
+open import univalent-combinatorics.standard-finite-types
+```
+
+
+
+## Idea
+
+The {{#concept "exclusive sum" Disambiguation="of types" Agda=exclusive-sum}} of
+two types `A` and `B`, is the type consisting of
+
+- elements of `A` together with a proof that `B` is
+ [empty](foundation-core.empty-types.md), or
+- elements of `B` together with a proof that `A` is empty.
+
+The
+{{#concept "exclusive sum" Disambiguation="of propositions" Agda=exclusive-sum-Prop}}
+of [propositions](foundation-core.propositions.md) `P` and `Q` is likewise the
+[proposition](foundation-core.propositions.md) that `P` holds and `Q` does not
+hold, or `P` does not hold and `Q` holds. The exclusive sum of two propositions
+is equivalent to the
+[exclusive disjunction](foundation.exclusive-disjunction.md), which is shown
+there.
+
+## Definitions
+
+### The exclusive sum of types
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ exclusive-sum : UU (l1 ⊔ l2)
+ exclusive-sum = (A × ¬ B) + (B × ¬ A)
+```
+
+### The exclusive sum of propositions
+
+```agda
+module _
+ {l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
+ where
+
+ exclusive-sum-Prop : Prop (l1 ⊔ l2)
+ exclusive-sum-Prop =
+ coproduct-Prop
+ ( P ∧ (¬' Q))
+ ( Q ∧ (¬' P))
+ ( λ p q → pr2 q (pr1 p))
+
+ type-exclusive-sum-Prop : UU (l1 ⊔ l2)
+ type-exclusive-sum-Prop = type-Prop exclusive-sum-Prop
+
+ abstract
+ is-prop-type-exclusive-sum-Prop : is-prop type-exclusive-sum-Prop
+ is-prop-type-exclusive-sum-Prop = is-prop-type-Prop exclusive-sum-Prop
+```
+
+### The canonical inclusion of the exclusive sum into the coproduct
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ map-coproduct-exclusive-sum : exclusive-sum A B → A + B
+ map-coproduct-exclusive-sum (inl (a , _)) = inl a
+ map-coproduct-exclusive-sum (inr (b , _)) = inr b
+```
+
+### The symmetric operation of exclusive sum
+
+```agda
+predicate-symmetric-exclusive-sum :
+ {l : Level} (p : unordered-pair (UU l)) → type-unordered-pair p → UU l
+predicate-symmetric-exclusive-sum p x =
+ ( element-unordered-pair p x) × (¬ (other-element-unordered-pair p x))
+
+symmetric-exclusive-sum : {l : Level} → symmetric-operation (UU l) (UU l)
+symmetric-exclusive-sum p =
+ Σ (type-unordered-pair p) (predicate-symmetric-exclusive-sum p)
+```
+
+### The symmetric operation of the exclusive sum of propositions
+
+```agda
+predicate-symmetric-exclusive-sum-Prop :
+ {l : Level} (p : unordered-pair (Prop l)) →
+ type-unordered-pair p → UU l
+predicate-symmetric-exclusive-sum-Prop p =
+ predicate-symmetric-exclusive-sum (map-unordered-pair type-Prop p)
+
+type-symmetric-exclusive-sum-Prop :
+ {l : Level} → symmetric-operation (Prop l) (UU l)
+type-symmetric-exclusive-sum-Prop p =
+ symmetric-exclusive-sum (map-unordered-pair type-Prop p)
+
+all-elements-equal-type-symmetric-exclusive-sum-Prop :
+ {l : Level} (p : unordered-pair (Prop l)) →
+ all-elements-equal (type-symmetric-exclusive-sum-Prop p)
+all-elements-equal-type-symmetric-exclusive-sum-Prop (X , P) x y =
+ cases-is-prop-type-symmetric-exclusive-sum-Prop
+ ( has-decidable-equality-is-finite
+ ( is-finite-type-UU-Fin 2 X)
+ ( pr1 x)
+ ( pr1 y))
+ where
+ cases-is-prop-type-symmetric-exclusive-sum-Prop :
+ is-decidable (pr1 x = pr1 y) → x = y
+ cases-is-prop-type-symmetric-exclusive-sum-Prop (inl p) =
+ eq-pair-Σ
+ ( p)
+ ( eq-is-prop
+ ( is-prop-product-Prop
+ ( P (pr1 y))
+ ( neg-type-Prop
+ ( other-element-unordered-pair
+ ( map-unordered-pair (type-Prop) (X , P))
+ ( pr1 y)))))
+ cases-is-prop-type-symmetric-exclusive-sum-Prop (inr np) =
+ ex-falso
+ ( tr
+ ( λ z → ¬ (type-Prop (P z)))
+ ( compute-swap-2-Element-Type X (pr1 x) (pr1 y) np)
+ ( pr2 (pr2 x))
+ ( pr1 (pr2 y)))
+
+is-prop-type-symmetric-exclusive-sum-Prop :
+ {l : Level} (p : unordered-pair (Prop l)) →
+ is-prop (type-symmetric-exclusive-sum-Prop p)
+is-prop-type-symmetric-exclusive-sum-Prop p =
+ is-prop-all-elements-equal
+ ( all-elements-equal-type-symmetric-exclusive-sum-Prop p)
+
+symmetric-exclusive-sum-Prop :
+ {l : Level} → symmetric-operation (Prop l) (Prop l)
+pr1 (symmetric-exclusive-sum-Prop E) = type-symmetric-exclusive-sum-Prop E
+pr2 (symmetric-exclusive-sum-Prop E) =
+ is-prop-type-symmetric-exclusive-sum-Prop E
+```
+
+## Properties
+
+### The symmetric exclusive sum at a standard unordered pair
+
+```agda
+module _
+ {l : Level} {A B : UU l}
+ where
+
+ exclusive-sum-symmetric-exclusive-sum :
+ symmetric-exclusive-sum (standard-unordered-pair A B) → exclusive-sum A B
+ exclusive-sum-symmetric-exclusive-sum (pair (inl (inr _)) (p , nq)) =
+ inl
+ ( pair p
+ ( tr
+ ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
+ ( compute-swap-Fin-two-ℕ (zero-Fin 1))
+ ( nq)))
+ exclusive-sum-symmetric-exclusive-sum (pair (inr _) (q , np)) =
+ inr
+ ( pair
+ ( q)
+ ( tr
+ ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
+ ( compute-swap-Fin-two-ℕ (one-Fin 1))
+ ( np)))
+
+ symmetric-exclusive-sum-exclusive-sum :
+ exclusive-sum A B → symmetric-exclusive-sum (standard-unordered-pair A B)
+ pr1 (symmetric-exclusive-sum-exclusive-sum (inl (a , nb))) = (zero-Fin 1)
+ pr1 (pr2 (symmetric-exclusive-sum-exclusive-sum (inl (a , nb)))) = a
+ pr2 (pr2 (symmetric-exclusive-sum-exclusive-sum (inl (a , nb)))) =
+ tr
+ ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
+ ( inv (compute-swap-Fin-two-ℕ (zero-Fin 1)))
+ ( nb)
+ pr1 (symmetric-exclusive-sum-exclusive-sum (inr (na , b))) = (one-Fin 1)
+ pr1 (pr2 (symmetric-exclusive-sum-exclusive-sum (inr (b , na)))) = b
+ pr2 (pr2 (symmetric-exclusive-sum-exclusive-sum (inr (b , na)))) =
+ tr
+ ( λ t → ¬ (element-unordered-pair (standard-unordered-pair A B) t))
+ ( inv (compute-swap-Fin-two-ℕ (one-Fin 1)))
+ ( na)
+```
+
+```text
+ eq-equiv-Prop
+ ( ( ( equiv-coproduct
+ ( ( ( left-unit-law-coproduct (type-Prop (P ∧ (¬' Q)))) ∘e
+ ( equiv-coproduct
+ ( left-absorption-Σ
+ ( λ x →
+ ( type-Prop
+ ( pr2 (standard-unordered-pair P Q) (inl (inl x)))) ×
+ ( ¬ ( type-Prop
+ ( pr2
+ ( standard-unordered-pair P Q)
+ ( map-swap-2-Element-Type
+ ( pr1 (standard-unordered-pair P Q))
+ ( inl (inl x))))))))
+ ( ( equiv-product
+ ( id-equiv)
+ ( tr
+ ( λ b →
+ ( ¬ (type-Prop (pr2 (standard-unordered-pair P Q) b))) ≃
+ ( ¬ (type-Prop Q)))
+ ( inv (compute-swap-Fin-two-ℕ (zero-Fin ?)))
+ ( id-equiv))) ∘e
+ ( left-unit-law-Σ
+ ( λ y →
+ ( type-Prop
+ ( pr2 (standard-unordered-pair P Q) (inl (inr y)))) ×
+ ( ¬ ( type-Prop
+ ( pr2
+ ( standard-unordered-pair P Q)
+ ( map-swap-2-Element-Type
+ ( pr1 (standard-unordered-pair P Q))
+ ( inl (inr y))))))))))) ∘e
+ ( ( right-distributive-Σ-coproduct
+ ( Fin 0)
+ ( unit)
+ ( λ x →
+ ( type-Prop (pr2 (standard-unordered-pair P Q) (inl x))) ×
+ ( ¬ ( type-Prop
+ ( pr2
+ ( standard-unordered-pair P Q)
+ ( map-swap-2-Element-Type
+ ( pr1 (standard-unordered-pair P Q)) (inl x)))))))))
+ ( ( equiv-product
+ ( tr
+ ( λ b →
+ ¬ (type-Prop (pr2 (standard-unordered-pair P Q) b)) ≃
+ ¬ (type-Prop P))
+ ( inv (compute-swap-Fin-two-ℕ (one-Fin ?)))
+ ( id-equiv))
+ ( id-equiv)) ∘e
+ ( ( commutative-product) ∘e
+ ( left-unit-law-Σ
+ ( λ y →
+ ( type-Prop (pr2 (standard-unordered-pair P Q) (inr y))) ×
+ ( ¬ ( type-Prop
+ ( pr2
+ ( standard-unordered-pair P Q)
+ ( map-swap-2-Element-Type
+ ( pr1 (standard-unordered-pair P Q))
+ ( inr y)))))))))) ∘e
+ ( right-distributive-Σ-coproduct
+ ( Fin 1)
+ ( unit)
+ ( λ x →
+ ( type-Prop (pr2 (standard-unordered-pair P Q) x)) ×
+ ( ¬ ( type-Prop
+ ( pr2
+ ( standard-unordered-pair P Q)
+ ( map-swap-2-Element-Type
+ ( pr1 (standard-unordered-pair P Q))
+ ( x)))))))))
+```
+
+```agda
+module _
+ {l : Level} (P Q : Prop l)
+ where
+
+ exclusive-sum-symmetric-exclusive-sum-Prop :
+ type-hom-Prop
+ ( symmetric-exclusive-sum-Prop (standard-unordered-pair P Q))
+ ( exclusive-sum-Prop P Q)
+ exclusive-sum-symmetric-exclusive-sum-Prop (pair (inl (inr _)) (p , nq)) =
+ inl
+ ( pair p
+ ( tr
+ ( λ t →
+ ¬ ( type-Prop
+ ( element-unordered-pair (standard-unordered-pair P Q) t)))
+ ( compute-swap-Fin-two-ℕ (zero-Fin 1))
+ ( nq)))
+ exclusive-sum-symmetric-exclusive-sum-Prop (pair (inr _) (q , np)) =
+ inr
+ ( pair q
+ ( tr
+ ( λ t →
+ ¬ ( type-Prop
+ ( element-unordered-pair (standard-unordered-pair P Q) t)))
+ ( compute-swap-Fin-two-ℕ (one-Fin 1))
+ ( np)))
+
+ symmetric-exclusive-sum-exclusive-sum-Prop :
+ type-hom-Prop
+ ( exclusive-sum-Prop P Q)
+ ( symmetric-exclusive-sum-Prop (standard-unordered-pair P Q))
+ pr1 (symmetric-exclusive-sum-exclusive-sum-Prop (inl (p , nq))) = zero-Fin 1
+ pr1 (pr2 (symmetric-exclusive-sum-exclusive-sum-Prop (inl (p , nq)))) = p
+ pr2 (pr2 (symmetric-exclusive-sum-exclusive-sum-Prop (inl (p , nq)))) =
+ tr
+ ( λ t →
+ ¬ (type-Prop (element-unordered-pair (standard-unordered-pair P Q) t)))
+ ( inv (compute-swap-Fin-two-ℕ (zero-Fin 1)))
+ ( nq)
+ pr1 (symmetric-exclusive-sum-exclusive-sum-Prop (inr (q , np))) = one-Fin 1
+ pr1 (pr2 (symmetric-exclusive-sum-exclusive-sum-Prop (inr (q , np)))) = q
+ pr2 (pr2 (symmetric-exclusive-sum-exclusive-sum-Prop (inr (q , np)))) =
+ tr
+ ( λ t →
+ ¬ (type-Prop (element-unordered-pair (standard-unordered-pair P Q) t)))
+ ( inv (compute-swap-Fin-two-ℕ (one-Fin 1)))
+ ( np)
+
+eq-commmutative-exclusive-sum-exclusive-sum :
+ {l : Level} (P Q : Prop l) →
+ symmetric-exclusive-sum-Prop (standard-unordered-pair P Q) =
+ exclusive-sum-Prop P Q
+eq-commmutative-exclusive-sum-exclusive-sum P Q =
+ eq-iff
+ ( exclusive-sum-symmetric-exclusive-sum-Prop P Q)
+ ( symmetric-exclusive-sum-exclusive-sum-Prop P Q)
+```
+
+### The exclusive sum of decidable types is decidable
+
+```agda
+is-decidable-exclusive-sum :
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} →
+ is-decidable A → is-decidable B → is-decidable (exclusive-sum A B)
+is-decidable-exclusive-sum d e =
+ is-decidable-coproduct
+ ( is-decidable-product d (is-decidable-neg e))
+ ( is-decidable-product e (is-decidable-neg d))
+
+module _
+ {l1 l2 : Level} (P : Decidable-Prop l1) (Q : Decidable-Prop l2)
+ where
+
+ is-decidable-exclusive-sum-Decidable-Prop :
+ is-decidable
+ ( type-exclusive-sum-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q))
+ is-decidable-exclusive-sum-Decidable-Prop =
+ is-decidable-coproduct
+ ( is-decidable-product
+ ( is-decidable-Decidable-Prop P)
+ ( is-decidable-neg (is-decidable-Decidable-Prop Q)))
+ ( is-decidable-product
+ ( is-decidable-Decidable-Prop Q)
+ ( is-decidable-neg (is-decidable-Decidable-Prop P)))
+
+ exclusive-sum-Decidable-Prop : Decidable-Prop (l1 ⊔ l2)
+ pr1 exclusive-sum-Decidable-Prop =
+ type-exclusive-sum-Prop (prop-Decidable-Prop P) (prop-Decidable-Prop Q)
+ pr1 (pr2 exclusive-sum-Decidable-Prop) =
+ is-prop-type-exclusive-sum-Prop
+ ( prop-Decidable-Prop P)
+ ( prop-Decidable-Prop Q)
+ pr2 (pr2 exclusive-sum-Decidable-Prop) =
+ is-decidable-exclusive-sum-Decidable-Prop
+```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/existential-quantification.lagda.md b/src/foundation/existential-quantification.lagda.md
index a749f4825c..8fdf7327a3 100644
--- a/src/foundation/existential-quantification.lagda.md
+++ b/src/foundation/existential-quantification.lagda.md
@@ -9,12 +9,16 @@ module foundation.existential-quantification where
```agda
open import foundation.conjunction
open import foundation.dependent-pair-types
+open import foundation.functoriality-propositional-truncation
open import foundation.logical-equivalences
open import foundation.propositional-extensionality
open import foundation.propositional-truncations
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.functoriality-dependent-pair-types
open import foundation-core.identity-types
open import foundation-core.propositions
```
@@ -23,122 +27,322 @@ open import foundation-core.propositions
## Idea
-Given a family of propositions `P` over `A`, the existential quantification of
-`P` over `A` is the proposition `∃ A P` asserting that there is an element
-`a : A` such that `P a` holds. We use the propositional truncation to define the
-existential quantification, because the Curry-Howard interpretation of the
-existential quantification as `Σ A P` does not guarantee that existential
-quantifications are interpreted as propositions.
+Given a family of [propositions](foundation-core.propositions.md) `P` over `A`,
+the
+{{#concept "existential quantification" Disambiguation="on a subtype" WDID=Q773483 WD="existential quantification" Agda=exists}}
+of `P` over `A` is the proposition `∃ A P` asserting that there is an element
+`a : A` such that `P a` holds. We use the
+[propositional truncation](foundation.propositional-truncations.md) to define
+the existential quantification,
-## Definition
+```text
+ ∃ (x : A), (P x) := ║ Σ (x : A), (P x) ║₋₁
+```
+
+because the Curry-Howard interpretation of the existential quantification as
+`Σ A P` does not guarantee that existential quantifications are interpreted as
+propositions.
+
+The
+{{#concept "universal property" Disambiguation="of existential quantification" Agda=universal-property-exists}}
+of existential quantification states that it is the
+[least upper bound](order-theory.least-upper-bounds-large-posets.md) on the
+family of propositions `P` in the
+[locale of propositions](foundation.large-locale-of-propositions.md), by which
+we mean that for every proposition `Q` we have the
+[logical equivalence](foundation.logical-equivalences.md)
+
+```text
+ (∀ (x : A), (P x ⇒ Q)) ⇔ ((∃ (x : A), (P x)) ⇒ Q).
+```
+
+## Definitions
+
+### Existence of structure
+
+Given a [structure](foundation.structure.md) `B : A → 𝒰` on a type `A`, the
+propositional truncation
+
+```text
+ ║ Σ (x : A), (B x) ║₋₁
+```
+
+satisfies the universal property of the existential quantification
+
+```text
+ ∃ (x : A), ║ B x ║₋₁
+```
+
+and is thus equivalent to it. Therefore, we may reasonably call this
+construction the
+{{#concept "existential quantification" Disambiguation="structure" Agda=exists-structure-Prop}}
+of structure. It is important to keep in mind that this is not a generalization
+of the concept but rather a conflation, and should be read as the statement _the
+type of elements `x : A` equipped with `y : B x` is
+[inhabited](foundation.inhabited-types.md)_.
+
+Existence of structure is a widely occurring notion in univalent mathematics.
+For instance, the condition that an element `y : B` is in the
+[image](foundation.images.md) of a map `f : A → B` is formulated using existence
+of structure: The element `y` is in the image of `f` if the type of `x : A`
+equipped with an identification `f x = y` is inhabited.
-### Existential quantification for families of propositions
+Because subtypes are a special case of structure, and Agda can generally infer
+structures for us, we will continue to conflate the two in our formalizations
+for the benefit that we have to specify the subtype in our code less often. For
+instance, even though the introduction rule for existential quantification
+`intro-exists` is phrased in terms of existential quantification on structures,
+it equally applies to existential quantification on subtypes.
```agda
-exists-Prop :
- {l1 l2 : Level} (A : UU l1) (P : A → Prop l2) → Prop (l1 ⊔ l2)
-exists-Prop {l1} {l2} A P = trunc-Prop (Σ A (λ x → type-Prop (P x)))
+module _
+ {l1 l2 : Level} (A : UU l1) (B : A → UU l2)
+ where
+
+ exists-structure-Prop : Prop (l1 ⊔ l2)
+ exists-structure-Prop = trunc-Prop (Σ A B)
-exists :
- {l1 l2 : Level} (A : UU l1) (P : A → Prop l2) → UU (l1 ⊔ l2)
-exists A P = type-Prop (exists-Prop A P)
+ exists-structure : UU (l1 ⊔ l2)
+ exists-structure = type-Prop exists-structure-Prop
-abstract
- is-prop-exists :
- {l1 l2 : Level} (A : UU l1) (P : A → Prop l2) → is-prop (exists A P)
- is-prop-exists A P = is-prop-type-Prop (exists-Prop A P)
+ is-prop-exists-structure : is-prop exists-structure
+ is-prop-exists-structure = is-prop-type-Prop exists-structure-Prop
```
-### Existential quantification of arbitrary type families
+### Existential quantification
```agda
-∃-Prop :
- {l1 l2 : Level} (A : UU l1) (B : A → UU l2) → Prop (l1 ⊔ l2)
-∃-Prop A B = trunc-Prop (Σ A B)
+module _
+ {l1 l2 : Level} (A : UU l1) (P : A → Prop l2)
+ where
+
+ exists-Prop : Prop (l1 ⊔ l2)
+ exists-Prop = exists-structure-Prop A (type-Prop ∘ P)
-∃ :
- {l1 l2 : Level} (A : UU l1) (B : A → UU l2) → UU (l1 ⊔ l2)
-∃ A B = type-Prop (∃-Prop A B)
+ exists : UU (l1 ⊔ l2)
+ exists = type-Prop exists-Prop
-is-prop-∃ :
- {l1 l2 : Level} (A : UU l1) (B : A → UU l2) → is-prop (∃ A B)
-is-prop-∃ A B = is-prop-type-Prop (∃-Prop A B)
+ abstract
+ is-prop-exists : is-prop exists
+ is-prop-exists = is-prop-type-Prop exists-Prop
+
+ ∃ : Prop (l1 ⊔ l2)
+ ∃ = exists-Prop
+```
+
+### The introduction rule for existential quantification
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ intro-exists : (a : A) (b : B a) → exists-structure A B
+ intro-exists a b = unit-trunc-Prop (a , b)
+```
+
+**Note.** Even though the introduction rule is formalized in terms of
+existential quantification on structures, it equally applies to existential
+quantification on subtypes. This is because subtypes are a special case of
+structure. The benefit of this approach is that Agda can infer structures for
+us, but not generally subtypes.
+
+### The universal property of existential quantification
+
+```agda
+module _
+ {l1 l2 l3 : Level} (A : UU l1) (B : A → UU l2) (S : Prop l3)
+ where
+
+ universal-property-exists-structure : UUω
+ universal-property-exists-structure =
+ {l : Level} (Q : Prop l) →
+ (type-Prop S → type-Prop Q) ↔ ((x : A) → B x → type-Prop Q)
+
+module _
+ {l1 l2 l3 : Level} (A : UU l1) (P : A → Prop l2) (S : Prop l3)
+ where
+
+ universal-property-exists : UUω
+ universal-property-exists =
+ universal-property-exists-structure A (type-Prop ∘ P) S
```
## Properties
-### The introduction rule for existential quantification
+### The elimination rule of existential quantification
+
+The
+{{#concept "universal property" Disambiguation="of existential quantification"}}
+of existential quantification states `∃ A P` is the least upper bound on the
+predicate `P` in the locale of propositions.
```agda
-intro-exists :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) →
- (x : A) → type-Prop (P x) → exists A P
-intro-exists P x p = unit-trunc-Prop (pair x p)
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ ev-intro-exists :
+ {C : UU l3} → (exists-structure A B → C) → (x : A) → B x → C
+ ev-intro-exists H x p = H (intro-exists x p)
-intro-∃ :
- {l1 l2 : Level} {A : UU l1} {B : A → UU l2} (a : A) (b : B a) →
- ∃ A B
-intro-∃ a b = unit-trunc-Prop (pair a b)
+ elim-exists :
+ (Q : Prop l3) →
+ ((x : A) → B x → type-Prop Q) → (exists-structure A B → type-Prop Q)
+ elim-exists Q f = map-universal-property-trunc-Prop Q (ind-Σ f)
+
+ abstract
+ is-equiv-ev-intro-exists :
+ (Q : Prop l3) → is-equiv (ev-intro-exists {type-Prop Q})
+ is-equiv-ev-intro-exists Q =
+ is-equiv-has-converse
+ ( function-Prop (exists-structure A B) Q)
+ ( Π-Prop A (λ x → function-Prop (B x) Q))
+ ( elim-exists Q)
```
-### The elimination rule and the universal property of existential quantification
+### The existential quantification satisfies the universal property of existential quantification
```agda
-ev-intro-exists-Prop :
- {l1 l2 l3 : Level} {A : UU l1} (P : A → Prop l2) (Q : Prop l3) →
- type-hom-Prop (exists-Prop A P) Q → (x : A) → type-hom-Prop (P x) Q
-ev-intro-exists-Prop P Q H x p = H (intro-exists P x p)
-
-elim-exists-Prop :
- {l1 l2 l3 : Level} {A : UU l1} (P : A → Prop l2) (Q : Prop l3) →
- ((x : A) → type-hom-Prop (P x) Q) → type-hom-Prop (exists-Prop A P) Q
-elim-exists-Prop P Q f =
- map-universal-property-trunc-Prop Q (ind-Σ f)
-
-abstract
- is-equiv-ev-intro-exists-Prop :
- {l1 l2 l3 : Level} (A : UU l1) (P : A → Prop l2) (Q : Prop l3) →
- is-equiv (ev-intro-exists-Prop P Q)
- is-equiv-ev-intro-exists-Prop A P Q =
- is-equiv-is-prop
- ( is-prop-type-hom-Prop (exists-Prop A P) Q)
- ( is-prop-Π ((λ x → is-prop-type-hom-Prop (P x) Q)))
- ( elim-exists-Prop P Q)
-
-is-least-upper-bound-exists-Prop :
- {l1 l2 l3 : Level} {A : UU l1} (P : A → Prop l2) (Q : Prop l3) →
- ((a : A) → type-hom-Prop (P a) Q) ↔ type-hom-Prop (exists-Prop A P) Q
-pr1 (is-least-upper-bound-exists-Prop P Q) = elim-exists-Prop P Q
-pr2 (is-least-upper-bound-exists-Prop P Q) h a p = h (intro-∃ a p)
-```
-
-### Conjunction distributes over existential quatification
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ up-exists :
+ universal-property-exists-structure A B (exists-structure-Prop A B)
+ up-exists Q = (ev-intro-exists , elim-exists Q)
+```
+
+### Propositions that satisfy the universal property of a existential quantification are equivalent to the existential quantification
```agda
module _
- {l1 l2 l3 : Level} (P : Prop l1) {A : UU l2} (Q : A → Prop l3)
+ {l1 l2 l3 : Level} {A : UU l1} {B : A → UU l2} (Q : Prop l3)
+ (up-Q : universal-property-exists-structure A B Q)
+ where
+
+ forward-implication-iff-universal-property-exists :
+ exists-structure A B → type-Prop Q
+ forward-implication-iff-universal-property-exists =
+ elim-exists Q (forward-implication (up-Q Q) id)
+
+ backward-implication-iff-universal-property-exists :
+ type-Prop Q → exists-structure A B
+ backward-implication-iff-universal-property-exists =
+ backward-implication (up-Q (exists-structure-Prop A B)) intro-exists
+
+ iff-universal-property-exists :
+ exists-structure A B ↔ type-Prop Q
+ iff-universal-property-exists =
+ ( forward-implication-iff-universal-property-exists ,
+ backward-implication-iff-universal-property-exists)
+```
+
+### Existential quantification of structure is the same as existential quantification over its propositional reflection
+
+We proceed by showing that the latter satisfies the universal property of the
+former.
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ universal-property-exists-structure-exists-trunc-Prop :
+ universal-property-exists-structure A B (exists-Prop A (trunc-Prop ∘ B))
+ universal-property-exists-structure-exists-trunc-Prop Q =
+ ( λ f a b → f (unit-trunc-Prop (a , unit-trunc-Prop b))) ,
+ ( λ f → rec-trunc-Prop Q (λ (a , |b|) → rec-trunc-Prop Q (f a) |b|))
+
+ compute-exists-trunc-Prop : exists-structure A B ↔ exists A (trunc-Prop ∘ B)
+ compute-exists-trunc-Prop =
+ iff-universal-property-exists
+ ( exists-Prop A (trunc-Prop ∘ B))
+ ( universal-property-exists-structure-exists-trunc-Prop)
+```
+
+### Taking the cartesian product with a proposition distributes over existential quantification of structures
+
+```agda
+module _
+ {l1 l2 l3 : Level} (P : Prop l1) {A : UU l2} {B : A → UU l3}
where
- iff-distributive-conjunction-exists-Prop :
- ( conjunction-Prop P (exists-Prop A Q)) ⇔
- ( exists-Prop A (λ a → conjunction-Prop P (Q a)))
- pr1 iff-distributive-conjunction-exists-Prop (p , e) =
- elim-exists-Prop Q
- ( exists-Prop A (λ a → conjunction-Prop P (Q a)))
- ( λ x q → intro-∃ x (p , q))
+ map-distributive-product-exists-structure :
+ type-Prop P × exists-structure A B →
+ exists-structure A (λ x → type-Prop P × B x)
+ map-distributive-product-exists-structure (p , e) =
+ elim-exists
+ ( exists-structure-Prop A (λ x → type-Prop P × B x))
+ ( λ x q → intro-exists x (p , q))
( e)
- pr2 iff-distributive-conjunction-exists-Prop =
- elim-exists-Prop
- ( λ x → conjunction-Prop P (Q x))
- ( conjunction-Prop P (exists-Prop A Q))
- ( λ x (p , q) → (p , intro-∃ x q))
-
- distributive-conjunction-exists-Prop :
- conjunction-Prop P (exists-Prop A Q) =
- exists-Prop A (λ a → conjunction-Prop P (Q a))
- distributive-conjunction-exists-Prop =
+
+ map-inv-distributive-product-exists-structure :
+ exists-structure A (λ x → type-Prop P × B x) →
+ type-Prop P × exists-structure A B
+ map-inv-distributive-product-exists-structure =
+ elim-exists
+ ( P ∧ exists-structure-Prop A B)
+ ( λ x (p , q) → (p , intro-exists x q))
+
+ iff-distributive-product-exists-structure :
+ ( type-Prop P × exists-structure A B) ↔
+ ( exists-structure A (λ x → type-Prop P × B x))
+ iff-distributive-product-exists-structure =
+ ( map-distributive-product-exists-structure ,
+ map-inv-distributive-product-exists-structure)
+
+ eq-distributive-product-exists-structure :
+ P ∧ exists-structure-Prop A B =
+ exists-structure-Prop A (λ x → type-Prop P × B x)
+ eq-distributive-product-exists-structure =
eq-iff'
- ( conjunction-Prop P (exists-Prop A Q))
- ( exists-Prop A (λ a → conjunction-Prop P (Q a)))
- ( iff-distributive-conjunction-exists-Prop)
+ ( P ∧ exists-structure-Prop A B)
+ ( exists-structure-Prop A (λ x → type-Prop P × B x))
+ ( iff-distributive-product-exists-structure)
```
+
+### Conjunction distributes over existential quantification
+
+```agda
+module _
+ {l1 l2 l3 : Level} (P : Prop l1) {A : UU l2} (Q : A → Prop l3)
+ where
+
+ map-distributive-conjunction-exists :
+ type-Prop (P ∧ (∃ A Q) ⇒ ∃ A (λ x → P ∧ Q x))
+ map-distributive-conjunction-exists =
+ map-distributive-product-exists-structure P
+
+ map-inv-distributive-conjunction-exists :
+ type-Prop (∃ A (λ x → P ∧ Q x) ⇒ P ∧ (∃ A Q))
+ map-inv-distributive-conjunction-exists =
+ map-inv-distributive-product-exists-structure P
+
+ iff-distributive-conjunction-exists :
+ type-Prop (P ∧ ∃ A Q ⇔ ∃ A (λ x → P ∧ Q x))
+ iff-distributive-conjunction-exists =
+ iff-distributive-product-exists-structure P
+
+ eq-distributive-conjunction-exists :
+ P ∧ (∃ A Q) = ∃ A (λ x → P ∧ Q x)
+ eq-distributive-conjunction-exists =
+ eq-distributive-product-exists-structure P
+```
+
+## See also
+
+- Existential quantification is the indexed counterpart to
+ [disjunction](foundation.disjunction.md)
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [existential quantifier](https://ncatlab.org/nlab/show/existential+quantifier)
+ at $n$Lab
diff --git a/src/foundation/fibered-equivalences.lagda.md b/src/foundation/fibered-equivalences.lagda.md
index 9b4a35d2c4..2b90a61064 100644
--- a/src/foundation/fibered-equivalences.lagda.md
+++ b/src/foundation/fibered-equivalences.lagda.md
@@ -11,6 +11,7 @@ open import foundation.dependent-pair-types
open import foundation.embeddings
open import foundation.equivalences
open import foundation.fibered-maps
+open import foundation.logical-equivalences
open import foundation.pullbacks
open import foundation.slice
open import foundation.universe-levels
@@ -333,7 +334,7 @@ module _
is-pullback (pr1 ihH) g (cone-fibered-map f g ihH) ≃
is-fibered-equiv-fibered-map f g ihH
equiv-is-fibered-equiv-is-pullback is-equiv-i =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-pullback (pr1 ihH) g (cone-fibered-map f g ihH))
( is-prop-is-fibered-equiv-fibered-map f g ihH)
( is-fibered-equiv-is-pullback is-equiv-i)
diff --git a/src/foundation/functoriality-propositional-truncation.lagda.md b/src/foundation/functoriality-propositional-truncation.lagda.md
index 8df3a04a2a..1c895736c6 100644
--- a/src/foundation/functoriality-propositional-truncation.lagda.md
+++ b/src/foundation/functoriality-propositional-truncation.lagda.md
@@ -10,6 +10,7 @@ module foundation.functoriality-propositional-truncation where
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.function-extensionality
+open import foundation.logical-equivalences
open import foundation.propositional-truncations
open import foundation.universe-levels
open import foundation.whiskering-homotopies-composition
@@ -116,7 +117,7 @@ abstract
(A ≃ B) → (type-trunc-Prop A ≃ type-trunc-Prop B)
pr1 (equiv-trunc-Prop e) = map-equiv-trunc-Prop e
pr2 (equiv-trunc-Prop e) =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-trunc-Prop)
( is-prop-type-trunc-Prop)
( map-inv-equiv-trunc-Prop e)
diff --git a/src/foundation/functoriality-set-truncation.lagda.md b/src/foundation/functoriality-set-truncation.lagda.md
index 136c3f4933..c0539827f8 100644
--- a/src/foundation/functoriality-set-truncation.lagda.md
+++ b/src/foundation/functoriality-set-truncation.lagda.md
@@ -11,8 +11,10 @@ open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.functoriality-truncation
open import foundation.images
+open import foundation.injective-maps
open import foundation.propositional-truncations
open import foundation.set-truncations
+open import foundation.sets
open import foundation.slice
open import foundation.surjective-maps
open import foundation.uniqueness-image
@@ -30,9 +32,7 @@ open import foundation-core.function-types
open import foundation-core.functoriality-dependent-pair-types
open import foundation-core.homotopies
open import foundation-core.identity-types
-open import foundation-core.injective-maps
open import foundation-core.propositions
-open import foundation-core.sets
open import foundation-core.truncation-levels
```
diff --git a/src/foundation/homotopy-preorder-of-types.lagda.md b/src/foundation/homotopy-preorder-of-types.lagda.md
new file mode 100644
index 0000000000..c6c1e1902e
--- /dev/null
+++ b/src/foundation/homotopy-preorder-of-types.lagda.md
@@ -0,0 +1,60 @@
+# The homotopy preorder of types
+
+```agda
+module
+ foundation.homotopy-preorder-of-types
+ where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.empty-types
+open import foundation.identity-types
+open import foundation.mere-functions
+open import foundation.propositional-truncations
+open import foundation.propositions
+open import foundation.sets
+open import foundation.universe-levels
+
+open import order-theory.large-preorders
+open import order-theory.posets
+open import order-theory.preorders
+```
+
+
+
+## Idea
+
+The {{#concept "homotopy preorder of types" Agda=Homotopy-Type-Large-Preorder}}
+is the [(large) preorder](order-theory.large-preorders.md) whose objects are
+types, and whose ordering relation is defined by
+[mere functions](foundation.mere-functions.md), i.e. by the
+[propositional truncation](foundation.propositional-truncations.md) of the
+function types:
+
+```text
+ A ≤ B := ║(A → B)║₋₁.
+```
+
+## Definitions
+
+### The large homotopy preorder of types
+
+```agda
+Homotopy-Type-Large-Preorder : Large-Preorder lsuc (_⊔_)
+Homotopy-Type-Large-Preorder =
+ λ where
+ .type-Large-Preorder l → UU l
+ .leq-prop-Large-Preorder → prop-mere-function
+ .refl-leq-Large-Preorder → refl-mere-function
+ .transitive-leq-Large-Preorder X Y Z → transitive-mere-function
+```
+
+### The small homotopy preorder of types
+
+```agda
+Homotopy-Type-Preorder : (l : Level) → Preorder (lsuc l) l
+Homotopy-Type-Preorder = preorder-Large-Preorder Homotopy-Type-Large-Preorder
+```
diff --git a/src/foundation/impredicative-encodings.lagda.md b/src/foundation/impredicative-encodings.lagda.md
index 274f2edc0e..dbf1136729 100644
--- a/src/foundation/impredicative-encodings.lagda.md
+++ b/src/foundation/impredicative-encodings.lagda.md
@@ -10,15 +10,18 @@ module foundation.impredicative-encodings where
open import foundation.conjunction
open import foundation.dependent-pair-types
open import foundation.disjunction
+open import foundation.empty-types
open import foundation.existential-quantification
+open import foundation.homotopies
open import foundation.logical-equivalences
open import foundation.negation
open import foundation.propositional-truncations
+open import foundation.universal-quantification
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.identity-types
open import foundation-core.propositions
open import foundation-core.sets
@@ -28,278 +31,278 @@ open import foundation-core.sets
## Idea
-By quantifying over all propositions in a universe, we can define all the
-logical operations. The idea is to use the fact that any proposition `P` is
-equivalent to the proposition `(Q : Prop l) → (P ⇒ Q) ⇒ Q`, which can be thought
-of as the least proposition `Q` containing `P`.
+By universally quantifying over all
+[propositions](foundation-core.propositions.md) in a universe, we can define all
+the logical operations. The idea is to use the fact that any proposition `P` is
+[equivalent](foundation.logical-equivalences.md) to the proposition
+`(Q : Prop l) → (P ⇒ Q) ⇒ Q`, which can be thought of as the least proposition
+`Q` containing `P`.
### The impredicative encoding of the propositional truncation
```agda
-impredicative-trunc-Prop :
- {l : Level} → UU l → Prop (lsuc l)
-impredicative-trunc-Prop {l} A =
- Π-Prop
- ( Prop l)
- ( λ Q → function-Prop (A → type-Prop Q) Q)
-
-type-impredicative-trunc-Prop :
- {l : Level} → UU l → UU (lsuc l)
-type-impredicative-trunc-Prop {l} A =
- type-Prop (impredicative-trunc-Prop A)
-
-map-impredicative-trunc-Prop :
- {l : Level} (A : UU l) →
- type-trunc-Prop A → type-impredicative-trunc-Prop A
-map-impredicative-trunc-Prop {l} A =
- map-universal-property-trunc-Prop
- ( impredicative-trunc-Prop A)
- ( λ x Q f → f x)
-
-map-inv-impredicative-trunc-Prop :
- {l : Level} (A : UU l) →
- type-impredicative-trunc-Prop A → type-trunc-Prop A
-map-inv-impredicative-trunc-Prop A H =
- H (trunc-Prop A) unit-trunc-Prop
-
-equiv-impredicative-trunc-Prop :
- {l : Level} (A : UU l) →
- type-trunc-Prop A ≃ type-impredicative-trunc-Prop A
-equiv-impredicative-trunc-Prop A =
- equiv-iff
- ( trunc-Prop A)
- ( impredicative-trunc-Prop A)
- ( map-impredicative-trunc-Prop A)
- ( map-inv-impredicative-trunc-Prop A)
+module _
+ {l : Level} (A : UU l)
+ where
+
+ impredicative-trunc-Prop : Prop (lsuc l)
+ impredicative-trunc-Prop =
+ ∀' (Prop l) (λ Q → function-Prop (A → type-Prop Q) Q)
+
+ type-impredicative-trunc-Prop : UU (lsuc l)
+ type-impredicative-trunc-Prop =
+ type-Prop impredicative-trunc-Prop
+
+ map-impredicative-trunc-Prop :
+ type-trunc-Prop A → type-impredicative-trunc-Prop
+ map-impredicative-trunc-Prop =
+ map-universal-property-trunc-Prop
+ ( impredicative-trunc-Prop)
+ ( λ x Q f → f x)
+
+ map-inv-impredicative-trunc-Prop :
+ type-impredicative-trunc-Prop → type-trunc-Prop A
+ map-inv-impredicative-trunc-Prop H =
+ H (trunc-Prop A) unit-trunc-Prop
+
+ equiv-impredicative-trunc-Prop :
+ type-trunc-Prop A ≃ type-impredicative-trunc-Prop
+ equiv-impredicative-trunc-Prop =
+ equiv-iff
+ ( trunc-Prop A)
+ ( impredicative-trunc-Prop)
+ ( map-impredicative-trunc-Prop)
+ ( map-inv-impredicative-trunc-Prop)
```
### The impredicative encoding of conjunction
```agda
-impredicative-conjunction-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → Prop (lsuc (l1 ⊔ l2))
-impredicative-conjunction-Prop {l1} {l2} P1 P2 =
- Π-Prop
- ( Prop (l1 ⊔ l2))
- ( λ Q → function-Prop (type-Prop P1 → (type-Prop P2 → type-Prop Q)) Q)
-
-type-impredicative-conjunction-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → UU (lsuc (l1 ⊔ l2))
-type-impredicative-conjunction-Prop P1 P2 =
- type-Prop (impredicative-conjunction-Prop P1 P2)
-
-map-impredicative-conjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-conjunction-Prop P1 P2 → type-impredicative-conjunction-Prop P1 P2
-map-impredicative-conjunction-Prop {l1} {l2} P1 P2 (pair p1 p2) Q f =
- f p1 p2
-
-map-inv-impredicative-conjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-impredicative-conjunction-Prop P1 P2 → type-conjunction-Prop P1 P2
-map-inv-impredicative-conjunction-Prop P1 P2 H =
- H (conjunction-Prop P1 P2) (λ p1 p2 → pair p1 p2)
-
-equiv-impredicative-conjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-conjunction-Prop P1 P2 ≃ type-impredicative-conjunction-Prop P1 P2
-equiv-impredicative-conjunction-Prop P1 P2 =
- equiv-iff
- ( conjunction-Prop P1 P2)
- ( impredicative-conjunction-Prop P1 P2)
- ( map-impredicative-conjunction-Prop P1 P2)
- ( map-inv-impredicative-conjunction-Prop P1 P2)
+module _
+ {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2)
+ where
+
+ impredicative-conjunction-Prop : Prop (lsuc (l1 ⊔ l2))
+ impredicative-conjunction-Prop =
+ ∀' (Prop (l1 ⊔ l2)) (λ Q → (P1 ⇒ P2 ⇒ Q) ⇒ Q)
+
+ type-impredicative-conjunction-Prop : UU (lsuc (l1 ⊔ l2))
+ type-impredicative-conjunction-Prop =
+ type-Prop impredicative-conjunction-Prop
+
+ map-product-impredicative-conjunction-Prop :
+ type-conjunction-Prop P1 P2 → type-impredicative-conjunction-Prop
+ map-product-impredicative-conjunction-Prop (p1 , p2) Q f = f p1 p2
+
+ map-inv-product-impredicative-conjunction-Prop :
+ type-impredicative-conjunction-Prop → type-conjunction-Prop P1 P2
+ map-inv-product-impredicative-conjunction-Prop H = H (P1 ∧ P2) pair
+
+ equiv-product-impredicative-conjunction-Prop :
+ type-conjunction-Prop P1 P2 ≃ type-impredicative-conjunction-Prop
+ equiv-product-impredicative-conjunction-Prop =
+ equiv-iff
+ ( P1 ∧ P2)
+ ( impredicative-conjunction-Prop)
+ ( map-product-impredicative-conjunction-Prop)
+ ( map-inv-product-impredicative-conjunction-Prop)
```
### The impredicative encoding of disjunction
```agda
-impredicative-disjunction-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → Prop (lsuc (l1 ⊔ l2))
-impredicative-disjunction-Prop {l1} {l2} P1 P2 =
- Π-Prop
- ( Prop (l1 ⊔ l2))
- ( λ Q →
- function-Prop
- ( type-implication-Prop P1 Q)
- ( function-Prop (type-implication-Prop P2 Q) Q))
-
-type-impredicative-disjunction-Prop :
- {l1 l2 : Level} → Prop l1 → Prop l2 → UU (lsuc (l1 ⊔ l2))
-type-impredicative-disjunction-Prop P1 P2 =
- type-Prop (impredicative-disjunction-Prop P1 P2)
-
-map-impredicative-disjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-disjunction-Prop P1 P2 → type-impredicative-disjunction-Prop P1 P2
-map-impredicative-disjunction-Prop {l1} {l2} P1 P2 =
- map-universal-property-trunc-Prop
- ( impredicative-disjunction-Prop P1 P2)
- ( rec-coproduct
- ( λ x Q f1 f2 → f1 x)
- ( λ y Q f1 f2 → f2 y))
-
-map-inv-impredicative-disjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-impredicative-disjunction-Prop P1 P2 → type-disjunction-Prop P1 P2
-map-inv-impredicative-disjunction-Prop P1 P2 H =
- H ( disjunction-Prop P1 P2)
- ( inl-disjunction-Prop P1 P2)
- ( inr-disjunction-Prop P1 P2)
-
-equiv-impredicative-disjunction-Prop :
- {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2) →
- type-disjunction-Prop P1 P2 ≃ type-impredicative-disjunction-Prop P1 P2
-equiv-impredicative-disjunction-Prop P1 P2 =
- equiv-iff
- ( disjunction-Prop P1 P2)
- ( impredicative-disjunction-Prop P1 P2)
- ( map-impredicative-disjunction-Prop P1 P2)
- ( map-inv-impredicative-disjunction-Prop P1 P2)
+module _
+ {l1 l2 : Level} (P1 : Prop l1) (P2 : Prop l2)
+ where
+
+ impredicative-disjunction-Prop : Prop (lsuc (l1 ⊔ l2))
+ impredicative-disjunction-Prop =
+ ∀' (Prop (l1 ⊔ l2)) (λ Q → (P1 ⇒ Q) ⇒ (P2 ⇒ Q) ⇒ Q)
+
+ type-impredicative-disjunction-Prop : UU (lsuc (l1 ⊔ l2))
+ type-impredicative-disjunction-Prop =
+ type-Prop impredicative-disjunction-Prop
+
+ map-impredicative-disjunction-Prop :
+ type-disjunction-Prop P1 P2 → type-impredicative-disjunction-Prop
+ map-impredicative-disjunction-Prop =
+ map-universal-property-trunc-Prop
+ ( impredicative-disjunction-Prop)
+ ( rec-coproduct (λ x Q f1 f2 → f1 x) (λ y Q f1 f2 → f2 y))
+
+ map-inv-impredicative-disjunction-Prop :
+ type-impredicative-disjunction-Prop → type-disjunction-Prop P1 P2
+ map-inv-impredicative-disjunction-Prop H =
+ H (P1 ∨ P2) (inl-disjunction) (inr-disjunction)
+
+ equiv-impredicative-disjunction-Prop :
+ type-disjunction-Prop P1 P2 ≃ type-impredicative-disjunction-Prop
+ equiv-impredicative-disjunction-Prop =
+ equiv-iff
+ ( P1 ∨ P2)
+ ( impredicative-disjunction-Prop)
+ ( map-impredicative-disjunction-Prop)
+ ( map-inv-impredicative-disjunction-Prop)
+```
+
+### The impredicative encoding of the empty type
+
+```agda
+impredicative-empty-Prop : Prop (lsuc lzero)
+impredicative-empty-Prop = ∀' (Prop lzero) (λ P → P)
+
+type-impredicative-empty-Prop : UU (lsuc lzero)
+type-impredicative-empty-Prop = type-Prop impredicative-empty-Prop
+
+is-empty-impredicative-empty-Prop :
+ is-empty type-impredicative-empty-Prop
+is-empty-impredicative-empty-Prop e = e empty-Prop
+
+equiv-impredicative-empty-Prop :
+ empty ≃ type-impredicative-empty-Prop
+equiv-impredicative-empty-Prop =
+ equiv-is-empty id is-empty-impredicative-empty-Prop
```
### The impredicative encoding of negation
```agda
-impredicative-neg-Prop :
- {l : Level} → UU l → Prop (lsuc l)
-impredicative-neg-Prop {l} A =
- Π-Prop (Prop l) (λ Q → function-Prop A Q)
-
-type-impredicative-neg-Prop :
- {l : Level} → UU l → UU (lsuc l)
-type-impredicative-neg-Prop A =
- type-Prop (impredicative-neg-Prop A)
-
-map-impredicative-neg-Prop :
- {l : Level} (A : UU l) →
- ¬ A → type-impredicative-neg-Prop A
-map-impredicative-neg-Prop A f Q a = ex-falso (f a)
-
-map-inv-impredicative-neg-Prop :
- {l : Level} (A : UU l) →
- type-impredicative-neg-Prop A → ¬ A
-map-inv-impredicative-neg-Prop A H a = H (neg-Prop' A) a a
-
-equiv-impredicative-neg-Prop :
- {l : Level} (A : UU l) →
- ¬ A ≃ type-impredicative-neg-Prop A
-equiv-impredicative-neg-Prop A =
- equiv-iff
- ( neg-Prop' A)
- ( impredicative-neg-Prop A)
- ( map-impredicative-neg-Prop A)
- ( map-inv-impredicative-neg-Prop A)
+module _
+ {l : Level} (A : UU l)
+ where
+
+ impredicative-neg-Prop : Prop (lsuc l)
+ impredicative-neg-Prop =
+ Π-Prop (Prop l) (λ Q → function-Prop A Q)
+
+ type-impredicative-neg-Prop : UU (lsuc l)
+ type-impredicative-neg-Prop = type-Prop impredicative-neg-Prop
+
+ map-impredicative-neg-Prop : ¬ A → type-impredicative-neg-Prop
+ map-impredicative-neg-Prop f Q a = ex-falso (f a)
+
+ map-inv-impredicative-neg-Prop : type-impredicative-neg-Prop → ¬ A
+ map-inv-impredicative-neg-Prop H a = H (neg-type-Prop A) a a
+
+ equiv-impredicative-neg-Prop : ¬ A ≃ type-impredicative-neg-Prop
+ equiv-impredicative-neg-Prop =
+ equiv-iff
+ ( neg-type-Prop A)
+ ( impredicative-neg-Prop)
+ ( map-impredicative-neg-Prop)
+ ( map-inv-impredicative-neg-Prop)
```
### The impredicative encoding of existential quantification
```agda
-impredicative-exists-Prop :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) → Prop (lsuc (l1 ⊔ l2))
-impredicative-exists-Prop {l1} {l2} {A} P =
- Π-Prop
- ( Prop (l1 ⊔ l2))
- ( λ Q → function-Prop ((x : A) → type-Prop (P x) → type-Prop Q) Q)
-
-type-impredicative-exists-Prop :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) → UU (lsuc (l1 ⊔ l2))
-type-impredicative-exists-Prop P =
- type-Prop (impredicative-exists-Prop P)
-
-map-impredicative-exists-Prop :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) →
- exists A P → type-impredicative-exists-Prop P
-map-impredicative-exists-Prop {l1} {l2} {A} P =
- map-universal-property-trunc-Prop
- ( impredicative-exists-Prop P)
- ( ind-Σ (λ x y Q h → h x y))
-
-map-inv-impredicative-exists-Prop :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) →
- type-impredicative-exists-Prop P → exists A P
-map-inv-impredicative-exists-Prop {A = A} P H =
- H ( exists-Prop A P)
- ( λ x y → unit-trunc-Prop (pair x y))
-
-equiv-impredicative-exists-Prop :
- {l1 l2 : Level} {A : UU l1} (P : A → Prop l2) →
- exists A P ≃ type-impredicative-exists-Prop P
-equiv-impredicative-exists-Prop {A = A} P =
- equiv-iff
- ( exists-Prop A P)
- ( impredicative-exists-Prop P)
- ( map-impredicative-exists-Prop P)
- ( map-inv-impredicative-exists-Prop P)
+module _
+ {l1 l2 : Level} {A : UU l1} (P : A → Prop l2)
+ where
+
+ impredicative-exists-Prop : Prop (lsuc (l1 ⊔ l2))
+ impredicative-exists-Prop =
+ ∀' (Prop (l1 ⊔ l2)) (λ Q → (∀' A (λ x → P x ⇒ Q)) ⇒ Q)
+
+ type-impredicative-exists-Prop : UU (lsuc (l1 ⊔ l2))
+ type-impredicative-exists-Prop =
+ type-Prop impredicative-exists-Prop
+
+ map-impredicative-exists-Prop :
+ exists A P → type-impredicative-exists-Prop
+ map-impredicative-exists-Prop =
+ map-universal-property-trunc-Prop
+ ( impredicative-exists-Prop)
+ ( ind-Σ (λ x y Q h → h x y))
+
+ map-inv-impredicative-exists-Prop :
+ type-impredicative-exists-Prop → exists A P
+ map-inv-impredicative-exists-Prop H =
+ H (∃ A P) (λ x y → unit-trunc-Prop (x , y))
+
+ equiv-impredicative-exists-Prop :
+ exists A P ≃ type-impredicative-exists-Prop
+ equiv-impredicative-exists-Prop =
+ equiv-iff
+ ( ∃ A P)
+ ( impredicative-exists-Prop)
+ ( map-impredicative-exists-Prop)
+ ( map-inv-impredicative-exists-Prop)
```
### The impredicative encoding of the based identity type of a set
```agda
-impredicative-based-id-Prop :
- {l : Level} (A : Set l) (a x : type-Set A) → Prop (lsuc l)
-impredicative-based-id-Prop {l} A a x =
- Π-Prop (type-Set A → Prop l) (λ Q → hom-Prop (Q a) (Q x))
-
-type-impredicative-based-id-Prop :
- {l : Level} (A : Set l) (a x : type-Set A) → UU (lsuc l)
-type-impredicative-based-id-Prop A a x =
- type-Prop (impredicative-based-id-Prop A a x)
-
-map-impredicative-based-id-Prop :
- {l : Level} (A : Set l) (a x : type-Set A) →
- a = x → type-impredicative-based-id-Prop A a x
-map-impredicative-based-id-Prop A a .a refl Q p = p
-
-map-inv-impredicative-based-id-Prop :
- {l : Level} (A : Set l) (a x : type-Set A) →
- type-impredicative-based-id-Prop A a x → a = x
-map-inv-impredicative-based-id-Prop A a x H =
- H (λ x → pair (a = x) (is-set-type-Set A a x)) refl
-
-equiv-impredicative-based-id-Prop :
- {l : Level} (A : Set l) (a x : type-Set A) →
- (a = x) ≃ type-impredicative-based-id-Prop A a x
-equiv-impredicative-based-id-Prop A a x =
- equiv-iff
- ( pair (a = x) (is-set-type-Set A a x))
- ( impredicative-based-id-Prop A a x)
- ( map-impredicative-based-id-Prop A a x)
- ( map-inv-impredicative-based-id-Prop A a x)
+module _
+ {l : Level} (A : Set l) (a x : type-Set A)
+ where
+
+ impredicative-based-id-Prop : Prop (lsuc l)
+ impredicative-based-id-Prop = ∀' (type-Set A → Prop l) (λ Q → Q a ⇒ Q x)
+
+ type-impredicative-based-id-Prop : UU (lsuc l)
+ type-impredicative-based-id-Prop = type-Prop impredicative-based-id-Prop
+
+ map-impredicative-based-id-Prop :
+ a = x → type-impredicative-based-id-Prop
+ map-impredicative-based-id-Prop refl Q p = p
+
+ map-inv-impredicative-based-id-Prop :
+ type-impredicative-based-id-Prop → a = x
+ map-inv-impredicative-based-id-Prop H = H (Id-Prop A a) refl
+
+ equiv-impredicative-based-id-Prop :
+ (a = x) ≃ type-impredicative-based-id-Prop
+ equiv-impredicative-based-id-Prop =
+ equiv-iff
+ ( Id-Prop A a x)
+ ( impredicative-based-id-Prop)
+ ( map-impredicative-based-id-Prop)
+ ( map-inv-impredicative-based-id-Prop)
```
### The impredicative encoding of Martin-Löf's identity type of a set
```agda
-impredicative-id-Prop :
- {l : Level} (A : Set l) (x y : type-Set A) → Prop (lsuc l)
-impredicative-id-Prop {l} A x y =
- Π-Prop (type-Set A → type-Set A → Prop l)
- (λ Q → function-Prop ((a : type-Set A) → type-Prop (Q a a)) (Q x y))
-
-type-impredicative-id-Prop :
- {l : Level} (A : Set l) (x y : type-Set A) → UU (lsuc l)
-type-impredicative-id-Prop A x y =
- type-Prop (impredicative-id-Prop A x y)
-
-map-impredicative-id-Prop :
- {l : Level} (A : Set l) (x y : type-Set A) →
- x = y → type-impredicative-id-Prop A x y
-map-impredicative-id-Prop A x .x refl Q r = r x
-
-map-inv-impredicative-id-Prop :
- {l : Level} (A : Set l) (x y : type-Set A) →
- type-impredicative-id-Prop A x y → x = y
-map-inv-impredicative-id-Prop A x y H =
- H (λ a b → pair (a = b) (is-set-type-Set A a b)) (λ a → refl)
-
-equiv-impredicative-id-Prop :
- {l : Level} (A : Set l) (x y : type-Set A) →
- (x = y) ≃ type-impredicative-id-Prop A x y
-equiv-impredicative-id-Prop A x y =
- equiv-iff
- ( pair (x = y) (is-set-type-Set A x y))
- ( impredicative-id-Prop A x y)
- ( map-impredicative-id-Prop A x y)
- ( map-inv-impredicative-id-Prop A x y)
+module _
+ {l : Level} (A : Set l) (x y : type-Set A)
+ where
+
+ impredicative-id-Prop : Prop (lsuc l)
+ impredicative-id-Prop =
+ ∀'
+ ( type-Set A → type-Set A → Prop l)
+ ( λ Q → (∀' (type-Set A) (λ a → Q a a)) ⇒ Q x y)
+
+ type-impredicative-id-Prop : UU (lsuc l)
+ type-impredicative-id-Prop = type-Prop impredicative-id-Prop
+
+ map-impredicative-id-Prop :
+ x = y → type-impredicative-id-Prop
+ map-impredicative-id-Prop refl Q r = r x
+
+ map-inv-impredicative-id-Prop :
+ type-impredicative-id-Prop → x = y
+ map-inv-impredicative-id-Prop H = H (Id-Prop A) (refl-htpy)
+
+ equiv-impredicative-id-Prop :
+ (x = y) ≃ type-impredicative-id-Prop
+ equiv-impredicative-id-Prop =
+ equiv-iff
+ ( Id-Prop A x y)
+ ( impredicative-id-Prop)
+ ( map-impredicative-id-Prop)
+ ( map-inv-impredicative-id-Prop)
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [Constructing coproduct types and boolean types from universes](https://mathoverflow.net/questions/457904/constructing-coproduct-types-and-boolean-types-from-universes)
+ at mathoverflow
diff --git a/src/foundation/impredicative-universes.lagda.md b/src/foundation/impredicative-universes.lagda.md
index 8699e535dc..c1f2ed3f59 100644
--- a/src/foundation/impredicative-universes.lagda.md
+++ b/src/foundation/impredicative-universes.lagda.md
@@ -17,7 +17,8 @@ open import foundation-core.small-types
## Idea
-A universe `U` is impredicative if the type of propositions in `U` is `U`-small.
+A universe `𝒰` is {{#concept "impredicative"}} if the type of
+[propositions](foundation-core.propositions.md) in `𝒰` is `𝒰`-small.
## Definition
@@ -25,3 +26,7 @@ A universe `U` is impredicative if the type of propositions in `U` is `U`-small.
is-impredicative-UU : (l : Level) → UU (lsuc l)
is-impredicative-UU l = is-small l (Prop l)
```
+
+## See also
+
+- [Impredicative encodings of the logical operations](foundation.impredicative-encodings.md)
diff --git a/src/foundation/inhabited-types.lagda.md b/src/foundation/inhabited-types.lagda.md
index 6360251748..874d306572 100644
--- a/src/foundation/inhabited-types.lagda.md
+++ b/src/foundation/inhabited-types.lagda.md
@@ -46,7 +46,7 @@ is-inhabited-Prop X = trunc-Prop X
is-inhabited : {l : Level} → UU l → UU l
is-inhabited X = type-Prop (is-inhabited-Prop X)
-is-property-is-inhabited : {l : Level} → (X : UU l) → is-prop (is-inhabited X)
+is-property-is-inhabited : {l : Level} (X : UU l) → is-prop (is-inhabited X)
is-property-is-inhabited X = is-prop-type-Prop (is-inhabited-Prop X)
Inhabited-Type : (l : Level) → UU (lsuc l)
diff --git a/src/foundation/injective-maps.lagda.md b/src/foundation/injective-maps.lagda.md
index 5eaf79f193..38991f22b2 100644
--- a/src/foundation/injective-maps.lagda.md
+++ b/src/foundation/injective-maps.lagda.md
@@ -9,10 +9,17 @@ open import foundation-core.injective-maps public
Imports
```agda
+open import foundation.dependent-pair-types
+open import foundation.logical-equivalences
open import foundation.universe-levels
+open import foundation-core.embeddings
open import foundation-core.empty-types
+open import foundation-core.identity-types
open import foundation-core.negation
+open import foundation-core.propositional-maps
+open import foundation-core.propositions
+open import foundation-core.sets
```
@@ -49,3 +56,51 @@ is-injective-is-empty :
is-empty A → is-injective f
is-injective-is-empty f is-empty-A {x} = ex-falso (is-empty-A x)
```
+
+### Any injective map between sets is an embedding
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ abstract
+ is-emb-is-injective' :
+ (is-set-A : is-set A) (is-set-B : is-set B) (f : A → B) →
+ is-injective f → is-emb f
+ is-emb-is-injective' is-set-A is-set-B f is-injective-f x y =
+ is-equiv-has-converse-is-prop
+ ( is-set-A x y)
+ ( is-set-B (f x) (f y))
+ ( is-injective-f)
+
+ is-emb-is-injective :
+ {f : A → B} → is-set B → is-injective f → is-emb f
+ is-emb-is-injective {f} H I =
+ is-emb-is-injective' (is-set-is-injective H I) H f I
+
+ is-prop-map-is-injective :
+ {f : A → B} → is-set B → is-injective f → is-prop-map f
+ is-prop-map-is-injective {f} H I =
+ is-prop-map-is-emb (is-emb-is-injective H I)
+```
+
+### For a map between sets, being injective is a property
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ is-prop-is-injective :
+ is-set A → (f : A → B) → is-prop (is-injective f)
+ is-prop-is-injective H f =
+ is-prop-implicit-Π
+ ( λ x →
+ is-prop-implicit-Π
+ ( λ y → is-prop-function-type (H x y)))
+
+ is-injective-Prop : is-set A → (A → B) → Prop (l1 ⊔ l2)
+ pr1 (is-injective-Prop H f) = is-injective f
+ pr2 (is-injective-Prop H f) = is-prop-is-injective H f
+```
diff --git a/src/foundation/intersections-subtypes.lagda.md b/src/foundation/intersections-subtypes.lagda.md
index 33b3583dbb..3ff9cc3f3b 100644
--- a/src/foundation/intersections-subtypes.lagda.md
+++ b/src/foundation/intersections-subtypes.lagda.md
@@ -14,6 +14,7 @@ open import foundation.large-locale-of-subtypes
open import foundation.powersets
open import foundation.universe-levels
+open import foundation-core.decidable-propositions
open import foundation-core.propositions
open import foundation-core.subtypes
diff --git a/src/foundation/isolated-elements.lagda.md b/src/foundation/isolated-elements.lagda.md
index c42a2f870c..ebf38b94f7 100644
--- a/src/foundation/isolated-elements.lagda.md
+++ b/src/foundation/isolated-elements.lagda.md
@@ -17,9 +17,11 @@ open import foundation.dependent-pair-types
open import foundation.embeddings
open import foundation.fundamental-theorem-of-identity-types
open import foundation.identity-types
+open import foundation.injective-maps
open import foundation.maybe
open import foundation.negated-equality
open import foundation.negation
+open import foundation.sets
open import foundation.type-arithmetic-unit-type
open import foundation.unit-type
open import foundation.universe-levels
@@ -32,9 +34,7 @@ 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.propositions
-open import foundation-core.sets
open import foundation-core.subtypes
open import foundation-core.torsorial-type-families
open import foundation-core.transport-along-identifications
@@ -214,7 +214,7 @@ module _
is-prop-is-isolated :
{l1 : Level} {A : UU l1} (a : A) → is-prop (is-isolated a)
is-prop-is-isolated a =
- is-prop-is-inhabited
+ is-prop-has-element
( λ H → is-prop-Π (is-prop-is-decidable ∘ is-prop-eq-isolated-element a H))
is-isolated-Prop :
@@ -257,7 +257,7 @@ module _
( is-emb-inclusion-isolated-element A)
( is-emb-is-injective
( is-set-isolated-element A)
- ( λ {star} {star} p → refl))
+ ( λ p → refl))
emb-point-isolated-element : unit ↪ A
pr1 emb-point-isolated-element = point-isolated-element
diff --git a/src/foundation/large-locale-of-propositions.lagda.md b/src/foundation/large-locale-of-propositions.lagda.md
index 233f214be3..816cd18c3d 100644
--- a/src/foundation/large-locale-of-propositions.lagda.md
+++ b/src/foundation/large-locale-of-propositions.lagda.md
@@ -9,8 +9,10 @@ module foundation.large-locale-of-propositions where
```agda
open import foundation.conjunction
open import foundation.existential-quantification
+open import foundation.logical-equivalences
open import foundation.propositional-extensionality
open import foundation.unit-type
+open import foundation.universal-property-cartesian-product-types
open import foundation.universe-levels
open import foundation-core.function-types
@@ -30,8 +32,17 @@ open import order-theory.top-elements-large-posets
## Idea
-The types of [propositions](foundation-core.propositions.md) `Prop l` combined
-form a [large locale](order-theory.large-locales.md).
+The [large locale](order-theory.large-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
+[large poset](order-theory.large-posets.md) the structure of a
+[large meet-semilattice](order-theory.large-meet-semilattices.md), and
+[existential quantification](foundation.existential-quantification.md) gives it
+the structure of a [large suplattice](order-theory.large-suplattices.md).
+
+**Note.** The collection of all propositions is large because we do not assume
+[propositional resizing](foundation.propositional-resizing.md).
## Definitions
@@ -60,8 +71,8 @@ has-meets-Prop-Large-Locale :
has-meets-Large-Poset Prop-Large-Poset
meet-has-meets-Large-Poset has-meets-Prop-Large-Locale = conjunction-Prop
is-greatest-binary-lower-bound-meet-has-meets-Large-Poset
- has-meets-Prop-Large-Locale =
- iff-universal-property-conjunction-Prop
+ has-meets-Prop-Large-Locale P Q R =
+ is-greatest-binary-lower-bound-conjunction-Prop P Q R
```
### The largest element in the large poset of propositions
@@ -96,10 +107,10 @@ is-large-suplattice-Prop-Large-Locale :
is-large-suplattice-Large-Poset lzero Prop-Large-Poset
sup-has-least-upper-bound-family-of-elements-Large-Poset
( is-large-suplattice-Prop-Large-Locale {I = I} P) =
- exists-Prop I P
+ ∃ I P
is-least-upper-bound-sup-has-least-upper-bound-family-of-elements-Large-Poset
- ( is-large-suplattice-Prop-Large-Locale {I = I} P) =
- is-least-upper-bound-exists-Prop P
+ ( is-large-suplattice-Prop-Large-Locale {I = I} P) R =
+ inv-iff (up-exists R)
```
### The large frame of propositions
@@ -113,7 +124,7 @@ is-large-meet-semilattice-Large-Frame Prop-Large-Frame =
is-large-suplattice-Large-Frame Prop-Large-Frame =
is-large-suplattice-Prop-Large-Locale
distributive-meet-sup-Large-Frame Prop-Large-Frame =
- distributive-conjunction-exists-Prop
+ eq-distributive-conjunction-exists
```
### The large locale of propositions
@@ -122,3 +133,7 @@ distributive-meet-sup-Large-Frame Prop-Large-Frame =
Prop-Large-Locale : Large-Locale lsuc (_⊔_) lzero
Prop-Large-Locale = Prop-Large-Frame
```
+
+## See also
+
+- [Propositional resizing](foundation.propositional-resizing.md)
diff --git a/src/foundation/law-of-excluded-middle.lagda.md b/src/foundation/law-of-excluded-middle.lagda.md
index d28bb30393..d82d58a594 100644
--- a/src/foundation/law-of-excluded-middle.lagda.md
+++ b/src/foundation/law-of-excluded-middle.lagda.md
@@ -22,7 +22,7 @@ open import univalent-combinatorics.2-element-types
## Idea
-The **law of excluded middle** asserts that any
+The {{#concept "law of excluded middle" Agda=LEM}} asserts that any
[proposition](foundation-core.propositions.md) `P` is
[decidable](foundation.decidable-types.md).
diff --git a/src/foundation/lawveres-fixed-point-theorem.lagda.md b/src/foundation/lawveres-fixed-point-theorem.lagda.md
index 00323ca065..bbe7fbd83e 100644
--- a/src/foundation/lawveres-fixed-point-theorem.lagda.md
+++ b/src/foundation/lawveres-fixed-point-theorem.lagda.md
@@ -21,22 +21,26 @@ open import foundation-core.identity-types
## Idea
-Lawvere's fixed point theorem asserts that if there is a surjective map
+{{#concept "Lawvere's fixed point theorem" Agda=fixed-point-theorem-Lawvere}}
+asserts that if there is a [surjective map](foundation.surjective-maps.md)
`A → (A → B)`, then any map `B → B` must have a fixed point.
## Theorem
```agda
-abstract
- fixed-point-theorem-Lawvere :
- {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → A → B} →
- is-surjective f → (h : B → B) → ∃ B (λ b → h b = b)
- fixed-point-theorem-Lawvere {A = A} {B} {f} H h =
- apply-universal-property-trunc-Prop
- ( H g)
- ( ∃-Prop B (λ b → h b = b))
- ( λ p → intro-∃ (f (pr1 p) (pr1 p)) (inv (htpy-eq (pr2 p) (pr1 p))))
- where
- g : A → B
- g a = h (f a a)
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2} {f : A → A → B}
+ where
+
+ abstract
+ fixed-point-theorem-Lawvere :
+ is-surjective f → (h : B → B) → exists-structure B (λ b → h b = b)
+ fixed-point-theorem-Lawvere H h =
+ apply-universal-property-trunc-Prop
+ ( H g)
+ ( exists-structure-Prop B (λ b → h b = b))
+ ( λ (x , p) → intro-exists (f x x) (inv (htpy-eq p x)))
+ where
+ g : A → B
+ g a = h (f a a)
```
diff --git a/src/foundation/limited-principle-of-omniscience.lagda.md b/src/foundation/limited-principle-of-omniscience.lagda.md
index bd00f84e57..5b6d52ed2e 100644
--- a/src/foundation/limited-principle-of-omniscience.lagda.md
+++ b/src/foundation/limited-principle-of-omniscience.lagda.md
@@ -11,6 +11,7 @@ open import elementary-number-theory.natural-numbers
open import foundation.disjunction
open import foundation.existential-quantification
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.identity-types
@@ -24,17 +25,19 @@ open import univalent-combinatorics.standard-finite-types
## Statement
-The **limited principle of omniscience** (LPO) asserts that for every sequence
-`f : ℕ → Fin 2` there either exists an `n` such that `f n = 1` or for all `n`
-we have `f n = 0`.
+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`
+there either [exists](foundation.existential-quantification.md) an `n` such that
+`f n = 1` or for all `n` we have `f n = 0`.
```agda
LPO : UU lzero
LPO =
(f : ℕ → Fin 2) →
type-disjunction-Prop
- ( ∃-Prop ℕ (λ n → f n = one-Fin 1))
- ( Π-Prop ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))
+ ( ∃ ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (one-Fin 1)))
+ ( ∀' ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))
```
## See also
diff --git a/src/foundation/logical-equivalences.lagda.md b/src/foundation/logical-equivalences.lagda.md
index 8c72c70a96..b03f4e4a5b 100644
--- a/src/foundation/logical-equivalences.lagda.md
+++ b/src/foundation/logical-equivalences.lagda.md
@@ -32,20 +32,44 @@ open import foundation-core.torsorial-type-families
## Idea
-**Logical equivalences** between two types `A` and `B` consist of a map `A → B`
-and a map `B → A`. The type of logical equivalences between types is the
-Curry-Howard interpretation of logical equivalences between
+{{#concept "Logical equivalences" Agda=}} between two types `A` and `B` consist
+of a map `A → B` and a map `B → A`. The type of logical equivalences between
+types is the Curry-Howard interpretation of logical equivalences between
[propositions](foundation-core.propositions.md).
## Definition
+### The structure on a map of being a logical equivalence
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ has-converse : (A → B) → UU (l1 ⊔ l2)
+ has-converse f = B → A
+
+ is-prop-has-converse :
+ is-prop A → (f : A → B) → is-prop (has-converse f)
+ is-prop-has-converse is-prop-A f = is-prop-function-type is-prop-A
+
+has-converse-Prop :
+ {l1 l2 : Level} (A : Prop l1) {B : UU l2} → (type-Prop A → B) → Prop (l1 ⊔ l2)
+has-converse-Prop A f =
+ ( has-converse f ,
+ is-prop-has-converse (is-prop-type-Prop A) f)
+```
+
### Logical equivalences between types
```agda
+iff : {l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2)
+iff A B = Σ (A → B) has-converse
+
infix 6 _↔_
_↔_ : {l1 l2 : Level} → UU l1 → UU l2 → UU (l1 ⊔ l2)
-A ↔ B = (A → B) × (B → A)
+_↔_ = iff
module _
{l1 l2 : Level} {A : UU l1} {B : UU l2} (H : A ↔ B)
@@ -80,8 +104,15 @@ module _
infix 6 _⇔_
- _⇔_ : UU (l1 ⊔ l2)
- _⇔_ = type-iff-Prop
+ _⇔_ : Prop (l1 ⊔ l2)
+ _⇔_ = iff-Prop
+```
+
+### The identity logical equivalence
+
+```agda
+id-iff : {l1 : Level} {A : UU l1} → A ↔ A
+id-iff = (id , id)
```
### Composition of logical equivalences
@@ -136,13 +167,44 @@ module _
### Logical equivalences between propositions induce equivalences
```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ abstract
+ is-equiv-has-converse-is-prop :
+ is-prop A → is-prop B → {f : A → B} → (B → A) → is-equiv f
+ is-equiv-has-converse-is-prop is-prop-A is-prop-B {f} g =
+ is-equiv-is-invertible
+ ( g)
+ ( λ y → eq-is-prop is-prop-B)
+ ( λ x → eq-is-prop is-prop-A)
+
+ abstract
+ equiv-iff-is-prop : is-prop A → is-prop B → (A → B) → (B → A) → A ≃ B
+ pr1 (equiv-iff-is-prop is-prop-A is-prop-B f g) = f
+ pr2 (equiv-iff-is-prop is-prop-A is-prop-B f g) =
+ is-equiv-has-converse-is-prop is-prop-A is-prop-B g
+
module _
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2)
where
- equiv-iff' : (P ⇔ Q) → (type-Prop P ≃ type-Prop Q)
- pr1 (equiv-iff' t) = pr1 t
- pr2 (equiv-iff' t) = is-equiv-is-prop (pr2 P) (pr2 Q) (pr2 t)
+ abstract
+ is-equiv-has-converse :
+ {f : type-Prop P → type-Prop Q} → (type-Prop Q → type-Prop P) → is-equiv f
+ is-equiv-has-converse =
+ is-equiv-has-converse-is-prop
+ ( is-prop-type-Prop P)
+ ( is-prop-type-Prop Q)
+
+ equiv-iff' : type-Prop (P ⇔ Q) → (type-Prop P ≃ type-Prop Q)
+ pr1 (equiv-iff' t) = forward-implication t
+ pr2 (equiv-iff' t) =
+ is-equiv-has-converse-is-prop
+ ( is-prop-type-Prop P)
+ ( is-prop-type-Prop Q)
+ ( backward-implication t)
equiv-iff :
(type-Prop P → type-Prop Q) → (type-Prop Q → type-Prop P) →
@@ -176,7 +238,7 @@ compute-fiber-iff-equiv {A = A} {B} (f , g) =
### Two equal propositions are logically equivalent
```agda
-iff-eq : {l1 : Level} {P Q : Prop l1} → P = Q → P ⇔ Q
+iff-eq : {l1 : Level} {P Q : Prop l1} → P = Q → type-Prop (P ⇔ Q)
pr1 (iff-eq refl) = id
pr2 (iff-eq refl) = id
```
@@ -189,14 +251,14 @@ abstract
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
is-equiv (equiv-iff' P Q)
is-equiv-equiv-iff P Q =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-iff-Prop P Q)
( is-prop-type-equiv-Prop P Q)
( iff-equiv)
equiv-equiv-iff :
{l1 l2 : Level} (P : Prop l1) (Q : Prop l2) →
- (P ⇔ Q) ≃ (type-Prop P ≃ type-Prop Q)
+ (type-Prop P ↔ type-Prop Q) ≃ (type-Prop P ≃ type-Prop Q)
pr1 (equiv-equiv-iff P Q) = equiv-iff' P Q
pr2 (equiv-equiv-iff P Q) = is-equiv-equiv-iff P Q
```
diff --git a/src/foundation/mere-functions.lagda.md b/src/foundation/mere-functions.lagda.md
new file mode 100644
index 0000000000..4597357ccd
--- /dev/null
+++ b/src/foundation/mere-functions.lagda.md
@@ -0,0 +1,101 @@
+# Mere functions
+
+```agda
+module foundation.mere-functions where
+```
+
+Imports
+
+```agda
+open import foundation.propositional-truncations
+open import foundation.universe-levels
+
+open import foundation-core.function-types
+open import foundation-core.propositions
+```
+
+
+
+## Idea
+
+The type of
+{{#concept "mere functions" Disambiguation="of types" Agda=mere-function}} from
+`A` to `B` is the
+[propositional truncation](foundation.propositional-truncations.md) of the type
+of maps from `A` to `B`.
+
+```text
+ mere-function A B := ║(A → B)║₋₁
+```
+
+## Definitions
+
+### Mere functions between types
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ prop-mere-function : Prop (l1 ⊔ l2)
+ prop-mere-function = trunc-Prop (A → B)
+
+ mere-function : UU (l1 ⊔ l2)
+ mere-function = type-Prop prop-mere-function
+
+ is-prop-mere-function : is-prop mere-function
+ is-prop-mere-function = is-prop-type-Prop prop-mere-function
+```
+
+### The evaluation map on mere functions
+
+If we have a mere function from `A` to `B` and `A` is inhabited, then `B` is
+inhabited.
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ ev-mere-function' : (mere-function A B) → A → ║ B ║₋₁
+ ev-mere-function' |f| a =
+ rec-trunc-Prop (trunc-Prop B) (λ f → unit-trunc-Prop (f a)) |f|
+
+ ev-mere-function : (mere-function A B) → ║ A ║₋₁ → ║ B ║₋₁
+ ev-mere-function |f| |a| =
+ rec-trunc-Prop (trunc-Prop B) (ev-mere-function' |f|) (|a|)
+```
+
+### Mere functions form a reflexive relation
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ refl-mere-function : mere-function A A
+ refl-mere-function = unit-trunc-Prop id
+```
+
+### Mere functions form a transitive relation
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ where
+
+ transitive-mere-function :
+ mere-function B C → mere-function A B → mere-function A C
+ transitive-mere-function |g| =
+ rec-trunc-Prop
+ ( prop-mere-function A C)
+ ( λ f →
+ rec-trunc-Prop
+ ( prop-mere-function A C)
+ ( λ g → unit-trunc-Prop (g ∘ f))
+ ( |g|))
+```
+
+## See also
+
+- [Mere logical equivalences](foundation.mere-logical-equivalences.md)
diff --git a/src/foundation/mere-logical-equivalences.lagda.md b/src/foundation/mere-logical-equivalences.lagda.md
new file mode 100644
index 0000000000..5bf9ebe781
--- /dev/null
+++ b/src/foundation/mere-logical-equivalences.lagda.md
@@ -0,0 +1,181 @@
+# Mere logical equivalences
+
+```agda
+module foundation.mere-logical-equivalences where
+```
+
+Imports
+
+```agda
+open import foundation.conjunction
+open import foundation.dependent-pair-types
+open import foundation.inhabited-types
+open import foundation.logical-equivalences
+open import foundation.mere-functions
+open import foundation.propositional-truncations
+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
+```
+
+
+
+## Idea
+
+Two types `A` and `B` are
+{{#concept "merely logically equivalent" Disambiguation="types" Agda=mere-iff}}
+if the type of [logical equivalences](foundation.logical-equivalences.md)
+between `A` and `B` is [inhabited](foundation.inhabited-types.md).
+
+```text
+ mere-iff A B := ║(A ↔ B)║₋₁
+```
+
+## Definitions
+
+### Mere logical equivalence of types
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ prop-mere-iff : Prop (l1 ⊔ l2)
+ prop-mere-iff = trunc-Prop (A ↔ B)
+
+ mere-iff : UU (l1 ⊔ l2)
+ mere-iff = type-Prop prop-mere-iff
+
+ is-prop-mere-iff : is-prop mere-iff
+ is-prop-mere-iff = is-prop-type-Prop prop-mere-iff
+```
+
+## Properties
+
+### Mere logical equivalence is a reflexive relation
+
+```agda
+module _
+ {l : Level} (A : UU l)
+ where
+
+ refl-mere-iff : mere-iff A A
+ refl-mere-iff = unit-trunc-Prop id-iff
+```
+
+### Mere logical equivalence is a transitive relation
+
+```agda
+module _
+ {l1 l2 l3 : Level} {A : UU l1} {B : UU l2} {C : UU l3}
+ where
+
+ transitive-mere-iff : mere-iff B C → mere-iff A B → mere-iff A C
+ transitive-mere-iff |g| =
+ rec-trunc-Prop
+ ( prop-mere-iff A C)
+ ( λ f →
+ rec-trunc-Prop
+ ( prop-mere-iff A C)
+ ( λ g → unit-trunc-Prop (g ∘iff f))
+ ( |g|))
+```
+
+### Mere logical equivalence is a symmetric relation
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ symmetric-mere-iff : mere-iff A B → mere-iff B A
+ symmetric-mere-iff =
+ rec-trunc-Prop (prop-mere-iff B A) (unit-trunc-Prop ∘ inv-iff)
+```
+
+### Merely logically equivalent types are coinhabited
+
+If `A` and `B` are merely logically equivalent then they are
+[coinhabited](foundation.coinhabited-pairs-of-types.md).
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : UU l2)
+ where
+
+ ev-forward-mere-iff' : mere-iff A B → A → is-inhabited B
+ ev-forward-mere-iff' |f| a =
+ rec-trunc-Prop
+ ( trunc-Prop B)
+ ( λ f → unit-trunc-Prop (forward-implication f a))
+ ( |f|)
+
+ ev-forward-mere-iff : mere-iff A B → is-inhabited A → is-inhabited B
+ ev-forward-mere-iff |f| =
+ rec-trunc-Prop (trunc-Prop B) (ev-forward-mere-iff' |f|)
+
+ ev-backward-mere-iff' : mere-iff A B → B → is-inhabited A
+ ev-backward-mere-iff' |f| b =
+ rec-trunc-Prop
+ ( trunc-Prop A)
+ ( λ f → unit-trunc-Prop (backward-implication f b))
+ ( |f|)
+
+ ev-backward-mere-iff : mere-iff A B → is-inhabited B → is-inhabited A
+ ev-backward-mere-iff |f| =
+ rec-trunc-Prop (trunc-Prop A) (ev-backward-mere-iff' |f|)
+
+ is-coinhabited-mere-iff : mere-iff A B → is-inhabited A ↔ is-inhabited B
+ is-coinhabited-mere-iff |f| =
+ ( ev-forward-mere-iff |f| , ev-backward-mere-iff |f|)
+```
+
+### Merely logically equivalent types have bidirectional mere functions
+
+If `A` and `B` are merely logically equivalent then we have a mere function from
+`A` to `B` and a mere function from `B` to `A`.
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ forward-mere-function-mere-iff : mere-iff A B → mere-function A B
+ forward-mere-function-mere-iff =
+ rec-trunc-Prop
+ ( prop-mere-function A B)
+ ( unit-trunc-Prop ∘ forward-implication)
+
+ backward-mere-function-mere-iff : mere-iff A B → mere-function B A
+ backward-mere-function-mere-iff =
+ rec-trunc-Prop
+ ( prop-mere-function B A)
+ ( unit-trunc-Prop ∘ backward-implication)
+```
+
+### Mere logical equivalence is equivalent to having bidirectional mere functions
+
+For all types we have the equivalence
+
+```text
+ (mere-iff A B) ≃ (mere-function A B × mere-function B A).
+```
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : UU l2}
+ where
+
+ compute-mere-iff :
+ mere-iff A B ≃ mere-function A B × mere-function B A
+ compute-mere-iff = equiv-product-trunc-Prop (A → B) (B → A)
+```
+
+## See also
+
+- [Mere functions](foundation.mere-functions.md)
+- [Coinhabitedness](foundation.coinhabited-pairs-of-types.md) is a related but
+ weaker notion than mere logical equivalence.
diff --git a/src/foundation/negation.lagda.md b/src/foundation/negation.lagda.md
index aa4a880ec6..adbbb19004 100644
--- a/src/foundation/negation.lagda.md
+++ b/src/foundation/negation.lagda.md
@@ -34,12 +34,19 @@ type `A` is the type `A → empty`.
is-prop-neg : {l : Level} {A : UU l} → is-prop (¬ A)
is-prop-neg {A = A} = is-prop-function-type is-prop-empty
-neg-Prop' : {l1 : Level} → UU l1 → Prop l1
-pr1 (neg-Prop' A) = ¬ A
-pr2 (neg-Prop' A) = is-prop-neg
+neg-type-Prop : {l1 : Level} → UU l1 → Prop l1
+neg-type-Prop A = ¬ A , is-prop-neg
neg-Prop : {l1 : Level} → Prop l1 → Prop l1
-neg-Prop P = neg-Prop' (type-Prop P)
+neg-Prop P = neg-type-Prop (type-Prop P)
+
+type-neg-Prop : {l1 : Level} → Prop l1 → UU l1
+type-neg-Prop P = type-Prop (neg-Prop P)
+
+infix 25 ¬'_
+
+¬'_ : {l1 : Level} → Prop l1 → Prop l1
+¬'_ = neg-Prop
```
### Reductio ad absurdum
@@ -57,8 +64,8 @@ equiv-neg :
(X ≃ Y) → (¬ X ≃ ¬ Y)
equiv-neg {l1} {l2} {X} {Y} e =
equiv-iff'
- ( neg-Prop' X)
- ( neg-Prop' Y)
+ ( neg-type-Prop X)
+ ( neg-type-Prop Y)
( pair (map-neg (map-inv-equiv e)) (map-neg (map-equiv e)))
```
@@ -66,14 +73,24 @@ equiv-neg {l1} {l2} {X} {Y} e =
```agda
no-fixed-points-neg :
- {l : Level} (A : UU l) → ¬ (A ↔ (¬ A))
-no-fixed-points-neg A (pair f g) =
+ {l : Level} (A : UU l) → ¬ (A ↔ ¬ A)
+no-fixed-points-neg A (f , g) =
( λ (h : ¬ A) → h (g h)) (λ (a : A) → f a a)
-```
-```agda
abstract
no-fixed-points-neg-Prop :
- {l1 : Level} (P : Prop l1) → ¬ (P ⇔ neg-Prop P)
+ {l : Level} (P : Prop l) → ¬ (type-Prop P ↔ ¬ (type-Prop P))
no-fixed-points-neg-Prop P = no-fixed-points-neg (type-Prop P)
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [negation](https://ncatlab.org/nlab/show/negation) at $n$Lab
+- [Negation](https://en.wikipedia.org/wiki/Negation) at Wikipedia
diff --git a/src/foundation/partitions.lagda.md b/src/foundation/partitions.lagda.md
index 414d0edeca..11658b834b 100644
--- a/src/foundation/partitions.lagda.md
+++ b/src/foundation/partitions.lagda.md
@@ -8,6 +8,7 @@ module foundation.partitions where
```agda
open import foundation.action-on-identifications-functions
+open import foundation.conjunction
open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.embeddings
@@ -512,9 +513,7 @@ module _
share-common-element-block-partition-Prop :
(C : block-partition P) → Prop (l1 ⊔ l2)
share-common-element-block-partition-Prop C =
- ∃-Prop A
- ( λ a →
- is-in-block-partition P B a × is-in-block-partition P C a)
+ ∃ A (λ a → subtype-block-partition P B a ∧ subtype-block-partition P C a)
share-common-element-block-partition :
(C : block-partition P) → UU (l1 ⊔ l2)
diff --git a/src/foundation/path-split-maps.lagda.md b/src/foundation/path-split-maps.lagda.md
index 857a1071b5..0255429411 100644
--- a/src/foundation/path-split-maps.lagda.md
+++ b/src/foundation/path-split-maps.lagda.md
@@ -12,6 +12,7 @@ open import foundation-core.path-split-maps public
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.iterated-dependent-product-types
+open import foundation.logical-equivalences
open import foundation.universe-levels
open import foundation-core.contractible-types
@@ -47,7 +48,7 @@ module _
is-equiv-is-path-split-is-equiv :
(f : A → B) → is-equiv (is-path-split-is-equiv f)
is-equiv-is-path-split-is-equiv f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-property-is-equiv f)
( is-prop-is-path-split f)
( is-equiv-is-path-split f)
@@ -60,7 +61,7 @@ module _
is-equiv-is-equiv-is-path-split :
(f : A → B) → is-equiv (is-equiv-is-path-split f)
is-equiv-is-equiv-is-path-split f =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-path-split f)
( is-property-is-equiv f)
( is-path-split-is-equiv f)
diff --git a/src/foundation/preunivalent-type-families.lagda.md b/src/foundation/preunivalent-type-families.lagda.md
index 6a4a7e9f2a..0a71d755dd 100644
--- a/src/foundation/preunivalent-type-families.lagda.md
+++ b/src/foundation/preunivalent-type-families.lagda.md
@@ -14,16 +14,16 @@ open import foundation.embeddings
open import foundation.equivalence-injective-type-families
open import foundation.faithful-maps
open import foundation.function-types
+open import foundation.injective-maps
open import foundation.preunivalence
open import foundation.retractions
+open import foundation.sets
open import foundation.subuniverses
open import foundation.transport-along-identifications
open import foundation.universe-levels
open import foundation-core.equivalences
open import foundation-core.identity-types
-open import foundation-core.injective-maps
-open import foundation-core.sets
open import foundation-core.univalence
```
diff --git a/src/foundation/propositional-extensionality.lagda.md b/src/foundation/propositional-extensionality.lagda.md
index 4a65faa25d..a92f6d6176 100644
--- a/src/foundation/propositional-extensionality.lagda.md
+++ b/src/foundation/propositional-extensionality.lagda.md
@@ -37,8 +37,15 @@ open import foundation-core.torsorial-type-families
## Idea
-Propositional extensionality characterizes identifications of propositions. It
-asserts that for any two propositions `P` and `Q`, we have `Id P Q ≃ (P ⇔ Q)`.
+{{#concept "Propositional extensionality" Agda=propositional-extensionality}}
+characterizes [identifications](foundation-core.identity-types.md) of
+[propositions](foundation-core.propositions.md). It asserts that for any two
+propositions `P` and `Q`, we have `(P = Q) ≃ (P ⇔ Q)`.
+
+**Note.** While we derive propositional extensionality from the
+[univalence axiom](foundation-core.univalence.md), it is a strictly weaker
+principle, meaning that the principle of propositional extensionality does not
+imply univalence.
## Properties
@@ -51,7 +58,7 @@ module _
abstract
is-torsorial-iff :
- (P : Prop l1) → is-torsorial (λ (Q : Prop l1) → P ⇔ Q)
+ (P : Prop l1) → is-torsorial (λ (Q : Prop l1) → type-Prop P ↔ type-Prop Q)
is-torsorial-iff P =
is-contr-equiv
( Σ (Prop l1) (λ Q → type-Prop P ≃ type-Prop Q))
@@ -66,16 +73,14 @@ module _
abstract
is-equiv-iff-eq : (P Q : Prop l1) → is-equiv (iff-eq {l1} {P} {Q})
is-equiv-iff-eq P =
- fundamental-theorem-id
- ( is-torsorial-iff P)
- ( λ Q → iff-eq {P = P} {Q})
+ fundamental-theorem-id (is-torsorial-iff P) (λ Q → iff-eq {P = P} {Q})
propositional-extensionality :
- (P Q : Prop l1) → (P = Q) ≃ (P ⇔ Q)
+ (P Q : Prop l1) → (P = Q) ≃ (type-Prop P ↔ type-Prop Q)
pr1 (propositional-extensionality P Q) = iff-eq
pr2 (propositional-extensionality P Q) = is-equiv-iff-eq P Q
- eq-iff' : (P Q : Prop l1) → P ⇔ Q → P = Q
+ eq-iff' : (P Q : Prop l1) → type-Prop P ↔ type-Prop Q → P = Q
eq-iff' P Q = map-inv-is-equiv (is-equiv-iff-eq P Q)
eq-iff :
@@ -84,7 +89,7 @@ module _
eq-iff {P} {Q} f g = eq-iff' P Q (pair f g)
eq-equiv-Prop :
- {P Q : Prop l1} → (type-Prop P ≃ type-Prop Q) → P = Q
+ {P Q : Prop l1} → type-Prop P ≃ type-Prop Q → P = Q
eq-equiv-Prop e =
eq-iff (map-equiv e) (map-inv-equiv e)
@@ -93,11 +98,10 @@ module _
equiv-eq-Prop {P} refl = id-equiv
is-torsorial-equiv-Prop :
- (P : Prop l1) →
- is-torsorial (λ Q → type-Prop P ≃ type-Prop Q)
+ (P : Prop l1) → is-torsorial (λ Q → type-Prop P ≃ type-Prop Q)
is-torsorial-equiv-Prop P =
is-contr-equiv'
- ( Σ (Prop l1) (λ Q → P ⇔ Q))
+ ( Σ (Prop l1) (λ Q → type-Prop P ↔ type-Prop Q))
( equiv-tot (equiv-equiv-iff P))
( is-torsorial-iff P)
```
@@ -106,7 +110,7 @@ module _
```agda
is-set-type-Prop : {l : Level} → is-set (Prop l)
-is-set-type-Prop {l} P Q =
+is-set-type-Prop P Q =
is-prop-equiv
( propositional-extensionality P Q)
( is-prop-iff-Prop P Q)
@@ -120,7 +124,7 @@ pr2 (Prop-Set l) = is-set-type-Prop
```agda
is-univalent-type-Prop : {l : Level} → is-univalent (type-Prop {l})
-is-univalent-type-Prop {l} P =
+is-univalent-type-Prop P =
fundamental-theorem-id
( is-torsorial-equiv-Prop P)
( λ Q → equiv-tr type-Prop)
@@ -134,7 +138,7 @@ abstract
{l1 : Level} → is-torsorial (λ (P : Prop l1) → type-Prop P)
is-torsorial-true-Prop {l1} =
is-contr-equiv
- ( Σ (Prop l1) (λ P → raise-unit-Prop l1 ⇔ P))
+ ( Σ (Prop l1) (λ P → raise-unit l1 ↔ type-Prop P))
( equiv-tot
( λ P →
inv-equiv
@@ -144,7 +148,7 @@ abstract
( type-Prop P)) ∘e
( right-unit-law-product-is-contr
( is-contr-Π
- ( λ x →
+ ( λ _ →
is-proof-irrelevant-is-prop
( is-prop-raise-unit)
( raise-star)))))))
@@ -156,10 +160,10 @@ abstract
```agda
abstract
is-torsorial-false-Prop :
- {l1 : Level} → is-torsorial (λ (P : Prop l1) → type-Prop (neg-Prop P))
+ {l1 : Level} → is-torsorial (λ (P : Prop l1) → ¬ (type-Prop P))
is-torsorial-false-Prop {l1} =
is-contr-equiv
- ( Σ (Prop l1) (λ P → raise-empty-Prop l1 ⇔ P))
+ ( Σ (Prop l1) (λ P → raise-empty l1 ↔ type-Prop P))
( equiv-tot
( λ P →
inv-equiv
@@ -172,3 +176,15 @@ abstract
( type-Prop P))))))
( is-torsorial-iff (raise-empty-Prop l1))
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [propositional extensionality](https://ncatlab.org/nlab/show/propositional+extensionality)
+ at $n$Lab.
diff --git a/src/foundation/propositional-resizing.lagda.md b/src/foundation/propositional-resizing.lagda.md
index b3be0891b7..46d48494aa 100644
--- a/src/foundation/propositional-resizing.lagda.md
+++ b/src/foundation/propositional-resizing.lagda.md
@@ -18,10 +18,13 @@ open import foundation-core.subtypes
## Idea
-We say that there is propositional resizing for propositions of universe levels
-`l1` and `l2` if there is a type `Ω : UU l1` equipped with a subtype `Q` such
-that for each proposition `P` of universe level `l2` there is an element `u : Ω`
-such that `Q u ≃ P`.
+We say that a universe `𝒱` satisfies `𝒰`-small
+{{#concept "propositional resizing"}} 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
@@ -31,3 +34,14 @@ propositional-resizing l1 l2 =
Σ ( Σ (UU l1) (subtype l1))
( λ Ω → (P : Prop l2) → Σ (pr1 Ω) (λ u → type-equiv-Prop (pr2 Ω u) P))
```
+
+## See also
+
+- [The large locale of propositions](foundation.large-locale-of-propositions.md)
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/propositional-truncations.lagda.md b/src/foundation/propositional-truncations.lagda.md
index 8f12d80321..1181cc2572 100644
--- a/src/foundation/propositional-truncations.lagda.md
+++ b/src/foundation/propositional-truncations.lagda.md
@@ -10,6 +10,7 @@ module foundation.propositional-truncations where
open import foundation.action-on-identifications-functions
open import foundation.dependent-pair-types
open import foundation.functoriality-cartesian-product-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.truncations
open import foundation.universal-property-propositional-truncation
@@ -20,7 +21,6 @@ open import foundation-core.contractible-types
open import foundation-core.coproduct-types
open import foundation-core.equivalences
open import foundation-core.function-types
-open import foundation-core.homotopies
open import foundation-core.identity-types
open import foundation-core.precomposition-dependent-functions
open import foundation-core.precomposition-functions
@@ -59,8 +59,16 @@ all-elements-equal-type-trunc-Prop {l} {A} =
trunc-Prop : {l : Level} → UU l → Prop l
trunc-Prop = trunc neg-one-𝕋
+
+║_║₋₁ : {l : Level} → UU l → UU l
+║_║₋₁ = type-trunc-Prop
```
+**Notation.** The [box drawings double vertical](https://codepoints.net/U+2551)
+symbol `║` in the propositional truncation notation `║_║₋₁` can be inserted with
+`agda-input` using the escape sequence `\--=` and selecting the second item in
+the list.
+
## Properties
### The condition in the induction principle is a property
@@ -245,7 +253,7 @@ module _
abstract
is-equiv-map-idempotent-trunc-Prop : is-equiv map-idempotent-trunc-Prop
is-equiv-map-idempotent-trunc-Prop =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-trunc-Prop)
( is-prop-type-trunc-Prop)
( unit-trunc-Prop)
@@ -259,7 +267,7 @@ module _
is-equiv-map-inv-idempotent-trunc-Prop :
is-equiv (unit-trunc-Prop {A = type-trunc-Prop A})
is-equiv-map-inv-idempotent-trunc-Prop =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-trunc-Prop)
( is-prop-type-trunc-Prop)
( map-idempotent-trunc-Prop)
@@ -356,8 +364,8 @@ abstract
is-equiv-map-distributive-trunc-product-Prop :
{l1 l2 : Level} {A : UU l1} {B : UU l2} →
is-equiv (map-distributive-trunc-product-Prop {A = A} {B = B})
- is-equiv-map-distributive-trunc-product-Prop {l1} {l2} {A} {B} =
- is-equiv-is-prop
+ is-equiv-map-distributive-trunc-product-Prop =
+ is-equiv-has-converse-is-prop
( is-prop-type-trunc-Prop)
( is-prop-product is-prop-type-trunc-Prop is-prop-type-trunc-Prop)
( map-inv-distributive-trunc-product-Prop)
@@ -373,8 +381,8 @@ abstract
is-equiv-map-inv-distributive-trunc-product-Prop :
{l1 l2 : Level} {A : UU l1} {B : UU l2} →
is-equiv (map-inv-distributive-trunc-product-Prop {A = A} {B = B})
- is-equiv-map-inv-distributive-trunc-product-Prop {l1} {l2} {A} {B} =
- is-equiv-is-prop
+ is-equiv-map-inv-distributive-trunc-product-Prop =
+ is-equiv-has-converse-is-prop
( is-prop-product is-prop-type-trunc-Prop is-prop-type-trunc-Prop)
( is-prop-type-trunc-Prop)
( map-distributive-trunc-product-Prop)
@@ -411,7 +419,7 @@ module _
is-equiv-map-trunc-Prop-diagonal-coproduct :
is-equiv map-trunc-Prop-diagonal-coproduct
is-equiv-map-trunc-Prop-diagonal-coproduct =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
is-prop-type-trunc-Prop
is-prop-type-trunc-Prop
map-inv-trunc-Prop-diagonal-coproduct
@@ -419,21 +427,32 @@ module _
is-equiv-map-inv-trunc-Prop-diagonal-coproduct :
is-equiv map-inv-trunc-Prop-diagonal-coproduct
is-equiv-map-inv-trunc-Prop-diagonal-coproduct =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
is-prop-type-trunc-Prop
is-prop-type-trunc-Prop
map-trunc-Prop-diagonal-coproduct
equiv-trunc-Prop-diagonal-coproduct :
- (type-trunc-Prop (A + A)) ≃ type-trunc-Prop A
+ type-trunc-Prop (A + A) ≃ type-trunc-Prop A
pr1 equiv-trunc-Prop-diagonal-coproduct = map-trunc-Prop-diagonal-coproduct
pr2 equiv-trunc-Prop-diagonal-coproduct =
is-equiv-map-trunc-Prop-diagonal-coproduct
inv-equiv-trunc-Prop-diagonal-coproduct :
- (type-trunc-Prop A) ≃ type-trunc-Prop (A + A)
+ type-trunc-Prop A ≃ type-trunc-Prop (A + A)
pr1 inv-equiv-trunc-Prop-diagonal-coproduct =
map-inv-trunc-Prop-diagonal-coproduct
pr2 inv-equiv-trunc-Prop-diagonal-coproduct =
is-equiv-map-inv-trunc-Prop-diagonal-coproduct
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [bracket type](https://ncatlab.org/nlab/show/bracket+type) at $n$Lab
diff --git a/src/foundation/propositions.lagda.md b/src/foundation/propositions.lagda.md
index e0f0a8094c..5a9092e007 100644
--- a/src/foundation/propositions.lagda.md
+++ b/src/foundation/propositions.lagda.md
@@ -11,6 +11,7 @@ open import foundation-core.propositions public
```agda
open import foundation.contractible-types
open import foundation.dependent-pair-types
+open import foundation.logical-equivalences
open import foundation.retracts-of-types
open import foundation.universe-levels
@@ -69,9 +70,67 @@ abstract
equiv-is-prop-equiv : {l1 l2 : Level} {A : UU l1} {B : UU l2} →
A ≃ B → is-prop A ≃ is-prop B
equiv-is-prop-equiv {A = A} {B = B} e =
- equiv-prop
- (is-prop-is-prop A)
- (is-prop-is-prop B)
- (is-prop-equiv' e)
- (is-prop-equiv e)
+ equiv-iff-is-prop
+ ( is-prop-is-prop A)
+ ( is-prop-is-prop B)
+ ( is-prop-equiv' e)
+ ( is-prop-equiv e)
```
+
+## See also
+
+### Operations on propositions
+
+There is a wide range of operations on propositions due to the rich structure of
+intuitionistic logic. Below we give a structured overview of a notable selection
+of such operations and their notation in the library.
+
+The list is split into two sections, the first consists of operations that
+generalize to arbitrary types and even sufficiently nice
+[subuniverses](foundation.subuniverses.md), such as
+$n$-[types](foundation-core.truncated-types.md).
+
+| Name | Operator on types | Operator on propositions/subtypes |
+| ----------------------------------------------------------- | ----------------- | --------------------------------- |
+| [Dependent sum](foundation.dependent-pair-types.md) | `Σ` | `Σ-Prop` |
+| [Dependent product](foundation.dependent-function-types.md) | `Π` | `Π-Prop` |
+| [Functions](foundation-core.function-types.md) | `→` | `⇒` |
+| [Logical equivalence](foundation.logical-equivalences.md) | `↔` | `⇔` |
+| [Product](foundation-core.cartesian-product-types.md) | `×` | `product-Prop` |
+| [Join](synthetic-homotopy-theory.joins-of-types.md) | `*` | `join-Prop` |
+| [Exclusive sum](foundation.exclusive-sum.md) | `exclusive-sum` | `exclusive-sum-Prop` |
+| [Coproduct](foundation-core.coproduct-types.md) | `+` | _N/A_ |
+
+Note that for many operations in the second section, there is an equivalent
+operation on propositions in the first.
+
+| Name | Operator on types | Operator on propositions/subtypes |
+| ---------------------------------------------------------------------------- | --------------------------- | ---------------------------------------- |
+| [Initial object](foundation-core.empty-types.md) | `empty` | `empty-Prop` |
+| [Terminal object](foundation.unit-type.md) | `unit` | `unit-Prop` |
+| [Existential quantification](foundation.existential-quantification.md) | `exists-structure` | `∃` |
+| [Unique existential quantification](foundation.uniqueness-quantification.md) | `uniquely-exists-structure` | `∃!` |
+| [Universal quantification](foundation.universal-quantification.md) | | `∀'` (equivalent to `Π-Prop`) |
+| [Conjunction](foundation.conjunction.md) | | `∧` (equivalent to `product-Prop`) |
+| [Disjunction](foundation.disjunction.md) | `disjunction-type` | `∨` (equivalent to `join-Prop`) |
+| [Exclusive disjunction](foundation.exclusive-disjunction.md) | `xor-type` | `⊻` (equivalent to `exclusive-sum-Prop`) |
+| [Negation](foundation.negation.md) | `¬` | `¬'` |
+| [Double negation](foundation.double-negation.md) | `¬¬` | `¬¬'` |
+
+We can also organize these operations by indexed and binary variants, giving us
+the following table:
+
+| Name | Binary | Indexed |
+| ---------------------- | ------ | ------- |
+| Product | `×` | `Π` |
+| Conjunction | `∧` | `∀'` |
+| Constructive existence | `+` | `Σ` |
+| Existence | `∨` | `∃` |
+| Unique existence | `⊻` | `∃!` |
+
+### Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/retractions.lagda.md b/src/foundation/retractions.lagda.md
index 43253b5f8c..fe2da2d349 100644
--- a/src/foundation/retractions.lagda.md
+++ b/src/foundation/retractions.lagda.md
@@ -9,7 +9,6 @@ open import foundation-core.retractions public
Imports
```agda
-open import foundation.action-on-identifications-functions
open import foundation.coslice
open import foundation.dependent-pair-types
open import foundation.retracts-of-types
@@ -20,7 +19,6 @@ 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.injective-maps
```
diff --git a/src/foundation/set-presented-types.lagda.md b/src/foundation/set-presented-types.lagda.md
index ee7d3b9858..e6f75f27db 100644
--- a/src/foundation/set-presented-types.lagda.md
+++ b/src/foundation/set-presented-types.lagda.md
@@ -7,11 +7,11 @@ module foundation.set-presented-types where
Imports
```agda
+open import foundation.equivalences
open import foundation.existential-quantification
open import foundation.set-truncations
open import foundation.universe-levels
-open import foundation-core.equivalences
open import foundation-core.function-types
open import foundation-core.propositions
open import foundation-core.sets
@@ -21,12 +21,18 @@ open import foundation-core.sets
## Idea
-A type `A` is said to be set presented if there exists a map `f : X → A` from a
-set `X` such that the composite `X → A → type-trunc-Set A` is an equivalence.
+A type `A` is said to be
+{{#concept "set presented" Agda=has-set-presentation-Prop}} if there
+[exists](foundation.existential-quantification.md) a map `f : X → A` from a
+[set](foundation-core.sets.md) `X` such that the composite
+`X → A → type-trunc-Set A` is an [equivalence](foundation.equivalences.md).
```agda
-has-set-presentation-Prop :
- {l1 l2 : Level} (A : Set l1) (B : UU l2) → Prop (l1 ⊔ l2)
-has-set-presentation-Prop A B =
- ∃-Prop (type-Set A → B) (λ f → is-equiv (unit-trunc-Set ∘ f))
+module _
+ {l1 l2 : Level} (A : Set l1) (B : UU l2)
+ where
+
+ has-set-presentation-Prop : Prop (l1 ⊔ l2)
+ has-set-presentation-Prop =
+ ∃ (type-Set A → B) (λ f → is-equiv-Prop (unit-trunc-Set ∘ f))
```
diff --git a/src/foundation/set-truncations.lagda.md b/src/foundation/set-truncations.lagda.md
index df3386684a..154cc2d2b5 100644
--- a/src/foundation/set-truncations.lagda.md
+++ b/src/foundation/set-truncations.lagda.md
@@ -49,30 +49,38 @@ open import foundation-core.truncation-levels
## Idea
-The **set truncation** of a type `A` is a map `η : A → trunc-Set A` that
-satisfies
+The {{#concept "set truncation" Agda=trunc-Set}} of a type `A` is a map
+`η : A → trunc-Set A` that satisfies
[the universal property of set truncations](foundation.universal-property-set-truncation.md).
-## Definition
+## Definitions
```agda
+trunc-Set : {l : Level} → UU l → Set l
+trunc-Set = trunc zero-𝕋
+
type-trunc-Set : {l : Level} → UU l → UU l
type-trunc-Set = type-trunc zero-𝕋
is-set-type-trunc-Set : {l : Level} {A : UU l} → is-set (type-trunc-Set A)
is-set-type-trunc-Set = is-trunc-type-trunc
-trunc-Set : {l : Level} → UU l → Set l
-trunc-Set = trunc zero-𝕋
-
unit-trunc-Set : {l : Level} {A : UU l} → A → type-trunc-Set A
unit-trunc-Set = unit-trunc
is-set-truncation-trunc-Set :
{l1 : Level} (A : UU l1) → is-set-truncation (trunc-Set A) unit-trunc-Set
is-set-truncation-trunc-Set A = is-truncation-trunc
+
+║_║₀ : {l : Level} → UU l → UU l
+║_║₀ = type-trunc-Set
```
+**Notation.** The [box drawings double vertical](https://codepoints.net/U+2551)
+symbol `║` in the set truncation notation `║_║₀` can be inserted with
+`agda-input` using the escape sequence `\--=` and selecting the second item in
+the list.
+
## Properties
### The dependent universal property of set truncations
diff --git a/src/foundation/sets.lagda.md b/src/foundation/sets.lagda.md
index 15ad33b747..32edfb667b 100644
--- a/src/foundation/sets.lagda.md
+++ b/src/foundation/sets.lagda.md
@@ -11,6 +11,7 @@ open import foundation-core.sets public
```agda
open import foundation.contractible-types
open import foundation.dependent-pair-types
+open import foundation.logical-equivalences
open import foundation.subuniverses
open import foundation.truncated-types
open import foundation.univalent-type-families
@@ -21,6 +22,7 @@ open import foundation-core.cartesian-product-types
open import foundation-core.embeddings
open import foundation-core.equivalences
open import foundation-core.identity-types
+open import foundation-core.injective-maps
open import foundation-core.precomposition-functions
open import foundation-core.propositional-maps
open import foundation-core.propositions
diff --git a/src/foundation/slice.lagda.md b/src/foundation/slice.lagda.md
index 146aa27710..0a41b6dd82 100644
--- a/src/foundation/slice.lagda.md
+++ b/src/foundation/slice.lagda.md
@@ -14,6 +14,7 @@ open import foundation.function-extensionality
open import foundation.fundamental-theorem-of-identity-types
open import foundation.homotopies
open import foundation.homotopy-induction
+open import foundation.logical-equivalences
open import foundation.structure-identity-principle
open import foundation.type-arithmetic-dependent-pair-types
open import foundation.univalence
@@ -246,11 +247,12 @@ module _
(h : hom-slice f g) →
is-equiv (pr1 h) ≃
is-fiberwise-equiv (map-equiv (equiv-fiberwise-hom-hom-slice f g) h)
- α h = equiv-prop
- ( is-property-is-equiv _)
- ( is-prop-Π (λ _ → is-property-is-equiv _))
- ( is-fiberwise-equiv-fiberwise-equiv-equiv-slice h)
- ( is-equiv-hom-slice-is-fiberwise-equiv-fiberwise-hom-hom-slice h)
+ α h =
+ equiv-iff-is-prop
+ ( is-property-is-equiv _)
+ ( is-prop-Π (λ _ → is-property-is-equiv _))
+ ( is-fiberwise-equiv-fiberwise-equiv-equiv-slice h)
+ ( is-equiv-hom-slice-is-fiberwise-equiv-fiberwise-hom-hom-slice h)
equiv-equiv-slice-fiberwise-equiv :
fiberwise-equiv (fiber f) (fiber g) ≃ equiv-slice f g
diff --git a/src/foundation/subsingleton-induction.lagda.md b/src/foundation/subsingleton-induction.lagda.md
index 1e6b1f0c7d..7007417be8 100644
--- a/src/foundation/subsingleton-induction.lagda.md
+++ b/src/foundation/subsingleton-induction.lagda.md
@@ -21,11 +21,12 @@ open import foundation-core.sections
## Idea
-**Subsingleton induction** on a type `A` is a slight variant of
-[singleton induction](foundation.singleton-induction.md) which in turn is a
-principle analogous to induction for the [unit type](foundation.unit-type.md).
-Subsingleton induction uses the observation that a type equipped with an element
-is [contractible](foundation-core.contractible-types.md) if and only if it is a
+{{#concept "Subsingleton induction" Agda=ind-subsingleton}} on a type `A` is a
+slight variant of [singleton induction](foundation.singleton-induction.md) which
+in turn is a principle analogous to induction for the
+[unit type](foundation.unit-type.md). Subsingleton induction uses the
+observation that a type equipped with an element is
+[contractible](foundation-core.contractible-types.md) if and only if it is a
[proposition](foundation-core.propositions.md).
Subsingleton induction states that given a type family `B` over `A`, to
@@ -95,3 +96,10 @@ abstract
{l1 : Level} (A : UU l1) → ({l2 : Level} → is-subsingleton l2 A) → is-prop A
is-prop-is-subsingleton A S = is-prop-ind-subsingleton A (pr1 ∘ S)
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/subterminal-types.lagda.md b/src/foundation/subterminal-types.lagda.md
index 317b7132b6..85ab7093da 100644
--- a/src/foundation/subterminal-types.lagda.md
+++ b/src/foundation/subterminal-types.lagda.md
@@ -23,8 +23,10 @@ open import foundation-core.propositions
## Idea
-A type is said to be subterminal if it embeds into the unit type. A type is
-subterminal if and only if it is a proposition.
+A type is said to be {{#concept "subterminal" Agda=is-subterminal}} if it
+[embeds](foundation-core.embeddings.md) into the
+[unit type](foundation.unit-type.md). A type is subterminal if and only if it is
+a [proposition](foundation-core.propositions.md).
## Definition
@@ -82,3 +84,10 @@ module _
is-proof-irrelevant-is-subterminal H =
is-proof-irrelevant-all-elements-equal (eq-is-subterminal H)
```
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
diff --git a/src/foundation/subtypes.lagda.md b/src/foundation/subtypes.lagda.md
index 33ae55ad1a..b181a9a10d 100644
--- a/src/foundation/subtypes.lagda.md
+++ b/src/foundation/subtypes.lagda.md
@@ -147,7 +147,7 @@ module _
{l2 l3 : Level} (P : subtype l2 A) (Q : subtype l3 A) → P ⊆ Q → Q ⊆ P →
(x : A) → is-in-subtype P x ≃ is-in-subtype Q x
equiv-antisymmetric-leq-subtype P Q H K x =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-in-subtype P x)
( is-prop-is-in-subtype Q x)
( H x)
diff --git a/src/foundation/symmetric-difference.lagda.md b/src/foundation/symmetric-difference.lagda.md
index 806c2bbce1..13308cf30a 100644
--- a/src/foundation/symmetric-difference.lagda.md
+++ b/src/foundation/symmetric-difference.lagda.md
@@ -10,8 +10,7 @@ module foundation.symmetric-difference where
open import foundation.decidable-subtypes
open import foundation.decidable-types
open import foundation.dependent-pair-types
-open import foundation.exclusive-disjunction
-open import foundation.identity-types hiding (inv)
+open import foundation.exclusive-sum
open import foundation.intersections-subtypes
open import foundation.universe-levels
@@ -19,6 +18,7 @@ open import foundation-core.coproduct-types
open import foundation-core.decidable-propositions
open import foundation-core.equivalences
open import foundation-core.function-types
+open import foundation-core.identity-types
open import foundation-core.propositions
open import foundation-core.subtypes
open import foundation-core.transport-along-identifications
@@ -41,12 +41,14 @@ module _
symmetric-difference-subtype :
subtype l1 X → subtype l2 X → subtype (l1 ⊔ l2) X
- symmetric-difference-subtype P Q x = xor-Prop (P x) (Q x)
+ symmetric-difference-subtype P Q x =
+ exclusive-sum-Prop (P x) (Q x)
symmetric-difference-decidable-subtype :
decidable-subtype l1 X → decidable-subtype l2 X →
decidable-subtype (l1 ⊔ l2) X
- symmetric-difference-decidable-subtype P Q x = xor-Decidable-Prop (P x) (Q x)
+ symmetric-difference-decidable-subtype P Q x =
+ exclusive-sum-Decidable-Prop (P x) (Q x)
```
## Properties
diff --git a/src/foundation/truncated-types.lagda.md b/src/foundation/truncated-types.lagda.md
index 3adb80b12a..d2248f9dab 100644
--- a/src/foundation/truncated-types.lagda.md
+++ b/src/foundation/truncated-types.lagda.md
@@ -13,6 +13,7 @@ open import elementary-number-theory.natural-numbers
open import foundation.dependent-pair-types
open import foundation.equivalences
+open import foundation.logical-equivalences
open import foundation.subtype-identity-principle
open import foundation.truncation-levels
open import foundation.univalence
@@ -104,7 +105,7 @@ module _
equiv-is-trunc-equiv : A ≃ B → is-trunc k A ≃ is-trunc k B
equiv-is-trunc-equiv e =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-trunc k A)
( is-prop-is-trunc k B)
( is-trunc-equiv' k A e)
diff --git a/src/foundation/unions-subtypes.lagda.md b/src/foundation/unions-subtypes.lagda.md
index d726dd99da..f1bd0aee75 100644
--- a/src/foundation/unions-subtypes.lagda.md
+++ b/src/foundation/unions-subtypes.lagda.md
@@ -37,7 +37,7 @@ module _
where
union-subtype : subtype l1 X → subtype l2 X → subtype (l1 ⊔ l2) X
- union-subtype P Q x = disjunction-Prop (P x) (Q x)
+ union-subtype P Q x = (P x) ∨ (Q x)
```
### Unions of decidable subtypes
diff --git a/src/foundation/unique-existence.lagda.md b/src/foundation/unique-existence.lagda.md
deleted file mode 100644
index a60d5ee118..0000000000
--- a/src/foundation/unique-existence.lagda.md
+++ /dev/null
@@ -1,26 +0,0 @@
-# Unique existence
-
-```agda
-module foundation.unique-existence where
-```
-
-Imports
-
-```agda
-open import foundation.universe-levels
-
-open import foundation-core.torsorial-type-families
-```
-
-
-
-## Idea
-
-Unique existence `∃! A B` is defined as `Σ A B` being contractible.
-
-## Definition
-
-```agda
-∃! : {l1 l2 : Level} → (A : UU l1) → (A → UU l2) → UU (l1 ⊔ l2)
-∃! A B = is-torsorial B
-```
diff --git a/src/foundation/uniqueness-quantification.lagda.md b/src/foundation/uniqueness-quantification.lagda.md
new file mode 100644
index 0000000000..dbbce2a023
--- /dev/null
+++ b/src/foundation/uniqueness-quantification.lagda.md
@@ -0,0 +1,133 @@
+# Uniqueness quantification
+
+```agda
+module foundation.uniqueness-quantification where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.torsorial-type-families
+open import foundation.universe-levels
+
+open import foundation-core.contractible-types
+open import foundation-core.function-types
+open import foundation-core.identity-types
+open import foundation-core.propositions
+```
+
+
+
+## Idea
+
+Given a predicate `P : A → Prop` we say there
+{{#concept "uniquely exists" Disambiguation="in a subtype" WDID=Q2502253 WD="uniqueness quantification" Agda=∃!}}
+_an `x : A` satisfying `P`_, if the [subtype](foundation-core.subtypes.md)
+`Σ (x : A), (P x)` is [contractible](foundation-core.contractible-types.md).
+
+More generally, given a [structure](foundation.structure.md) `B : A → 𝒰` we say
+there
+{{#concept "uniquely exists" Disambiguation="structure" Agda=uniquely-exists-structure}}
+_an `x : A` and a `y : B x`_, if the
+[total type](foundation.dependent-pair-types.md) `Σ (x : A), (B x)` is
+contractible.
+
+Note that the unique existence of structure is defined in the exact same way as
+the concept of
+[torsorial type families](foundation-core.torsorial-type-families.md). Whether
+to use the concept of unique existence of a structure or the concept of
+torsorial type family is dependent on the context. Torsoriality is used often to
+emphasize the relation of the type family to the identity type, whereas
+uniqueness of structure is used to emphasize the uniqueness of elements equipped
+with further structure. For example, we tend to use unique existence in
+combination with universal properties, in order to conclude that a certain map
+equipped with some homotopies exists uniquely.
+
+As a special case of uniqueness quantification, we recover
+[exclusive disjunction](foundation.exclusive-disjunction.md) when the indexing
+type is a [2-element type](univalent-combinatorics.2-element-types.md).
+Concretely, we have the equivalence
+
+```text
+ ∃! (t : bool), (P t) ≐ is-contr (Σ (t : bool), (P t))
+ ≃ is-contr ((P false) + (P true))
+ ≐ P false ⊻ P true.
+```
+
+## Definitions
+
+### Unique existence of structure
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (B : A → UU l2)
+ where
+
+ uniquely-exists-structure-Prop : Prop (l1 ⊔ l2)
+ uniquely-exists-structure-Prop = is-torsorial-Prop B
+
+ uniquely-exists-structure : UU (l1 ⊔ l2)
+ uniquely-exists-structure = type-Prop uniquely-exists-structure-Prop
+
+ is-prop-uniquely-exists-structure : is-prop uniquely-exists-structure
+ is-prop-uniquely-exists-structure =
+ is-prop-type-Prop uniquely-exists-structure-Prop
+```
+
+### Unique existence in a subtype
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (P : A → Prop l2)
+ where
+
+ uniquely-exists-Prop : Prop (l1 ⊔ l2)
+ uniquely-exists-Prop = uniquely-exists-structure-Prop A (type-Prop ∘ P)
+
+ uniquely-exists : UU (l1 ⊔ l2)
+ uniquely-exists = type-Prop uniquely-exists-Prop
+
+ is-prop-uniquely-exists : is-prop uniquely-exists
+ is-prop-uniquely-exists = is-prop-type-Prop uniquely-exists-Prop
+
+ ∃! : Prop (l1 ⊔ l2)
+ ∃! = uniquely-exists-Prop
+```
+
+### Components of unique existence
+
+```agda
+module _
+ {l1 l2 : Level} {A : UU l1} {B : A → UU l2}
+ where
+
+ pair-uniquely-exists : uniquely-exists-structure A B → Σ A B
+ pair-uniquely-exists = center
+
+ pr1-uniquely-exists : uniquely-exists-structure A B → A
+ pr1-uniquely-exists = pr1 ∘ pair-uniquely-exists
+
+ pr2-uniquely-exists :
+ (p : uniquely-exists-structure A B) → B (pr1-uniquely-exists p)
+ pr2-uniquely-exists = pr2 ∘ pair-uniquely-exists
+
+ contraction-uniquely-exists :
+ (p : uniquely-exists-structure A B) →
+ (q : Σ A B) → pair-uniquely-exists p = q
+ contraction-uniquely-exists = contraction
+```
+
+## See also
+
+- Unique existence is the indexed counterpart to
+ [exclusive disjunction](foundation.exclusive-disjunction.md).
+- A different name for _unique existence_ is
+ [torsoriality](foundation.torsorial-type-families.md).
+
+## External links
+
+- [uniqueness quantifier](https://ncatlab.org/nlab/show/uniqueness+quantifier)
+ at $n$Lab
+- [Uniqueness quantification](https://en.wikipedia.org/wiki/Uniqueness_quantification)
+ at Wikipedia
diff --git a/src/foundation/universal-property-cartesian-product-types.lagda.md b/src/foundation/universal-property-cartesian-product-types.lagda.md
index e2ff4ee524..fe91b41474 100644
--- a/src/foundation/universal-property-cartesian-product-types.lagda.md
+++ b/src/foundation/universal-property-cartesian-product-types.lagda.md
@@ -9,6 +9,7 @@ module foundation.universal-property-cartesian-product-types where
```agda
open import foundation.cones-over-cospan-diagrams
open import foundation.dependent-pair-types
+open import foundation.logical-equivalences
open import foundation.standard-pullbacks
open import foundation.unit-type
open import foundation.universal-property-dependent-pair-types
@@ -74,6 +75,10 @@ module _
pr1 (map-inv-up-product (f , g) x) = f x
pr2 (map-inv-up-product (f , g) x) = g x
+ iff-up-product :
+ ((x : X) → A x × B x) ↔ ((x : X) → A x) × ((x : X) → B x)
+ iff-up-product = (map-up-product , map-inv-up-product)
+
up-product : is-equiv map-up-product
up-product =
is-equiv-is-invertible (map-inv-up-product) (refl-htpy) (refl-htpy)
diff --git a/src/foundation/universal-property-image.lagda.md b/src/foundation/universal-property-image.lagda.md
index 1e1636519d..0d8de3ea76 100644
--- a/src/foundation/universal-property-image.lagda.md
+++ b/src/foundation/universal-property-image.lagda.md
@@ -104,7 +104,7 @@ abstract
{ B : UU l3} (i : B ↪ X) (q : hom-slice f (map-emb i)) →
is-image' f i q → is-image f i q
is-image-is-image' f i q up' C j =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-hom-slice (map-emb i) j)
( is-prop-hom-slice f j)
( up' C j)
diff --git a/src/foundation/universal-property-propositional-truncation.lagda.md b/src/foundation/universal-property-propositional-truncation.lagda.md
index bd74f3e7d0..61fbe14f5b 100644
--- a/src/foundation/universal-property-propositional-truncation.lagda.md
+++ b/src/foundation/universal-property-propositional-truncation.lagda.md
@@ -11,6 +11,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.functoriality-cartesian-product-types
+open import foundation.logical-equivalences
open import foundation.precomposition-functions-into-subuniverses
open import foundation.subtype-identity-principle
open import foundation.unit-type
@@ -165,7 +166,7 @@ abstract
extension-property-propositional-truncation P f →
is-propositional-truncation P f
is-propositional-truncation-extension-property P f up-P Q =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-Π (λ x → is-prop-type-Prop Q))
( is-prop-Π (λ x → is-prop-type-Prop Q))
( up-P Q)
@@ -183,7 +184,7 @@ abstract
is-propositional-truncation P' f' →
is-equiv h
is-equiv-is-ptruncation-is-ptruncation P P' f f' h H is-ptr-P is-ptr-P' =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-Prop P)
( is-prop-type-Prop P')
( map-inv-is-equiv (is-ptr-P' P) f)
@@ -278,7 +279,7 @@ abstract
{l1 l2 : Level} {A : UU l1} (P : Prop l2) (f : A → type-Prop P) →
(g : type-Prop P → A) → is-propositional-truncation P f
is-propositional-truncation-has-section {A = A} P f g Q =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-function-type (is-prop-type-Prop Q))
( is-prop-function-type (is-prop-type-Prop Q))
( λ h → h ∘ g)
@@ -358,7 +359,7 @@ abstract
( is-equiv-comp
( λ h a a' → h a (f' a'))
( λ h a p' → h (f a) p')
- ( is-ptr-f (pair (type-hom-Prop P' Q) (is-prop-type-hom-Prop P' Q)))
+ ( is-ptr-f (pair (type-hom-Prop P' Q) (is-prop-hom-Prop P' Q)))
( is-equiv-map-Π-is-fiberwise-equiv
( λ a → is-ptr-f' Q)))
```
diff --git a/src/foundation/universal-property-set-quotients.lagda.md b/src/foundation/universal-property-set-quotients.lagda.md
index a4b7eb03fc..5e8ef52d96 100644
--- a/src/foundation/universal-property-set-quotients.lagda.md
+++ b/src/foundation/universal-property-set-quotients.lagda.md
@@ -20,7 +20,9 @@ open import foundation.function-extensionality
open import foundation.fundamental-theorem-of-identity-types
open import foundation.identity-types
open import foundation.images
+open import foundation.injective-maps
open import foundation.locally-small-types
+open import foundation.logical-equivalences
open import foundation.propositional-extensionality
open import foundation.propositional-truncations
open import foundation.reflecting-maps-equivalence-relations
@@ -43,7 +45,6 @@ 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.propositional-maps
open import foundation-core.propositions
open import foundation-core.small-types
@@ -527,7 +528,7 @@ module _
P-Prop : (b : type-Set B) (x : type-Set X) → Prop (l1 ⊔ l3 ⊔ l)
P-Prop b x =
- ∃-Prop A
+ exists-structure-Prop A
( λ a →
( map-reflecting-map-equivalence-relation R f a = x) ×
( map-reflecting-map-equivalence-relation R q a = b))
@@ -838,7 +839,7 @@ module _
pr1 equiv-equalizes-equal-values-extension-along-surjection-Set =
equalizes-equal-values-extension-along-surjection-Set
pr2 equiv-equalizes-equal-values-extension-along-surjection-Set =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-extension-along-surjection-Set f C g)
( is-prop-equalizes-equal-values-surjection-Set)
( extension-equalizes-equal-values-surjection-Set)
diff --git a/src/foundation/universal-quantification.lagda.md b/src/foundation/universal-quantification.lagda.md
new file mode 100644
index 0000000000..e4b0bd024f
--- /dev/null
+++ b/src/foundation/universal-quantification.lagda.md
@@ -0,0 +1,151 @@
+# Universal quantification
+
+```agda
+module foundation.universal-quantification where
+```
+
+Imports
+
+```agda
+open import foundation.dependent-pair-types
+open import foundation.evaluation-functions
+open import foundation.logical-equivalences
+open import foundation.propositional-truncations
+open import foundation.universe-levels
+
+open import foundation-core.equivalences
+open import foundation-core.function-types
+open import foundation-core.propositions
+```
+
+
+
+## Idea
+
+Given a type `A` and a [subtype](foundation-core.subtypes.md) `P : A → Prop`,
+the
+{{#concept "universal quantification" Disambiguation="on a subtype" WDID=Q126695 WD="universal quantification"}}
+
+```text
+ ∀ (x : A), (P x)
+```
+
+is the [proposition](foundation-core.propositions.md) that there exists a proof
+of `P x` for every `x` in `A`.
+
+The
+{{#concept "universal property" Disambiguation="of universal quantification" Agda=universal-property-for-all}}
+of universal quantification states that it is the
+[greatest lower bound](order-theory.greatest-lower-bounds-large-posets.md) on
+the family of propositions `P` in the
+[locale of propositions](foundation.large-locale-of-propositions.md), by which
+we mean that for every proposition `Q` we have the
+[logical equivalence](foundation.logical-equivalences.md)
+
+```text
+ (∀ (a : A), (R → P a)) ↔ (R → ∀ (a : A), (P a))
+```
+
+**Notation.** Because of syntactic limitations of the Agda language, we cannot
+use `∀` for the universal quantification in formalizations, and instead use
+`∀'`.
+
+## Definitions
+
+### Universal quantification
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (P : A → Prop l2)
+ where
+
+ for-all-Prop : Prop (l1 ⊔ l2)
+ for-all-Prop = Π-Prop A P
+
+ type-for-all-Prop : UU (l1 ⊔ l2)
+ type-for-all-Prop = type-Prop for-all-Prop
+
+ is-prop-for-all-Prop : is-prop type-for-all-Prop
+ is-prop-for-all-Prop = is-prop-type-Prop for-all-Prop
+
+ for-all : UU (l1 ⊔ l2)
+ for-all = type-for-all-Prop
+
+ ∀' : Prop (l1 ⊔ l2)
+ ∀' = for-all-Prop
+```
+
+### The universal property of universal quantification
+
+The
+{{#concept "universal property" Disambiguation="of universal quantification" Agda=universal-property-for-all}}
+of the universal quantification `∀ (a : A), (P a)` states that for every
+proposition `R`, the canonical map
+
+```text
+ (∀ (a : A), (R → P a)) → (R → ∀ (a : A), (P a))
+```
+
+is a [logical equivalence](foundation.logical-equivalences.md). Indeed, this
+holds for any type `R`.
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (P : A → Prop l2)
+ where
+
+ universal-property-for-all : {l3 : Level} (S : Prop l3) → UUω
+ universal-property-for-all S =
+ {l : Level} (R : Prop l) →
+ type-Prop ((∀' A (λ a → R ⇒ P a)) ⇔ (R ⇒ S))
+
+ ev-for-all :
+ {l : Level} {B : UU l} →
+ for-all A (λ a → function-Prop B (P a)) → B → for-all A P
+ ev-for-all f r a = f a r
+```
+
+## Properties
+
+### Universal quantification satisfies its universal property
+
+```agda
+module _
+ {l1 l2 : Level} (A : UU l1) (P : A → Prop l2)
+ where
+
+ map-up-for-all :
+ {l : Level} {B : UU l} →
+ (B → for-all A P) → for-all A (λ a → function-Prop B (P a))
+ map-up-for-all f a r = f r a
+
+ is-equiv-ev-for-all :
+ {l : Level} {B : UU l} → is-equiv (ev-for-all A P {B = B})
+ is-equiv-ev-for-all {B = B} =
+ is-equiv-has-converse
+ ( ∀' A (λ a → function-Prop B (P a)))
+ ( function-Prop B (∀' A P))
+ ( map-up-for-all)
+
+ up-for-all : universal-property-for-all A P (∀' A P)
+ up-for-all R = (ev-for-all A P , map-up-for-all)
+```
+
+## See also
+
+- Universal quantification is the indexed counterpart to
+ [conjunction](foundation.conjunction.md).
+
+## Table of files about propositional logic
+
+The following table gives an overview of basic constructions in propositional
+logic and related considerations.
+
+{{#include tables/propositional-logic.md}}
+
+## External links
+
+- [universal quantifier](https://ncatlab.org/nlab/show/universal+quantifier) at
+ $n$Lab
+- [Universal quantification](https://en.wikipedia.org/wiki/Universal_quantification)
+ at Wikipedia
diff --git a/src/foundation/unordered-pairs.lagda.md b/src/foundation/unordered-pairs.lagda.md
index fcc7e11206..3a6a00c0ad 100644
--- a/src/foundation/unordered-pairs.lagda.md
+++ b/src/foundation/unordered-pairs.lagda.md
@@ -107,7 +107,9 @@ module _
is-in-unordered-pair-Prop :
{l : Level} {A : UU l} (p : unordered-pair A) (a : A) → Prop l
is-in-unordered-pair-Prop p a =
- ∃-Prop (type-unordered-pair p) (λ x → element-unordered-pair p x = a)
+ exists-structure-Prop
+ ( type-unordered-pair p)
+ ( λ x → element-unordered-pair p x = a)
is-in-unordered-pair :
{l : Level} {A : UU l} (p : unordered-pair A) (a : A) → UU l
diff --git a/src/foundation/weak-limited-principle-of-omniscience.lagda.md b/src/foundation/weak-limited-principle-of-omniscience.lagda.md
index ff5da85ae7..b4a805bc24 100644
--- a/src/foundation/weak-limited-principle-of-omniscience.lagda.md
+++ b/src/foundation/weak-limited-principle-of-omniscience.lagda.md
@@ -11,6 +11,7 @@ open import elementary-number-theory.natural-numbers
open import foundation.disjunction
open import foundation.negation
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.propositions
@@ -23,17 +24,23 @@ open import univalent-combinatorics.standard-finite-types
## Statement
-The **Weak Limited Principle of Omniscience** asserts that for any sequence
-`f : ℕ → Fin 2` either `f n = 0` for all `n : ℕ` or not. In particular, it is a
-restricted form of the law of excluded middle.
+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
+[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)))))
+
WLPO : UU lzero
-WLPO =
- (f : ℕ → Fin 2) →
- type-disjunction-Prop
- ( Π-Prop ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1)))
- ( neg-Prop (Π-Prop ℕ (λ n → Id-Prop (Fin-Set 2) (f n) (zero-Fin 1))))
+WLPO = type-Prop WLPO-Prop
```
## See also
diff --git a/src/group-theory/cayleys-theorem.lagda.md b/src/group-theory/cayleys-theorem.lagda.md
index 8a1459ba57..6b7594ab67 100644
--- a/src/group-theory/cayleys-theorem.lagda.md
+++ b/src/group-theory/cayleys-theorem.lagda.md
@@ -12,6 +12,7 @@ open import foundation.embeddings
open import foundation.equivalence-extensionality
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.universe-levels
open import group-theory.embeddings-groups
diff --git a/src/group-theory/commutator-subgroups.lagda.md b/src/group-theory/commutator-subgroups.lagda.md
index 9309698b09..2316c49c7e 100644
--- a/src/group-theory/commutator-subgroups.lagda.md
+++ b/src/group-theory/commutator-subgroups.lagda.md
@@ -129,7 +129,7 @@ module _
( generating-subset-commutator-subgroup-Group (conjugation-Group G x y))
( λ where
( (y , z) , refl) →
- intro-∃
+ intro-exists
( conjugation-Group G x y , conjugation-Group G x z)
( inv (distributive-conjugation-commutator-Group G x y z)))
diff --git a/src/group-theory/cyclic-groups.lagda.md b/src/group-theory/cyclic-groups.lagda.md
index b51af13775..08c862b1fc 100644
--- a/src/group-theory/cyclic-groups.lagda.md
+++ b/src/group-theory/cyclic-groups.lagda.md
@@ -51,9 +51,7 @@ module _
is-cyclic-prop-Group : Prop l1
is-cyclic-prop-Group =
- ∃-Prop
- ( type-Group G)
- ( is-generating-element-Group G)
+ ∃ (type-Group G) (is-generating-element-prop-Group G)
is-cyclic-Group : UU l1
is-cyclic-Group = type-Prop is-cyclic-prop-Group
@@ -136,13 +134,17 @@ module _
is-cyclic-has-generating-element-Group :
is-cyclic-Group G →
- {l : Level} → ∃ (type-Group G) (λ g → is-emb-ev-element-hom-Group' G g l)
- is-cyclic-has-generating-element-Group H =
+ {l : Level} →
+ exists-structure
+ ( type-Group G)
+ ( λ g → is-emb-ev-element-hom-Group' G g l)
+ is-cyclic-has-generating-element-Group H {l} =
apply-universal-property-trunc-Prop H
- ( ∃-Prop (type-Group G) (λ g → is-emb-ev-element-hom-Group' G g _))
+ ( exists-structure-Prop
+ ( type-Group G)
+ ( λ g → is-emb-ev-element-hom-Group' G g l))
( λ (g , u) →
- intro-∃ g
- ( is-emb-ev-element-is-generating-element-Group G g u))
+ intro-exists g (is-emb-ev-element-is-generating-element-Group G g u))
```
## See also
diff --git a/src/group-theory/elements-of-finite-order-groups.lagda.md b/src/group-theory/elements-of-finite-order-groups.lagda.md
index 083748e2e4..de05cf476b 100644
--- a/src/group-theory/elements-of-finite-order-groups.lagda.md
+++ b/src/group-theory/elements-of-finite-order-groups.lagda.md
@@ -47,10 +47,9 @@ module _
has-finite-order-element-prop-Group : Prop l1
has-finite-order-element-prop-Group =
- ∃-Prop
- ( nonzero-ℤ)
+ ∃ ( nonzero-ℤ)
( λ k →
- has-same-elements-Subgroup ℤ-Group
+ has-same-elements-prop-Subgroup ℤ-Group
( subgroup-element-Group ℤ-Group (int-nonzero-ℤ k))
( subgroup-order-element-Group G x))
```
diff --git a/src/group-theory/integer-powers-of-elements-groups.lagda.md b/src/group-theory/integer-powers-of-elements-groups.lagda.md
index 1e105e8922..47de6edf13 100644
--- a/src/group-theory/integer-powers-of-elements-groups.lagda.md
+++ b/src/group-theory/integer-powers-of-elements-groups.lagda.md
@@ -78,7 +78,7 @@ module _
is-integer-power-of-element-prop-Group :
(x y : type-Group G) → Prop l
is-integer-power-of-element-prop-Group x y =
- ∃-Prop ℤ (λ k → integer-power-Group G k x = y)
+ exists-structure-Prop ℤ (λ k → integer-power-Group G k x = y)
is-integer-power-of-element-Group :
(x y : type-Group G) → UU l
diff --git a/src/group-theory/isomorphisms-group-actions.lagda.md b/src/group-theory/isomorphisms-group-actions.lagda.md
index 9337d23248..49635eee09 100644
--- a/src/group-theory/isomorphisms-group-actions.lagda.md
+++ b/src/group-theory/isomorphisms-group-actions.lagda.md
@@ -16,6 +16,7 @@ open import foundation.equivalences
open import foundation.function-extensionality
open import foundation.functoriality-dependent-pair-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.subtypes
open import foundation.torsorial-type-families
@@ -196,7 +197,7 @@ module _
is-equiv-is-equiv-hom-is-iso-action-Group :
is-equiv is-equiv-hom-is-iso-action-Group
is-equiv-is-equiv-hom-is-iso-action-Group =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-is-iso-Large-Precategory
( action-Group-Large-Precategory G) {X = X} {Y = Y} f)
( is-property-is-equiv (map-hom-action-Group G X Y f))
@@ -205,7 +206,7 @@ module _
is-equiv-is-iso-is-equiv-hom-action-Group :
is-equiv is-iso-is-equiv-hom-action-Group
is-equiv-is-iso-is-equiv-hom-action-Group =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-property-is-equiv (map-hom-action-Group G X Y f))
( is-prop-is-iso-Large-Precategory
( action-Group-Large-Precategory G) {X = X} {Y = Y} f)
diff --git a/src/group-theory/nontrivial-groups.lagda.md b/src/group-theory/nontrivial-groups.lagda.md
index 83a4e7c5a2..725296c895 100644
--- a/src/group-theory/nontrivial-groups.lagda.md
+++ b/src/group-theory/nontrivial-groups.lagda.md
@@ -49,7 +49,7 @@ module _
is-nontrivial-prop-Group : Prop l1
is-nontrivial-prop-Group =
- ∃-Prop (type-Group G) (λ g → unit-Group G ≠ g)
+ exists-structure-Prop (type-Group G) (λ g → unit-Group G ≠ g)
is-nontrivial-Group : UU l1
is-nontrivial-Group =
@@ -70,7 +70,7 @@ module _
is-not-trivial-prop-Group : Prop l1
is-not-trivial-prop-Group =
- neg-Prop' ((x : type-Group G) → unit-Group G = x)
+ neg-type-Prop ((x : type-Group G) → unit-Group G = x)
is-not-trivial-Group : UU l1
is-not-trivial-Group =
@@ -138,10 +138,7 @@ module _
( f)
( forward-implication
( iff-eq (ap (λ T → subset-Subgroup G T x) α))
- ( inr-disjunction-Prop
- ( Id-Prop (set-Group G) _ _)
- ( P)
- ( p))))
+ ( inr-disjunction p)))
( λ q →
map-left-unit-law-disjunction-is-empty-Prop
( Id-Prop (set-Group G) _ _)
@@ -149,10 +146,7 @@ module _
( f)
( backward-implication
( iff-eq (ap (λ T → subset-Subgroup G T x) α))
- ( inr-disjunction-Prop
- ( Id-Prop (set-Group G) _ _)
- ( Q)
- ( q))))))
+ ( inr-disjunction q)))))
```
### If the map `subgroup-Prop G : Prop lzero → Subgroup l1 G` is an embedding, then `G` is not a trivial group
diff --git a/src/group-theory/normal-closures-subgroups.lagda.md b/src/group-theory/normal-closures-subgroups.lagda.md
index b978b2b76a..704d72a7b2 100644
--- a/src/group-theory/normal-closures-subgroups.lagda.md
+++ b/src/group-theory/normal-closures-subgroups.lagda.md
@@ -75,7 +75,7 @@ module _
generating-subset-normal-closure-Subgroup : subset-Group (l1 ⊔ l2) G
generating-subset-normal-closure-Subgroup x =
- ∃-Prop
+ exists-structure-Prop
( type-Group G)
( λ y → fiber (conjugation-Group G y ∘ inclusion-Subgroup G H) x)
diff --git a/src/group-theory/powers-of-elements-monoids.lagda.md b/src/group-theory/powers-of-elements-monoids.lagda.md
index 4cc4d89acd..66c4e4f41b 100644
--- a/src/group-theory/powers-of-elements-monoids.lagda.md
+++ b/src/group-theory/powers-of-elements-monoids.lagda.md
@@ -58,7 +58,7 @@ module _
is-power-of-element-prop-Monoid : (x y : type-Monoid M) → Prop l
is-power-of-element-prop-Monoid x y =
- ∃-Prop ℕ (λ n → power-Monoid M n x = y)
+ exists-structure-Prop ℕ (λ n → power-Monoid M n x = y)
is-power-of-element-Monoid : (x y : type-Monoid M) → UU l
is-power-of-element-Monoid x y =
diff --git a/src/group-theory/quotient-groups.lagda.md b/src/group-theory/quotient-groups.lagda.md
index f4c13fd08f..0f695c97ba 100644
--- a/src/group-theory/quotient-groups.lagda.md
+++ b/src/group-theory/quotient-groups.lagda.md
@@ -19,6 +19,7 @@ open import foundation.function-types
open import foundation.functoriality-dependent-pair-types
open import foundation.functoriality-set-quotients
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.reflecting-maps-equivalence-relations
open import foundation.set-quotients
@@ -584,7 +585,7 @@ module _
( equivalence-relation-congruence-Normal-Subgroup G N)
( set-Group H))
( λ f →
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-preserves-mul-Group (quotient-Group G N) H f)
( is-prop-preserves-mul-Group G H
( f ∘ map-quotient-hom-Group G N))
diff --git a/src/group-theory/subgroups-concrete-groups.lagda.md b/src/group-theory/subgroups-concrete-groups.lagda.md
index 187b71f58b..eb86e36730 100644
--- a/src/group-theory/subgroups-concrete-groups.lagda.md
+++ b/src/group-theory/subgroups-concrete-groups.lagda.md
@@ -244,7 +244,7 @@ is-set-subgroup-Concrete-Group G X Y =
( transitive-action-subgroup-Concrete-Group G Y)
( pr1 e)
( pr1 f)
- ( intro-∃
+ ( intro-exists
( unit-coset-subgroup-Concrete-Group G X)
( pr2 e ∙ inv (pr2 f)))))))
```
diff --git a/src/group-theory/subgroups.lagda.md b/src/group-theory/subgroups.lagda.md
index 3b2d7daa93..98631e94e6 100644
--- a/src/group-theory/subgroups.lagda.md
+++ b/src/group-theory/subgroups.lagda.md
@@ -671,11 +671,7 @@ module _
contains-unit-subgroup-Prop :
contains-unit-subset-Group G subset-subgroup-Prop
- contains-unit-subgroup-Prop =
- inl-disjunction-Prop
- ( Id-Prop (set-Group G) (unit-Group G) (unit-Group G))
- ( P)
- ( refl)
+ contains-unit-subgroup-Prop = inl-disjunction refl
is-closed-under-multiplication-subgroup-Prop' :
(x y : type-Group G) →
diff --git a/src/group-theory/substitution-functor-group-actions.lagda.md b/src/group-theory/substitution-functor-group-actions.lagda.md
index ef24eda041..d8d73e7e03 100644
--- a/src/group-theory/substitution-functor-group-actions.lagda.md
+++ b/src/group-theory/substitution-functor-group-actions.lagda.md
@@ -135,7 +135,7 @@ module _
( equivalence-relation-obj-left-adjoint-subst-action-Group X)
( h , x)
( h' , x') =
- ∃-Prop
+ exists-structure-Prop
( type-Group G)
( λ g →
( Id (mul-Group H (map-hom-Group G H f g) h) h') ×
@@ -143,7 +143,7 @@ module _
pr1
( pr2 (equivalence-relation-obj-left-adjoint-subst-action-Group X))
( h , x) =
- intro-∃
+ intro-exists
( unit-Group G)
( pair
( ( ap (mul-Group' H h) (preserves-unit-hom-Group G H f)) ∙
@@ -162,7 +162,7 @@ module _
( h' , x')
( h , x))
( λ (g , p , q) →
- intro-∃
+ intro-exists
( inv-Group G g)
( ( ( ap
( mul-Group' H h')
@@ -191,7 +191,7 @@ module _
( h , x)
( h'' , x''))
( λ (g' , p' , q') →
- intro-∃
+ intro-exists
( mul-Group G g' g)
( ( ( ap
( mul-Group' H h)
diff --git a/src/group-theory/torsion-elements-groups.lagda.md b/src/group-theory/torsion-elements-groups.lagda.md
index 5fbdd3842e..9b7a129185 100644
--- a/src/group-theory/torsion-elements-groups.lagda.md
+++ b/src/group-theory/torsion-elements-groups.lagda.md
@@ -54,7 +54,7 @@ module _
is-torsion-element-prop-Group : Prop l1
is-torsion-element-prop-Group =
- ∃-Prop
+ exists-structure-Prop
( nonzero-ℤ)
( λ k → integer-power-Group G (int-nonzero-ℤ k) x = unit-Group G)
@@ -88,7 +88,7 @@ module _
is-torsion-element-unit-Group : is-torsion-element-Group G (unit-Group G)
is-torsion-element-unit-Group =
- intro-∃
+ intro-exists
( one-nonzero-ℤ)
( integer-power-unit-Group G one-ℤ)
diff --git a/src/group-theory/torsion-free-groups.lagda.md b/src/group-theory/torsion-free-groups.lagda.md
index 4269e473a5..f207fffb76 100644
--- a/src/group-theory/torsion-free-groups.lagda.md
+++ b/src/group-theory/torsion-free-groups.lagda.md
@@ -17,6 +17,7 @@ open import foundation.equivalences
open import foundation.existential-quantification
open import foundation.fundamental-theorem-of-identity-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.singleton-subtypes
@@ -149,7 +150,7 @@ module _
is-torsion-free-has-unique-torsion-element-Group :
has-unique-torsion-element-Group G → is-torsion-free-Group G
is-torsion-free-has-unique-torsion-element-Group H x k p =
- ap pr1 (eq-is-contr H {x , intro-∃ k p} {unit-torsion-element-Group G})
+ ap pr1 (eq-is-contr H {x , intro-exists k p} {unit-torsion-element-Group G})
abstract
has-unique-torsion-element-is-torsion-free-Group :
@@ -158,11 +159,10 @@ module _
fundamental-theorem-id'
( λ where x refl → is-torsion-element-unit-Group G)
( λ x →
- is-equiv-is-prop
- ( is-set-type-Group G _ _)
+ is-equiv-has-converse-is-prop
+ ( is-set-type-Group G (unit-Group G) x)
( is-prop-is-torsion-element-Group G x)
- ( elim-exists-Prop
- ( λ k → Id-Prop (set-Group G) _ _)
- ( Id-Prop (set-Group G) _ _)
- ( λ k p → inv (H x k p))))
+ ( elim-exists
+ ( Id-Prop (set-Group G) (unit-Group G) x)
+ ( λ k p → inv (H x k p))))
```
diff --git a/src/group-theory/transitive-concrete-group-actions.lagda.md b/src/group-theory/transitive-concrete-group-actions.lagda.md
index 5b56c26696..f6af5d47b3 100644
--- a/src/group-theory/transitive-concrete-group-actions.lagda.md
+++ b/src/group-theory/transitive-concrete-group-actions.lagda.md
@@ -278,7 +278,8 @@ module _
( f)
htpy-exists-equiv-transitive-action-Concrete-Group :
- ∃ ( type-transitive-action-Concrete-Group G X)
+ exists-structure
+ ( type-transitive-action-Concrete-Group G X)
( λ x →
map-equiv-transitive-action-Concrete-Group G X Y e x =
map-equiv-transitive-action-Concrete-Group G X Y f x) →
diff --git a/src/higher-group-theory/cyclic-higher-groups.lagda.md b/src/higher-group-theory/cyclic-higher-groups.lagda.md
index 34d954d1f2..d9067fb971 100644
--- a/src/higher-group-theory/cyclic-higher-groups.lagda.md
+++ b/src/higher-group-theory/cyclic-higher-groups.lagda.md
@@ -60,7 +60,7 @@ module _
is-cyclic-prop-∞-Group : Prop (l1 ⊔ lsuc l2)
is-cyclic-prop-∞-Group =
- ∃-Prop
+ exists-structure-Prop
( type-∞-Group G)
( λ g → (H : ∞-Group l2) → is-emb (ev-element-∞-Group G H g))
diff --git a/src/linear-algebra/vectors.lagda.md b/src/linear-algebra/vectors.lagda.md
index 0fd59958df..a870ba74c9 100644
--- a/src/linear-algebra/vectors.lagda.md
+++ b/src/linear-algebra/vectors.lagda.md
@@ -51,7 +51,7 @@ infixr 10 _∷_
data vec {l : Level} (A : UU l) : ℕ → UU l where
empty-vec : vec A zero-ℕ
- _∷_ : ∀ {n} → A → vec A n → vec A (succ-ℕ n)
+ _∷_ : {n : ℕ} → A → vec A n → vec A (succ-ℕ n)
module _
{l : Level} {A : UU l}
diff --git a/src/lists/sorted-vectors.lagda.md b/src/lists/sorted-vectors.lagda.md
index 2d7921986a..dbb2353d4f 100644
--- a/src/lists/sorted-vectors.lagda.md
+++ b/src/lists/sorted-vectors.lagda.md
@@ -100,7 +100,7 @@ module _
```agda
is-least-element-vec-is-leq-head-sorted-vec :
{n : ℕ}
- (x : type-Decidable-Total-Order X)
+ (x : type-Decidable-Total-Order X)
(v : vec (type-Decidable-Total-Order X) (succ-ℕ n)) →
is-sorted-vec v → leq-Decidable-Total-Order X x (head-vec v) →
is-least-element-vec x v
@@ -109,16 +109,16 @@ module _
is-least-element-vec-is-leq-head-sorted-vec {succ-ℕ n} x (y ∷ v) s p =
p ,
( is-least-element-vec-is-leq-head-sorted-vec
+ ( x)
+ ( v)
+ ( is-sorted-tail-is-sorted-vec (y ∷ v) s)
+ ( transitive-leq-Decidable-Total-Order
+ ( X)
( x)
- ( v)
- ( is-sorted-tail-is-sorted-vec (y ∷ v) s)
- ( transitive-leq-Decidable-Total-Order
- ( X)
- ( x)
- ( y)
- ( head-vec v)
- ( is-leq-head-head-tail-is-sorted-vec (y ∷ v) s)
- ( p)))
+ ( y)
+ ( head-vec v)
+ ( is-leq-head-head-tail-is-sorted-vec (y ∷ v) s)
+ ( p)))
```
### An equivalent definition of being sorted
@@ -187,7 +187,7 @@ module _
(fv : functional-vec (type-Decidable-Total-Order X) n) →
Prop l2
is-least-element-functional-vec-Prop n x fv =
- Π-Prop (Fin n) λ k → leq-Decidable-Total-Order-Prop X x (fv k)
+ Π-Prop (Fin n) (λ k → leq-Decidable-Total-Order-Prop X x (fv k))
is-least-element-functional-vec :
(n : ℕ)
@@ -217,10 +217,10 @@ module _
is-least-element-vec-is-least-element-functional-vec (succ-ℕ n) x fv p =
(p (inr star)) ,
( is-least-element-vec-is-least-element-functional-vec
- ( n)
- ( x)
- ( tail-functional-vec n fv)
- ( p ∘ inl))
+ ( n)
+ ( x)
+ ( tail-functional-vec n fv)
+ ( p ∘ inl))
is-least-element-functional-vec-is-least-element-vec :
(n : ℕ)
@@ -256,9 +256,9 @@ module _
( x)
( functional-vec-vec n v ∘ map-equiv a)
( is-least-element-permute-functional-vec
- ( n)
- ( x)
- ( functional-vec-vec n v)
- ( a)
- ( is-least-element-functional-vec-is-least-element-vec n x v p))
+ ( n)
+ ( x)
+ ( functional-vec-vec n v)
+ ( a)
+ ( is-least-element-functional-vec-is-least-element-vec n x v p))
```
diff --git a/src/order-theory/directed-families.lagda.md b/src/order-theory/directed-families.lagda.md
index ec0e563f2a..e82194883e 100644
--- a/src/order-theory/directed-families.lagda.md
+++ b/src/order-theory/directed-families.lagda.md
@@ -8,10 +8,12 @@ module order-theory.directed-families where
```agda
open import foundation.cartesian-product-types
+open import foundation.conjunction
open import foundation.dependent-pair-types
open import foundation.existential-quantification
open import foundation.inhabited-types
open import foundation.propositions
+open import foundation.universal-quantification
open import foundation.universe-levels
open import order-theory.posets
@@ -32,15 +34,15 @@ is-directed-family-Poset-Prop :
{l1 l2 l3 : Level} (P : Poset l1 l2) (I : Inhabited-Type l3)
(x : type-Inhabited-Type I → type-Poset P) → Prop (l2 ⊔ l3)
is-directed-family-Poset-Prop P I x =
- Π-Prop
+ ∀'
( type-Inhabited-Type I)
( λ i →
- Π-Prop
+ ∀'
( type-Inhabited-Type I)
( λ j →
- ∃-Prop
- ( type-Inhabited-Type I)
- ( λ k → leq-Poset P (x i) (x k) × leq-Poset P (x j) (x k))))
+ ∃ ( type-Inhabited-Type I)
+ ( λ k →
+ leq-Poset-Prop P (x i) (x k) ∧ leq-Poset-Prop P (x j) (x k))))
is-directed-family-Poset :
{l1 l2 l3 : Level} (P : Poset l1 l2) (I : Inhabited-Type l3)
diff --git a/src/order-theory/galois-connections.lagda.md b/src/order-theory/galois-connections.lagda.md
index 111bc3090a..166d4d22a0 100644
--- a/src/order-theory/galois-connections.lagda.md
+++ b/src/order-theory/galois-connections.lagda.md
@@ -390,7 +390,7 @@ module _
(G H : Galois-Connection P Q) →
is-equiv (htpy-upper-adjoint-htpy-lower-adjoint-Galois-Connection G H)
is-equiv-htpy-upper-adjoint-htpy-lower-adjoint-Galois-Connection G H =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-htpy-lower-adjoint-Galois-Connection G H)
( is-prop-htpy-upper-adjoint-Galois-Connection P Q G H)
( htpy-lower-adjoint-htpy-upper-adjoint-Galois-Connection G H)
@@ -399,7 +399,7 @@ module _
(G H : Galois-Connection P Q) →
is-equiv (htpy-lower-adjoint-htpy-upper-adjoint-Galois-Connection G H)
is-equiv-htpy-lower-adjoint-htpy-upper-adjoint-Galois-Connection G H =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-htpy-upper-adjoint-Galois-Connection P Q G H)
( is-prop-htpy-lower-adjoint-Galois-Connection G H)
( htpy-upper-adjoint-htpy-lower-adjoint-Galois-Connection G H)
diff --git a/src/order-theory/greatest-lower-bounds-posets.lagda.md b/src/order-theory/greatest-lower-bounds-posets.lagda.md
index 8a1f391c47..c6afa32f62 100644
--- a/src/order-theory/greatest-lower-bounds-posets.lagda.md
+++ b/src/order-theory/greatest-lower-bounds-posets.lagda.md
@@ -76,8 +76,8 @@ module _
{x y : type-Poset P} →
is-greatest-binary-lower-bound-Poset P a b x →
is-binary-lower-bound-Poset P a b y → leq-Poset P y x
- forward-implication-is-greatest-binary-lower-bound-Poset H =
- forward-implication (H _)
+ forward-implication-is-greatest-binary-lower-bound-Poset {x} {y} H =
+ forward-implication (H y)
backward-implication-is-greatest-binary-lower-bound-Poset :
{x y : type-Poset P} →
@@ -100,9 +100,9 @@ module _
{x : type-Poset P} →
is-greatest-binary-lower-bound-Poset P a b x →
is-binary-lower-bound-Poset P a b x
- is-binary-lower-bound-is-greatest-binary-lower-bound-Poset H =
+ is-binary-lower-bound-is-greatest-binary-lower-bound-Poset {x} H =
backward-implication-is-greatest-binary-lower-bound-Poset H
- ( refl-leq-Poset P _)
+ ( refl-leq-Poset P x)
leq-left-is-greatest-binary-lower-bound-Poset :
{x : type-Poset P} →
@@ -209,23 +209,25 @@ module _
{x y : type-Poset P} →
is-greatest-lower-bound-family-of-elements-Poset P a x →
((i : I) → leq-Poset P y (a i)) → leq-Poset P y x
- forward-implication-is-greatest-lower-bound-family-of-elements-Poset H =
- forward-implication (H _)
+ forward-implication-is-greatest-lower-bound-family-of-elements-Poset
+ { x} {y} H =
+ forward-implication (H y)
backward-implication-is-greatest-lower-bound-family-of-elements-Poset :
{x y : type-Poset P} →
is-greatest-lower-bound-family-of-elements-Poset P a x →
leq-Poset P y x → (i : I) → leq-Poset P y (a i)
- backward-implication-is-greatest-lower-bound-family-of-elements-Poset H =
- backward-implication (H _)
+ backward-implication-is-greatest-lower-bound-family-of-elements-Poset
+ {x} {y} H =
+ backward-implication (H y)
is-lower-bound-is-greatest-lower-bound-family-of-elements-Poset :
{x : type-Poset P} →
is-greatest-lower-bound-family-of-elements-Poset P a x →
is-lower-bound-family-of-elements-Poset P a x
- is-lower-bound-is-greatest-lower-bound-family-of-elements-Poset H =
+ is-lower-bound-is-greatest-lower-bound-family-of-elements-Poset {x} H =
backward-implication-is-greatest-lower-bound-family-of-elements-Poset H
- ( refl-leq-Poset P _)
+ ( refl-leq-Poset P x)
```
### The proposition that a family of elements has a greatest lower bound
@@ -280,12 +282,12 @@ module _
is-greatest-lower-bound-family-of-elements-Poset P a x →
is-greatest-lower-bound-family-of-elements-Poset P a y →
x = y
- eq-is-greatest-lower-bound-family-of-elements-Poset H K =
+ eq-is-greatest-lower-bound-family-of-elements-Poset {x} {y} H K =
ap
( pr1)
( all-elements-equal-has-greatest-lower-bound-family-of-elements-Poset
( P)
( a)
- ( _ , H)
- ( _ , K))
+ ( x , H)
+ ( y , K))
```
diff --git a/src/order-theory/inhabited-finite-total-orders.lagda.md b/src/order-theory/inhabited-finite-total-orders.lagda.md
index d9908e0043..1a8d6c8968 100644
--- a/src/order-theory/inhabited-finite-total-orders.lagda.md
+++ b/src/order-theory/inhabited-finite-total-orders.lagda.md
@@ -38,8 +38,8 @@ module _
is-inhabited-Total-Order-𝔽 : UU (l1 ⊔ l2)
is-inhabited-Total-Order-𝔽 = is-finite-Poset (poset-Total-Order-𝔽 P)
- is-prop-is-inhabited-Total-Order-𝔽 : is-prop is-inhabited-Total-Order-𝔽
- is-prop-is-inhabited-Total-Order-𝔽 =
+ is-property-is-inhabited-Total-Order-𝔽 : is-prop is-inhabited-Total-Order-𝔽
+ is-property-is-inhabited-Total-Order-𝔽 =
is-prop-is-finite-Poset (poset-Total-Order-𝔽 P)
is-finite-type-is-inhabited-Total-Order-𝔽 :
diff --git a/src/order-theory/large-posets.lagda.md b/src/order-theory/large-posets.lagda.md
index 8f6800de89..ebe0411daf 100644
--- a/src/order-theory/large-posets.lagda.md
+++ b/src/order-theory/large-posets.lagda.md
@@ -16,6 +16,7 @@ open import category-theory.precategories
open import foundation.dependent-pair-types
open import foundation.identity-types
open import foundation.large-binary-relations
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.universe-levels
@@ -172,7 +173,7 @@ module _
is-large-category-Large-Poset :
is-large-category-Large-Precategory large-precategory-Large-Poset
is-large-category-Large-Poset {l} x y =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-type-Large-Poset P x y)
( is-prop-iso-is-prop-hom-Precategory
( precategory-Large-Poset l)
diff --git a/src/order-theory/least-upper-bounds-posets.lagda.md b/src/order-theory/least-upper-bounds-posets.lagda.md
index 9052b87c52..7bddf850d8 100644
--- a/src/order-theory/least-upper-bounds-posets.lagda.md
+++ b/src/order-theory/least-upper-bounds-posets.lagda.md
@@ -76,8 +76,8 @@ module _
{x y : type-Poset P} →
is-least-binary-upper-bound-Poset P a b x →
is-binary-upper-bound-Poset P a b y → leq-Poset P x y
- forward-implication-is-least-binary-upper-bound-Poset H =
- forward-implication (H _)
+ forward-implication-is-least-binary-upper-bound-Poset {x} {y} H =
+ forward-implication (H y)
backward-implication-is-least-binary-upper-bound-Poset :
{x y : type-Poset P} →
@@ -100,9 +100,9 @@ module _
{x : type-Poset P} →
is-least-binary-upper-bound-Poset P a b x →
is-binary-upper-bound-Poset P a b x
- is-binary-upper-bound-is-least-binary-upper-bound-Poset H =
+ is-binary-upper-bound-is-least-binary-upper-bound-Poset {x} H =
backward-implication-is-least-binary-upper-bound-Poset H
- ( refl-leq-Poset P _)
+ ( refl-leq-Poset P x)
leq-left-is-least-binary-upper-bound-Poset :
{x : type-Poset P} →
@@ -209,23 +209,23 @@ module _
{x y : type-Poset P} →
is-least-upper-bound-family-of-elements-Poset P a x →
is-upper-bound-family-of-elements-Poset P a y → leq-Poset P x y
- forward-implication-is-least-upper-bound-family-of-elements-Poset H =
- forward-implication (H _)
+ forward-implication-is-least-upper-bound-family-of-elements-Poset {x} {y} H =
+ forward-implication (H y)
backward-implication-is-least-upper-bound-family-of-elements-Poset :
{x y : type-Poset P} →
is-least-upper-bound-family-of-elements-Poset P a x →
leq-Poset P x y → is-upper-bound-family-of-elements-Poset P a y
- backward-implication-is-least-upper-bound-family-of-elements-Poset H =
- backward-implication (H _)
+ backward-implication-is-least-upper-bound-family-of-elements-Poset {x} {y} H =
+ backward-implication (H y)
is-upper-bound-is-least-upper-bound-family-of-elements-Poset :
{x : type-Poset P} →
is-least-upper-bound-family-of-elements-Poset P a x →
is-upper-bound-family-of-elements-Poset P a x
- is-upper-bound-is-least-upper-bound-family-of-elements-Poset H =
+ is-upper-bound-is-least-upper-bound-family-of-elements-Poset {x} H =
backward-implication-is-least-upper-bound-family-of-elements-Poset H
- ( refl-leq-Poset P _)
+ ( refl-leq-Poset P x)
```
### The proposition that a family of elements has a least upper bound
@@ -280,12 +280,12 @@ module _
is-least-upper-bound-family-of-elements-Poset P a x →
is-least-upper-bound-family-of-elements-Poset P a y →
x = y
- eq-is-least-upper-bound-family-of-elements-Poset H K =
+ eq-is-least-upper-bound-family-of-elements-Poset {x} {y} H K =
ap
( pr1)
( all-elements-equal-has-least-upper-bound-family-of-elements-Poset
( P)
( a)
- ( _ , H)
- ( _ , K))
+ ( x , H)
+ ( y , K))
```
diff --git a/src/order-theory/posets.lagda.md b/src/order-theory/posets.lagda.md
index 488d467a33..e3e5d5e91b 100644
--- a/src/order-theory/posets.lagda.md
+++ b/src/order-theory/posets.lagda.md
@@ -16,6 +16,7 @@ open import foundation.cartesian-product-types
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.sets
open import foundation.universe-levels
@@ -155,7 +156,7 @@ module _
is-category-precategory-Poset : is-category-Precategory precategory-Poset
is-category-precategory-Poset x y =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-type-Poset X x y)
( is-prop-iso-is-prop-hom-Precategory precategory-Poset
( is-prop-leq-Poset X x y))
diff --git a/src/order-theory/total-preorders.lagda.md b/src/order-theory/total-preorders.lagda.md
index 1c744e1e93..fdc75e0459 100644
--- a/src/order-theory/total-preorders.lagda.md
+++ b/src/order-theory/total-preorders.lagda.md
@@ -36,7 +36,7 @@ module _
incident-Preorder-Prop : (x y : type-Preorder X) → Prop l2
incident-Preorder-Prop x y =
- disjunction-Prop (leq-Preorder-Prop X x y) (leq-Preorder-Prop X y x)
+ (leq-Preorder-Prop X x y) ∨ (leq-Preorder-Prop X y x)
incident-Preorder : (x y : type-Preorder X) → UU l2
incident-Preorder x y = type-Prop (incident-Preorder-Prop x y)
diff --git a/src/organic-chemistry/alkanes.lagda.md b/src/organic-chemistry/alkanes.lagda.md
index 807e7c425a..ae5bc42c05 100644
--- a/src/organic-chemistry/alkanes.lagda.md
+++ b/src/organic-chemistry/alkanes.lagda.md
@@ -24,5 +24,6 @@ not have any double or triple carbon-carbon bonds.
```agda
is-alkane-hydrocarbon : {l1 l2 : Level} → hydrocarbon l1 l2 → UU (l1 ⊔ l2)
-is-alkane-hydrocarbon H = ∀ c → is-saturated-carbon-hydrocarbon H c
+is-alkane-hydrocarbon H =
+ (c : vertex-hydrocarbon H) → is-saturated-carbon-hydrocarbon H c
```
diff --git a/src/organic-chemistry/ethane.lagda.md b/src/organic-chemistry/ethane.lagda.md
index c795c7937b..43f02201ec 100644
--- a/src/organic-chemistry/ethane.lagda.md
+++ b/src/organic-chemistry/ethane.lagda.md
@@ -22,6 +22,7 @@ open import foundation.identity-types
open import foundation.injective-maps
open import foundation.propositional-truncations
open import foundation.propositions
+open import foundation.sets
open import foundation.torsorial-type-families
open import foundation.transport-along-identifications
open import foundation.unit-type
diff --git a/src/orthogonal-factorization-systems/factorizations-of-maps-function-classes.lagda.md b/src/orthogonal-factorization-systems/factorizations-of-maps-function-classes.lagda.md
index 0a65bb2a43..31bf354643 100644
--- a/src/orthogonal-factorization-systems/factorizations-of-maps-function-classes.lagda.md
+++ b/src/orthogonal-factorization-systems/factorizations-of-maps-function-classes.lagda.md
@@ -8,7 +8,6 @@ module orthogonal-factorization-systems.factorizations-of-maps-function-classes
```agda
open import foundation.action-on-identifications-functions
-open import foundation.conjunction
open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equivalences
@@ -65,7 +64,7 @@ module _
is-function-class-factorization-Prop : Prop (lL ⊔ lR)
is-function-class-factorization-Prop =
- conjunction-Prop
+ product-Prop
( L (left-map-factorization F))
( R (right-map-factorization F))
diff --git a/src/orthogonal-factorization-systems/factorizations-of-maps-global-function-classes.lagda.md b/src/orthogonal-factorization-systems/factorizations-of-maps-global-function-classes.lagda.md
index d3c3e3b691..49d6546998 100644
--- a/src/orthogonal-factorization-systems/factorizations-of-maps-global-function-classes.lagda.md
+++ b/src/orthogonal-factorization-systems/factorizations-of-maps-global-function-classes.lagda.md
@@ -8,7 +8,6 @@ module orthogonal-factorization-systems.factorizations-of-maps-global-function-c
```agda
open import foundation.action-on-identifications-functions
-open import foundation.conjunction
open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equivalences
diff --git a/src/orthogonal-factorization-systems/function-classes.lagda.md b/src/orthogonal-factorization-systems/function-classes.lagda.md
index cacff136e9..db1a440376 100644
--- a/src/orthogonal-factorization-systems/function-classes.lagda.md
+++ b/src/orthogonal-factorization-systems/function-classes.lagda.md
@@ -16,6 +16,7 @@ open import foundation.equivalences
open import foundation.function-types
open import foundation.identity-types
open import foundation.iterated-dependent-product-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.pullbacks
open import foundation.subtypes
@@ -233,7 +234,7 @@ module _
is-equiv-has-identity-maps-has-equivalences-function-class :
is-equiv has-identity-maps-has-equivalences-function-class
is-equiv-has-identity-maps-has-equivalences-function-class =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-has-equivalences-function-class P)
( is-prop-has-identity-maps-function-class P)
( has-equivalences-has-identity-maps-function-class)
@@ -248,7 +249,7 @@ module _
is-equiv-has-equivalences-has-identity-maps-function-class :
is-equiv has-equivalences-has-identity-maps-function-class
is-equiv-has-equivalences-has-identity-maps-function-class =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-has-identity-maps-function-class P)
( is-prop-has-equivalences-function-class P)
( has-identity-maps-has-equivalences-function-class)
diff --git a/src/orthogonal-factorization-systems/local-types.lagda.md b/src/orthogonal-factorization-systems/local-types.lagda.md
index 07c0e3ca8c..38b13c18ae 100644
--- a/src/orthogonal-factorization-systems/local-types.lagda.md
+++ b/src/orthogonal-factorization-systems/local-types.lagda.md
@@ -21,6 +21,7 @@ open import foundation.function-types
open import foundation.functoriality-dependent-function-types
open import foundation.homotopies
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.postcomposition-functions
open import foundation.precomposition-dependent-functions
open import foundation.precomposition-functions
@@ -276,7 +277,7 @@ module _
(((y : Y) → A (f y)) → ((x : X) → A x)) →
is-local-dependent-type f A
is-local-dependent-type-is-prop A is-prop-A =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-Π is-prop-A)
( is-prop-Π (is-prop-A ∘ f))
diff --git a/src/orthogonal-factorization-systems/null-types.lagda.md b/src/orthogonal-factorization-systems/null-types.lagda.md
index 7b75cf5b95..f2cc5610e0 100644
--- a/src/orthogonal-factorization-systems/null-types.lagda.md
+++ b/src/orthogonal-factorization-systems/null-types.lagda.md
@@ -11,6 +11,7 @@ open import foundation.constant-maps
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.fibers-of-maps
+open import foundation.logical-equivalences
open import foundation.precomposition-functions
open import foundation.propositions
open import foundation.type-arithmetic-unit-type
@@ -77,7 +78,7 @@ module _
equiv-is-local-is-null : is-null Y A ≃ is-local (λ y → star) A
equiv-is-local-is-null =
- equiv-prop
+ equiv-iff-is-prop
( is-property-is-equiv (const Y A))
( is-property-is-equiv (precomp (λ y → star) A))
( is-local-is-null)
diff --git a/src/primitives/machine-integers.lagda.md b/src/primitives/machine-integers.lagda.md
index 7a285b84d7..84db3aa726 100644
--- a/src/primitives/machine-integers.lagda.md
+++ b/src/primitives/machine-integers.lagda.md
@@ -32,5 +32,5 @@ primitive
primWord64ToNat : Word64 → ℕ
primWord64FromNat : ℕ → Word64
primWord64ToNatInjective :
- ∀ a b → primWord64ToNat a = primWord64ToNat b → a = b
+ (a b : Word64) → primWord64ToNat a = primWord64ToNat b → a = b
```
diff --git a/src/real-numbers/dedekind-real-numbers.lagda.md b/src/real-numbers/dedekind-real-numbers.lagda.md
index 599a3c4926..0604f765e2 100644
--- a/src/real-numbers/dedekind-real-numbers.lagda.md
+++ b/src/real-numbers/dedekind-real-numbers.lagda.md
@@ -33,6 +33,7 @@ open import foundation.sets
open import foundation.subtypes
open import foundation.transport-along-identifications
open import foundation.truncated-types
+open import foundation.universal-quantification
open import foundation.universe-levels
open import foundation-core.truncation-levels
@@ -74,29 +75,15 @@ module _
is-dedekind-cut-Prop : Prop (l1 ⊔ l2)
is-dedekind-cut-Prop =
- product-Prop
- ( product-Prop (exists-Prop ℚ L) (exists-Prop ℚ U))
- ( product-Prop
- ( product-Prop
- ( Π-Prop ℚ
- ( λ q →
- iff-Prop
- ( L q)
- ( exists-Prop ℚ (λ r → product-Prop (le-ℚ-Prop q r) (L r)))))
- ( Π-Prop ℚ
- ( λ r →
- iff-Prop
- ( U r)
- ( exists-Prop ℚ (λ q → product-Prop (le-ℚ-Prop q r) (U q))))))
- ( product-Prop
- ( Π-Prop ℚ (λ q → neg-Prop (product-Prop (L q) (U q))))
- ( Π-Prop ℚ
- ( λ q →
- Π-Prop ℚ
- ( λ r →
- implication-Prop
- ( le-ℚ-Prop q r)
- ( disjunction-Prop (L q) (U r)))))))
+ conjunction-Prop
+ ( (∃ ℚ L) ∧ (∃ ℚ U))
+ ( conjunction-Prop
+ ( conjunction-Prop
+ ( ∀' ℚ ( λ q → L q ⇔ ∃ ℚ (λ r → le-ℚ-Prop q r ∧ L r)))
+ ( ∀' ℚ ( λ r → U r ⇔ ∃ ℚ (λ q → le-ℚ-Prop q r ∧ U q))))
+ ( conjunction-Prop
+ ( ∀' ℚ (λ q → ¬' (L q ∧ U q)))
+ ( ∀' ℚ (λ q → ∀' ℚ (λ r → le-ℚ-Prop q r ⇒ (L q ∨ U r))))))
is-dedekind-cut : UU (l1 ⊔ l2)
is-dedekind-cut = type-Prop is-dedekind-cut-Prop
@@ -141,13 +128,15 @@ module _
is-rounded-lower-cut-ℝ :
(q : ℚ) →
- is-in-lower-cut-ℝ q ↔ ∃ ℚ (λ r → (le-ℚ q r) × (is-in-lower-cut-ℝ r))
+ is-in-lower-cut-ℝ q ↔
+ exists ℚ (λ r → (le-ℚ-Prop q r) ∧ (lower-cut-ℝ r))
is-rounded-lower-cut-ℝ =
pr1 (pr1 (pr2 is-dedekind-cut-cut-ℝ))
is-rounded-upper-cut-ℝ :
(r : ℚ) →
- is-in-upper-cut-ℝ r ↔ ∃ ℚ (λ q → (le-ℚ q r) × (is-in-upper-cut-ℝ q))
+ is-in-upper-cut-ℝ r ↔
+ exists ℚ (λ q → (le-ℚ-Prop q r) ∧ (upper-cut-ℝ q))
is-rounded-upper-cut-ℝ =
pr2 (pr1 (pr2 is-dedekind-cut-cut-ℝ))
@@ -156,7 +145,8 @@ module _
pr1 (pr2 (pr2 is-dedekind-cut-cut-ℝ))
is-located-lower-upper-cut-ℝ :
- (q r : ℚ) → le-ℚ q r → (lower-cut-ℝ q) ∨ (upper-cut-ℝ r)
+ (q r : ℚ) → le-ℚ q r →
+ type-disjunction-Prop (lower-cut-ℝ q) (upper-cut-ℝ r)
is-located-lower-upper-cut-ℝ =
pr2 (pr2 (pr2 is-dedekind-cut-cut-ℝ))
@@ -182,16 +172,15 @@ abstract
is-set-Σ
( is-set-function-type (is-trunc-Truncated-Type neg-one-𝕋))
( λ x →
- ( is-set-Σ
+ is-set-Σ
( is-set-function-type (is-trunc-Truncated-Type neg-one-𝕋))
( λ y →
( is-set-is-prop
( is-prop-type-Prop
- ( is-dedekind-cut-Prop x y))))))
+ ( is-dedekind-cut-Prop x y)))))
ℝ-Set : (l : Level) → Set (lsuc l)
-pr1 (ℝ-Set l) = ℝ l
-pr2 (ℝ-Set l) = is-set-ℝ l
+ℝ-Set l = ℝ l , is-set-ℝ l
```
## Properties of lower/upper Dedekind cuts
@@ -288,13 +277,11 @@ module _
is-lower-complement-upper-cut-ℝ-Prop : (p q : ℚ) → Prop l
is-lower-complement-upper-cut-ℝ-Prop p q =
- product-Prop
- ( le-ℚ-Prop p q)
- ( neg-Prop ( upper-cut-ℝ x q))
+ ( le-ℚ-Prop p q) ∧ (¬' (upper-cut-ℝ x q))
lower-complement-upper-cut-ℝ : subtype l ℚ
lower-complement-upper-cut-ℝ p =
- exists-Prop ℚ (is-lower-complement-upper-cut-ℝ-Prop p)
+ ∃ ℚ (is-lower-complement-upper-cut-ℝ-Prop p)
```
```agda
@@ -305,8 +292,7 @@ module _
subset-lower-cut-lower-complement-upper-cut-ℝ :
lower-complement-upper-cut-ℝ x ⊆ lower-cut-ℝ x
subset-lower-cut-lower-complement-upper-cut-ℝ p =
- elim-exists-Prop
- ( is-lower-complement-upper-cut-ℝ-Prop x p)
+ elim-exists
( lower-cut-ℝ x p)
( λ q I →
map-right-unit-law-disjunction-is-empty-Prop
@@ -318,12 +304,10 @@ module _
subset-lower-complement-upper-cut-lower-cut-ℝ :
lower-cut-ℝ x ⊆ lower-complement-upper-cut-ℝ x
subset-lower-complement-upper-cut-lower-cut-ℝ p H =
- elim-exists-Prop
- ( λ q → product-Prop (le-ℚ-Prop p q) (lower-cut-ℝ x q))
+ elim-exists
( lower-complement-upper-cut-ℝ x p)
( λ q I →
intro-exists
- ( is-lower-complement-upper-cut-ℝ-Prop x p)
( q)
( map-product
( id)
@@ -350,13 +334,11 @@ module _
is-upper-complement-lower-cut-ℝ-Prop : (q p : ℚ) → Prop l
is-upper-complement-lower-cut-ℝ-Prop q p =
- product-Prop
- ( le-ℚ-Prop p q)
- ( neg-Prop ( lower-cut-ℝ x p))
+ (le-ℚ-Prop p q) ∧ (¬' (lower-cut-ℝ x p))
upper-complement-lower-cut-ℝ : subtype l ℚ
upper-complement-lower-cut-ℝ q =
- exists-Prop ℚ (is-upper-complement-lower-cut-ℝ-Prop q)
+ ∃ ℚ (is-upper-complement-lower-cut-ℝ-Prop q)
```
```agda
@@ -367,8 +349,7 @@ module _
subset-upper-cut-upper-complement-lower-cut-ℝ :
upper-complement-lower-cut-ℝ x ⊆ upper-cut-ℝ x
subset-upper-cut-upper-complement-lower-cut-ℝ q =
- elim-exists-Prop
- ( is-upper-complement-lower-cut-ℝ-Prop x q)
+ elim-exists
( upper-cut-ℝ x q)
( λ p I →
map-left-unit-law-disjunction-is-empty-Prop
@@ -380,12 +361,10 @@ module _
subset-upper-complement-lower-cut-upper-cut-ℝ :
upper-cut-ℝ x ⊆ upper-complement-lower-cut-ℝ x
subset-upper-complement-lower-cut-upper-cut-ℝ q H =
- elim-exists-Prop
- ( λ p → product-Prop (le-ℚ-Prop p q) (upper-cut-ℝ x p))
+ elim-exists
( upper-complement-lower-cut-ℝ x q)
( λ p I →
intro-exists
- ( is-upper-complement-lower-cut-ℝ-Prop x q)
( p)
( map-product
( id)
@@ -418,10 +397,9 @@ module _
( eq-lower-cut-lower-complement-upper-cut-ℝ x)
( eq-lower-cut-lower-complement-upper-cut-ℝ y)
( λ p →
- elim-exists-Prop
- ( is-lower-complement-upper-cut-ℝ-Prop x p)
+ elim-exists
( lower-complement-upper-cut-ℝ y p)
- ( λ q → intro-∃ q ∘ tot (λ _ K → K ∘ H q)))
+ ( λ q → intro-exists q ∘ tot (λ _ K → K ∘ H q)))
subset-upper-cut-lower-cut-ℝ :
lower-cut-ℝ x ⊆ lower-cut-ℝ y → upper-cut-ℝ y ⊆ upper-cut-ℝ x
@@ -431,10 +409,9 @@ module _
( eq-upper-cut-upper-complement-lower-cut-ℝ y)
( eq-upper-cut-upper-complement-lower-cut-ℝ x)
( λ q →
- elim-exists-Prop
- ( is-upper-complement-lower-cut-ℝ-Prop y q)
+ elim-exists
( upper-complement-lower-cut-ℝ x q)
- ( λ p → intro-∃ p ∘ tot (λ _ K → K ∘ H p)))
+ ( λ p → intro-exists p ∘ tot (λ _ K → K ∘ H p)))
module _
{l : Level} (x y : ℝ l)
diff --git a/src/real-numbers/rational-real-numbers.lagda.md b/src/real-numbers/rational-real-numbers.lagda.md
index eb34a051e8..777cecf09b 100644
--- a/src/real-numbers/rational-real-numbers.lagda.md
+++ b/src/real-numbers/rational-real-numbers.lagda.md
@@ -13,6 +13,7 @@ open import elementary-number-theory.strict-inequality-rational-numbers
open import foundation.action-on-identifications-functions
open import foundation.cartesian-product-types
+open import foundation.conjunction
open import foundation.dependent-pair-types
open import foundation.disjunction
open import foundation.embeddings
@@ -56,17 +57,15 @@ is-dedekind-cut-le-ℚ :
(λ (q : ℚ) → le-ℚ-Prop q x)
(λ (r : ℚ) → le-ℚ-Prop x r)
is-dedekind-cut-le-ℚ x =
- ( left-∃-le-ℚ x , right-∃-le-ℚ x) ,
+ ( exists-lesser-ℚ x , exists-greater-ℚ x) ,
( ( λ (q : ℚ) →
dense-le-ℚ q x ,
- elim-exists-Prop
- ( λ r → product-Prop ( le-ℚ-Prop q r) ( le-ℚ-Prop r x))
+ elim-exists
( le-ℚ-Prop q x)
( λ r (H , H') → transitive-le-ℚ q r x H' H)) ,
( λ (r : ℚ) →
α x r ∘ dense-le-ℚ x r ,
- elim-exists-Prop
- ( λ q → product-Prop ( le-ℚ-Prop q r) ( le-ℚ-Prop x q))
+ elim-exists
( le-ℚ-Prop x r)
( λ q (H , H') → transitive-le-ℚ x q r H H'))) ,
( λ (q : ℚ) (H , H') → asymmetric-le-ℚ q x H H') ,
@@ -74,20 +73,12 @@ is-dedekind-cut-le-ℚ x =
where
α :
(a b : ℚ) →
- ∃ ℚ (λ r → le-ℚ a r × le-ℚ r b) →
- ∃ ℚ (λ r → le-ℚ r b × le-ℚ a r)
+ exists ℚ (λ r → le-ℚ-Prop a r ∧ le-ℚ-Prop r b) →
+ exists ℚ (λ r → le-ℚ-Prop r b ∧ le-ℚ-Prop a r)
α a b =
- elim-exists-Prop
- ( ( λ r →
- product-Prop
- ( le-ℚ-Prop a r)
- ( le-ℚ-Prop r b)))
- ( exists-Prop ℚ
- ( λ r →
- product-Prop
- ( le-ℚ-Prop r b)
- ( le-ℚ-Prop a r)))
- ( λ r ( p , q) → intro-∃ r ( q , p))
+ elim-exists
+ ( ∃ ℚ (λ r → le-ℚ-Prop r b ∧ le-ℚ-Prop a r))
+ ( λ r ( p , q) → intro-exists r ( q , p))
```
### The canonical map from `ℚ` to `ℝ`
@@ -110,9 +101,7 @@ module _
is-rational-ℝ-Prop : Prop l
is-rational-ℝ-Prop =
- product-Prop
- ( neg-Prop (lower-cut-ℝ x p))
- ( neg-Prop (upper-cut-ℝ x p))
+ (¬' (lower-cut-ℝ x p)) ∧ (¬' (upper-cut-ℝ x p))
is-rational-ℝ : UU l
is-rational-ℝ = type-Prop is-rational-ℝ-Prop
@@ -130,21 +119,19 @@ all-eq-is-rational-ℝ x p q H H' =
left-case : le-ℚ p q → p = q
left-case I =
ex-falso
- ( elim-disjunction-Prop
- ( lower-cut-ℝ x p)
- ( upper-cut-ℝ x q)
+ ( elim-disjunction
( empty-Prop)
- ( pr1 H , pr2 H')
+ ( pr1 H)
+ ( pr2 H')
( is-located-lower-upper-cut-ℝ x p q I))
right-case : le-ℚ q p → p = q
right-case I =
ex-falso
- ( elim-disjunction-Prop
- ( lower-cut-ℝ x q)
- ( upper-cut-ℝ x p)
+ ( elim-disjunction
( empty-Prop)
- ( pr1 H' , pr2 H)
+ ( pr1 H')
+ ( pr2 H)
( is-located-lower-upper-cut-ℝ x q p I))
is-prop-rational-real : {l : Level} (x : ℝ l) → is-prop (Σ ℚ (is-rational-ℝ x))
@@ -204,12 +191,12 @@ eq-real-rational-is-rational-ℝ x q H =
( lower-cut-ℝ x)
( λ r →
pair
- ( λ I → elim-disjunction-Prop
- ( lower-cut-ℝ x r)
- ( upper-cut-ℝ x q)
- ( lower-cut-ℝ x r)
- ( id , λ H' → ex-falso (pr2 H H'))
- ( is-located-lower-upper-cut-ℝ x r q I))
+ ( λ I →
+ elim-disjunction
+ ( lower-cut-ℝ x r)
+ ( id)
+ ( λ H' → ex-falso (pr2 H H'))
+ ( is-located-lower-upper-cut-ℝ x r q I))
( trichotomy-le-ℚ r q
( λ I _ → I)
( λ E H' → ex-falso (pr1 (tr (is-rational-ℝ x) (inv E) H) H'))
diff --git a/src/ring-theory/nilpotent-elements-rings.lagda.md b/src/ring-theory/nilpotent-elements-rings.lagda.md
index 1c96f4d478..96f8512935 100644
--- a/src/ring-theory/nilpotent-elements-rings.lagda.md
+++ b/src/ring-theory/nilpotent-elements-rings.lagda.md
@@ -79,7 +79,7 @@ is-nilpotent-element-neg-Ring R {x} H =
apply-universal-property-trunc-Prop H
( is-nilpotent-element-ring-Prop R (neg-Ring R x))
( λ (n , p) →
- intro-∃ n
+ intro-exists n
( ( power-neg-Ring R n x) ∙
( ( ap (mul-Ring R (power-Ring R n (neg-one-Ring R))) p) ∙
( right-zero-law-mul-Ring R (power-Ring R n (neg-one-Ring R))))))
diff --git a/src/ring-theory/nilpotent-elements-semirings.lagda.md b/src/ring-theory/nilpotent-elements-semirings.lagda.md
index 190f920d60..434bdf7fb8 100644
--- a/src/ring-theory/nilpotent-elements-semirings.lagda.md
+++ b/src/ring-theory/nilpotent-elements-semirings.lagda.md
@@ -37,7 +37,7 @@ number `n` such that `x^n = 0`.
is-nilpotent-element-semiring-Prop :
{l : Level} (R : Semiring l) → type-Semiring R → Prop l
is-nilpotent-element-semiring-Prop R x =
- ∃-Prop ℕ (λ n → power-Semiring R n x = zero-Semiring R)
+ exists-structure-Prop ℕ (λ n → power-Semiring R n x = zero-Semiring R)
is-nilpotent-element-Semiring :
{l : Level} (R : Semiring l) → type-Semiring R → UU l
@@ -59,7 +59,7 @@ is-prop-is-nilpotent-element-Semiring R x =
is-nilpotent-zero-Semiring :
{l : Level} (R : Semiring l) →
is-nilpotent-element-Semiring R (zero-Semiring R)
-is-nilpotent-zero-Semiring R = intro-∃ 1 refl
+is-nilpotent-zero-Semiring R = intro-exists 1 refl
```
### If `x` and `y` commute and are both nilpotent, then `x + y` is nilpotent
@@ -77,7 +77,7 @@ is-nilpotent-add-Semiring R x y H f h =
apply-universal-property-trunc-Prop h
( is-nilpotent-element-semiring-Prop R (add-Semiring R x y))
( λ (m , q) →
- intro-∃
+ intro-exists
( n +ℕ m)
( ( is-linear-combination-power-add-Semiring R n m x y H) ∙
( ( ap-add-Semiring R
@@ -104,7 +104,7 @@ module _
apply-universal-property-trunc-Prop f
( is-nilpotent-element-semiring-Prop R (mul-Semiring R x y))
( λ (n , p) →
- intro-∃ n
+ intro-exists n
( ( distributive-power-mul-Semiring R n H) ∙
( ( ap (mul-Semiring' R (power-Semiring R n y)) p) ∙
( left-zero-law-mul-Semiring R
diff --git a/src/set-theory/countable-sets.lagda.md b/src/set-theory/countable-sets.lagda.md
index 71c87d6103..38263b07f1 100644
--- a/src/set-theory/countable-sets.lagda.md
+++ b/src/set-theory/countable-sets.lagda.md
@@ -71,7 +71,8 @@ module _
is-surjective-map-enumeration E = pr2 E
is-countable-Prop : Prop l
- is-countable-Prop = ∃-Prop (ℕ → Maybe (type-Set X)) is-surjective
+ is-countable-Prop =
+ ∃ (ℕ → Maybe (type-Set X)) (is-surjective-Prop)
is-countable : UU l
is-countable = type-Prop (is-countable-Prop)
@@ -94,7 +95,7 @@ module _
is-countable-Prop' : Prop (lsuc l ⊔ l)
is-countable-Prop' =
- ∃-Prop
+ exists-structure-Prop
( decidable-subtype l ℕ)
( λ P → type-decidable-subtype P ↠ type-Set X)
@@ -113,7 +114,7 @@ surjective map `f : ℕ → X`. Let us call the latter as "directly countable".
```agda
is-directly-countable-Prop : {l : Level} → Set l → Prop l
is-directly-countable-Prop X =
- ∃-Prop (ℕ → type-Set X) is-surjective
+ ∃ (ℕ → type-Set X) (is-surjective-Prop)
is-directly-countable : {l : Level} → Set l → UU l
is-directly-countable X = type-Prop (is-directly-countable-Prop X)
diff --git a/src/set-theory/cumulative-hierarchy.lagda.md b/src/set-theory/cumulative-hierarchy.lagda.md
index b908befe92..78520e6978 100644
--- a/src/set-theory/cumulative-hierarchy.lagda.md
+++ b/src/set-theory/cumulative-hierarchy.lagda.md
@@ -51,7 +51,7 @@ has-smaller-image :
{A : UU l1} {B : UU l2} {C : UU l3} →
(A → C) → (B → C) → UU (l1 ⊔ l2 ⊔ l3)
has-smaller-image {l1} {l2} {l3} {A} {B} {C} f g =
- (a : A) → ∃ B (λ b → g b = f a)
+ (a : A) → exists-structure B (λ b → g b = f a)
has-same-image :
{ l1 l2 l3 : Level} →
@@ -124,9 +124,9 @@ module _
( e : has-same-image f g)
( IH₁ : (a : A) → P (f a))
( IH₂ : (b : B) → P (g b)) →
- ( ( a : A) → ∃ B (λ b → Σ (f a = g b)
+ ( ( a : A) → exists-structure B (λ b → Σ (f a = g b)
( λ p → tr P p (IH₁ a) = IH₂ b))) →
- ( ( b : B) → ∃ A (λ a →
+ ( ( b : B) → exists-structure A (λ a →
Σ (g b = f a) (λ p → tr P p (IH₂ b) = IH₁ a))) →
tr P (set-ext-pseudo-cumulative-hierarchy V f g e) (ρ f IH₁) = ρ g IH₂) →
(x : type-pseudo-cumulative-hierarchy V) → P x
@@ -145,9 +145,9 @@ module _
( e : has-same-image f g)
( IH₁ : (a : A) → P (f a))
( IH₂ : (b : B) → P (g b)) →
- ( (a : A) → ∃ B (λ b → Σ (f a = g b)
+ ( (a : A) → exists-structure B (λ b → Σ (f a = g b)
( λ p → tr P p (IH₁ a) = IH₂ b))) →
- ( (b : B) → ∃ A (λ a → Σ (g b = f a)
+ ( (b : B) → exists-structure A (λ a → Σ (g b = f a)
(λ p → tr P p (IH₂ b) = IH₁ a))) →
tr P (set-ext-pseudo-cumulative-hierarchy V f g e) (ρ f IH₁) = ρ g IH₂) →
{ A : UU l1} (f : A → type-pseudo-cumulative-hierarchy V)
@@ -192,7 +192,7 @@ module _
```agda
underlying-function-cumulative-hierarchy :
(v : type-pseudo-cumulative-hierarchy V) →
- ∃ ( UU l1)
+ exists-structure ( UU l1)
( λ A →
Σ ( A → type-pseudo-cumulative-hierarchy V)
( λ f → set-pseudo-cumulative-hierarchy V f = v))
@@ -251,8 +251,8 @@ module _
( e : has-same-image f g)
( IH₁ : A → X)
( IH₂ : B → X) →
- ( (a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
- ( (b : B) → ∃ A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
+ ( (a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( (b : B) → exists-structure A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
ρ f IH₁ = ρ g IH₂) →
type-pseudo-cumulative-hierarchy V → X
recursion-principle-cumulative-hierarchy {l2} X σ ρ τ =
@@ -262,9 +262,9 @@ module _
{ A B : UU l1} (f : A → pr1 V) (g : B → pr1 V)
( e : has-same-image f g)
( IH₁ : (a : A) → X) (IH₂ : (b : B) → X) →
- ( (a : A) → ∃ B (λ b → Σ (f a = g b)
+ ( (a : A) → exists-structure B (λ b → Σ (f a = g b)
(λ p → tr (λ _ → X) p (IH₁ a) = IH₂ b))) →
- ( (b : B) → ∃ A (λ a → Σ (g b = f a)
+ ( (b : B) → exists-structure A (λ a → Σ (g b = f a)
(λ p → tr (λ _ → X) p (IH₂ b) = IH₁ a))) →
tr (λ _ → X) (pr2 (pr2 (pr2 V)) f g e) (ρ f IH₁) = ρ g IH₂
τ' {A} {B} f g e IH₁ IH₂ hIH₁ hIH₂ =
@@ -279,13 +279,17 @@ module _
set-pseudo-cumulative-hierarchy V f
= set-pseudo-cumulative-hierarchy V g
path-f-g = set-ext-pseudo-cumulative-hierarchy V f g e
- hIH₁' : (a : A) → ∃ B (λ b → Σ (f a = g b) (λ _ → IH₁ a = IH₂ b))
+ hIH₁' :
+ (a : A) →
+ exists-structure B (λ b → Σ (f a = g b) (λ _ → IH₁ a = IH₂ b))
hIH₁' a =
map-trunc-Prop
( λ (b , p , q) →
( b , p , (inv (tr-constant-type-family p _) ∙ q)))
( hIH₁ a)
- hIH₂' : (b : B) → ∃ A (λ a → Σ (g b = f a) (λ _ → IH₂ b = IH₁ a))
+ hIH₂' :
+ (b : B) →
+ exists-structure A (λ a → Σ (g b = f a) (λ _ → IH₂ b = IH₁ a))
hIH₂' b =
map-trunc-Prop
( λ (a , p , q) →
@@ -301,8 +305,8 @@ module _
( e : has-same-image f g)
( IH₁ : A → X)
( IH₂ : B → X) →
- ( ( a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
- ( ( b : B) → ∃ A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
+ ( ( a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( ( b : B) → exists-structure A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
ρ f IH₁ = ρ g IH₂) →
{A : UU l1} → (f : A → type-pseudo-cumulative-hierarchy V) → (IH : A → X) →
recursion-principle-cumulative-hierarchy X σ ρ τ
@@ -315,9 +319,9 @@ module _
{ A B : UU l1} (f : A → pr1 V) (g : B → pr1 V)
( e : has-same-image f g)
( IH₁ : (a : A) → X) (IH₂ : (b : B) → X) →
- ( ( a : A) → ∃ B (λ b → Σ (f a = g b)
+ ( ( a : A) → exists-structure B (λ b → Σ (f a = g b)
( λ p → tr (λ _ → X) p (IH₁ a) = IH₂ b))) →
- ( ( b : B) → ∃ A (λ a → Σ (g b = f a)
+ ( ( b : B) → exists-structure A (λ a → Σ (g b = f a)
( λ p → tr (λ _ → X) p (IH₂ b) = IH₁ a))) →
tr (λ _ → X) (pr2 (pr2 (pr2 V)) f g e) (ρ f IH₁) = ρ g IH₂
τ' {A} {B} f g e IH₁ IH₂ hIH₁ hIH₂ =
@@ -332,13 +336,17 @@ module _
set-pseudo-cumulative-hierarchy V f
= set-pseudo-cumulative-hierarchy V g
path-f-g = set-ext-pseudo-cumulative-hierarchy V f g e
- hIH₁' : (a : A) → ∃ B (λ b → Σ (f a = g b) (λ _ → IH₁ a = IH₂ b))
+ hIH₁' :
+ (a : A) →
+ exists-structure B (λ b → Σ (f a = g b) (λ _ → IH₁ a = IH₂ b))
hIH₁' a =
map-trunc-Prop
( λ (b , p , q) →
( b , p , (inv (tr-constant-type-family p _) ∙ q)))
( hIH₁ a)
- hIH₂' : (b : B) → ∃ A (λ a → Σ (g b = f a) (λ _ → IH₂ b = IH₁ a))
+ hIH₂' :
+ (b : B) →
+ exists-structure A (λ a → Σ (g b = f a) (λ _ → IH₂ b = IH₁ a))
hIH₂' b =
map-trunc-Prop
( λ (a , p , q) →
@@ -359,7 +367,7 @@ A simplification of the recursion principle, when the codomain is `Prop l2`.
( e : has-smaller-image f g)
( IH₁ : A → Prop l2)
( IH₂ : B → Prop l2) →
- ( (a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( (a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
type-Prop (ρ f IH₁) → type-Prop (ρ g IH₂)) →
type-pseudo-cumulative-hierarchy V → Prop l2
prop-recursion-principle-cumulative-hierarchy {l2} ρ τ =
@@ -371,8 +379,8 @@ A simplification of the recursion principle, when the codomain is `Prop l2`.
( g : B → type-pseudo-cumulative-hierarchy V) →
( e : has-same-image f g)
( IH₁ : A → Prop l2) (IH₂ : B → Prop l2) →
- ( (a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
- ( (b : B) → ∃ A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
+ ( (a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( (b : B) → exists-structure A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
ρ f IH₁ = ρ g IH₂
τ' f g (e₁ , e₂) IH₁ IH₂ hIH₁ hIH₂ =
eq-iff (τ f g e₁ IH₁ IH₂ hIH₁) (τ g f e₂ IH₂ IH₁ hIH₂)
@@ -388,7 +396,7 @@ A simplification of the recursion principle, when the codomain is `Prop l2`.
( e : has-smaller-image f g)
( IH₁ : A → Prop l2)
( IH₂ : B → Prop l2) →
- ( (a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( (a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
type-Prop (ρ f IH₁) → type-Prop (ρ g IH₂)) →
{ A : UU l1} → (f : A → type-pseudo-cumulative-hierarchy V) →
( IH : A → Prop l2) →
@@ -403,8 +411,8 @@ A simplification of the recursion principle, when the codomain is `Prop l2`.
( g : B → type-pseudo-cumulative-hierarchy V) →
( e : has-same-image f g)
( IH₁ : A → Prop l2) (IH₂ : B → Prop l2) →
- ( (a : A) → ∃ B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
- ( (b : B) → ∃ A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
+ ( (a : A) → exists-structure B (λ b → (f a = g b) × (IH₁ a = IH₂ b))) →
+ ( (b : B) → exists-structure A (λ a → (g b = f a) × (IH₂ b = IH₁ a))) →
ρ f IH₁ = ρ g IH₂
τ' f g (e₁ , e₂) IH₁ IH₂ hIH₁ hIH₂ =
eq-iff (τ f g e₁ IH₁ IH₂ hIH₁) (τ g f e₂ IH₂ IH₁ hIH₂)
@@ -452,16 +460,18 @@ needed.
Prop (lsuc l1)
∈-cumulative-hierarchy-Prop x =
simple-prop-recursion-principle-cumulative-hierarchy
- (λ {A} f → (∃ A (λ a → f a = x)) , is-prop-∃ _ _) e
+ ( λ {A} f → exists-structure-Prop A (λ a → f a = x))
+ ( e)
where
e :
{A B : UU l1} (f : A → type-pseudo-cumulative-hierarchy V)
( g : B → type-pseudo-cumulative-hierarchy V) →
( e : has-smaller-image f g) →
- ( ∃ A (λ a → f a = x) → ∃ B (λ b → g b = x))
+ ( exists-structure A (λ a → f a = x) →
+ exists-structure B (λ b → g b = x))
e {A} {B} f g s =
map-universal-property-trunc-Prop
- ( ∃-Prop B (λ b → g b = x))
+ ( exists-structure-Prop B (λ b → g b = x))
( λ ( a , p) →
map-trunc-Prop (λ (b , q) → (b , q ∙ p)) (s a))
@@ -476,7 +486,7 @@ needed.
( x : type-pseudo-cumulative-hierarchy V) {A : UU l1}
( f : A → type-pseudo-cumulative-hierarchy V) →
( ∈-cumulative-hierarchy x (set-pseudo-cumulative-hierarchy V f)) =
- ∃ A (λ a → f a = x)
+ exists-structure A (λ a → f a = x)
id-∈-cumulative-hierarchy x f =
ap pr1 (compute-simple-prop-recursion-principle-cumulative-hierarchy _ _ f)
@@ -485,7 +495,7 @@ needed.
{ A : UU l1}
{ f : A → type-pseudo-cumulative-hierarchy V} →
( ∈-cumulative-hierarchy x (set-pseudo-cumulative-hierarchy V f)) →
- ∃ A (λ a → f a = x)
+ exists-structure A (λ a → f a = x)
∈-cumulative-hierarchy-mere-preimage {x} {A} {f} =
tr id (id-∈-cumulative-hierarchy x f)
@@ -493,7 +503,7 @@ needed.
{ x : type-pseudo-cumulative-hierarchy V} →
{ A : UU l1}
{ f : A → type-pseudo-cumulative-hierarchy V} →
- ∃ A (λ a → f a = x) →
+ exists-structure A (λ a → f a = x) →
( ∈-cumulative-hierarchy x (set-pseudo-cumulative-hierarchy V f))
mere-preimage-∈-cumulative-hierarchy {x} {A} {f} =
tr id (inv (id-∈-cumulative-hierarchy x f))
@@ -549,7 +559,8 @@ needed.
( set-pseudo-cumulative-hierarchy V g)
has-smaller-image-⊆-cumulative-hierarchy {A} {B} f g s x m =
mere-preimage-∈-cumulative-hierarchy
- ( map-universal-property-trunc-Prop (∃-Prop B (λ b → g b = x))
+ ( map-universal-property-trunc-Prop
+ ( exists-structure-Prop B (λ b → g b = x))
( λ (a , p) →
map-trunc-Prop (λ (b , q) → (b , q ∙ p)) (s a))
( ∈-cumulative-hierarchy-mere-preimage m))
@@ -723,15 +734,17 @@ needed.
( x : type-pseudo-cumulative-hierarchy V) →
( r : type-pseudo-cumulative-hierarchy V →
type-pseudo-cumulative-hierarchy V) →
- ∃ ( type-pseudo-cumulative-hierarchy V)
+ exists-structure
+ ( type-pseudo-cumulative-hierarchy V)
( λ v →
( y : type-pseudo-cumulative-hierarchy V) →
∈-cumulative-hierarchy y v ↔
- ∃ ( type-pseudo-cumulative-hierarchy V)
+ exists-structure
+ ( type-pseudo-cumulative-hierarchy V)
( λ z → (∈-cumulative-hierarchy z x) × (y = r z)))
replacement-cumulative-hierarchy x r =
map-universal-property-trunc-Prop
- ( ∃-Prop (type-pseudo-cumulative-hierarchy V) _)
+ ( exists-structure-Prop (type-pseudo-cumulative-hierarchy V) _)
( λ where
( A , f , refl) →
unit-trunc-Prop
@@ -750,7 +763,7 @@ needed.
( λ H →
mere-preimage-∈-cumulative-hierarchy
( map-universal-property-trunc-Prop
- ( ∃-Prop A _)
+ ( exists-structure-Prop A _)
( λ where
( z , K , refl) →
map-trunc-Prop
diff --git a/src/structured-types/pointed-isomorphisms.lagda.md b/src/structured-types/pointed-isomorphisms.lagda.md
index 8c97d7ac80..f38534a262 100644
--- a/src/structured-types/pointed-isomorphisms.lagda.md
+++ b/src/structured-types/pointed-isomorphisms.lagda.md
@@ -13,6 +13,7 @@ open import foundation.contractible-types
open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.retractions
open import foundation.sections
@@ -379,7 +380,7 @@ module _
pr1 equiv-is-pointed-iso-is-pointed-equiv =
is-pointed-iso-is-pointed-equiv f
pr2 equiv-is-pointed-iso-is-pointed-equiv =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-property-is-equiv (map-pointed-map f))
( is-prop-is-pointed-iso f)
( is-pointed-equiv-is-pointed-iso f)
diff --git a/src/structured-types/small-pointed-types.lagda.md b/src/structured-types/small-pointed-types.lagda.md
index 11bbfa1986..c1ee5d36fa 100644
--- a/src/structured-types/small-pointed-types.lagda.md
+++ b/src/structured-types/small-pointed-types.lagda.md
@@ -112,7 +112,7 @@ module _
is-prop-is-pointed-small-Pointed-Type :
is-prop (is-pointed-small-Pointed-Type l2 X)
is-prop-is-pointed-small-Pointed-Type =
- is-prop-is-inhabited
+ is-prop-has-element
( λ (Y , e) →
is-prop-equiv'
( equiv-tot (λ Z → equiv-comp-pointed-equiv' e))
diff --git a/src/synthetic-homotopy-theory/joins-of-types.lagda.md b/src/synthetic-homotopy-theory/joins-of-types.lagda.md
index 6d1e096c00..36a860e028 100644
--- a/src/synthetic-homotopy-theory/joins-of-types.lagda.md
+++ b/src/synthetic-homotopy-theory/joins-of-types.lagda.md
@@ -17,6 +17,7 @@ open import foundation.equivalences
open import foundation.function-types
open import foundation.homotopies
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositions
open import foundation.type-arithmetic-cartesian-product-types
open import foundation.type-arithmetic-dependent-pair-types
@@ -318,13 +319,13 @@ module _
where
cocone-disjunction : cocone pr1 pr2 (type-disjunction-Prop A B)
- pr1 cocone-disjunction = inl-disjunction-Prop A B
- pr1 (pr2 cocone-disjunction) = inr-disjunction-Prop A B
+ pr1 cocone-disjunction = inl-disjunction
+ pr1 (pr2 cocone-disjunction) = inr-disjunction
pr2 (pr2 cocone-disjunction) (a , b) =
eq-is-prop'
- ( is-prop-type-disjunction-Prop A B)
- ( inl-disjunction-Prop A B a)
- ( inr-disjunction-Prop A B b)
+ ( is-prop-disjunction-Prop A B)
+ ( inl-disjunction a)
+ ( inr-disjunction b)
map-disjunction-join-Prop : type-join-Prop A B → type-disjunction-Prop A B
map-disjunction-join-Prop =
@@ -332,15 +333,16 @@ module _
map-join-disjunction-Prop : type-disjunction-Prop A B → type-join-Prop A B
map-join-disjunction-Prop =
- elim-disjunction-Prop A B
+ elim-disjunction
( join-Prop A B)
- ( inl-join-Prop A B , inr-join-Prop A B)
+ ( inl-join-Prop A B)
+ ( inr-join-Prop A B)
is-equiv-map-disjunction-join-Prop : is-equiv map-disjunction-join-Prop
is-equiv-map-disjunction-join-Prop =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-join-Prop A B)
- ( is-prop-type-disjunction-Prop A B)
+ ( is-prop-disjunction-Prop A B)
( map-join-disjunction-Prop)
equiv-disjunction-join-Prop :
@@ -350,8 +352,8 @@ module _
is-equiv-map-join-disjunction-Prop : is-equiv map-join-disjunction-Prop
is-equiv-map-join-disjunction-Prop =
- is-equiv-is-prop
- ( is-prop-type-disjunction-Prop A B)
+ is-equiv-has-converse-is-prop
+ ( is-prop-disjunction-Prop A B)
( is-prop-type-join-Prop A B)
( map-disjunction-join-Prop)
@@ -369,10 +371,10 @@ module _
( cocone-join)
( cocone-disjunction)
( map-disjunction-join-Prop)
- ( ( λ _ → eq-is-prop (is-prop-type-disjunction-Prop A B)) ,
- ( λ _ → eq-is-prop (is-prop-type-disjunction-Prop A B)) ,
+ ( ( λ _ → eq-is-prop (is-prop-disjunction-Prop A B)) ,
+ ( λ _ → eq-is-prop (is-prop-disjunction-Prop A B)) ,
( λ (a , b) → eq-is-contr
- ( is-prop-type-disjunction-Prop A B
+ ( is-prop-disjunction-Prop A B
( horizontal-map-cocone pr1 pr2
( cocone-map pr1 pr2
( cocone-join)
diff --git a/src/trees/bounded-multisets.lagda.md b/src/trees/bounded-multisets.lagda.md
index ada81ad1ab..ceda35ae57 100644
--- a/src/trees/bounded-multisets.lagda.md
+++ b/src/trees/bounded-multisets.lagda.md
@@ -12,6 +12,7 @@ open import elementary-number-theory.natural-numbers
open import foundation.dependent-pair-types
open import foundation.empty-types
open import foundation.existential-quantification
+open import foundation.propositions
open import foundation.universe-levels
open import trees.empty-multisets
@@ -51,6 +52,16 @@ module _
is-empty-𝕍 X
is-of-natural-height-𝕍 (succ-ℕ n) (tree-𝕎 X Y) =
(x : X) → is-of-natural-height-𝕍 n (Y x)
+
+ is-prop-is-of-natural-height-𝕍 :
+ (n : ℕ) (X : 𝕍 l) → is-prop (is-of-natural-height-𝕍 n X)
+ is-prop-is-of-natural-height-𝕍 zero-ℕ = is-property-is-empty-𝕍
+ is-prop-is-of-natural-height-𝕍 (succ-ℕ n) (tree-𝕎 X Y) =
+ is-prop-Π (λ x → is-prop-is-of-natural-height-𝕍 n (Y x))
+
+ is-of-natural-height-prop-𝕍 : ℕ → 𝕍 l → Prop l
+ is-of-natural-height-prop-𝕍 n X =
+ ( is-of-natural-height-𝕍 n X , is-prop-is-of-natural-height-𝕍 n X)
```
### Explicitly bounded multisets
@@ -80,7 +91,8 @@ data
{n : ℕ} {X : UU l} (Y : X → Bounded-𝕍 l n) → Bounded-𝕍 l (succ-ℕ n)
Bounded-𝕍' : (l : Level) → UU (lsuc l)
-Bounded-𝕍' l = Σ (𝕍 l) (λ X → ∃ ℕ (λ n → is-of-natural-height-𝕍 n X))
+Bounded-𝕍' l =
+ Σ (𝕍 l) (λ X → exists ℕ (λ n → is-of-natural-height-prop-𝕍 n X))
```
## Properties
diff --git a/src/trees/directed-trees.lagda.md b/src/trees/directed-trees.lagda.md
index 6af472ebe5..0198184e27 100644
--- a/src/trees/directed-trees.lagda.md
+++ b/src/trees/directed-trees.lagda.md
@@ -178,7 +178,8 @@ module _
where
is-proper-node-Directed-Tree-Prop : node-Directed-Tree T → Prop l1
- is-proper-node-Directed-Tree-Prop x = neg-Prop' (is-root-Directed-Tree T x)
+ is-proper-node-Directed-Tree-Prop x =
+ neg-type-Prop (is-root-Directed-Tree T x)
is-proper-node-Directed-Tree : node-Directed-Tree T → UU l1
is-proper-node-Directed-Tree x =
diff --git a/src/trees/empty-multisets.lagda.md b/src/trees/empty-multisets.lagda.md
index 880a5c680b..d31b3e8a6b 100644
--- a/src/trees/empty-multisets.lagda.md
+++ b/src/trees/empty-multisets.lagda.md
@@ -10,6 +10,7 @@ module trees.empty-multisets where
open import foundation.dependent-pair-types
open import foundation.empty-types
open import foundation.identity-types
+open import foundation.propositions
open import foundation.universe-levels
open import trees.elementhood-relation-w-types
@@ -35,6 +36,12 @@ module _
is-empty-𝕍 : 𝕍 l → UU l
is-empty-𝕍 (tree-𝕎 X Y) = is-empty X
+
+ is-property-is-empty-𝕍 : (X : 𝕍 l) → is-prop (is-empty-𝕍 X)
+ is-property-is-empty-𝕍 (tree-𝕎 X Y) = is-property-is-empty
+
+ is-empty-prop-𝕍 : 𝕍 l → Prop l
+ is-empty-prop-𝕍 X = is-empty-𝕍 X , is-property-is-empty-𝕍 X
```
### The predicate of being a multiset with no elements
diff --git a/src/trees/lower-types-w-types.lagda.md b/src/trees/lower-types-w-types.lagda.md
index 0f2c8176a7..5e283be7bb 100644
--- a/src/trees/lower-types-w-types.lagda.md
+++ b/src/trees/lower-types-w-types.lagda.md
@@ -42,5 +42,5 @@ module _
upper-bound-rank-inclusion-lower-𝕎 :
{x : 𝕎 A B} (y : lower-𝕎 x) → inclusion-lower-𝕎 y ≼-𝕎 x
upper-bound-rank-inclusion-lower-𝕎 (lower-tree-𝕎 g) y =
- intro-∃ y (upper-bound-rank-inclusion-lower-𝕎 (g y))
+ intro-exists y (upper-bound-rank-inclusion-lower-𝕎 (g y))
```
diff --git a/src/trees/ranks-of-elements-w-types.lagda.md b/src/trees/ranks-of-elements-w-types.lagda.md
index bc8589c340..26e9e63740 100644
--- a/src/trees/ranks-of-elements-w-types.lagda.md
+++ b/src/trees/ranks-of-elements-w-types.lagda.md
@@ -42,7 +42,7 @@ module _
_≼-𝕎-Prop_ : 𝕎 A B → 𝕎 A B → Prop (l1 ⊔ l2)
(tree-𝕎 x α) ≼-𝕎-Prop (tree-𝕎 y β) =
- Π-Prop (B x) (λ b → exists-Prop (B y) (λ c → (α b) ≼-𝕎-Prop (β c)))
+ Π-Prop (B x) (λ b → ∃ (B y) (λ c → (α b) ≼-𝕎-Prop (β c)))
_≼-𝕎_ : 𝕎 A B → 𝕎 A B → UU (l1 ⊔ l2)
x ≼-𝕎 y = type-Prop (x ≼-𝕎-Prop y)
@@ -63,7 +63,7 @@ module _
_≺-𝕎-Prop_ : 𝕎 A B → 𝕎 A B → Prop (l1 ⊔ l2)
x ≺-𝕎-Prop y =
- exists-Prop (Σ (𝕎 A B) (λ w → w ∈-𝕎 y)) (λ t → x ≼-𝕎-Prop (pr1 t))
+ ∃ (Σ (𝕎 A B) (λ w → w ∈-𝕎 y)) (λ t → x ≼-𝕎-Prop (pr1 t))
_≺-𝕎_ : 𝕎 A B → 𝕎 A B → UU (l1 ⊔ l2)
x ≺-𝕎 y = type-Prop (x ≺-𝕎-Prop y)
@@ -93,10 +93,9 @@ module _
Π-Prop
( u ∈-𝕎 x)
( λ H →
- exists-Prop
- ( 𝕎 A B)
+ ∃ ( 𝕎 A B)
( λ v →
- exists-Prop (v ∈-𝕎 y) (λ K → u ≼-𝕎-Prop v))))
+ ∃ (v ∈-𝕎 y) (λ K → u ≼-𝕎-Prop v))))
_strong-≼-𝕎_ : 𝕎 A B → 𝕎 A B → UU (l1 ⊔ l2)
x strong-≼-𝕎 y = type-Prop (x strong-≼-𝕎-Prop y)
@@ -126,11 +125,11 @@ module _
transitive-≼-𝕎 {tree-𝕎 x α} {tree-𝕎 y β} {tree-𝕎 z γ} H K a =
apply-universal-property-trunc-Prop
( H a)
- ( exists-Prop (B z) (λ c → (α a) ≼-𝕎-Prop (γ c)))
+ ( ∃ (B z) (λ c → (α a) ≼-𝕎-Prop (γ c)))
( λ t →
apply-universal-property-trunc-Prop
( K (pr1 t))
- ( exists-Prop (B z) (λ c → (α a) ≼-𝕎-Prop (γ c)))
+ ( ∃ (B z) (λ c → (α a) ≼-𝕎-Prop (γ c)))
( λ s →
unit-trunc-Prop
( pair
@@ -153,30 +152,23 @@ module _
strong-≼-≼-𝕎 : {x y : 𝕎 A B} → (x ≼-𝕎 y) → (x strong-≼-𝕎 y)
strong-≼-≼-𝕎 {tree-𝕎 x α} {tree-𝕎 y β} H .(α b) (pair b refl) =
apply-universal-property-trunc-Prop (H b)
- ( exists-Prop
- ( 𝕎 A B)
- ( (λ v → exists-Prop (v ∈-𝕎 tree-𝕎 y β) (λ hv → (α b) ≼-𝕎-Prop v))))
+ ( ∃ ( 𝕎 A B)
+ ( (λ v → ∃ (v ∈-𝕎 tree-𝕎 y β) (λ hv → (α b) ≼-𝕎-Prop v))))
( f)
where
f :
Σ (B y) (λ c → pr1 (α b ≼-𝕎-Prop β c)) →
exists
( 𝕎 A B)
- ( λ v → exists-Prop (v ∈-𝕎 tree-𝕎 y β) (λ hv → α b ≼-𝕎-Prop v))
+ ( λ v → ∃ (v ∈-𝕎 tree-𝕎 y β) (λ hv → α b ≼-𝕎-Prop v))
f (pair c K) =
- intro-exists
- ( λ v → exists-Prop (v ∈-𝕎 tree-𝕎 y β) (λ hv → α b ≼-𝕎-Prop v))
- ( β c)
- ( intro-exists
- ( λ hβc → α b ≼-𝕎-Prop β c)
- ( pair c refl)
- ( K))
+ intro-exists (β c) ( intro-exists (pair c refl) K)
≼-strong-≼-𝕎 : {x y : 𝕎 A B} → (x strong-≼-𝕎 y) → (x ≼-𝕎 y)
≼-strong-≼-𝕎 {tree-𝕎 x α} {tree-𝕎 y β} H b =
apply-universal-property-trunc-Prop
- ( H (α b) (pair b refl))
- ( exists-Prop (B y) (λ c → α b ≼-𝕎-Prop β c))
+ ( H (α b) (b , refl))
+ ( ∃ (B y) (λ c → α b ≼-𝕎-Prop β c))
( f)
where
f :
@@ -185,11 +177,13 @@ module _
exists (B y) (λ c → α b ≼-𝕎-Prop β c)
f (pair v K) =
apply-universal-property-trunc-Prop K
- ( exists-Prop (B y) (λ c → α b ≼-𝕎-Prop β c))
+ ( ∃ (B y) (λ c → α b ≼-𝕎-Prop β c))
( g)
where
- g : (v ∈-𝕎 tree-𝕎 y β) × (α b ≼-𝕎 v) → ∃ (B y) (λ c → α b ≼-𝕎 β c)
- g (pair (pair c p) M) = intro-∃ c (tr (λ t → α b ≼-𝕎 t) (inv p) M)
+ g :
+ (v ∈-𝕎 tree-𝕎 y β) × (α b ≼-𝕎 v) →
+ exists (B y) (λ c → α b ≼-𝕎-Prop β c)
+ g (pair (pair c p) M) = intro-exists c (tr (λ t → α b ≼-𝕎 t) (inv p) M)
```
### If `x ∈ y` then the rank of `x` is at most the rank of `y`
@@ -202,7 +196,6 @@ module _
≼-∈-𝕎 : {x y : 𝕎 A B} → (x ∈-𝕎 y) → (x ≼-𝕎 y)
≼-∈-𝕎 {tree-𝕎 x α} {tree-𝕎 y β} (pair v p) u =
intro-exists
- ( λ z → α u ≼-𝕎-Prop β z)
( v)
( tr
( λ t → α u ≼-𝕎 t)
@@ -307,7 +300,6 @@ module _
g : Σ (Σ (𝕎 A B) (λ w → w ∈-𝕎 z)) (λ t → y ≼-𝕎 pr1 t) → x ≺-𝕎 z
g (pair (pair v P) Q) =
intro-exists
- ( λ t → x ≼-𝕎-Prop (pr1 t))
( pair v P)
( transitive-≼-𝕎 {x = x} {w} {v} M
( transitive-≼-𝕎 {x = w} {y} {v} (≼-∈-𝕎 L) Q))
diff --git a/src/univalent-combinatorics/2-element-decidable-subtypes.lagda.md b/src/univalent-combinatorics/2-element-decidable-subtypes.lagda.md
index 1cc4c33057..31511a5324 100644
--- a/src/univalent-combinatorics/2-element-decidable-subtypes.lagda.md
+++ b/src/univalent-combinatorics/2-element-decidable-subtypes.lagda.md
@@ -34,6 +34,7 @@ open import foundation.negated-equality
open import foundation.negation
open import foundation.propositional-truncations
open import foundation.propositions
+open import foundation.sets
open import foundation.subtypes
open import foundation.transport-along-identifications
open import foundation.type-arithmetic-coproduct-types
diff --git a/src/univalent-combinatorics/2-element-subtypes.lagda.md b/src/univalent-combinatorics/2-element-subtypes.lagda.md
index b86c98de5c..782ef754ab 100644
--- a/src/univalent-combinatorics/2-element-subtypes.lagda.md
+++ b/src/univalent-combinatorics/2-element-subtypes.lagda.md
@@ -254,7 +254,7 @@ module _
Σ ( type-2-Element-Type X)
( λ x → ¬ ((y : type-2-Element-Type X) → Id (f x) (f y)))
first-element h =
- exists-not-not-forall-count (λ z → (w : type-2-Element-Type X) →
+ exists-not-not-for-all-count (λ z → (w : type-2-Element-Type X) →
Id (f z) (f w)) (λ z → {!!})
{!!} {!!}
two-elements-different-image :
diff --git a/src/univalent-combinatorics/2-element-types.lagda.md b/src/univalent-combinatorics/2-element-types.lagda.md
index 00a6a7d5dd..84068ec248 100644
--- a/src/univalent-combinatorics/2-element-types.lagda.md
+++ b/src/univalent-combinatorics/2-element-types.lagda.md
@@ -526,7 +526,7 @@ abstract
( ¬ (pr1 X))
( apply-universal-property-trunc-Prop
( pr2 X)
- ( double-negation-Prop' (pr1 X))
+ ( double-negation-type-Prop (pr1 X))
( λ e → intro-double-negation {l} (map-equiv e (zero-Fin 1))))
( d X))
```
diff --git a/src/univalent-combinatorics/binomial-types.lagda.md b/src/univalent-combinatorics/binomial-types.lagda.md
index bf406c5b33..67c3f5a8a2 100644
--- a/src/univalent-combinatorics/binomial-types.lagda.md
+++ b/src/univalent-combinatorics/binomial-types.lagda.md
@@ -271,7 +271,7 @@ abstract
( pair (raise-empty-Prop _) map-inv-raise)))) ∘e
( right-distributive-Σ-coproduct
( Σ (Prop _) type-Prop)
- ( Σ (Prop _) (¬ ∘ type-Prop))
+ ( Σ (Prop _) (¬_ ∘ type-Prop))
( ind-coproduct _
( λ Q →
mere-equiv (Maybe B) ((Σ A _) + (type-Prop (pr1 Q))))
diff --git a/src/univalent-combinatorics/counting-dependent-pair-types.lagda.md b/src/univalent-combinatorics/counting-dependent-pair-types.lagda.md
index d300a6adc9..ce45d38d86 100644
--- a/src/univalent-combinatorics/counting-dependent-pair-types.lagda.md
+++ b/src/univalent-combinatorics/counting-dependent-pair-types.lagda.md
@@ -52,8 +52,8 @@ Consider a type family `B` over `A`, and consider the following statements
3. The elements of `Σ A B` can be counted.
If 1 holds, then 2 holds if and only if 3 holds. Furthermore if 2 and 3 hold,
-then 1 holds if and only if the elements of `Σ A (¬ ∘ B)` can be counted, i.e.,
-if the elements in the complement of `B` can be counted.
+then 1 holds if and only if the elements of `Σ (x : A), ¬ (B x)` can be counted,
+i.e., if the elements in the complement of `B` can be counted.
## Proofs
@@ -168,12 +168,12 @@ count-base-count-Σ b e f =
```
More generally, if `Σ A B` and each `B x` can be counted, then `A` can be
-counted if and only if the type `Σ A (¬ ∘ B)` can be counted. However, to avoid
-having to invoke function extensionality, we show that if `Σ A B` and each `B x`
-can be counted, then `A` can be counted if and only if
+counted if and only if the type `Σ (x : A), ¬ (B x)` can be counted. However, to
+avoid having to invoke function extensionality, we show that if `Σ A B` and each
+`B x` can be counted, then `A` can be counted if and only if
```text
- count (Σ A (λ x → is-zero-ℕ (number-of-elements-count (f x)))),
+ count (Σ A (λ x → is-zero-ℕ (number-of-elements-count (f x)))),
```
where `f : (x : A) → count (B x)`. Thus, we have a precise characterization of
diff --git a/src/univalent-combinatorics/dependent-pair-types.lagda.md b/src/univalent-combinatorics/dependent-pair-types.lagda.md
index 4e25dec6af..932efc1d7a 100644
--- a/src/univalent-combinatorics/dependent-pair-types.lagda.md
+++ b/src/univalent-combinatorics/dependent-pair-types.lagda.md
@@ -155,7 +155,7 @@ abstract
( ( right-unit-law-Σ-is-contr
( λ x →
is-proof-irrelevant-is-prop
- ( is-prop-is-inhabited-or-empty (B x))
+ ( is-property-is-inhabited-or-empty (B x))
( is-inhabited-or-empty-is-finite (g x)))) ∘e
( inv-equiv
( left-distributive-Σ-coproduct A
diff --git a/src/univalent-combinatorics/embeddings-standard-finite-types.lagda.md b/src/univalent-combinatorics/embeddings-standard-finite-types.lagda.md
index 69b7455380..458ec1b433 100644
--- a/src/univalent-combinatorics/embeddings-standard-finite-types.lagda.md
+++ b/src/univalent-combinatorics/embeddings-standard-finite-types.lagda.md
@@ -18,6 +18,7 @@ open import foundation.embeddings
open import foundation.empty-types
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.unit-type
open import foundation.universe-levels
diff --git a/src/univalent-combinatorics/embeddings.lagda.md b/src/univalent-combinatorics/embeddings.lagda.md
index 625f1cd7da..5ec8168e06 100644
--- a/src/univalent-combinatorics/embeddings.lagda.md
+++ b/src/univalent-combinatorics/embeddings.lagda.md
@@ -19,6 +19,7 @@ open import foundation.empty-types
open import foundation.equivalences
open import foundation.propositional-truncations
open import foundation.propositions
+open import foundation.sets
open import foundation.unit-type
open import foundation.universe-levels
diff --git a/src/univalent-combinatorics/equality-standard-finite-types.lagda.md b/src/univalent-combinatorics/equality-standard-finite-types.lagda.md
index d8bc06357e..2ed5c9e278 100644
--- a/src/univalent-combinatorics/equality-standard-finite-types.lagda.md
+++ b/src/univalent-combinatorics/equality-standard-finite-types.lagda.md
@@ -21,6 +21,7 @@ open import foundation.empty-types
open import foundation.equivalences
open import foundation.functoriality-coproduct-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.negated-equality
open import foundation.negation
open import foundation.propositions
@@ -77,7 +78,7 @@ extensionality-Fin :
(x = y) ≃ (Eq-Fin k x y)
pr1 (extensionality-Fin k x y) = Eq-Fin-eq k
pr2 (extensionality-Fin k x y) =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-set-Fin k x y)
( is-prop-Eq-Fin k x y)
( eq-Eq-Fin k)
diff --git a/src/univalent-combinatorics/fibers-of-maps.lagda.md b/src/univalent-combinatorics/fibers-of-maps.lagda.md
index ad8da9ae05..914e7cabd5 100644
--- a/src/univalent-combinatorics/fibers-of-maps.lagda.md
+++ b/src/univalent-combinatorics/fibers-of-maps.lagda.md
@@ -18,6 +18,7 @@ open import foundation.equality-dependent-pair-types
open import foundation.equivalences
open import foundation.functoriality-dependent-pair-types
open import foundation.identity-types
+open import foundation.logical-equivalences
open import foundation.propositional-truncations
open import foundation.propositions
open import foundation.sections
@@ -146,7 +147,7 @@ equiv-is-finite-domain-is-finite-fiber :
(B : 𝔽 l2) (f : A → (type-𝔽 B)) →
((b : type-𝔽 B) → is-finite (fiber f b)) ≃ is-finite A
equiv-is-finite-domain-is-finite-fiber {A = A} B f =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-Π (λ b → is-prop-is-finite (fiber f b)))
( is-prop-is-finite A)
( λ P →
diff --git a/src/univalent-combinatorics/finite-types.lagda.md b/src/univalent-combinatorics/finite-types.lagda.md
index dcee18540b..5da3f667e3 100644
--- a/src/univalent-combinatorics/finite-types.lagda.md
+++ b/src/univalent-combinatorics/finite-types.lagda.md
@@ -25,6 +25,7 @@ open import foundation.functoriality-dependent-pair-types
open import foundation.functoriality-propositional-truncation
open import foundation.identity-types
open import foundation.inhabited-types
+open import foundation.logical-equivalences
open import foundation.mere-equivalences
open import foundation.propositional-truncations
open import foundation.propositions
@@ -510,7 +511,7 @@ compute-total-UU-Fin : {l : Level} → Σ ℕ (UU-Fin l) ≃ 𝔽 l
compute-total-UU-Fin =
( equiv-tot
( λ X →
- equiv-prop
+ equiv-iff-is-prop
( is-prop-has-finite-cardinality)
( is-prop-is-finite X)
( is-finite-has-finite-cardinality)
@@ -697,7 +698,7 @@ abstract
( has-finite-cardinality-is-finite H)
( pair n Q))
pr2 (equiv-has-cardinality-id-number-of-elements-is-finite X H n) =
- is-equiv-is-prop
+ is-equiv-has-converse-is-prop
( is-prop-type-trunc-Prop)
( is-set-ℕ (number-of-elements-is-finite H) n)
( λ p →
diff --git a/src/univalent-combinatorics/inequality-types-with-counting.lagda.md b/src/univalent-combinatorics/inequality-types-with-counting.lagda.md
index 0a0a2fcd3d..137bb6bd70 100644
--- a/src/univalent-combinatorics/inequality-types-with-counting.lagda.md
+++ b/src/univalent-combinatorics/inequality-types-with-counting.lagda.md
@@ -13,6 +13,7 @@ open import foundation.dependent-pair-types
open import foundation.equivalences
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.universe-levels
open import univalent-combinatorics.counting
diff --git a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md
index dd9d41741b..adc3eb2069 100644
--- a/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md
+++ b/src/univalent-combinatorics/kuratowsky-finite-sets.lagda.md
@@ -37,7 +37,7 @@ the set-quotient of a standard finite type.
```agda
is-kuratowsky-finite-set-Prop : {l : Level} → Set l → Prop l
is-kuratowsky-finite-set-Prop X =
- ∃-Prop ℕ (λ n → Fin n ↠ type-Set X)
+ exists-structure-Prop ℕ (λ n → Fin n ↠ type-Set X)
is-kuratowsky-finite-set : {l : Level} → Set l → UU l
is-kuratowsky-finite-set X = type-Prop (is-kuratowsky-finite-set-Prop X)
diff --git a/src/univalent-combinatorics/pi-finite-types.lagda.md b/src/univalent-combinatorics/pi-finite-types.lagda.md
index d530ac77ea..9f4a0b0633 100644
--- a/src/univalent-combinatorics/pi-finite-types.lagda.md
+++ b/src/univalent-combinatorics/pi-finite-types.lagda.md
@@ -143,9 +143,9 @@ mere-equiv-number-of-connected-components H =
is-π-finite-Prop : {l : Level} (k : ℕ) → UU l → Prop l
is-π-finite-Prop zero-ℕ X = has-finite-connected-components-Prop X
is-π-finite-Prop (succ-ℕ k) X =
- product-Prop ( is-π-finite-Prop zero-ℕ X)
- ( Π-Prop X
- ( λ x → Π-Prop X (λ y → is-π-finite-Prop k (Id x y))))
+ product-Prop
+ ( is-π-finite-Prop zero-ℕ X)
+ ( Π-Prop X (λ x → Π-Prop X (λ y → is-π-finite-Prop k (Id x y))))
is-π-finite : {l : Level} (k : ℕ) → UU l → UU l
is-π-finite k X = type-Prop (is-π-finite-Prop k X)
diff --git a/src/univalent-combinatorics/pigeonhole-principle.lagda.md b/src/univalent-combinatorics/pigeonhole-principle.lagda.md
index 97927d8ab5..8c52386070 100644
--- a/src/univalent-combinatorics/pigeonhole-principle.lagda.md
+++ b/src/univalent-combinatorics/pigeonhole-principle.lagda.md
@@ -24,6 +24,7 @@ open import foundation.negation
open import foundation.pairs-of-distinct-elements
open import foundation.propositional-truncations
open import foundation.propositions
+open import foundation.sets
open import foundation.unit-type
open import foundation.universe-levels
open import foundation.whiskering-homotopies-composition
diff --git a/src/univalent-combinatorics/repetitions-of-values.lagda.md b/src/univalent-combinatorics/repetitions-of-values.lagda.md
index 9fe994bec5..c38d663bcd 100644
--- a/src/univalent-combinatorics/repetitions-of-values.lagda.md
+++ b/src/univalent-combinatorics/repetitions-of-values.lagda.md
@@ -18,6 +18,7 @@ open import foundation.identity-types
open import foundation.injective-maps
open import foundation.negated-equality
open import foundation.negation
+open import foundation.sets
open import univalent-combinatorics.decidable-dependent-function-types
open import univalent-combinatorics.decidable-propositions
@@ -48,7 +49,7 @@ repetition-of-values-is-not-injective-Fin k l f N =
where
u : Σ (Fin k) (λ x → ¬ ((y : Fin k) → f x = f y → x = y))
u =
- exists-not-not-forall-Fin k
+ exists-not-not-for-all-Fin k
( λ x →
is-decidable-Π-Fin k
( λ y →
@@ -61,7 +62,7 @@ repetition-of-values-is-not-injective-Fin k l f N =
H : ¬ ((y : Fin k) → f x = f y → x = y)
H = pr2 u
v : Σ (Fin k) (λ y → ¬ (f x = f y → x = y))
- v = exists-not-not-forall-Fin k
+ v = exists-not-not-for-all-Fin k
( λ y →
is-decidable-function-type
( has-decidable-equality-Fin l (f x) (f y))
@@ -72,7 +73,7 @@ repetition-of-values-is-not-injective-Fin k l f N =
K : ¬ (f x = f y → x = y)
K = pr2 v
w : (f x = f y) × (x ≠ y)
- w = exists-not-not-forall-count
+ w = exists-not-not-for-all-count
( λ _ → Id x y)
( λ _ →
has-decidable-equality-Fin k x y)
diff --git a/src/univalent-combinatorics/retracts-of-finite-types.lagda.md b/src/univalent-combinatorics/retracts-of-finite-types.lagda.md
index bc30866f9d..f12270336a 100644
--- a/src/univalent-combinatorics/retracts-of-finite-types.lagda.md
+++ b/src/univalent-combinatorics/retracts-of-finite-types.lagda.md
@@ -18,6 +18,7 @@ open import foundation.functoriality-propositional-truncation
open import foundation.injective-maps
open import foundation.retractions
open import foundation.retracts-of-types
+open import foundation.sets
open import foundation.universe-levels
open import univalent-combinatorics.counting
diff --git a/src/univalent-combinatorics/sequences-finite-types.lagda.md b/src/univalent-combinatorics/sequences-finite-types.lagda.md
index b9d113365b..e9cd8106af 100644
--- a/src/univalent-combinatorics/sequences-finite-types.lagda.md
+++ b/src/univalent-combinatorics/sequences-finite-types.lagda.md
@@ -24,6 +24,7 @@ open import foundation.pairs-of-distinct-elements
open import foundation.repetitions-of-values
open import foundation.repetitions-sequences
open import foundation.sequences
+open import foundation.sets
open import foundation.universe-levels
open import univalent-combinatorics.counting
diff --git a/src/univalent-combinatorics/sigma-decompositions.lagda.md b/src/univalent-combinatorics/sigma-decompositions.lagda.md
index 1d1a8ff667..f45a3477c8 100644
--- a/src/univalent-combinatorics/sigma-decompositions.lagda.md
+++ b/src/univalent-combinatorics/sigma-decompositions.lagda.md
@@ -18,6 +18,7 @@ open import foundation.functoriality-dependent-pair-types
open import foundation.homotopies
open import foundation.identity-types
open import foundation.inhabited-types
+open import foundation.logical-equivalences
open import foundation.precomposition-functions
open import foundation.propositions
open import foundation.relaxed-sigma-decompositions
@@ -437,7 +438,7 @@ module _
type-Prop (is-finite-displayed-Σ-Decomposition A
(map-equiv equiv-displayed-fibered-Σ-Decomposition D))
equiv-is-finite-displayed-fibered-Σ-Decomposition =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-type-Prop (is-finite-fibered-Σ-Decomposition A D))
( is-prop-type-Prop
( is-finite-displayed-Σ-Decomposition A
diff --git a/src/univalent-combinatorics/skipping-element-standard-finite-types.lagda.md b/src/univalent-combinatorics/skipping-element-standard-finite-types.lagda.md
index 1ab1535c34..3495723917 100644
--- a/src/univalent-combinatorics/skipping-element-standard-finite-types.lagda.md
+++ b/src/univalent-combinatorics/skipping-element-standard-finite-types.lagda.md
@@ -16,6 +16,7 @@ open import foundation.embeddings
open import foundation.equality-coproduct-types
open import foundation.identity-types
open import foundation.injective-maps
+open import foundation.sets
open import foundation.unit-type
open import univalent-combinatorics.standard-finite-types
diff --git a/src/univalent-combinatorics/surjective-maps.lagda.md b/src/univalent-combinatorics/surjective-maps.lagda.md
index f4950acb3d..231a339513 100644
--- a/src/univalent-combinatorics/surjective-maps.lagda.md
+++ b/src/univalent-combinatorics/surjective-maps.lagda.md
@@ -17,6 +17,7 @@ open import foundation.decidable-equality
open import foundation.decidable-types
open import foundation.dependent-pair-types
open import foundation.equivalences
+open import foundation.logical-equivalences
open import foundation.propositional-truncations
open import foundation.propositions
open import foundation.type-arithmetic-dependent-pair-types
@@ -119,7 +120,7 @@ module _
is-finite X ≃
( has-decidable-equality X × type-trunc-Prop (Σ ℕ (λ n → Fin n ↠ X)))
is-finite-iff-∃-surjection-has-decidable-equality =
- equiv-prop
+ equiv-iff-is-prop
( is-prop-is-finite X)
( is-prop-product is-prop-has-decidable-equality is-prop-type-trunc-Prop)
( λ fin-X → is-finite-if-∃-surjection-has-decidable-equality fin-X)
diff --git a/tables/propositional-logic.md b/tables/propositional-logic.md
new file mode 100644
index 0000000000..85f1cc601c
--- /dev/null
+++ b/tables/propositional-logic.md
@@ -0,0 +1,26 @@
+| Concept | File |
+| ---------------------------------------------------- | ----------------------------------------------------------------------------------------------------------------------- |
+| Propositions (foundation-core) | [`foundation-core.propositions`](foundation-core.propositions.md) |
+| Propositions (foundation) | [`foundation.propositions`](foundation.propositions.md) |
+| Subterminal types | [`foundation.subterminal-types`](foundation.subterminal-types.md) |
+| Subsingleton induction | [`foundation.subsingleton-induction`](foundation.subsingleton-induction.md) |
+| Empty types (foundation-core) | [`foundation-core.empty-types`](foundation-core.empty-types.md) |
+| Empty types (foundation) | [`foundation.empty-types`](foundation.empty-types.md) |
+| Unit type | [`foundation.unit-type`](foundation.unit-type.md) |
+| Logical equivalences | [`foundation.logical-equivalences`](foundation.logical-equivalences.md) |
+| Propositional extensionality | [`foundation.propositional-extensionality`](foundation.propositional-extensionality.md) |
+| Mere logical equivalences | [`foundation.mere-logical-equivalences`](foundation.mere-logical-equivalences.md) |
+| Conjunction | [`foundation.conjunction`](foundation.conjunction.md) |
+| Disjunction | [`foundation.disjunction`](foundation.disjunction.md) |
+| Exclusive disjunction | [`foundation.exclusive-disjunction`](foundation.exclusive-disjunction.md) |
+| Existential quantification | [`foundation.existential-quantification`](foundation.existential-quantification.md) |
+| Uniqueness quantification | [`foundation.uniqueness-quantification`](foundation.uniqueness-quantification.md) |
+| Universal quantification | [`foundation.universal-quantification`](foundation.universal-quantification.md) |
+| Negation | [`foundation.negation`](foundation.negation.md) |
+| Double negation | [`foundation.double-negation`](foundation.double-negation.md) |
+| Propositional truncations | [`foundation.propositional-truncations`](foundation.propositional-truncations.md) |
+| Universal property of propositional truncations | [`foundation.universal-property-propositional-truncation`](foundation.universal-property-propositional-truncation.md) |
+| The induction principle of propositional truncations | [`foundation.induction-principle-propositional-truncation`](foundation.induction-principle-propositional-truncation.md) |
+| Functoriality of propositional truncations | [`foundation.functoriality-propositional-truncations`](foundation.functoriality-propositional-truncation.md) |
+| Propositional resizing | [`foundation.propositional-resizing`](foundation.propositional-resizing.md) |
+| Impredicative encodings of the logical operations | [`foundation.impredicative-encodings`](foundation.impredicative-encodings.md) |