cat/src/Cat/Exponential.agda

40 lines
1.4 KiB
Agda
Raw Normal View History

module Cat.Exponential where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cat.Category
open import Cat.Product
open Category
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open HasProducts hasProducts
open Product hiding (obj)
private
_×p_ : (A B : Object ) Object
_×p_ A B = Product.obj (product A B)
module _ (B C : Object ) where
IsExponential : (Cᴮ : Object ) [ Cᴮ ×p B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object ) (f : [ A ×p B , C ])
∃![ f~ ] ( [ eval parallelProduct f~ (Category.𝟙 )] f)
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : Object
eval : [ obj ×p B , C ]
{{isExponential}} : IsExponential obj eval
-- If I make this an instance-argument then the instance resolution
-- algorithm goes into an infinite loop. Why?
exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts
transpose : (A : Object ) [ A ×p B , C ] [ A , obj ]
transpose A f = proj₁ (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field
exponent : (A B : Object ) Exponential A B