cat/src/Cat/Category/Properties.agda

125 lines
4.6 KiB
Agda
Raw Normal View History

2018-01-24 15:38:28 +00:00
{-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category.Properties where
open import Agda.Primitive
open import Data.Product
2018-01-30 10:19:48 +00:00
open import Cubical
open import Cat.Category
open import Cat.Functor
open import Cat.Categories.Sets
open import Cat.Equality
open Equality.Data.Product
module _ { ' : Level} { : Category '} { A B : Category.Object } {X : Category.Object } (f : Category.Arrow A B) where
open Category
open IsCategory (isCategory)
2018-02-02 14:33:54 +00:00
iso-is-epi : Isomorphism f Epimorphism {X = X} f
iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym (proj₁ ident)
g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv)
g₀ (f f-) ≡⟨ assoc
(g₀ f) f- ≡⟨ cong (λ φ φ f-) eq
(g₁ f) f- ≡⟨ sym assoc
g₁ (f f-) ≡⟨ cong (_∘_ g₁) right-inv
g₁ 𝟙 ≡⟨ proj₁ ident
g₁
2018-02-02 14:33:54 +00:00
iso-is-mono : Isomorphism f Monomorphism {X = X} f
iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (proj₂ ident)
𝟙 g₀ ≡⟨ cong (λ φ φ g₀) (sym left-inv)
(f- f) g₀ ≡⟨ sym assoc
f- (f g₀) ≡⟨ cong (_∘_ f-) eq
f- (f g₁) ≡⟨ assoc
(f- f) g₁ ≡⟨ cong (λ φ φ g₁) left-inv
𝟙 g₁ ≡⟨ proj₂ ident
g₁
2018-02-02 14:33:54 +00:00
iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso
{-
epi-mono-is-not-iso : { '} ¬ (( : Category {} {'}) {A B X : Object } (f : Arrow A B ) Epimorphism { = } {X = X} f Monomorphism { = } {X = X} f Isomorphism { = } f)
epi-mono-is-not-iso f =
let k = f {!!} {!!} {!!} {!!}
in {!!}
-}
open import Cat.Category
open Category
open import Cat.Functor
open Functor
2018-02-02 14:33:54 +00:00
-- module _ { : Level} { : Category }
-- {isSObj : isSet ( .Object)}
-- {isz2 : ∀ {} → {A B : Set } → isSet (Sets [ A , B ])} where
-- -- open import Cat.Categories.Cat using (Cat)
-- open import Cat.Categories.Fun
-- open import Cat.Categories.Sets
-- -- module Cat = Cat.Categories.Cat
-- open Exponential
-- private
-- Cat = Cat
-- prshf = presheaf { = }
-- module = IsCategory ( .isCategory)
-- -- Exp : Set (lsuc (lsuc ))
-- -- Exp = Exponential (Cat (lsuc ) )
-- -- Sets (Opposite )
-- _⇑_ : (A B : Cat .Object) → Cat .Object
-- A ⇑ B = (exponent A B) .obj
-- where
-- open HasExponentials (Cat.hasExponentials )
-- module _ {A B : .Object} (f : .Arrow A B) where
-- :func→: : NaturalTransformation (prshf A) (prshf B)
-- :func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .assoc
-- module _ {c : .Object} where
-- eqTrans : (λ _ → Transformation (prshf c) (prshf c))
-- [ (λ _ x → [ .𝟙 ∘ x ]) ≡ identityTrans (prshf c) ]
-- eqTrans = funExt λ x → funExt λ x → .ident .proj₂
-- eqNat : (λ i → Natural (prshf c) (prshf c) (eqTrans i))
-- [(λ _ → funExt (λ _ → .assoc)) ≡ identityNatural (prshf c)]
-- eqNat = λ i {A} {B} f →
-- let
-- open IsCategory (Sets .isCategory)
-- lemm : (Sets [ eqTrans i B ∘ prshf c .func→ f ]) ≡
-- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
-- lemm = {!!}
-- lem : (λ _ → Sets [ Functor.func* (prshf c) A , prshf c .func* B ])
-- [ Sets [ eqTrans i B ∘ prshf c .func→ f ]
-- ≡ Sets [ prshf c .func→ f ∘ eqTrans i A ] ]
-- lem
-- = isz2 _ _ lemm _ i
-- -- (Sets [ eqTrans i B ∘ prshf c .func→ f ])
-- -- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
-- -- lemm
-- -- _ i
-- in
-- lem
-- -- eqNat = λ {A} {B} i [B,A] i' [A,c] →
-- -- let
-- -- k : [ {!!} , {!!} ]
-- -- k = [A,c]
-- -- in {! [ ? ∘ ? ]!}
-- :ident: : (:func→: ( .𝟙 {c})) (Fun .𝟙 {o = prshf c})
-- :ident: = Σ≡ eqTrans eqNat
-- yoneda : Functor (Fun { = Opposite } {𝔻 = Sets {}})
-- yoneda = record
-- { func* = prshf
-- ; func→ = :func→:
-- ; isFunctor = record
-- { ident = :ident:
-- ; distrib = {!!}
-- }
-- }