Do not export helpers in Fun

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-20 16:50:03 +01:00
parent 629115661b
commit 31257a4d97
3 changed files with 46 additions and 42 deletions

View file

@ -69,6 +69,7 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
open RawCategory RawFun open RawCategory RawFun
open Univalence (λ {A} {B} {f} isIdentity {A} {B} {f}) open Univalence (λ {A} {B} {f} isIdentity {A} {B} {f})
private
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where
module A = Functor A module A = Functor A
module B = Functor B module B = Functor B

View file

@ -328,6 +328,8 @@ module _ { : Level} where
SetsHasProducts = record { product = product } SetsHasProducts = record { product = product }
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
open Category
-- Covariant Presheaf -- Covariant Presheaf
Representable : Set (a lsuc b) Representable : Set (a lsuc b)
Representable = Functor (𝓢𝓮𝓽 b) Representable = Functor (𝓢𝓮𝓽 b)
@ -336,8 +338,6 @@ module _ {a b : Level} ( : Category a b) where
Presheaf : Set (a lsuc b) Presheaf : Set (a lsuc b)
Presheaf = Functor (opposite ) (𝓢𝓮𝓽 b) Presheaf = Functor (opposite ) (𝓢𝓮𝓽 b)
open Category
-- The "co-yoneda" embedding. -- The "co-yoneda" embedding.
representable : Category.Object Representable representable : Category.Object Representable
representable A = record representable A = record

View file

@ -28,9 +28,12 @@ module _ { : Level} { : Category } where
private private
𝓢 = Sets 𝓢 = Sets
open Fun (opposite ) 𝓢 open Fun (opposite ) 𝓢
presheaf = Cat.Categories.Sets.presheaf
module = Category module = Category
presheaf : .Object Presheaf
presheaf = Cat.Categories.Sets.presheaf
module _ {A B : .Object} (f : [ A , B ]) where module _ {A B : .Object} (f : [ A , B ]) where
fmap : Transformation (presheaf A) (presheaf B) fmap : Transformation (presheaf A) (presheaf B)
fmap C x = [ f x ] fmap C x = [ f x ]