Mercurial > hg > Members > kono > Proof > category
annotate freyd2.agda @ 642:53f2a11474ee
on going ..
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 03 Jul 2017 08:41:01 +0900 |
parents | c65d08d85092 |
children | 5eb746702732 |
rev | line source |
---|---|
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
1 open import Category -- https://github.com/konn/category-agda |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
2 open import Level |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
3 open import Category.Sets renaming ( _o_ to _*_ ) |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
4 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
5 module freyd2 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
6 where |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
7 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
8 open import HomReasoning |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
9 open import cat-utility |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
10 open import Relation.Binary.Core |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
11 open import Function |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
12 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
13 ---------- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
14 -- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
15 -- a : Obj A |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
16 -- U : A → Sets |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
17 -- U ⋍ Hom (a,-) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
18 -- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
19 |
617
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
20 -- maybe this is a part of local smallness |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
21 postulate ≈-≡ : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) {a b : Obj A } { x y : Hom A a b } → (x≈y : A [ x ≈ y ]) → x ≡ y |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
22 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
23 import Relation.Binary.PropositionalEquality |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
24 -- Extensionality a b = {A : Set a} {B : A → Set b} {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g → ( λ x → f x ≡ λ x → g x ) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
25 postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Relation.Binary.PropositionalEquality.Extensionality c₂ c₂ |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
26 |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
27 |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
28 ---- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
29 -- |
617
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
30 -- Hom ( a, - ) is Object mapping in Yoneda Functor |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
31 -- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
32 ---- |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
33 |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
34 open NTrans |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
35 open Functor |
498
c981a2f0642f
UpreseveLimit detailing
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
497
diff
changeset
|
36 open Limit |
c981a2f0642f
UpreseveLimit detailing
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
497
diff
changeset
|
37 open IsLimit |
c981a2f0642f
UpreseveLimit detailing
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
497
diff
changeset
|
38 open import Category.Cat |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
39 |
616 | 40 Yoneda : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) (a : Obj A) → Functor A (Sets {c₂}) |
41 Yoneda {c₁} {c₂} {ℓ} A a = record { | |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
42 FObj = λ b → Hom A a b |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
43 ; FMap = λ {x} {y} (f : Hom A x y ) → λ ( g : Hom A a x ) → A [ f o g ] -- f : Hom A x y → Hom Sets (Hom A a x ) (Hom A a y) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
44 ; isFunctor = record { |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
45 identity = λ {b} → extensionality A ( λ x → lemma-y-obj1 {b} x ) ; |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
46 distr = λ {a} {b} {c} {f} {g} → extensionality A ( λ x → lemma-y-obj2 a b c f g x ) ; |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
47 ≈-cong = λ eq → extensionality A ( λ x → lemma-y-obj3 x eq ) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
48 } |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
49 } where |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
50 lemma-y-obj1 : {b : Obj A } → (x : Hom A a b) → A [ id1 A b o x ] ≡ x |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
51 lemma-y-obj1 {b} x = let open ≈-Reasoning A in ≈-≡ A idL |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
52 lemma-y-obj2 : (a₁ b c : Obj A) (f : Hom A a₁ b) (g : Hom A b c ) → (x : Hom A a a₁ )→ |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
53 A [ A [ g o f ] o x ] ≡ (Sets [ _[_o_] A g o _[_o_] A f ]) x |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
54 lemma-y-obj2 a₁ b c f g x = let open ≈-Reasoning A in ≈-≡ A ( begin |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
55 A [ A [ g o f ] o x ] |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
56 ≈↑⟨ assoc ⟩ |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
57 A [ g o A [ f o x ] ] |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
58 ≈⟨⟩ |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
59 ( λ x → A [ g o x ] ) ( ( λ x → A [ f o x ] ) x ) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
60 ∎ ) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
61 lemma-y-obj3 : {b c : Obj A} {f g : Hom A b c } → (x : Hom A a b ) → A [ f ≈ g ] → A [ f o x ] ≡ A [ g o x ] |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
62 lemma-y-obj3 {_} {_} {f} {g} x eq = let open ≈-Reasoning A in ≈-≡ A ( begin |
497
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
63 A [ f o x ] |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
64 ≈⟨ resp refl-hom eq ⟩ |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
65 A [ g o x ] |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
66 ∎ ) |
e8b85a05a6b2
add if U is iso to representable functor then preserve limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
67 |
609 | 68 -- Representable U ≈ Hom(A,-) |
502 | 69 |
609 | 70 record Representable { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) ( U : Functor A (Sets {c₂}) ) (a : Obj A) : Set (suc ℓ ⊔ (suc (suc c₂) ⊔ suc c₁ )) where |
502 | 71 field |
72 -- FObj U x : A → Set | |
609 | 73 -- FMap U f = Set → Set (locally small) |
502 | 74 -- λ b → Hom (a,b) : A → Set |
75 -- λ f → Hom (a,-) = λ b → Hom a b | |
76 | |
616 | 77 repr→ : NTrans A (Sets {c₂}) U (Yoneda A a ) |
78 repr← : NTrans A (Sets {c₂}) (Yoneda A a) U | |
79 reprId→ : {x : Obj A} → Sets [ Sets [ TMap repr→ x o TMap repr← x ] ≈ id1 (Sets {c₂}) (FObj (Yoneda A a) x )] | |
609 | 80 reprId← : {x : Obj A} → Sets [ Sets [ TMap repr← x o TMap repr→ x ] ≈ id1 (Sets {c₂}) (FObj U x)] |
608 | 81 |
609 | 82 open Representable |
608 | 83 open import freyd |
502 | 84 |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
85 _↓_ : { c₁ c₂ ℓ : Level} { c₁' c₂' ℓ' : Level} { A : Category c₁ c₂ ℓ } { B : Category c₁' c₂' ℓ' } |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
86 → ( F G : Functor A B ) |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
87 → Category (c₂' ⊔ c₁) (ℓ' ⊔ c₂) ℓ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
88 _↓_ { c₁} {c₂} {ℓ} {c₁'} {c₂'} {ℓ'} { A } { B } F G = CommaCategory |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
89 where open import Comma1 F G |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
90 |
641 | 91 natf : { c₁ c₂ ℓ : Level} { c₁' c₂' ℓ' : Level} { A : Category c₁ c₂ ℓ } { B : Category c₁' c₂' ℓ' } |
92 → { F G : Functor A B } | |
93 → Functor A B → Functor A (F ↓ G) → Functor A B | |
94 natf { c₁} {c₂} {ℓ} {c₁'} {c₂'} {ℓ'} { A } { B } {F} {G} H I = nat-f H I | |
95 where open import Comma1 F G | |
96 | |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
97 open import freyd |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
98 open SmallFullSubcategory |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
99 open Complete |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
100 open PreInitial |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
101 open HasInitialObject |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
102 open import Comma1 |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
103 open CommaObj |
609 | 104 open LimitPreserve |
608 | 105 |
609 | 106 -- Representable Functor U preserve limit , so K{*}↓U is complete |
610 | 107 -- |
616 | 108 -- Yoneda A b = λ a → Hom A a b : Functor A Sets |
617
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
109 -- : Functor Sets A |
610 | 110 |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
111 YonedaFpreserveLimit0 : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (I : Category c₁ c₂ ℓ) |
612 | 112 (b : Obj A ) |
610 | 113 (Γ : Functor I A) (limita : Limit A I Γ) → |
616 | 114 IsLimit Sets I (Yoneda A b ○ Γ) (FObj (Yoneda A b) (a0 limita)) (LimitNat A I Sets Γ (a0 limita) (t0 limita) (Yoneda A b)) |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
115 YonedaFpreserveLimit0 {c₁} {c₂} {ℓ} A I b Γ lim = record { |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
116 limit = λ a t → ψ a t |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
117 ; t0f=t = λ {a t i} → t0f=t0 a t i |
614
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
118 ; limit-uniqueness = λ {b} {t} {f} t0f=t → limit-uniqueness0 {b} {t} {f} t0f=t |
610 | 119 } where |
616 | 120 hat0 : NTrans I Sets (K Sets I (FObj (Yoneda A b) (a0 lim))) (Yoneda A b ○ Γ) |
121 hat0 = LimitNat A I Sets Γ (a0 lim) (t0 lim) (Yoneda A b) | |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
122 haa0 : Obj Sets |
616 | 123 haa0 = FObj (Yoneda A b) (a0 lim) |
124 ta : (a : Obj Sets) ( x : a ) ( t : NTrans I Sets (K Sets I a) (Yoneda A b ○ Γ)) → NTrans I A (K A I b ) Γ | |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
125 ta a x t = record { TMap = λ i → (TMap t i ) x ; isNTrans = record { commute = commute1 } } where |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
126 commute1 : {a₁ b₁ : Obj I} {f : Hom I a₁ b₁} → |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
127 A [ A [ FMap Γ f o TMap t a₁ x ] ≈ A [ TMap t b₁ x o FMap (K A I b) f ] ] |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
128 commute1 {a₁} {b₁} {f} = let open ≈-Reasoning A in begin |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
129 FMap Γ f o TMap t a₁ x |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
130 ≈⟨⟩ |
616 | 131 ( ( FMap (Yoneda A b ○ Γ ) f ) * TMap t a₁ ) x |
611
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
132 ≈⟨ ≡-≈ ( cong (λ k → k x ) (IsNTrans.commute (isNTrans t)) ) ⟩ |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
133 ( TMap t b₁ * ( FMap (K Sets I a) f ) ) x |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
134 ≈⟨⟩ |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
135 ( TMap t b₁ * id1 Sets a ) x |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
136 ≈⟨⟩ |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
137 TMap t b₁ x |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
138 ≈↑⟨ idR ⟩ |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
139 TMap t b₁ x o id1 A b |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
140 ≈⟨⟩ |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
141 TMap t b₁ x o FMap (K A I b) f |
b1b5c6b4c570
natural transformation in representable functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
610
diff
changeset
|
142 ∎ |
616 | 143 ψ : (X : Obj Sets) ( t : NTrans I Sets (K Sets I X) (Yoneda A b ○ Γ)) |
144 → Hom Sets X (FObj (Yoneda A b) (a0 lim)) | |
145 ψ X t x = FMap (Yoneda A b) (limit (isLimit lim) b (ta X x t )) (id1 A b ) | |
146 t0f=t0 : (a : Obj Sets ) ( t : NTrans I Sets (K Sets I a) (Yoneda A b ○ Γ)) (i : Obj I) | |
147 → Sets [ Sets [ TMap (LimitNat A I Sets Γ (a0 lim) (t0 lim) (Yoneda A b)) i o ψ a t ] ≈ TMap t i ] | |
612 | 148 t0f=t0 a t i = let open ≈-Reasoning A in extensionality A ( λ x → ≈-≡ A ( begin |
616 | 149 ( Sets [ TMap (LimitNat A I Sets Γ (a0 lim) (t0 lim) (Yoneda A b)) i o ψ a t ] ) x |
612 | 150 ≈⟨⟩ |
616 | 151 FMap (Yoneda A b) ( TMap (t0 lim) i) (FMap (Yoneda A b) (limit (isLimit lim) b (ta a x t )) (id1 A b )) |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
152 ≈⟨⟩ -- FMap (Hom A b ) f g = A [ f o g ] |
613 | 153 TMap (t0 lim) i o (limit (isLimit lim) b (ta a x t ) o id1 A b ) |
154 ≈⟨ cdr idR ⟩ | |
155 TMap (t0 lim) i o limit (isLimit lim) b (ta a x t ) | |
156 ≈⟨ t0f=t (isLimit lim) ⟩ | |
157 TMap (ta a x t) i | |
158 ≈⟨⟩ | |
612 | 159 TMap t i x |
160 ∎ )) | |
616 | 161 limit-uniqueness0 : {a : Obj Sets} {t : NTrans I Sets (K Sets I a) (Yoneda A b ○ Γ)} {f : Hom Sets a (FObj (Yoneda A b) (a0 lim))} → |
162 ({i : Obj I} → Sets [ Sets [ TMap (LimitNat A I Sets Γ (a0 lim) (t0 lim) (Yoneda A b)) i o f ] ≈ TMap t i ]) → | |
614
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
163 Sets [ ψ a t ≈ f ] |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
164 limit-uniqueness0 {a} {t} {f} t0f=t = let open ≈-Reasoning A in extensionality A ( λ x → ≈-≡ A ( begin |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
165 ψ a t x |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
166 ≈⟨⟩ |
616 | 167 FMap (Yoneda A b) (limit (isLimit lim) b (ta a x t )) (id1 A b ) |
614
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
168 ≈⟨⟩ |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
169 limit (isLimit lim) b (ta a x t ) o id1 A b |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
170 ≈⟨ idR ⟩ |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
171 limit (isLimit lim) b (ta a x t ) |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
172 ≈⟨ limit-uniqueness (isLimit lim) ( λ {i} → ≡-≈ ( cong ( λ g → g x )( t0f=t {i} ))) ⟩ |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
173 f x |
e6be03d94284
Representational Functor preserve limit done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
613
diff
changeset
|
174 ∎ )) |
610 | 175 |
609 | 176 |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
177 YonedaFpreserveLimit : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (I : Category c₁ c₂ ℓ) |
616 | 178 (b : Obj A ) → LimitPreserve A I Sets (Yoneda A b) |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
179 YonedaFpreserveLimit A I b = record { |
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
180 preserve = λ Γ lim → YonedaFpreserveLimit0 A I b Γ lim |
610 | 181 } |
609 | 182 |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
183 |
608 | 184 -- K{*}↓U has preinitial full subcategory if U is representable |
609 | 185 -- if U is representable, K{*}↓U has initial Object ( so it has preinitial full subcategory ) |
608 | 186 |
617
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
187 open CommaHom |
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
188 |
627 | 189 data * {c : Level} : Set c where |
190 OneObj : * | |
191 | |
609 | 192 KUhasInitialObj : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) |
608 | 193 (a : Obj A ) |
628 | 194 → HasInitialObject ( K (Sets) A * ↓ (Yoneda A a) ) ( record { obj = a ; hom = λ x → id1 A a } ) |
621 | 195 KUhasInitialObj {c₁} {c₂} {ℓ} A a = record { |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
196 initial = λ b → initial0 b |
636 | 197 ; uniqueness = λ f → unique f |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
198 } where |
621 | 199 commaCat : Category (c₂ ⊔ c₁) c₂ ℓ |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
200 commaCat = K Sets A * ↓ Yoneda A a |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
201 initObj : Obj (K Sets A * ↓ Yoneda A a) |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
202 initObj = record { obj = a ; hom = λ x → id1 A a } |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
203 comm2 : (b : Obj commaCat) ( x : * ) → ( Sets [ FMap (Yoneda A a) (hom b OneObj) o (λ x₁ → id1 A a) ] ) x ≡ hom b x |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
204 comm2 b OneObj = let open ≈-Reasoning A in ≈-≡ A ( begin |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
205 ( Sets [ FMap (Yoneda A a) (hom b OneObj) o (λ x₁ → id1 A a) ] ) OneObj |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
206 ≈⟨⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
207 FMap (Yoneda A a) (hom b OneObj) (id1 A a) |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
208 ≈⟨⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
209 hom b OneObj o id1 A a |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
210 ≈⟨ idR ⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
211 hom b OneObj |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
212 ∎ ) |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
213 comm1 : (b : Obj commaCat) → Sets [ Sets [ FMap (Yoneda A a) (hom b OneObj) o hom initObj ] ≈ Sets [ hom b o FMap (K Sets A *) (hom b OneObj) ] ] |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
214 comm1 b = let open ≈-Reasoning Sets in begin |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
215 FMap (Yoneda A a) (hom b OneObj) o ( λ x → id1 A a ) |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
216 ≈⟨ extensionality A ( λ x → comm2 b x ) ⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
217 hom b |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
218 ≈⟨⟩ |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
219 hom b o FMap (K Sets A *) (hom b OneObj) |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
220 ∎ |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
221 initial0 : (b : Obj commaCat) → |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
222 Hom commaCat initObj b |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
223 initial0 b = record { |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
224 arrow = hom b OneObj |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
225 ; comm = comm1 b } |
625 | 226 -- what is comm f ? |
227 comm-f : (b : Obj (K Sets A * ↓ (Yoneda A a))) (f : Hom (K Sets A * ↓ Yoneda A a) initObj b) | |
228 → Sets [ Sets [ FMap (Yoneda A a) (arrow f) o ( λ x → id1 A a ) ] | |
229 ≈ Sets [ hom b o FMap (K Sets A *) (arrow f) ] ] | |
230 comm-f b f = comm f | |
636 | 231 unique : {b : Obj (K Sets A * ↓ Yoneda A a)} (f : Hom (K Sets A * ↓ Yoneda A a) initObj b) |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
232 → (K Sets A * ↓ Yoneda A a) [ f ≈ initial0 b ] |
636 | 233 unique {b} f = let open ≈-Reasoning A in begin |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
234 arrow f |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
235 ≈↑⟨ idR ⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
236 arrow f o id1 A a |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
237 ≈⟨⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
238 ( Sets [ FMap (Yoneda A a) (arrow f) o id1 Sets (FObj (Yoneda A a) a) ] ) (id1 A a) |
625 | 239 ≈⟨⟩ |
240 ( Sets [ FMap (Yoneda A a) (arrow f) o ( λ x → id1 A a ) ] ) OneObj | |
241 ≈⟨ ≡-≈ ( cong (λ k → k OneObj ) (comm f )) ⟩ | |
624
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
242 ( Sets [ hom b o FMap (K Sets A *) (arrow f) ] ) OneObj |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
243 ≈⟨⟩ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
244 hom b OneObj |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
245 ∎ |
9b9bc1e076f3
introduce one element set
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
623
diff
changeset
|
246 |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
247 |
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
248 |
633 | 249 -- A is complete and K{*}↓U has preinitial full subcategory then U is representable |
615
a45c32ceca97
initial Object's arrow found
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
614
diff
changeset
|
250 |
617
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
251 open SmallFullSubcategory |
34540494fdcf
initital obj uniquness done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
616
diff
changeset
|
252 open PreInitial |
626 | 253 |
638 | 254 -- if U preserve limit, K{*}↓U has initial object from freyd.agda |
255 | |
626 | 256 ≡-cong = Relation.Binary.PropositionalEquality.cong |
257 | |
638 | 258 |
259 ub : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (U : Functor A (Sets {c₂}) )(b : Obj A) (x : FObj U b ) | |
260 → Hom Sets (FObj (K Sets A *) b) (FObj U b) | |
261 ub A U b x OneObj = x | |
262 ob : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (U : Functor A (Sets {c₂}) )(b : Obj A) (x : FObj U b ) | |
263 → Obj ( K Sets A * ↓ U) | |
264 ob A U b x = record { obj = b ; hom = ub A U b x} | |
265 fArrow : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (U : Functor A (Sets {c₂}) ) {a b : Obj A} (f : Hom A a b) (x : FObj U a ) | |
266 → Hom ( K Sets A * ↓ U) ( ob A U a x ) (ob A U b (FMap U f x) ) | |
267 fArrow A U {a} {b} f x = record { arrow = f ; comm = fArrowComm a b f x } | |
268 where | |
269 fArrowComm1 : (a b : Obj A) (f : Hom A a b) (x : FObj U a ) → (y : * ) → FMap U f ( ub A U a x y ) ≡ ub A U b (FMap U f x) y | |
270 fArrowComm1 a b f x OneObj = refl | |
271 fArrowComm : (a b : Obj A) (f : Hom A a b) (x : FObj U a ) → | |
272 Sets [ Sets [ FMap U f o hom (ob A U a x) ] ≈ Sets [ hom (ob A U b (FMap U f x)) o FMap (K Sets A *) f ] ] | |
273 fArrowComm a b f x = extensionality Sets ( λ y → begin | |
274 ( Sets [ FMap U f o hom (ob A U a x) ] ) y | |
275 ≡⟨⟩ | |
276 FMap U f ( hom (ob A U a x) y ) | |
277 ≡⟨⟩ | |
278 FMap U f ( ub A U a x y ) | |
279 ≡⟨ fArrowComm1 a b f x y ⟩ | |
280 ub A U b (FMap U f x) y | |
281 ≡⟨⟩ | |
282 hom (ob A U b (FMap U f x)) y | |
283 ∎ ) where | |
284 open import Relation.Binary.PropositionalEquality | |
285 open ≡-Reasoning | |
286 | |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
287 UpreserveLimit : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (I : Category c₁ c₂ ℓ) ( comp : Complete A I ) |
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
288 (U : Functor A (Sets {c₂}) ) |
641 | 289 (SFS : SmallFullSubcategory (K (Sets {c₂}) A * ↓ U) ) |
290 (PI : PreInitial (K (Sets) A * ↓ U) (SFSF SFS)) | |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
291 → LimitPreserve A I Sets U |
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
292 UpreserveLimit A I comp U SFS PI = record { |
638 | 293 preserve = λ Γ lim → limitInSets Γ lim |
294 } where | |
295 limitInSets : (Γ : Functor I A) (lim : Limit A I Γ) → | |
296 IsLimit Sets I (U ○ Γ) (FObj U (a0 lim)) (LimitNat A I Sets Γ (a0 lim) (t0 lim) U) | |
297 limitInSets Γ lim = record { | |
298 limit = λ a t → ψ a t | |
299 ; t0f=t = λ {a t i} → t0f=t0 {a} {t} {i} | |
300 ; limit-uniqueness = λ {b} {t} {f} t0f=t → limit-uniqueness0 {b} {t} {f} t0f=t | |
301 } where | |
642 | 302 revUub : (pi : FObj U (obj (preinitialObj PI)) ) → pi ≡ (hom (preinitialObj PI) OneObj) |
303 revUub _ = {!!} | |
304 revU' : (a : Obj (K Sets A * ↓ U)) | |
305 → Sets [ Sets [ FMap U ( arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a })) ) o hom (preinitialObj PI) ] ≈ hom a ] | |
306 revU' a = let open ≈-Reasoning Sets in begin | |
307 FMap U ( arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a }))) o hom (preinitialObj PI) | |
308 ≈⟨ comm (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a })) ⟩ | |
309 hom a | |
310 ∎ | |
641 | 311 revU : (a : Obj (K Sets A * ↓ U)) |
312 → Sets [ FMap U ( arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a })) ) ≈ ( λ upi → hom a OneObj ) ] | |
642 | 313 revU a = extensionality Sets ( λ (upi : FObj U (obj (preinitialObj PI)) ) → ( begin |
314 FMap U ( arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a }))) upi | |
641 | 315 ≡⟨ {!!} ⟩ |
316 FMap U ( arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a }))) (hom (preinitialObj PI) OneObj) | |
317 ≡⟨ ≡-cong ( λ k → k OneObj ) ( comm (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) a }))) ⟩ | |
318 hom a OneObj | |
319 ∎ ) ) where | |
320 open import Relation.Binary.PropositionalEquality | |
321 open ≡-Reasoning | |
639 | 322 tacomm0 : (a : Obj Sets) ( t : NTrans I Sets (K Sets I a) (U ○ Γ) ) (x : a) {y : Obj I} {z : Obj I} {f : Hom I y z} |
323 → Sets [ Sets [ FMap (U ○ Γ) f o TMap t y ] ≈ Sets [ TMap t z o FMap ( K Sets I a ) f ] ] | |
324 tacomm0 a t x {y} {z} {f} = IsNTrans.commute ( isNTrans t ) {y} {z} {f} | |
638 | 325 tacomm : (a : Obj Sets) ( t : NTrans I Sets (K Sets I a) (U ○ Γ) ) (x : a) {y : Obj I} {z : Obj I} {f : Hom I y z} |
326 → A [ A [ FMap Γ f o arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ y) (TMap t y x))})) ] ≈ | |
327 A [ arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ z) (TMap t z x))} )) | |
328 o FMap (K A I (obj (preinitialObj PI))) f ] ] | |
329 tacomm a t x {y} {z} {f} = let open ≈-Reasoning A in begin | |
330 FMap Γ f o arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ y) (TMap t y x))})) | |
331 ≈⟨ {!!} ⟩ | |
640 | 332 arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ z) (FMap U (FMap Γ f) (TMap t y x)))} )) |
333 ≈⟨ {!!} ⟩ | |
638 | 334 arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ z) (TMap t z x))} )) |
335 ≈↑⟨ idR ⟩ | |
336 arrow (SFSFMap← SFS (preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ z) (TMap t z x))} )) | |
337 o FMap (K A I (obj (preinitialObj PI))) f | |
338 ∎ | |
339 ta : (a : Obj Sets) ( t : NTrans I Sets (K Sets I a) (U ○ Γ) ) (x : a) → NTrans I A (K A I (obj (preinitialObj PI))) Γ | |
340 ta a t x = record { TMap = λ (a : Obj I ) → | |
341 arrow ( SFSFMap← SFS ( preinitialArrow PI {FObj (SFSF SFS) (ob A U (FObj Γ a) (TMap t a x))} ) ) | |
639 | 342 ; isNTrans = record { commute = λ {y} {z} {f} → tacomm a t x {y} {z} {f} }} |
638 | 343 ψ : (a : Obj Sets) → NTrans I Sets (K Sets I a) (U ○ Γ) → Hom Sets a (FObj U (a0 lim)) |
344 ψ a t x = FMap U (limit (isLimit lim) (obj (preinitialObj PI)) (ta a t x)) ( hom (preinitialObj PI) OneObj ) | |
345 t0f=t0 : {a : Obj Sets} {t : NTrans I Sets (K Sets I a) (U ○ Γ)} {i : Obj I} → | |
346 Sets [ Sets [ TMap (LimitNat A I Sets Γ (a0 lim) (t0 lim) U) i o ψ a t ] ≈ TMap t i ] | |
347 t0f=t0 {a} {t} = {!!} | |
348 limit-uniqueness0 : {a : Obj Sets} {t : NTrans I Sets (K Sets I a) (U ○ Γ)} {f : Hom Sets a (FObj U (a0 lim))} → | |
349 ({i : Obj I} → Sets [ Sets [ TMap (LimitNat A I Sets Γ (a0 lim) (t0 lim) U) i o f ] ≈ TMap t i ]) → Sets [ ψ a t ≈ f ] | |
350 limit-uniqueness0 {a} {t} {f} = {!!} | |
635
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
351 |
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
352 -- if K{*}↓U has initial Obj, U is representable |
f7cc0ec00e05
introduce U preserving
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
634
diff
changeset
|
353 |
636 | 354 UisRepresentable : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) |
626 | 355 (U : Functor A (Sets {c₂}) ) |
636 | 356 ( i : Obj ( K (Sets) A * ↓ U) ) |
357 (In : HasInitialObject ( K (Sets) A * ↓ U) i ) | |
358 → Representable A U (obj i) | |
359 UisRepresentable A U i In = record { | |
627 | 360 repr→ = record { TMap = tmap1 ; isNTrans = record { commute = comm1 } } |
626 | 361 ; repr← = record { TMap = tmap2 ; isNTrans = record { commute = comm2 } } |
638 | 362 ; reprId→ = iso→ |
363 ; reprId← = iso← | |
626 | 364 } where |
638 | 365 comm11 : (a b : Obj A) (f : Hom A a b) (y : FObj U a ) → |
366 ( Sets [ FMap (Yoneda A (obj i)) f o ( λ x → arrow (initial In (ob A U a x))) ] ) y | |
367 ≡ (Sets [ ( λ x → arrow (initial In (ob A U b x))) o FMap U f ] ) y | |
368 comm11 a b f y = begin | |
369 ( Sets [ FMap (Yoneda A (obj i)) f o ( λ x → arrow (initial In (ob A U a x))) ] ) y | |
631 | 370 ≡⟨⟩ |
638 | 371 A [ f o arrow (initial In (ob A U a y)) ] |
631 | 372 ≡⟨⟩ |
638 | 373 A [ arrow ( fArrow A U f y ) o arrow (initial In (ob A U a y)) ] |
374 ≡⟨ ≈-≡ A ( uniqueness In {ob A U b (FMap U f y) } (( K Sets A * ↓ U) [ fArrow A U f y o initial In (ob A U a y)] ) ) ⟩ | |
375 arrow (initial In (ob A U b (FMap U f y) )) | |
629 | 376 ≡⟨⟩ |
638 | 377 (Sets [ ( λ x → arrow (initial In (ob A U b x))) o FMap U f ] ) y |
629 | 378 ∎ where |
379 open import Relation.Binary.PropositionalEquality | |
380 open ≡-Reasoning | |
636 | 381 tmap1 : (b : Obj A) → Hom Sets (FObj U b) (FObj (Yoneda A (obj i)) b) |
638 | 382 tmap1 b x = arrow ( initial In (ob A U b x ) ) |
636 | 383 comm1 : {a b : Obj A} {f : Hom A a b} → Sets [ Sets [ FMap (Yoneda A (obj i)) f o tmap1 a ] ≈ Sets [ tmap1 b o FMap U f ] ] |
626 | 384 comm1 {a} {b} {f} = let open ≈-Reasoning Sets in begin |
636 | 385 FMap (Yoneda A (obj i)) f o tmap1 a |
629 | 386 ≈⟨⟩ |
638 | 387 FMap (Yoneda A (obj i)) f o ( λ x → arrow (initial In ( ob A U a x ))) |
629 | 388 ≈⟨ extensionality Sets ( λ y → comm11 a b f y ) ⟩ |
638 | 389 ( λ x → arrow (initial In (ob A U b x))) o FMap U f |
629 | 390 ≈⟨⟩ |
626 | 391 tmap1 b o FMap U f |
392 ∎ | |
636 | 393 comm21 : (a b : Obj A) (f : Hom A a b) ( y : Hom A (obj i) a ) → |
394 (Sets [ FMap U f o (λ x → FMap U x (hom i OneObj))] ) y ≡ | |
395 (Sets [ ( λ x → (FMap U x ) (hom i OneObj)) o (λ x → A [ f o x ] ) ] ) y | |
626 | 396 comm21 a b f y = begin |
636 | 397 FMap U f ( FMap U y (hom i OneObj)) |
398 ≡⟨ ≡-cong ( λ k → k (hom i OneObj)) ( sym ( IsFunctor.distr (isFunctor U ) ) ) ⟩ | |
399 (FMap U (A [ f o y ] ) ) (hom i OneObj) | |
626 | 400 ∎ where |
401 open import Relation.Binary.PropositionalEquality | |
402 open ≡-Reasoning | |
636 | 403 tmap2 : (b : Obj A) → Hom Sets (FObj (Yoneda A (obj i)) b) (FObj U b) |
404 tmap2 b x = ( FMap U x ) ( hom i OneObj ) | |
626 | 405 comm2 : {a b : Obj A} {f : Hom A a b} → Sets [ Sets [ FMap U f o tmap2 a ] ≈ |
636 | 406 Sets [ tmap2 b o FMap (Yoneda A (obj i)) f ] ] |
626 | 407 comm2 {a} {b} {f} = let open ≈-Reasoning Sets in begin |
408 FMap U f o tmap2 a | |
409 ≈⟨⟩ | |
636 | 410 FMap U f o ( λ x → ( FMap U x ) ( hom i OneObj ) ) |
626 | 411 ≈⟨ extensionality Sets ( λ y → comm21 a b f y ) ⟩ |
636 | 412 ( λ x → ( FMap U x ) ( hom i OneObj ) ) o ( λ x → A [ f o x ] ) |
413 ≈⟨⟩ | |
414 ( λ x → ( FMap U x ) ( hom i OneObj ) ) o FMap (Yoneda A (obj i)) f | |
415 ≈⟨⟩ | |
416 tmap2 b o FMap (Yoneda A (obj i)) f | |
637
946ea019a2e7
if K{*}↓U has initial Obj, U is representable done.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
636
diff
changeset
|
417 ∎ |
946ea019a2e7
if K{*}↓U has initial Obj, U is representable done.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
636
diff
changeset
|
418 iso0 : ( x : Obj A) ( y : Hom A (obj i ) x ) ( z : * ) |
638 | 419 → ( Sets [ FMap U y o hom i ] ) z ≡ ( Sets [ ub A U x (FMap U y (hom i OneObj)) o FMap (K Sets A *) y ] ) z |
637
946ea019a2e7
if K{*}↓U has initial Obj, U is representable done.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
636
diff
changeset
|
420 iso0 x y OneObj = refl |
636 | 421 iso→ : {x : Obj A} → Sets [ Sets [ tmap1 x o tmap2 x ] ≈ id1 Sets (FObj (Yoneda A (obj i)) x) ] |
422 iso→ {x} = let open ≈-Reasoning A in extensionality Sets ( λ ( y : Hom A (obj i ) x ) → ≈-≡ A ( begin | |
423 ( Sets [ tmap1 x o tmap2 x ] ) y | |
626 | 424 ≈⟨⟩ |
638 | 425 arrow ( initial In (ob A U x (( FMap U y ) ( hom i OneObj ) ))) |
637
946ea019a2e7
if K{*}↓U has initial Obj, U is representable done.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
636
diff
changeset
|
426 ≈↑⟨ uniqueness In (record { arrow = y ; comm = extensionality Sets ( λ (z : * ) → iso0 x y z ) } ) ⟩ |
636 | 427 y |
428 ∎ )) | |
429 iso← : {x : Obj A} → Sets [ Sets [ tmap2 x o tmap1 x ] ≈ id1 Sets (FObj U x) ] | |
430 iso← {x} = extensionality Sets ( λ (y : FObj U x ) → ( begin | |
431 ( Sets [ tmap2 x o tmap1 x ] ) y | |
432 ≡⟨⟩ | |
638 | 433 ( FMap U ( arrow ( initial In (ob A U x y ) )) ) ( hom i OneObj ) |
434 ≡⟨ ≡-cong (λ k → k OneObj) ( comm ( initial In (ob A U x y ) )) ⟩ | |
435 hom (ob A U x y) OneObj | |
637
946ea019a2e7
if K{*}↓U has initial Obj, U is representable done.
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
636
diff
changeset
|
436 ≡⟨⟩ |
636 | 437 y |
438 ∎ ) ) where | |
439 open import Relation.Binary.PropositionalEquality | |
440 open ≡-Reasoning |