view limit-to.agda @ 353:d255a929815f

list representation of TwoCat Hom
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Wed, 01 Apr 2015 17:37:04 +0900
parents f589e71875ea
children 2a83718f50a0
line wrap: on
line source

open import Category -- https://github.com/konn/category-agda                                                                                     
open import Level

module limit-to {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) 
  where

open import cat-utility
open import HomReasoning
open import Relation.Binary.Core
open import Category.Sets
open Functor


-- If we have limit then we have equalizer                                                                                                                                                                  
---  two objects category
---
---          f
---       ------>
---     0         1
---       ------>
---          g

data  _∨_  {c₁} (A B : Set c₁): Set c₁ where
        or1 :  A → A ∨ B
        or2 :  B → A ∨ B


record TwoSet {c₁} (A : Set c₁): Set (suc c₁) where
    field
            a0 : A
            a1 : A
            element : (x : A) → ( x ≡ a0) ∨ (x ≡ a1 )
            

data Two {c₁}  : Set c₁ where
   t0 : Two 
   t1 : Two 


infixr 40 _::_
data  List {a} (A : Set a)  : Set a where
   [] : List A
   _::_ : A -> List A -> List A

-- record TwoCat {ℓ} (S : Sets {ℓ}) : Set ℓ where
--    field
--         twoSet : TwoSet S
--         hom : TwoSet S → TwoSet S → Set _


twocat : {c₁ c₂ ℓ : Level} → (id-0 id-1 f g : Set  c₂) → Category c₁ _ ℓ
twocat {c₁} {c₂} {ℓ} id-0 id-1 f g = record {
    Obj  = Two {c₁} ;
    Hom = λ a b →  List (Set  c₂) ;
    _o_ =  {!!} ;
    _≈_ = {!!} ;
    Id  = {!!} ;
    isCategory  = record {
            isEquivalence = {!!} ;
            identityL  = {!!} ;
            identityR  = {!!} ;
            o-resp-≈  = {!!} ;
            associative  = {!!}
       } 
   } where
        hom : Two {c₁} → Two {c₁} → List (Set  c₂)
        hom t0 t0 =  id-0 :: []
        hom t0 t1 =  f :: g :: []
        hom t1 t0 =  []
        hom t1 t1 =  id-1 :: []

open Limit

lim-to-equ : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (I : Category c₁ c₂ ℓ)
      (two : {!!}  ) 
      (lim : ( Γ : Functor I A ) → { a0 : Obj A } { u : NTrans I A ( K A I a0 ) Γ } → Limit A I Γ a0 u ) -- completeness
        →  {a b c : Obj A} (f g : Hom A a b)  → (e : Hom A c a ) → (fe=ge : A [ A [ f o e ] ≈ A [ g o e ] ] ) → Equalizer A e f g
lim-to-equ {c₁} A I two lim {a} {b} {c} f g e fe=ge = record {
        fe=ge =  fe=ge
        ; k = λ {d} h fh=gh → k {d} h fh=gh
        ; ek=h = λ {d} {h} {fh=gh} → {!!}
        ; uniqueness = λ {d} {h} {fh=gh} {k'} → {!!}
     } where
         Γobj :  Two {c₁} → Obj A
         Γobj t0 = a
         Γobj t1 = b
         Γmap :  Two {c₁} → Hom A a b
         Γmap t0 = f
         Γmap t1 = g
         Γ : Functor I A
         Γ = record {
            FObj  = λ x → Γobj {!!} ;
            FMap  = λ f → {!!} ;
            isFunctor = record {
                     ≈-cong  = {!!} ; 
                    identity = {!!} ;
                    distr = {!!}
            }
          }
         nat : (d : Obj A) → NTrans I A (K A I d) Γ
         nat d = record {
            TMap = λ x → {!!} ;
            isNTrans = record {
                commute = {!!}
            }
          }
         k : {d : Obj A}  (h : Hom A d a) → A [ A [ f  o  h ] ≈ A [ g  o h ] ] → Hom A d c
         k {d} h fh=gh  = limit (lim Γ {c} {nat c}) d (nat d)