cat/src/Cat/Category/Exponential.agda

43 lines
1.4 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Cat.Category.Exponential where
open import Cat.Prelude hiding (_×_)
open import Cat.Category
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.identity ] f)
IsExponential : (Cᴮ : Object) [ Cᴮ × B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Object) (f : [ A × B , C ])
∃![ f~ ] ( [ eval f~ |×| Category.identity ] 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 = fst (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
open Category
open Exponential public
field
exponent : (A B : Object) Exponential A B
_⇑_ : (A B : Object) Object
A B = (exponent A B) .obj