Mercurial > hg > Members > kono > Proof > category
view freyd.agda @ 441:61550782be4a
preinital full subcategory done
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Tue, 30 Aug 2016 15:11:17 +0900 |
parents | ff36c500962e |
children | 87600d338337 |
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 -- codomain 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 ] -- 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 open Complete open Equalizer 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 ) (comp : Complete A A) (SFS : SmallFullSubcategory A F FMap← ) → (PI : PreInitial A F ) → { a0 : Obj A } → HasInitialObject A (limit-c comp F) initialFromPreInitialFullSubcategory A F FMap← comp SFS PI = record { initial = initialArrow ; uniqueness = λ a f → lemma1 a f } where lim : ( Γ : Functor A A ) → Limit A A Γ (limit-c comp Γ) (limit-u comp Γ) lim Γ = isLimit comp Γ equ : {a b : Obj A} → (f g : Hom A a b) → Equalizer A (equalizer-e comp f g ) f g equ f g = isEqualizer comp f g a0 = limit-c comp F u = limit-u comp F ee : {a b : Obj A} → {f g : Hom A a b} → Obj A ee {a} {b} {f} {g} = equalizer-p comp f g ep : {a b : Obj A} → {f g : Hom A a b} → Hom A (ee {a} {b} {f} {g} ) a ep {a} {b} {f} {g} = equalizer-e comp f g f : {a : Obj A} -> Hom A a0 (FObj F (preinitialObj PI {a} ) ) f {a} = TMap u (preinitialObj PI {a} ) initialArrow : ∀( a : Obj A ) → Hom A a0 a initialArrow a = A [ preinitialArrow PI {a} o f ] equ-fi : { a : Obj A} -> {f' : Hom A a0 a} -> Equalizer A ep ( A [ preinitialArrow PI {a} o f ] ) f' equ-fi {a} {f'} = equ ( A [ preinitialArrow PI {a} o f ] ) f' e=id : {e : Hom A a0 a0} -> ( {c : Obj A} -> A [ A [ TMap u c o e ] ≈ TMap u c ] ) -> A [ e ≈ id1 A a0 ] e=id {e} uce=uc = let open ≈-Reasoning (A) in begin e ≈↑⟨ limit-uniqueness (lim F) e ( \{i} -> uce=uc ) ⟩ limit (lim F) a0 u ≈⟨ limit-uniqueness (lim F) (id1 A a0) ( \{i} -> idR ) ⟩ id1 A a0 ∎ kfuc=uc : {c k1 : Obj A} -> {p : Hom A k1 a0} -> A [ A [ TMap u c o A [ p o A [ preinitialArrow PI {k1} o TMap u (preinitialObj PI) ] ] ] ≈ TMap u c ] kfuc=uc {c} {k1} {p} = let open ≈-Reasoning (A) in begin TMap u c o ( p o ( preinitialArrow PI {k1} o TMap u (preinitialObj PI) )) ≈⟨ cdr assoc ⟩ TMap u c o ((p o preinitialArrow PI) o TMap u (preinitialObj PI)) ≈⟨ assoc ⟩ (TMap u c o ( p o ( preinitialArrow PI {k1} ))) o TMap u (preinitialObj PI) ≈↑⟨ car ( full→ SFS ) ⟩ FMap F (FMap← (TMap u c o p o preinitialArrow PI)) o TMap u (preinitialObj PI) ≈⟨ nat u ⟩ TMap u c o FMap (K A A (limit-c comp F)) (FMap← (TMap u c o p o preinitialArrow PI)) ≈⟨⟩ TMap u c o id1 A (limit-c comp F) ≈⟨ idR ⟩ TMap u c ∎ kfuc=1 : {k1 : Obj A} -> {p : Hom A k1 a0} -> A [ A [ p o A [ preinitialArrow PI {k1} o TMap u (preinitialObj PI) ] ] ≈ id1 A a0 ] kfuc=1 {k1} {p} = e=id ( kfuc=uc ) -- if equalizer has right inverse, f = g lemma2 : (a b : Obj A) {c : Obj A} ( f g : Hom A a b ) {e : Hom A c a } {e' : Hom A a c } ( equ : Equalizer A e f g ) (inv-e : A [ A [ e o e' ] ≈ id1 A a ] ) -> A [ f ≈ g ] lemma2 a b {c} f g {e} {e'} equ inv-e = let open ≈-Reasoning (A) in let open Equalizer in begin f ≈↑⟨ idR ⟩ f o id1 A a ≈↑⟨ cdr inv-e ⟩ f o ( e o e' ) ≈⟨ assoc ⟩ ( f o e ) o e' ≈⟨ car ( fe=ge equ ) ⟩ ( g o e ) o e' ≈↑⟨ assoc ⟩ g o ( e o e' ) ≈⟨ cdr inv-e ⟩ g o id1 A a ≈⟨ idR ⟩ g ∎ 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 f ≈⟨ lemma2 a0 a (A [ preinitialArrow PI {a} o f ]) f' {ep {a0} {a} {A [ preinitialArrow PI {a} o f ]} {f'} } (equ-fi ) (kfuc=1 {ee} {ep} ) ⟩ f' ∎ )