Mercurial > hg > Members > kono > Proof > category
annotate limit-to.agda @ 742:20f2700a481c
nat-ε
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Tue, 05 Dec 2017 01:57:41 +0900 |
parents | 917e51be9bbf |
children | 06388660995b |
rev | line source |
---|---|
415
dd086f5fb29f
same conflict again ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
414
diff
changeset
|
1 open import Category -- https://github.com/konn/category-agda |
350 | 2 open import Level |
3 | |
403 | 4 module limit-to where |
350 | 5 |
6 open import cat-utility | |
7 open import HomReasoning | |
8 open import Relation.Binary.Core | |
9 | |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
10 open import discrete |
427 | 11 |
365 | 12 |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
13 --- Equalizer from Limit ( 2→A IdnexFunctor Γ and IndexNat : K → Γ) |
458 | 14 --- |
15 --- | |
387 | 16 --- f |
431 | 17 --- e -----→ |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
18 --- c -----→ a b A |
431 | 19 --- ^ / -----→ |
387 | 20 --- |k h g |
415
dd086f5fb29f
same conflict again ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
414
diff
changeset
|
21 --- | / |
426 | 22 --- | / ^ |
23 --- | / | | |
24 --- |/ | Γ | |
25 --- d _ | | |
432
688066ac172e
add another method as a comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
431
diff
changeset
|
26 --- |\ | |
688066ac172e
add another method as a comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
431
diff
changeset
|
27 --- \ K af |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
28 --- \ -----→ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
29 --- \ t0 t1 I |
431 | 30 --- -----→ |
426 | 31 --- ag |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
32 --- |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
33 --- |
387 | 34 |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
35 open Complete |
350 | 36 open Limit |
487 | 37 open IsLimit |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
38 open NTrans |
352 | 39 |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
40 -- Functor Γ : TwoCat → A |
424
4329525f5085
fix imit-to-edu for MaybeCat / TwoCat
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
423
diff
changeset
|
41 |
461 | 42 IndexFunctor : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) ( a b : Obj A) ( f g : Hom A a b ) → Functor (TwoCat {c₁} {c₂}) A |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
43 IndexFunctor {c₁} {c₂} {ℓ} A a b f g = record { |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
44 FObj = λ a → fobj a |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
45 ; FMap = λ {a} {b} f → fmap {a} {b} f |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
46 ; isFunctor = record { |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
47 identity = λ{x} → identity x |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
48 ; distr = λ {a} {b} {c} {f} {g} → distr1 {a} {b} {c} {f} {g} |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
49 ; ≈-cong = λ {a} {b} {c} {f} → ≈-cong {a} {b} {c} {f} |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
50 } |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
51 } where |
461 | 52 T = TwoCat {c₁} {c₂} |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
53 fobj : Obj T → Obj A |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
54 fobj t0 = a |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
55 fobj t1 = b |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
56 fmap : {x y : Obj T } → (Hom T x y ) → Hom A (fobj x) (fobj y) |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
57 fmap {t0} {t0} id-t0 = id1 A a |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
58 fmap {t1} {t1} id-t1 = id1 A b |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
59 fmap {t0} {t1} arrow-f = f |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
60 fmap {t0} {t1} arrow-g = g |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
61 ≈-cong : {a : Obj T} {b : Obj T} {f g : Hom T a b} → T [ f ≈ g ] → A [ fmap f ≈ fmap g ] |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
62 ≈-cong {a} {b} {f} {.f} refl = let open ≈-Reasoning A in refl-hom |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
63 identity : (x : Obj T ) → A [ fmap (id1 T x) ≈ id1 A (fobj x) ] |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
64 identity t0 = let open ≈-Reasoning A in refl-hom |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
65 identity t1 = let open ≈-Reasoning A in refl-hom |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
66 distr1 : {a : Obj T} {b : Obj T} {c : Obj T} {f : Hom T a b} {g : Hom T b c} → |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
67 A [ fmap (T [ g o f ]) ≈ A [ fmap g o fmap f ] ] |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
68 distr1 {t0} {t0} {t0} {id-t0 } { id-t0 } = let open ≈-Reasoning A in sym-hom idL |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
69 distr1 {t1} {t1} {t1} { id-t1 } { id-t1 } = let open ≈-Reasoning A in begin |
467 | 70 id b |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
71 ≈↑⟨ idL ⟩ |
467 | 72 id b o id b |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
73 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
74 distr1 {t0} {t0} {t1} { id-t0 } { arrow-f } = let open ≈-Reasoning A in begin |
462 | 75 fmap (T [ arrow-f o id-t0 ] ) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
76 ≈⟨⟩ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
77 fmap arrow-f |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
78 ≈↑⟨ idR ⟩ |
467 | 79 fmap arrow-f o id a |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
80 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
81 distr1 {t0} {t0} {t1} { id-t0 } { arrow-g } = let open ≈-Reasoning A in begin |
462 | 82 fmap (T [ arrow-g o id-t0 ] ) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
83 ≈⟨⟩ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
84 fmap arrow-g |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
85 ≈↑⟨ idR ⟩ |
467 | 86 fmap arrow-g o id a |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
87 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
88 distr1 {t0} {t1} {t1} { arrow-f } { id-t1 } = let open ≈-Reasoning A in begin |
462 | 89 fmap (T [ id-t1 o arrow-f ] ) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
90 ≈⟨⟩ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
91 fmap arrow-f |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
92 ≈↑⟨ idL ⟩ |
467 | 93 id b o fmap arrow-f |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
94 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
95 distr1 {t0} {t1} {t1} { arrow-g } { id-t1 } = let open ≈-Reasoning A in begin |
462 | 96 fmap (T [ id-t1 o arrow-g ] ) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
97 ≈⟨⟩ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
98 fmap arrow-g |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
99 ≈↑⟨ idL ⟩ |
467 | 100 id b o fmap arrow-g |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
101 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
102 |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
103 --- Nat for Limit |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
104 -- |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
105 -- Nat : K → IndexFunctor |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
106 -- |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
107 |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
108 open Functor |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
109 |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
110 IndexNat : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
111 → {a b : Obj A} (f g : Hom A a b ) |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
112 (d : Obj A) → (h : Hom A d a ) → A [ A [ f o h ] ≈ A [ g o h ] ] → |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
113 NTrans TwoCat A (K TwoCat A d) (IndexFunctor {c₁} {c₂} {ℓ} A a b f g) |
460 | 114 IndexNat {c₁} {c₂} {ℓ} A {a} {b} f g d h fh=gh = record { |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
115 TMap = λ x → nmap x d h ; |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
116 isNTrans = record { |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
117 commute = λ {x} {y} {f'} → commute1 {x} {y} {f'} d h fh=gh |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
118 } |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
119 } where |
461 | 120 I = TwoCat {c₁} {c₂} |
429
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
121 Γ : Functor I A |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
122 Γ = IndexFunctor {c₁} {c₂} {ℓ} A a b f g |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
123 nmap : (x : Obj I ) ( d : Obj (A) ) (h : Hom A d a ) → Hom A (FObj (K I A d) x) (FObj Γ x) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
124 nmap t0 d h = h |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
125 nmap t1 d h = A [ f o h ] |
431 | 126 commute1 : {x y : Obj I} {f' : Hom I x y} (d : Obj A) (h : Hom A d a ) → A [ A [ f o h ] ≈ A [ g o h ] ] |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
127 → A [ A [ FMap Γ f' o nmap x d h ] ≈ A [ nmap y d h o FMap (K I A d) f' ] ] |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
128 commute1 {t0} {t1} {arrow-f} d h fh=gh = let open ≈-Reasoning A in begin |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
129 f o h |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
130 ≈↑⟨ idR ⟩ |
467 | 131 (f o h ) o id d |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
132 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
133 commute1 {t0} {t1} {arrow-g} d h fh=gh = let open ≈-Reasoning A in begin |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
134 g o h |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
135 ≈↑⟨ fh=gh ⟩ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
136 f o h |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
137 ≈↑⟨ idR ⟩ |
467 | 138 (f o h ) o id d |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
139 ∎ |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
140 commute1 {t0} {t0} {id-t0} d h fh=gh = let open ≈-Reasoning A in begin |
467 | 141 id a o h |
429
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
142 ≈⟨ idL ⟩ |
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
143 h |
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
144 ≈↑⟨ idR ⟩ |
467 | 145 h o id d |
429
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
146 ∎ |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
147 commute1 {t1} {t1} {id-t1} d h fh=gh = let open ≈-Reasoning A in begin |
467 | 148 id b o ( f o h ) |
429
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
149 ≈⟨ idL ⟩ |
428 | 150 f o h |
429
02eefa110eae
nat commute in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
428
diff
changeset
|
151 ≈↑⟨ idR ⟩ |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
152 ( f o h ) o id d |
428 | 153 ∎ |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
154 |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
155 |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
156 equlimit : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) {a b : Obj A} → (f g : Hom A a b) (lim : Limit TwoCat A (IndexFunctor A a b f g) ) → |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
157 Hom A (a0 lim) a |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
158 equlimit A {a} {b} f g lim = TMap (Limit.t0 lim) discrete.t0 |
460 | 159 |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
160 lim-to-equ : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
161 → {a b : Obj A} (f g : Hom A a b ) |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
162 (lim : Limit TwoCat A (IndexFunctor A a b f g) ) |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
163 → IsEqualizer A (equlimit A f g lim) f g |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
164 lim-to-equ {c₁} {c₂} {ℓ} A {a} {b} f g lim = record { |
601
2e7b5a777984
prove fe=ge in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
508
diff
changeset
|
165 fe=ge = fe=ge0 |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
166 ; k = λ {d} h fh=gh → k {d} h fh=gh |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
167 ; ek=h = λ {d} {h} {fh=gh} → ek=h d h fh=gh |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
168 ; uniqueness = λ {d} {h} {fh=gh} {k'} → uniquness d h fh=gh k' |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
169 } where |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
170 I : Category c₁ c₂ c₂ |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
171 I = TwoCat |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
172 Γ : Functor I A |
461 | 173 Γ = IndexFunctor A a b f g |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
174 e : Hom A (a0 lim) a |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
175 e = equlimit A f g lim |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
176 c : Obj A |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
177 c = a0 lim |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
178 inat : (d : Obj A) (h : Hom A d a) → A [ A [ f o h ] ≈ A [ g o h ] ] → NTrans I A (K I A d) Γ |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
179 inat = IndexNat A f g |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
180 fe=ge0 : A [ A [ f o (equlimit A f g lim ) ] ≈ A [ g o (equlimit A f g lim ) ] ] |
601
2e7b5a777984
prove fe=ge in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
508
diff
changeset
|
181 fe=ge0 = let open ≈-Reasoning A in begin |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
182 f o (equlimit A f g lim ) |
601
2e7b5a777984
prove fe=ge in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
508
diff
changeset
|
183 ≈⟨⟩ |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
184 FMap Γ arrow-f o TMap (Limit.t0 lim) discrete.t0 |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
185 ≈⟨ IsNTrans.commute ( isNTrans (Limit.t0 lim)) {discrete.t0} {discrete.t1} {arrow-f} ⟩ |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
186 TMap (Limit.t0 lim) discrete.t1 o FMap (K (TwoCat {c₁} {c₂} ) A (a0 lim)) id-t0 |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
187 ≈↑⟨ IsNTrans.commute ( isNTrans (Limit.t0 lim)) {discrete.t0} {discrete.t1} {arrow-g} ⟩ |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
188 FMap Γ arrow-g o TMap (Limit.t0 lim) discrete.t0 |
601
2e7b5a777984
prove fe=ge in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
508
diff
changeset
|
189 ≈⟨⟩ |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
190 g o (equlimit A f g lim ) |
601
2e7b5a777984
prove fe=ge in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
508
diff
changeset
|
191 ∎ |
350 | 192 k : {d : Obj A} (h : Hom A d a) → A [ A [ f o h ] ≈ A [ g o h ] ] → Hom A d c |
487 | 193 k {d} h fh=gh = limit (isLimit lim) d (inat d h fh=gh ) |
431 | 194 ek=h : (d : Obj A ) (h : Hom A d a ) → ( fh=gh : A [ A [ f o h ] ≈ A [ g o h ] ] ) → A [ A [ e o k h fh=gh ] ≈ h ] |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
195 ek=h d h fh=gh = let open ≈-Reasoning A in begin |
430 | 196 e o k h fh=gh |
460 | 197 ≈⟨⟩ |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
198 TMap (Limit.t0 lim) discrete.t0 o k h fh=gh |
487 | 199 ≈⟨ t0f=t (isLimit lim) {d} {inat d h fh=gh } {discrete.t0} ⟩ |
484
fcae3025d900
fix Limit pu a0 and t0 in record definition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
474
diff
changeset
|
200 TMap (inat d h fh=gh) discrete.t0 |
460 | 201 ≈⟨⟩ |
430 | 202 h |
203 ∎ | |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
204 uniq-nat : {i : Obj I} → (d : Obj A ) (h : Hom A d a ) ( k' : Hom A d c ) |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
205 ( fh=gh : A [ A [ f o h ] ≈ A [ g o h ] ]) → A [ A [ e o k' ] ≈ h ] → |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
206 A [ A [ TMap (Limit.t0 lim) i o k' ] ≈ TMap (inat d h fh=gh) i ] |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
207 uniq-nat {t0} d h k' fh=gh ek'=h = let open ≈-Reasoning A in begin |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
208 TMap (Limit.t0 lim) discrete.t0 o k' |
430 | 209 ≈⟨⟩ |
210 e o k' | |
211 ≈⟨ ek'=h ⟩ | |
212 h | |
213 ≈⟨⟩ | |
484
fcae3025d900
fix Limit pu a0 and t0 in record definition
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
474
diff
changeset
|
214 TMap (inat d h fh=gh) discrete.t0 |
430 | 215 ∎ |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
216 uniq-nat {t1} d h k' fh=gh ek'=h = let open ≈-Reasoning A in begin |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
217 TMap (Limit.t0 lim) t1 o k' |
460 | 218 ≈↑⟨ car (idR) ⟩ |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
219 ( TMap (Limit.t0 lim) t1 o id c ) o k' |
460 | 220 ≈⟨⟩ |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
221 ( TMap (Limit.t0 lim) t1 o FMap (K I A c) arrow-f ) o k' |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
222 ≈↑⟨ car ( nat1 (Limit.t0 lim) arrow-f ) ⟩ |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
223 ( FMap Γ arrow-f o TMap (Limit.t0 lim) discrete.t0 ) o k' |
430 | 224 ≈⟨⟩ |
225 (f o e ) o k' | |
226 ≈↑⟨ assoc ⟩ | |
227 f o ( e o k' ) | |
228 ≈⟨ cdr ek'=h ⟩ | |
229 f o h | |
230 ≈⟨⟩ | |
466 | 231 TMap (inat d h fh=gh) t1 |
430 | 232 ∎ |
431 | 233 uniquness : (d : Obj A ) (h : Hom A d a ) → ( fh=gh : A [ A [ f o h ] ≈ A [ g o h ] ] ) → |
372
b4855a3ebd34
add more lemma in limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
371
diff
changeset
|
234 ( k' : Hom A d c ) |
457
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
235 → A [ A [ e o k' ] ≈ h ] → A [ k h fh=gh ≈ k' ] |
0ba86e29f492
limit-to and discrete clean up
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
443
diff
changeset
|
236 uniquness d h fh=gh k' ek'=h = let open ≈-Reasoning A in begin |
430 | 237 k h fh=gh |
495 | 238 ≈⟨ limit-uniqueness (isLimit lim) ( λ{i} → uniq-nat {i} d h k' fh=gh ek'=h ) ⟩ |
430 | 239 k' |
240 ∎ | |
368
b18585089d2e
add more parameter to nat in lim-to-equ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
367
diff
changeset
|
241 |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
242 |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
243 --- Product from Limit ( given Discrete→A functor Γ and pnat : K → Γ) |
469
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
244 |
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
245 open DiscreteHom |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
246 |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
247 plimit : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (S : Set c₁) |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
248 → ( Γ : Functor (DiscreteCat S ) A ) → (lim : Limit ( DiscreteCat S ) A Γ ) → Obj A |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
249 plimit A S Γ lim = a0 lim |
469
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
250 |
472
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
251 discrete-identity : { c₁ : Level} { S : Set c₁} { a : DiscreteObj {c₁} S } → (f : DiscreteHom a a ) → (DiscreteCat S) [ f ≈ id1 (DiscreteCat S) a ] |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
252 discrete-identity f = refl |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
253 |
474 | 254 pnat : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (S : Set c₁) |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
255 → (Γ : Functor (DiscreteCat S) A ) |
474 | 256 → {q : Obj A } ( qi : (i : Obj ( DiscreteCat S)) → Hom A q (FObj Γ i) ) |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
257 → NTrans (DiscreteCat S )A (K (DiscreteCat S) A q) Γ |
474 | 258 pnat A S Γ {q} qi = record { |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
259 TMap = qi ; |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
260 isNTrans = record { commute = λ {a} {b} {f} → commute {a} {b} {f} } |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
261 } where |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
262 commute : {a b : Obj (DiscreteCat S) } {f : Hom (DiscreteCat S) a b} → |
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
263 A [ A [ FMap Γ f o qi a ] ≈ A [ qi b o FMap (K (DiscreteCat S) A q) f ] ] |
469
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
264 commute {a} {b} {f} with discrete f |
472
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
265 commute {a} {.a} {f} | refl = let open ≈-Reasoning A in begin |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
266 FMap Γ f o qi a |
472
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
267 ≈⟨ car ( fcong Γ (discrete-identity f )) ⟩ |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
268 FMap Γ (id1 (DiscreteCat S) a ) o qi a |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
269 ≈⟨ car ( IsFunctor.identity (isFunctor Γ) ) ⟩ |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
270 id1 A (FObj Γ a) o qi a |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
271 ≈⟨ idL ⟩ |
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
272 qi a |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
273 ≈↑⟨ idR ⟩ |
472
f3d6d0275a0a
discrete equality as a dom equality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
469
diff
changeset
|
274 qi a o id q |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
275 ≈⟨⟩ |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
276 qi a o FMap (K (DiscreteCat S) A q) f |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
277 ∎ |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
278 |
469
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
279 lim-to-product : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) ( S : Set c₁ ) |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
280 → ( Γ : Functor (DiscreteCat S) A ) |
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
281 → (lim : Limit (DiscreteCat S) A Γ ) |
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
282 → IProduct (Obj (DiscreteCat S)) A (FObj Γ) |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
283 lim-to-product A S Γ lim = record { |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
284 iprod = plimit A S Γ lim |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
285 ; pi = λ i → TMap (Limit.t0 lim) i |
508
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
495
diff
changeset
|
286 ; isIProduct = record { |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
287 iproduct = λ {q} qi → iproduct {q} qi ; |
691
917e51be9bbf
change argument of Limit and K
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
670
diff
changeset
|
288 pif=q = λ {q} {qi} {i} → pif=q {q} qi {i} ; |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
289 ip-uniqueness = λ {q } { h } → ip-uniqueness {q} {h} ; |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
290 ip-cong = λ {q } { qi } { qi' } qi=qi' → ip-cong {q} {qi} {qi'} qi=qi' |
508
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
495
diff
changeset
|
291 } |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
292 } where |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
293 D = DiscreteCat S |
469
65ab0da524b8
discrete f ≡ refl should be passed, but it doesn't
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
468
diff
changeset
|
294 I = Obj ( DiscreteCat S ) |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
295 ai = λ i → FObj Γ i |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
296 p = a0 lim |
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
297 pi = λ i → TMap (Limit.t0 lim) i |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
298 iproduct : {q : Obj A} → ( qi : (i : I) → Hom A q (ai i) ) → Hom A q p |
487 | 299 iproduct {q} qi = limit (isLimit lim) q (pnat A S Γ qi ) |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
300 pif=q : {q : Obj A} → ( qi : (i : I) → Hom A q (ai i) ) → ∀ { i : I } → A [ A [ ( pi i ) o ( iproduct qi ) ] ≈ (qi i) ] |
487 | 301 pif=q {q} qi {i} = t0f=t (isLimit lim) {q} {pnat A S Γ qi } {i} |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
302 ipu : {i : Obj D} → (q : Obj A) (h : Hom A q p ) → A [ A [ TMap (Limit.t0 lim) i o h ] ≈ A [ pi i o h ] ] |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
303 ipu {i} q h = let open ≈-Reasoning A in refl-hom |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
304 ip-uniqueness : {q : Obj A} { h : Hom A q p } → A [ iproduct ( λ (i : I) → A [ (pi i) o h ] ) ≈ h ] |
495 | 305 ip-uniqueness {q} {h} = limit-uniqueness (isLimit lim) {q} {pnat A S Γ (λ i → A [ pi i o h ] )} (ipu q h) |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
306 ipc : {q : Obj A} → { qi : (i : I) → Hom A q (ai i) } → { qi' : (i : I) → Hom A q (ai i) } |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
307 → (i : I ) → A [ qi i ≈ qi' i ] → |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
308 A [ A [ TMap (Limit.t0 lim) i o iproduct qi' ] ≈ TMap (pnat A S Γ qi) i ] |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
309 ipc {q} {qi} {qi'} i qi=qi' = let open ≈-Reasoning A in begin |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
310 TMap (Limit.t0 lim) i o iproduct qi' |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
311 ≈⟨⟩ |
670
99065a1e56ea
remove comp from limit-to
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
662
diff
changeset
|
312 TMap (Limit.t0 lim) i o limit (isLimit lim) q (pnat A S Γ qi' ) |
487 | 313 ≈⟨ t0f=t (isLimit lim) {q} {pnat A S Γ qi'} {i} ⟩ |
474 | 314 TMap (pnat A S Γ qi') i |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
315 ≈⟨⟩ |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
316 qi' i |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
317 ≈↑⟨ qi=qi' ⟩ |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
318 qi i |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
319 ≈⟨⟩ |
474 | 320 TMap (pnat A S Γ qi) i |
468
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
321 ∎ |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
322 ip-cong : {q : Obj A} → { qi : (i : I) → Hom A q (ai i) } → { qi' : (i : I) → Hom A q (ai i) } |
c375d8f93a2c
discrete category and product from a limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
467
diff
changeset
|
323 → ( ∀ (i : I ) → A [ qi i ≈ qi' i ] ) → A [ iproduct qi ≈ iproduct qi' ] |
495 | 324 ip-cong {q} {qi} {qi'} qi=qi' = limit-uniqueness (isLimit lim) {q} {pnat A S Γ qi} (λ {i} → ipc {q} {qi} {qi'} i (qi=qi' i)) |