diff --git a/src/Cat/Categories/Fun.agda b/src/Cat/Categories/Fun.agda index 4c020cb..9bbb438 100644 --- a/src/Cat/Categories/Fun.agda +++ b/src/Cat/Categories/Fun.agda @@ -69,54 +69,55 @@ module Fun {ℓc ℓc' ℓd ℓd' : Level} (ℂ : Category ℓc ℓc') (𝔻 : C open RawCategory RawFun open Univalence (λ {A} {B} {f} → isIdentity {A} {B} {f}) - module _ {A B : Functor ℂ 𝔻} where - module A = Functor A - module B = Functor B - module _ (p : A ≡ B) where - omapP : A.omap ≡ B.omap - omapP i = Functor.omap (p i) + private + module _ {A B : Functor ℂ 𝔻} where + module A = Functor A + module B = Functor B + module _ (p : A ≡ B) where + omapP : A.omap ≡ B.omap + omapP i = Functor.omap (p i) - coerceAB : ∀ {X} → 𝔻 [ A.omap X , A.omap X ] ≡ 𝔻 [ A.omap X , B.omap X ] - coerceAB {X} = cong (λ φ → 𝔻 [ A.omap X , φ X ]) omapP + coerceAB : ∀ {X} → 𝔻 [ A.omap X , A.omap X ] ≡ 𝔻 [ A.omap X , B.omap X ] + coerceAB {X} = cong (λ φ → 𝔻 [ A.omap X , φ X ]) omapP - -- The transformation will be the identity on 𝔻. Such an arrow has the - -- type `A.omap A → A.omap A`. Which we can coerce to have the type - -- `A.omap → B.omap` since `A` and `B` are equal. - coe𝟙 : Transformation A B - coe𝟙 X = coe coerceAB 𝔻.𝟙 + -- The transformation will be the identity on 𝔻. Such an arrow has the + -- type `A.omap A → A.omap A`. Which we can coerce to have the type + -- `A.omap → B.omap` since `A` and `B` are equal. + coe𝟙 : Transformation A B + coe𝟙 X = coe coerceAB 𝔻.𝟙 - module _ {a b : ℂ.Object} (f : ℂ [ a , b ]) where - nat' : 𝔻 [ coe𝟙 b ∘ A.fmap f ] ≡ 𝔻 [ B.fmap f ∘ coe𝟙 a ] - nat' = begin - (𝔻 [ coe𝟙 b ∘ A.fmap f ]) ≡⟨ {!!} ⟩ - (𝔻 [ B.fmap f ∘ coe𝟙 a ]) ∎ + module _ {a b : ℂ.Object} (f : ℂ [ a , b ]) where + nat' : 𝔻 [ coe𝟙 b ∘ A.fmap f ] ≡ 𝔻 [ B.fmap f ∘ coe𝟙 a ] + nat' = begin + (𝔻 [ coe𝟙 b ∘ A.fmap f ]) ≡⟨ {!!} ⟩ + (𝔻 [ B.fmap f ∘ coe𝟙 a ]) ∎ - transs : (i : I) → Transformation A (p i) - transs = {!!} + transs : (i : I) → Transformation A (p i) + transs = {!!} - natt : (i : I) → Natural A (p i) {!!} - natt = {!!} + natt : (i : I) → Natural A (p i) {!!} + natt = {!!} - t : Natural A B coe𝟙 - t = coe c (identityNatural A) - where - c : Natural A A (identityTrans A) ≡ Natural A B coe𝟙 - c = begin - Natural A A (identityTrans A) ≡⟨ (λ x → {!natt ?!}) ⟩ - Natural A B coe𝟙 ∎ - -- cong (λ φ → {!Natural A A (identityTrans A)!}) {!!} + t : Natural A B coe𝟙 + t = coe c (identityNatural A) + where + c : Natural A A (identityTrans A) ≡ Natural A B coe𝟙 + c = begin + Natural A A (identityTrans A) ≡⟨ (λ x → {!natt ?!}) ⟩ + Natural A B coe𝟙 ∎ + -- cong (λ φ → {!Natural A A (identityTrans A)!}) {!!} - k : Natural A A (identityTrans A) → Natural A B coe𝟙 - k n {a} {b} f = res - where - res : (𝔻 [ coe𝟙 b ∘ A.fmap f ]) ≡ (𝔻 [ B.fmap f ∘ coe𝟙 a ]) - res = {!!} + k : Natural A A (identityTrans A) → Natural A B coe𝟙 + k n {a} {b} f = res + where + res : (𝔻 [ coe𝟙 b ∘ A.fmap f ]) ≡ (𝔻 [ B.fmap f ∘ coe𝟙 a ]) + res = {!!} - nat : Natural A B coe𝟙 - nat = nat' + nat : Natural A B coe𝟙 + nat = nat' - fromEq : NaturalTransformation A B - fromEq = coe𝟙 , nat + fromEq : NaturalTransformation A B + fromEq = coe𝟙 , nat module _ {A B : Functor ℂ 𝔻} where obverse : A ≡ B → A ≅ B diff --git a/src/Cat/Categories/Sets.agda b/src/Cat/Categories/Sets.agda index e7e2024..f5af047 100644 --- a/src/Cat/Categories/Sets.agda +++ b/src/Cat/Categories/Sets.agda @@ -328,6 +328,8 @@ module _ {ℓ : Level} where SetsHasProducts = record { product = product } module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where + open Category ℂ + -- Covariant Presheaf Representable : Set (ℓa ⊔ lsuc ℓb) Representable = Functor ℂ (𝓢𝓮𝓽 ℓb) @@ -336,8 +338,6 @@ module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where Presheaf : Set (ℓa ⊔ lsuc ℓb) Presheaf = Functor (opposite ℂ) (𝓢𝓮𝓽 ℓb) - open Category ℂ - -- The "co-yoneda" embedding. representable : Category.Object ℂ → Representable representable A = record diff --git a/src/Cat/Category/Yoneda.agda b/src/Cat/Category/Yoneda.agda index d00a3ec..9447023 100644 --- a/src/Cat/Category/Yoneda.agda +++ b/src/Cat/Category/Yoneda.agda @@ -28,9 +28,12 @@ module _ {ℓ : Level} {ℂ : Category ℓ ℓ} where private 𝓢 = Sets ℓ open Fun (opposite ℂ) 𝓢 - presheaf = Cat.Categories.Sets.presheaf ℂ + module ℂ = Category ℂ + presheaf : ℂ.Object → Presheaf ℂ + presheaf = Cat.Categories.Sets.presheaf ℂ + module _ {A B : ℂ.Object} (f : ℂ [ A , B ]) where fmap : Transformation (presheaf A) (presheaf B) fmap C x = ℂ [ f ∘ x ]