Mercurial > hg > Members > kono > Proof > category
annotate cat-utility.agda @ 312:702adc45704f
is this right direction?
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 05 Jan 2014 23:37:12 +0900 |
parents | d6a6dd305da2 |
children | c483374f542b |
rev | line source |
---|---|
56 | 1 module cat-utility where |
2 | |
3 -- Shinji KONO <kono@ie.u-ryukyu.ac.jp> | |
4 | |
87 | 5 open import Category -- https://github.com/konn/category-agda |
6 open import Level | |
7 --open import Category.HomReasoning | |
8 open import HomReasoning | |
56 | 9 |
87 | 10 open Functor |
56 | 11 |
87 | 12 id1 : ∀{c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (a : Obj A ) → Hom A a a |
13 id1 A a = (Id {_} {_} {_} {A} a) | |
253 | 14 -- We cannot make A implicit |
56 | 15 |
87 | 16 record IsUniversalMapping {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') |
17 ( U : Functor B A ) | |
18 ( F : Obj A → Obj B ) | |
19 ( η : (a : Obj A) → Hom A a ( FObj U (F a) ) ) | |
20 ( _* : { a : Obj A}{ b : Obj B} → ( Hom A a (FObj U b) ) → Hom B (F a ) b ) | |
21 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
22 field | |
101 | 23 universalMapping : {a : Obj A} { b : Obj B } → { f : Hom A a (FObj U b) } → |
24 A [ A [ FMap U ( f * ) o η a ] ≈ f ] | |
25 uniquness : {a : Obj A} { b : Obj B } → { f : Hom A a (FObj U b) } → { g : Hom B (F a) b } → | |
26 A [ A [ FMap U g o η a ] ≈ f ] → B [ f * ≈ g ] | |
56 | 27 |
87 | 28 record UniversalMapping {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') |
29 ( U : Functor B A ) | |
30 ( F : Obj A → Obj B ) | |
31 ( η : (a : Obj A) → Hom A a ( FObj U (F a) ) ) | |
32 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
33 infixr 11 _* | |
34 field | |
35 _* : { a : Obj A}{ b : Obj B} → ( Hom A a (FObj U b) ) → Hom B (F a ) b | |
36 isUniversalMapping : IsUniversalMapping A B U F η _* | |
56 | 37 |
268 | 38 record coIsUniversalMapping {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') |
39 ( F : Functor A B ) | |
40 ( U : Obj B → Obj A ) | |
41 ( ε : (b : Obj B) → Hom B ( FObj F (U b) ) b ) | |
42 ( _*' : { b : Obj B}{ a : Obj A} → ( Hom B (FObj F a) b ) → Hom A a (U b ) ) | |
43 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
44 field | |
45 couniversalMapping : {b : Obj B} { a : Obj A } → { f : Hom B (FObj F a) b } → | |
46 B [ B [ ε b o FMap F ( f *' ) ] ≈ f ] | |
47 couniquness : {b : Obj B} { a : Obj A } → { f : Hom B (FObj F a) b } → { g : Hom A a (U b) } → | |
48 B [ B [ ε b o FMap F g ] ≈ f ] → A [ f *' ≈ g ] | |
49 | |
50 record coUniversalMapping {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') | |
51 ( F : Functor A B ) | |
52 ( U : Obj B → Obj A ) | |
53 ( ε : (b : Obj B) → Hom B ( FObj F (U b) ) b ) | |
54 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
55 infixr 11 _*' | |
56 field | |
57 _*' : { b : Obj B}{ a : Obj A} → ( Hom B (FObj F a) b ) → Hom A a (U b ) | |
58 iscoUniversalMapping : coIsUniversalMapping A B F U ε _*' | |
59 | |
87 | 60 open NTrans |
61 open import Category.Cat | |
62 record IsAdjunction {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') | |
63 ( U : Functor B A ) | |
64 ( F : Functor A B ) | |
65 ( η : NTrans A A identityFunctor ( U ○ F ) ) | |
66 ( ε : NTrans B B ( F ○ U ) identityFunctor ) | |
67 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
68 field | |
69 adjoint1 : { b : Obj B } → | |
70 A [ A [ ( FMap U ( TMap ε b )) o ( TMap η ( FObj U b )) ] ≈ id1 A (FObj U b) ] | |
71 adjoint2 : {a : Obj A} → | |
72 B [ B [ ( TMap ε ( FObj F a )) o ( FMap F ( TMap η a )) ] ≈ id1 B (FObj F a) ] | |
56 | 73 |
87 | 74 record Adjunction {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) (B : Category c₁' c₂' ℓ') |
75 ( U : Functor B A ) | |
76 ( F : Functor A B ) | |
77 ( η : NTrans A A identityFunctor ( U ○ F ) ) | |
78 ( ε : NTrans B B ( F ○ U ) identityFunctor ) | |
79 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
80 field | |
81 isAdjunction : IsAdjunction A B U F η ε | |
202
58ee98bbb333
remove an extensionality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
176
diff
changeset
|
82 U-functor = U |
58ee98bbb333
remove an extensionality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
176
diff
changeset
|
83 F-functor = F |
58ee98bbb333
remove an extensionality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
176
diff
changeset
|
84 Eta = η |
58ee98bbb333
remove an extensionality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
176
diff
changeset
|
85 Epsiron = ε |
56 | 86 |
87 | |
87 | 88 record IsMonad {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) |
89 ( T : Functor A A ) | |
90 ( η : NTrans A A identityFunctor T ) | |
91 ( μ : NTrans A A (T ○ T) T) | |
92 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where | |
93 field | |
94 assoc : {a : Obj A} → A [ A [ TMap μ a o TMap μ ( FObj T a ) ] ≈ A [ TMap μ a o FMap T (TMap μ a) ] ] | |
95 unity1 : {a : Obj A} → A [ A [ TMap μ a o TMap η ( FObj T a ) ] ≈ Id {_} {_} {_} {A} (FObj T a) ] | |
96 unity2 : {a : Obj A} → A [ A [ TMap μ a o (FMap T (TMap η a ))] ≈ Id {_} {_} {_} {A} (FObj T a) ] | |
56 | 97 |
87 | 98 record Monad {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) (T : Functor A A) (η : NTrans A A identityFunctor T) (μ : NTrans A A (T ○ T) T) |
99 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where | |
100 field | |
101 isMonad : IsMonad A T η μ | |
130 | 102 -- g ○ f = μ(c) T(g) f |
103 join : { a b : Obj A } → { c : Obj A } → | |
104 ( Hom A b ( FObj T c )) → ( Hom A a ( FObj T b)) → Hom A a ( FObj T c ) | |
105 join {_} {_} {c} g f = A [ TMap μ c o A [ FMap T g o f ] ] | |
106 | |
56 | 107 |
87 | 108 Functor*Nat : {c₁ c₂ ℓ c₁' c₂' ℓ' c₁'' c₂'' ℓ'' : Level} (A : Category c₁ c₂ ℓ) {B : Category c₁' c₂' ℓ'} (C : Category c₁'' c₂'' ℓ'') |
300 | 109 (F : Functor B C) → { G H : Functor A B } → ( n : NTrans A B G H ) → NTrans A C (F ○ G) (F ○ H) |
87 | 110 Functor*Nat A {B} C F {G} {H} n = record { |
300 | 111 TMap = λ a → FMap F (TMap n a) |
87 | 112 ; isNTrans = record { |
130 | 113 commute = commute |
87 | 114 } |
115 } where | |
130 | 116 commute : {a b : Obj A} {f : Hom A a b} |
87 | 117 → C [ C [ (FMap F ( FMap H f )) o ( FMap F (TMap n a)) ] ≈ C [ (FMap F (TMap n b )) o (FMap F (FMap G f)) ] ] |
130 | 118 commute {a} {b} {f} = let open ≈-Reasoning (C) in |
87 | 119 begin |
120 (FMap F ( FMap H f )) o ( FMap F (TMap n a)) | |
121 ≈⟨ sym (distr F) ⟩ | |
122 FMap F ( B [ (FMap H f) o TMap n a ]) | |
123 ≈⟨ IsFunctor.≈-cong (isFunctor F) ( nat n ) ⟩ | |
124 FMap F ( B [ (TMap n b ) o FMap G f ] ) | |
125 ≈⟨ distr F ⟩ | |
126 (FMap F (TMap n b )) o (FMap F (FMap G f)) | |
127 ∎ | |
56 | 128 |
87 | 129 Nat*Functor : {c₁ c₂ ℓ c₁' c₂' ℓ' c₁'' c₂'' ℓ'' : Level} (A : Category c₁ c₂ ℓ) {B : Category c₁' c₂' ℓ'} (C : Category c₁'' c₂'' ℓ'') |
300 | 130 { G H : Functor B C } → ( n : NTrans B C G H ) → (F : Functor A B) → NTrans A C (G ○ F) (H ○ F) |
87 | 131 Nat*Functor A {B} C {G} {H} n F = record { |
300 | 132 TMap = λ a → TMap n (FObj F a) |
87 | 133 ; isNTrans = record { |
130 | 134 commute = commute |
87 | 135 } |
136 } where | |
130 | 137 commute : {a b : Obj A} {f : Hom A a b} |
87 | 138 → C [ C [ ( FMap H (FMap F f )) o ( TMap n (FObj F a)) ] ≈ C [ (TMap n (FObj F b )) o (FMap G (FMap F f)) ] ] |
130 | 139 commute {a} {b} {f} = IsNTrans.commute ( isNTrans n) |
56 | 140 |
87 | 141 -- T ≃ (U_R ○ F_R) |
142 -- μ = U_R ε F_R | |
143 -- nat-ε | |
144 -- nat-η -- same as η but has different types | |
84 | 145 |
87 | 146 record MResolution {c₁ c₂ ℓ c₁' c₂' ℓ' : Level} (A : Category c₁ c₂ ℓ) ( B : Category c₁' c₂' ℓ' ) |
147 ( T : Functor A A ) | |
94
4fa718e4fd77
Comparison Functor constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
91
diff
changeset
|
148 -- { η : NTrans A A identityFunctor T } |
4fa718e4fd77
Comparison Functor constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
91
diff
changeset
|
149 -- { μ : NTrans A A (T ○ T) T } |
4fa718e4fd77
Comparison Functor constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
91
diff
changeset
|
150 -- { M : Monad A T η μ } |
87 | 151 ( UR : Functor B A ) ( FR : Functor A B ) |
152 { ηR : NTrans A A identityFunctor ( UR ○ FR ) } | |
153 { εR : NTrans B B ( FR ○ UR ) identityFunctor } | |
154 { μR : NTrans A A ( (UR ○ FR) ○ ( UR ○ FR )) ( UR ○ FR ) } | |
155 ( Adj : Adjunction A B UR FR ηR εR ) | |
156 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁' ⊔ c₂' ⊔ ℓ' )) where | |
157 field | |
158 T=UF : T ≃ (UR ○ FR) | |
300 | 159 μ=UεF : {x : Obj A } → A [ TMap μR x ≈ FMap UR ( TMap εR ( FObj FR x ) ) ] |
160 -- ηR=η : {x : Obj A } → A [ TMap ηR x ≈ TMap η x ] -- We need T → UR FR conversion | |
161 -- μR=μ : {x : Obj A } → A [ TMap μR x ≈ TMap μ x ] | |
86 | 162 |
88 | 163 |
260 | 164 record Equalizer { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) {c a b : Obj A} (e : Hom A c a) (f g : Hom A a b) : Set (ℓ ⊔ (c₁ ⊔ c₂)) where |
165 field | |
166 fe=ge : A [ A [ f o e ] ≈ A [ g o e ] ] | |
167 k : {d : Obj A} (h : Hom A d a) → A [ A [ f o h ] ≈ A [ g o h ] ] → Hom A d c | |
168 ek=h : {d : Obj A} → ∀ {h : Hom A d a} → {eq : A [ A [ f o h ] ≈ A [ g o h ] ] } → A [ A [ e o k {d} h eq ] ≈ h ] | |
169 uniqueness : {d : Obj A} → ∀ {h : Hom A d a} → {eq : A [ A [ f o h ] ≈ A [ g o h ] ] } → {k' : Hom A d c } → | |
170 A [ A [ e o k' ] ≈ h ] → A [ k {d} h eq ≈ k' ] | |
171 equalizer : Hom A c a | |
172 equalizer = e | |
173 | |
174 -- | |
175 -- Product | |
176 -- | |
264 | 177 -- c |
178 -- f | g | |
179 -- |f×g | |
180 -- v | |
300 | 181 -- a <-------- ab ---------→ b |
264 | 182 -- π1 π2 |
260 | 183 |
184 | |
185 record Product { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) (a b ab : Obj A) | |
186 ( π1 : Hom A ab a ) | |
187 ( π2 : Hom A ab b ) | |
188 : Set (ℓ ⊔ (c₁ ⊔ c₂)) where | |
189 field | |
190 _×_ : {c : Obj A} ( f : Hom A c a ) → ( g : Hom A c b ) → Hom A c ab | |
191 π1fxg=f : {c : Obj A} { f : Hom A c a } → { g : Hom A c b } → A [ A [ π1 o ( f × g ) ] ≈ f ] | |
192 π2fxg=g : {c : Obj A} { f : Hom A c a } → { g : Hom A c b } → A [ A [ π2 o ( f × g ) ] ≈ g ] | |
193 uniqueness : {c : Obj A} { h : Hom A c ab } → A [ ( A [ π1 o h ] ) × ( A [ π2 o h ] ) ≈ h ] | |
264 | 194 ×-cong : {c : Obj A} { f f' : Hom A c a } → { g g' : Hom A c b } → A [ f ≈ f' ] → A [ g ≈ g' ] → A [ f × g ≈ f' × g' ] |
260 | 195 axb : Obj A |
196 axb = ab | |
265 | 197 pi1 : Hom A ab a |
198 pi1 = π1 | |
199 pi2 : Hom A ab b | |
200 pi2 = π2 | |
260 | 201 |
202 -- Pullback | |
203 -- f | |
300 | 204 -- a ------→ c |
260 | 205 -- ^ ^ |
206 -- π1 | |g | |
207 -- | | | |
300 | 208 -- ab ------→ b |
260 | 209 -- ^ π2 |
210 -- | | |
211 -- d | |
212 -- | |
213 record Pullback { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) (a b c ab : Obj A) | |
214 ( f : Hom A a c ) ( g : Hom A b c ) | |
215 ( π1 : Hom A ab a ) ( π2 : Hom A ab b ) | |
216 : Set (ℓ ⊔ (c₁ ⊔ c₂)) where | |
217 field | |
218 commute : A [ A [ f o π1 ] ≈ A [ g o π2 ] ] | |
219 p : { d : Obj A } → { π1' : Hom A d a } { π2' : Hom A d b } → A [ A [ f o π1' ] ≈ A [ g o π2' ] ] → Hom A d ab | |
220 π1p=π1 : { d : Obj A } → { π1' : Hom A d a } { π2' : Hom A d b } → { eq : A [ A [ f o π1' ] ≈ A [ g o π2' ] ] } | |
221 → A [ A [ π1 o p eq ] ≈ π1' ] | |
222 π2p=π2 : { d : Obj A } → { π1' : Hom A d a } { π2' : Hom A d b } → { eq : A [ A [ f o π1' ] ≈ A [ g o π2' ] ] } | |
223 → A [ A [ π2 o p eq ] ≈ π2' ] | |
224 uniqueness : { d : Obj A } → ( p' : Hom A d ab ) → { π1' : Hom A d a } { π2' : Hom A d b } → { eq : A [ A [ f o π1' ] ≈ A [ g o π2' ] ] } | |
225 → { π1p=π1' : A [ A [ π1 o p' ] ≈ π1' ] } | |
226 → { π2p=π2' : A [ A [ π2 o p' ] ≈ π2' ] } | |
227 → A [ p eq ≈ p' ] | |
228 axb : Obj A | |
229 axb = ab | |
265 | 230 pi1 : Hom A ab a |
231 pi1 = π1 | |
232 pi2 : Hom A ab b | |
233 pi2 = π2 | |
312
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
234 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
235 -- |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
236 -- Limit |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
237 -- |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
238 ----- |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
239 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
240 -- Constancy Functor |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
241 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
242 K : { c₁' c₂' ℓ' : Level} (A : Category c₁' c₂' ℓ') { c₁'' c₂'' ℓ'' : Level} ( I : Category c₁'' c₂'' ℓ'' ) |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
243 → ( a : Obj A ) → Functor I A |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
244 K A I a = record { |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
245 FObj = λ i → a ; |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
246 FMap = λ f → id1 A a ; |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
247 isFunctor = let open ≈-Reasoning (A) in record { |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
248 ≈-cong = λ f=g → refl-hom |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
249 ; identity = refl-hom |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
250 ; distr = sym idL |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
251 } |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
252 } |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
253 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
254 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
255 record Limit { c₁' c₂' ℓ' : Level} { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) ( I : Category c₁' c₂' ℓ' ) ( Γ : Functor I A ) |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
256 ( a0 : Obj A ) ( t0 : NTrans I A ( K A I a0 ) Γ ) : Set (suc (c₁' ⊔ c₂' ⊔ ℓ' ⊔ c₁ ⊔ c₂ ⊔ ℓ )) where |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
257 field |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
258 limit : ( a : Obj A ) → ( t : NTrans I A ( K A I a ) Γ ) → Hom A a a0 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
259 t0f=t : { a : Obj A } → { t : NTrans I A ( K A I a ) Γ } → ∀ { i : Obj I } → |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
260 A [ A [ TMap t0 i o limit a t ] ≈ TMap t i ] |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
261 limit-uniqueness : { a : Obj A } → { t : NTrans I A ( K A I a ) Γ } → { f : Hom A a a0 } → ( ∀ { i : Obj I } → |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
262 A [ A [ TMap t0 i o f ] ≈ TMap t i ] ) → A [ limit a t ≈ f ] |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
263 A0 : Obj A |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
264 A0 = a0 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
265 T0 : NTrans I A ( K A I a0 ) Γ |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
266 T0 = t0 |
702adc45704f
is this right direction?
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
300
diff
changeset
|
267 |