2018-03-21 13:39:56 +00:00
|
|
|
|
-- | Custom prelude for this module
|
|
|
|
|
module Cat.Prelude where
|
|
|
|
|
|
|
|
|
|
open import Agda.Primitive public
|
|
|
|
|
-- FIXME Use:
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open import Agda.Builtin.Sigma public
|
2018-03-21 13:39:56 +00:00
|
|
|
|
-- Rather than
|
|
|
|
|
open import Data.Product public
|
|
|
|
|
renaming (∃! to ∃!≈)
|
2018-04-05 08:41:56 +00:00
|
|
|
|
using (_×_ ; Σ-syntax ; swap)
|
2018-03-21 13:39:56 +00:00
|
|
|
|
|
2018-04-11 09:10:33 +00:00
|
|
|
|
open import Function using (_∘_ ; _∘′_ ; _$_ ; case_of_ ; flip) public
|
|
|
|
|
|
|
|
|
|
idFun : ∀ {a} (A : Set a) → A → A
|
|
|
|
|
idFun A a = a
|
2018-03-21 13:39:56 +00:00
|
|
|
|
|
|
|
|
|
open import Cubical public
|
|
|
|
|
-- FIXME rename `gradLemma` to `fromIsomorphism` - perhaps just use wrapper
|
|
|
|
|
-- module.
|
|
|
|
|
open import Cubical.GradLemma
|
|
|
|
|
using (gradLemma)
|
|
|
|
|
public
|
|
|
|
|
open import Cubical.NType
|
2018-03-21 14:01:31 +00:00
|
|
|
|
using (⟨-2⟩ ; ⟨-1⟩ ; ⟨0⟩ ; TLevel ; HasLevel)
|
2018-03-21 13:39:56 +00:00
|
|
|
|
public
|
|
|
|
|
open import Cubical.NType.Properties
|
|
|
|
|
using
|
|
|
|
|
( lemPropF ; lemSig ; lemSigP ; isSetIsProp
|
2018-04-05 18:41:14 +00:00
|
|
|
|
; propPi ; propPiImpl ; propHasLevel ; setPi ; propSet
|
2018-04-06 16:27:24 +00:00
|
|
|
|
; propSig ; equivPreservesNType)
|
2018-03-21 13:39:56 +00:00
|
|
|
|
public
|
2018-03-22 13:27:16 +00:00
|
|
|
|
|
|
|
|
|
propIsContr : {ℓ : Level} → {A : Set ℓ} → isProp (isContr A)
|
|
|
|
|
propIsContr = propHasLevel ⟨-2⟩
|
|
|
|
|
|
2018-04-03 13:23:11 +00:00
|
|
|
|
open import Cubical.Sigma using (setSig ; sigPresSet ; sigPresNType) public
|
2018-03-21 14:01:31 +00:00
|
|
|
|
|
|
|
|
|
module _ (ℓ : Level) where
|
|
|
|
|
-- FIXME Ask if we can push upstream.
|
|
|
|
|
-- A redefinition of `Cubical.Universe` with an explicit parameter
|
|
|
|
|
_-type : TLevel → Set (lsuc ℓ)
|
|
|
|
|
n -type = Σ (Set ℓ) (HasLevel n)
|
|
|
|
|
|
|
|
|
|
hSet : Set (lsuc ℓ)
|
|
|
|
|
hSet = ⟨0⟩ -type
|
|
|
|
|
|
|
|
|
|
Prop : Set (lsuc ℓ)
|
|
|
|
|
Prop = ⟨-1⟩ -type
|
2018-03-21 13:39:56 +00:00
|
|
|
|
|
|
|
|
|
-----------------
|
|
|
|
|
-- * Utilities --
|
|
|
|
|
-----------------
|
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
-- | Unique existentials.
|
2018-03-21 13:39:56 +00:00
|
|
|
|
∃! : ∀ {a b} {A : Set a}
|
|
|
|
|
→ (A → Set b) → Set (a ⊔ b)
|
|
|
|
|
∃! = ∃!≈ _≡_
|
|
|
|
|
|
|
|
|
|
∃!-syntax : ∀ {a b} {A : Set a} → (A → Set b) → Set (a ⊔ b)
|
2018-03-27 12:18:13 +00:00
|
|
|
|
∃!-syntax = ∃!
|
2018-03-21 13:39:56 +00:00
|
|
|
|
|
|
|
|
|
syntax ∃!-syntax (λ x → B) = ∃![ x ] B
|
2018-03-21 13:47:01 +00:00
|
|
|
|
|
2018-04-03 10:40:20 +00:00
|
|
|
|
module _ {ℓa ℓb} {A : Set ℓa} {P : A → Set ℓb} (f g : ∃! P) where
|
2018-04-05 08:41:56 +00:00
|
|
|
|
open Σ (snd f) renaming (snd to u)
|
2018-04-03 10:40:20 +00:00
|
|
|
|
|
2018-04-05 08:41:56 +00:00
|
|
|
|
∃-unique : fst f ≡ fst g
|
|
|
|
|
∃-unique = u (fst (snd g))
|
2018-04-03 10:40:20 +00:00
|
|
|
|
|
2018-03-21 13:47:01 +00:00
|
|
|
|
module _ {ℓa ℓb : Level} {A : Set ℓa} {B : A → Set ℓb} {a b : Σ A B}
|
2018-04-05 08:41:56 +00:00
|
|
|
|
(fst≡ : (λ _ → A) [ fst a ≡ fst b ])
|
|
|
|
|
(snd≡ : (λ i → B (fst≡ i)) [ snd a ≡ snd b ]) where
|
2018-03-21 13:47:01 +00:00
|
|
|
|
|
|
|
|
|
Σ≡ : a ≡ b
|
2018-04-05 08:41:56 +00:00
|
|
|
|
fst (Σ≡ i) = fst≡ i
|
|
|
|
|
snd (Σ≡ i) = snd≡ i
|
2018-04-06 14:54:00 +00:00
|
|
|
|
|
|
|
|
|
import Relation.Binary
|
|
|
|
|
open Relation.Binary public using (Preorder ; Transitive ; IsEquivalence ; Rel)
|
|
|
|
|
|
|
|
|
|
equalityIsEquivalence : {ℓ : Level} {A : Set ℓ} → IsEquivalence {A = A} _≡_
|
|
|
|
|
IsEquivalence.refl equalityIsEquivalence = refl
|
|
|
|
|
IsEquivalence.sym equalityIsEquivalence = sym
|
|
|
|
|
IsEquivalence.trans equalityIsEquivalence = trans
|
|
|
|
|
|
|
|
|
|
IsPreorder
|
|
|
|
|
: {a ℓ : Level} {A : Set a}
|
|
|
|
|
→ (_∼_ : Rel A ℓ) -- The relation.
|
|
|
|
|
→ Set _
|
|
|
|
|
IsPreorder _~_ = Relation.Binary.IsPreorder _≡_ _~_
|