Delete equality module
This commit is contained in:
parent
183906dc8c
commit
29f45d1426
|
@ -5,9 +5,6 @@ module Cat.Categories.Cat where
|
||||||
|
|
||||||
open import Cat.Prelude renaming (proj₁ to fst ; proj₂ to snd)
|
open import Cat.Prelude renaming (proj₁ to fst ; proj₂ to snd)
|
||||||
|
|
||||||
open import Cubical
|
|
||||||
open import Cubical.Sigma
|
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Category.Functor
|
open import Cat.Category.Functor
|
||||||
open import Cat.Category.Product
|
open import Cat.Category.Product
|
||||||
|
@ -15,9 +12,6 @@ open import Cat.Category.Exponential hiding (_×_ ; product)
|
||||||
open import Cat.Category.NaturalTransformation
|
open import Cat.Category.NaturalTransformation
|
||||||
open import Cat.Categories.Fun
|
open import Cat.Categories.Fun
|
||||||
|
|
||||||
open import Cat.Equality
|
|
||||||
open Equality.Data.Product
|
|
||||||
|
|
||||||
-- The category of categories
|
-- The category of categories
|
||||||
module _ (ℓ ℓ' : Level) where
|
module _ (ℓ ℓ' : Level) where
|
||||||
private
|
private
|
||||||
|
|
|
@ -1,21 +1,18 @@
|
||||||
{-# OPTIONS --allow-unsolved-metas #-}
|
{-# OPTIONS --allow-unsolved-metas #-}
|
||||||
module Cat.Categories.Cube where
|
module Cat.Categories.Cube where
|
||||||
|
|
||||||
|
open import Cat.Prelude
|
||||||
open import Level
|
open import Level
|
||||||
open import Data.Bool hiding (T)
|
open import Data.Bool hiding (T)
|
||||||
open import Data.Sum hiding ([_,_])
|
open import Data.Sum hiding ([_,_])
|
||||||
open import Data.Unit
|
open import Data.Unit
|
||||||
open import Data.Empty
|
open import Data.Empty
|
||||||
open import Data.Product
|
|
||||||
open import Cubical
|
|
||||||
open import Function
|
open import Function
|
||||||
open import Relation.Nullary
|
open import Relation.Nullary
|
||||||
open import Relation.Nullary.Decidable
|
open import Relation.Nullary.Decidable
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Category.Functor
|
open import Cat.Category.Functor
|
||||||
open import Cat.Equality
|
|
||||||
open Equality.Data.Product
|
|
||||||
|
|
||||||
-- See chapter 1 for a discussion on how presheaf categories are CwF's.
|
-- See chapter 1 for a discussion on how presheaf categories are CwF's.
|
||||||
|
|
||||||
|
|
|
@ -1,17 +1,10 @@
|
||||||
{-# OPTIONS --allow-unsolved-metas #-}
|
{-# OPTIONS --allow-unsolved-metas #-}
|
||||||
module Cat.Categories.Fam where
|
module Cat.Categories.Fam where
|
||||||
|
|
||||||
open import Agda.Primitive
|
open import Cat.Prelude
|
||||||
open import Data.Product
|
|
||||||
import Function
|
import Function
|
||||||
|
|
||||||
open import Cubical
|
|
||||||
open import Cubical.Universe
|
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Equality
|
|
||||||
|
|
||||||
open Equality.Data.Product
|
|
||||||
|
|
||||||
module _ (ℓa ℓb : Level) where
|
module _ (ℓa ℓb : Level) where
|
||||||
private
|
private
|
||||||
|
|
|
@ -2,14 +2,10 @@
|
||||||
|
|
||||||
module Cat.Category.Yoneda where
|
module Cat.Category.Yoneda where
|
||||||
|
|
||||||
open import Agda.Primitive
|
open import Cat.Prelude
|
||||||
open import Data.Product
|
|
||||||
open import Cubical
|
|
||||||
open import Cubical.NType.Properties
|
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Category.Functor
|
open import Cat.Category.Functor
|
||||||
open import Cat.Equality
|
|
||||||
|
|
||||||
open import Cat.Categories.Fun
|
open import Cat.Categories.Fun
|
||||||
open import Cat.Categories.Sets hiding (presheaf)
|
open import Cat.Categories.Sets hiding (presheaf)
|
||||||
|
|
|
@ -1,22 +0,0 @@
|
||||||
{-# OPTIONS --cubical #-}
|
|
||||||
-- Defines equality-principles for data-types from the standard library.
|
|
||||||
|
|
||||||
module Cat.Equality where
|
|
||||||
|
|
||||||
open import Level
|
|
||||||
open import Cubical
|
|
||||||
|
|
||||||
-- _[_≡_] = PathP
|
|
||||||
|
|
||||||
module Equality where
|
|
||||||
module Data where
|
|
||||||
module Product where
|
|
||||||
open import Data.Product
|
|
||||||
|
|
||||||
module _ {ℓa ℓb : Level} {A : Set ℓa} {B : A → Set ℓb} {a b : Σ A B}
|
|
||||||
(proj₁≡ : (λ _ → A) [ proj₁ a ≡ proj₁ b ])
|
|
||||||
(proj₂≡ : (λ i → B (proj₁≡ i)) [ proj₂ a ≡ proj₂ b ]) where
|
|
||||||
|
|
||||||
Σ≡ : a ≡ b
|
|
||||||
proj₁ (Σ≡ i) = proj₁≡ i
|
|
||||||
proj₂ (Σ≡ i) = proj₂≡ i
|
|
|
@ -24,6 +24,8 @@ open import Cubical.NType.Properties
|
||||||
( lemPropF ; lemSig ; lemSigP ; isSetIsProp
|
( lemPropF ; lemSig ; lemSigP ; isSetIsProp
|
||||||
; propPi ; propHasLevel ; setPi ; propSet)
|
; propPi ; propHasLevel ; setPi ; propSet)
|
||||||
public
|
public
|
||||||
|
open import Cubical.Sigma using (setSig) public
|
||||||
|
open import Cubical.Universe using (hSet) public
|
||||||
|
|
||||||
-----------------
|
-----------------
|
||||||
-- * Utilities --
|
-- * Utilities --
|
||||||
|
@ -38,3 +40,11 @@ open import Cubical.NType.Properties
|
||||||
∃!-syntax = ∃
|
∃!-syntax = ∃
|
||||||
|
|
||||||
syntax ∃!-syntax (λ x → B) = ∃![ x ] B
|
syntax ∃!-syntax (λ x → B) = ∃![ x ] B
|
||||||
|
|
||||||
|
module _ {ℓa ℓb : Level} {A : Set ℓa} {B : A → Set ℓb} {a b : Σ A B}
|
||||||
|
(proj₁≡ : (λ _ → A) [ proj₁ a ≡ proj₁ b ])
|
||||||
|
(proj₂≡ : (λ i → B (proj₁≡ i)) [ proj₂ a ≡ proj₂ b ]) where
|
||||||
|
|
||||||
|
Σ≡ : a ≡ b
|
||||||
|
proj₁ (Σ≡ i) = proj₁≡ i
|
||||||
|
proj₂ (Σ≡ i) = proj₂≡ i
|
||||||
|
|
Loading…
Reference in a new issue