Mercurial > hg > Members > kono > Proof > category
view freyd.agda @ 350:c483374f542b
try equalizer from limit
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Wed, 24 Dec 2014 12:00:16 +0900 |
parents | 0d7fa6fc5979 |
children | 9f014f34b988 |
line wrap: on
line source
open import Category -- https://github.com/konn/category-agda open import Level module freyd {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) where open import cat-utility open import HomReasoning open import Relation.Binary.Core open Functor -- C is small full subcategory of A ( C is image of F ) record SmallFullSubcategory {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (F : Functor A A ) ( FMap← : { a b : Obj A } → Hom A (FObj F a) (FObj F b ) → Hom A a b ) : Set (suc ℓ ⊔ (suc c₁ ⊔ suc c₂)) where field ≈→≡ : {a b : Obj A } → { x y : Hom A (FObj F a) (FObj F b) } → (x≈y : A [ FMap F x ≈ FMap F y ]) → FMap F x ≡ FMap F y -- co-comain of FMap is local small full→ : { a b : Obj A } { x : Hom A (FObj F a) (FObj F b) } → A [ FMap F ( FMap← x ) ≈ x ] full← : { a b : Obj A } { x : Hom A a b } → A [ FMap← ( FMap F x ) ≈ x ] -- pre-initial record PreInitial {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (F : Functor A A ) : Set (suc ℓ ⊔ (suc c₁ ⊔ suc c₂)) where field preinitialObj : ∀{ a : Obj A } → Obj A preinitialArrow : ∀{ a : Obj A } → Hom A ( FObj F (preinitialObj {a} )) a -- initial object record HasInitialObject {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (i : Obj A) : Set (suc ℓ ⊔ (suc c₁ ⊔ suc c₂)) where field initial : ∀( a : Obj A ) → Hom A i a uniqueness : ( a : Obj A ) → ( f : Hom A i a ) → A [ f ≈ initial a ] -- A complete catagory has initial object if it has PreInitial-SmallFullSubcategory open NTrans open Limit open SmallFullSubcategory open PreInitial initialFromPreInitialFullSubcategory : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (F : Functor A A ) ( FMap← : { a b : Obj A } → Hom A (FObj F a) (FObj F b ) → Hom A a b ) (lim : ( Γ : Functor A A ) → { a0 : Obj A } { u : NTrans A A ( K A A a0 ) Γ } → Limit A A Γ a0 u ) -- completeness (SFS : SmallFullSubcategory A F FMap← ) → (PI : PreInitial A F ) → { a0 : Obj A } → { u : (a : Obj A) → NTrans A A (K A A a) F } → HasInitialObject A a0 initialFromPreInitialFullSubcategory A F FMap← lim SFS PI {a0} {u} = record { initial = initialArrow ; uniqueness = λ a f → lemma1 a f } where initialArrow : ∀( a : Obj A ) → Hom A a0 a initialArrow a = A [ preinitialArrow PI {a} o limit (lim F {FObj F (preinitialObj PI)} {u (FObj F (preinitialObj PI))} ) a0 (u a0 ) ] lemma1 : (a : Obj A) (f : Hom A a0 a) → A [ f ≈ initialArrow a ] lemma1 a f = let open ≈-Reasoning (A) in sym ( begin initialArrow a ≈⟨⟩ preinitialArrow PI {a} o limit (lim F) a0 (u a0) ≈⟨ {!!} ⟩ f ∎ )