Mercurial > hg > Members > kono > Proof > category
view CCCGraph1.agda @ 842:fa9d5d2b965d
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Thu, 02 Apr 2020 12:11:22 +0900 |
parents | 9fa1bf29fbf4 |
children | a73acfdef643 |
line wrap: on
line source
open import Level open import Category module CCCgraph1 where open import HomReasoning open import cat-utility open import Relation.Binary.PropositionalEquality hiding ( [_] ) open import CCC open import graph module ccc-from-graph {c₁ c₂ : Level} (G : Graph {c₁} {c₂} ) where open import Relation.Binary.PropositionalEquality hiding ( [_] ) open import Relation.Binary.Core open Graph data Objs : Set (c₁ ⊔ c₂) where atom : (vertex G) → Objs ⊤ : Objs _∧_ : Objs → Objs → Objs _<=_ : Objs → Objs → Objs data Arrow : Objs → Objs → Set (c₁ ⊔ c₂) where --- case i arrow : {a b : vertex G} → (edge G) a b → Arrow (atom a) (atom b) ○ : (a : Objs ) → Arrow a ⊤ π : {a b : Objs } → Arrow ( a ∧ b ) a π' : {a b : Objs } → Arrow ( a ∧ b ) b ε : {a b : Objs } → Arrow ((a <= b) ∧ b ) a _* : {a b c : Objs } → Arrow (c ∧ b ) a → Arrow c ( a <= b ) --- case v data Arrows : (b c : Objs ) → Set ( c₁ ⊔ c₂ ) where id : ( a : Objs ) → Arrows a a --- case i <_,_> : {a b c : Objs } → Arrows c a → Arrows c b → Arrows c (a ∧ b) --- case iii iv : {b c d : Objs } ( f : Arrow d c ) ( g : Arrows b d ) → Arrows b c -- cas iv _・_ : {a b c : Objs } (f : Arrows b c ) → (g : Arrows a b) → Arrows a c id a ・ g = g < f , g > ・ h = < f ・ h , g ・ h > iv f (id _) ・ h = iv f h iv (○ a) g ・ h = iv (○ _) (id _) iv π < g , g₁ > ・ h = g ・ h iv π' < g , g₁ > ・ h = g₁ ・ h iv ε < g , g₁ > ・ h = iv ε < g ・ h , g₁ ・ h > iv (f *) < g , g₁ > ・ h = iv (f *) < g ・ h , g₁ ・ h > iv f (iv f₁ g) ・ h = iv f ( (iv f₁ g) ・ h ) _==_ : {a b : Objs } → ( x y : Arrows a b ) → Set (c₁ ⊔ c₂) _==_ {a} {b} x y = ( x ・ id a ) ≡ ( y ・ id a ) PL : Category (c₁ ⊔ c₂) (c₁ ⊔ c₂) (c₁ ⊔ c₂) PL = record { Obj = Objs; Hom = λ a b → Arrows a b ; _o_ = λ{a} {b} {c} x y → x ・ y ; _≈_ = λ x y → x == y ; Id = λ{a} → id a ; isCategory = record { isEquivalence = record {refl = refl ; trans = trans ; sym = sym } ; identityL = λ {a b f} → identityL {a} {b} {f} ; identityR = λ {a b f} → identityR {a} {b} {f} ; o-resp-≈ = λ {a b c f g h i} → o-resp-≈ {a} {b} {c} {f} {g} {h} {i} ; associative = λ{a b c d f g h } → associative f g h } } where identityL : {A B : Objs} {f : Arrows A B} → (id B ・ f) == f identityL {_} {_} {id a} = refl identityL {a} {b} {< f , f₁ >} = refl identityL {_} {_} {iv f f₁} = refl identityR : {A B : Objs} {f : Arrows A B} → (f ・ id A) == f identityR {a} {_} {id a} = refl identityR {a} {b} {< f , g >} = cong₂ ( λ j k → < j , k > ) ( identityR {_} {_} {f} ) ( identityR {_} {_} {g} ) identityR {a} {b} {iv x (id a)} = refl identityR {a} {b} {iv π < f , f₁ >} = identityR {a} {b} {f} identityR {a} {b} {iv π' < f , f₁ >} = identityR {a} {b} {f₁} identityR {a} {⊤} {iv (○ .(_ ∧ _)) < f , f₁ >} = refl identityR {a} {b} {iv ε < f , f₁ >} = cong ( λ k → iv ε k ) ( identityR {_} {_} {< f , f₁ >} ) identityR {a} {_} {iv (x *) < f , f₁ >} = cong ( λ k → iv (x *) k ) ( identityR {_} {_} {< f , f₁ >} ) identityR {a} {b} {iv x (iv f f₁)} = {!!} -- cong ( λ k → iv x k ) ( identityR {_} {_} {iv f f₁} ) o-resp-≈ : {A B C : Objs} {f g : Arrows A B} {h i : Arrows B C} → f == g → h == i → (h ・ f) == (i ・ g) o-resp-≈ f=g h=i = {!!} associative : {a b c d : Objs} (f : Arrows c d) (g : Arrows b c) (h : Arrows a b) → (f ・ (g ・ h)) == ((f ・ g) ・ h) associative (id a) g h = refl associative (< f , f1 > ) g h = cong₂ ( λ j k → < j , k > ) (associative f g h) (associative f1 g h) associative (iv x f) g h = {!!} -- cong ( λ k → iv x k ) ( associative f g h )