cat/src/Cat/Category/Exponential.agda

40 lines
1.4 KiB
Agda
Raw Normal View History

2018-02-05 13:59:53 +00:00
module Cat.Category.Exponential where
open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cat.Category
2018-02-05 13:59:53 +00:00
open import Cat.Category.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 ])
2018-02-05 15:35:33 +00:00
∃![ f~ ] ( [ eval 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