cat/src/Cat/Category/Exponential.agda

43 lines
1.4 KiB
Agda
Raw Normal View History

2018-02-05 13:59:53 +00:00
module Cat.Category.Exponential where
2018-03-21 13:39:56 +00:00
open import Cat.Prelude hiding (_×_)
open import Cat.Category
2018-02-05 13:59:53 +00:00
open import Cat.Category.Product
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open Category
open HasProducts hasProducts public
module _ (B C : Object) where
record IsExponential'
(Cᴮ : Object)
(eval : [ Cᴮ × B , C ]) : Set ( ') where
field
uniq
: (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.𝟙 ] f)
IsExponential : (Cᴮ : Object) [ Cᴮ × B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object) (f : [ A × 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 × B , C ]
{{isExponential}} : IsExponential obj eval
transpose : (A : Object) [ A × B , C ] [ A , obj ]
transpose A f = proj₁ (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
open Category
2018-02-23 09:53:11 +00:00
open Exponential public
field
exponent : (A B : Object) Exponential A B
_⇑_ : (A B : Object) Object
A B = (exponent A B) .obj