Mercurial > hg > Members > kono > Proof > category
annotate src/CCCSets.agda @ 1027:5ae0290c34b4
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 29 Mar 2021 21:57:12 +0900 |
parents | 7916bde5e57b |
children | 28569574e3cf |
rev | line source |
---|---|
999 | 1 {-# OPTIONS --allow-unsolved-metas #-} |
2 module CCCSets where | |
3 | |
4 open import Level | |
5 open import Category | |
6 open import HomReasoning | |
7 open import cat-utility | |
8 open import Data.Product renaming (_×_ to _/\_ ) hiding ( <_,_> ) | |
9 open import Category.Constructions.Product | |
10 open import Relation.Binary.PropositionalEquality hiding ( [_] ) | |
11 open import CCC | |
12 | |
13 open Functor | |
14 | |
15 -- ccc-1 : Hom A a 1 ≅ {*} | |
16 -- ccc-2 : Hom A c (a × b) ≅ (Hom A c a ) × ( Hom A c b ) | |
17 -- ccc-3 : Hom A a (c ^ b) ≅ Hom A (a × b) c | |
18 | |
19 open import Category.Sets | |
20 | |
21 -- Sets is a CCC | |
22 | |
1018 | 23 open import SetsCompleteness hiding (ki1) |
24 | |
25 -- import Axiom.Extensionality.Propositional | |
26 -- postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Axiom.Extensionality.Propositional.Extensionality c₂ c₂ | |
999 | 27 |
28 data One {c : Level } : Set c where | |
1020 | 29 ! : One -- () in Haskell ( or any one object set ) |
999 | 30 |
31 sets : {c : Level } → CCC (Sets {c}) | |
32 sets = record { | |
33 1 = One | |
1020 | 34 ; ○ = λ _ → λ _ → ! |
999 | 35 ; _∧_ = _∧_ |
36 ; <_,_> = <,> | |
37 ; π = π | |
38 ; π' = π' | |
39 ; _<=_ = _<=_ | |
40 ; _* = _* | |
41 ; ε = ε | |
42 ; isCCC = isCCC | |
43 } where | |
44 1 : Obj Sets | |
45 1 = One | |
46 ○ : (a : Obj Sets ) → Hom Sets a 1 | |
1020 | 47 ○ a = λ _ → ! |
999 | 48 _∧_ : Obj Sets → Obj Sets → Obj Sets |
49 _∧_ a b = a /\ b | |
50 <,> : {a b c : Obj Sets } → Hom Sets c a → Hom Sets c b → Hom Sets c ( a ∧ b) | |
51 <,> f g = λ x → ( f x , g x ) | |
52 π : {a b : Obj Sets } → Hom Sets (a ∧ b) a | |
53 π {a} {b} = proj₁ | |
54 π' : {a b : Obj Sets } → Hom Sets (a ∧ b) b | |
55 π' {a} {b} = proj₂ | |
56 _<=_ : (a b : Obj Sets ) → Obj Sets | |
57 a <= b = b → a | |
58 _* : {a b c : Obj Sets } → Hom Sets (a ∧ b) c → Hom Sets a (c <= b) | |
59 f * = λ x → λ y → f ( x , y ) | |
60 ε : {a b : Obj Sets } → Hom Sets ((a <= b ) ∧ b) a | |
61 ε {a} {b} = λ x → ( proj₁ x ) ( proj₂ x ) | |
62 isCCC : CCC.IsCCC Sets 1 ○ _∧_ <,> π π' _<=_ _* ε | |
63 isCCC = record { | |
64 e2 = e2 | |
65 ; e3a = λ {a} {b} {c} {f} {g} → e3a {a} {b} {c} {f} {g} | |
66 ; e3b = λ {a} {b} {c} {f} {g} → e3b {a} {b} {c} {f} {g} | |
67 ; e3c = e3c | |
68 ; π-cong = π-cong | |
69 ; e4a = e4a | |
70 ; e4b = e4b | |
71 ; *-cong = *-cong | |
72 } where | |
73 e2 : {a : Obj Sets} {f : Hom Sets a 1} → Sets [ f ≈ ○ a ] | |
74 e2 {a} {f} = extensionality Sets ( λ x → e20 x ) | |
75 where | |
76 e20 : (x : a ) → f x ≡ ○ a x | |
77 e20 x with f x | |
1020 | 78 e20 x | ! = refl |
999 | 79 e3a : {a b c : Obj Sets} {f : Hom Sets c a} {g : Hom Sets c b} → |
80 Sets [ ( Sets [ π o ( <,> f g) ] ) ≈ f ] | |
81 e3a = refl | |
82 e3b : {a b c : Obj Sets} {f : Hom Sets c a} {g : Hom Sets c b} → | |
83 Sets [ Sets [ π' o ( <,> f g ) ] ≈ g ] | |
84 e3b = refl | |
85 e3c : {a b c : Obj Sets} {h : Hom Sets c (a ∧ b)} → | |
86 Sets [ <,> (Sets [ π o h ]) (Sets [ π' o h ]) ≈ h ] | |
87 e3c = refl | |
88 π-cong : {a b c : Obj Sets} {f f' : Hom Sets c a} {g g' : Hom Sets c b} → | |
89 Sets [ f ≈ f' ] → Sets [ g ≈ g' ] → Sets [ <,> f g ≈ <,> f' g' ] | |
90 π-cong refl refl = refl | |
91 e4a : {a b c : Obj Sets} {h : Hom Sets (c ∧ b) a} → | |
92 Sets [ Sets [ ε o <,> (Sets [ h * o π ]) π' ] ≈ h ] | |
93 e4a = refl | |
94 e4b : {a b c : Obj Sets} {k : Hom Sets c (a <= b)} → | |
95 Sets [ (Sets [ ε o <,> (Sets [ k o π ]) π' ]) * ≈ k ] | |
96 e4b = refl | |
97 *-cong : {a b c : Obj Sets} {f f' : Hom Sets (a ∧ b) c} → | |
98 Sets [ f ≈ f' ] → Sets [ f * ≈ f' * ] | |
99 *-cong refl = refl | |
100 | |
1020 | 101 open import Relation.Nullary |
102 open import Data.Empty | |
1023 | 103 open import equalizer |
999 | 104 |
1020 | 105 data Bool { c : Level } : Set c where |
106 true : Bool | |
107 false : Bool | |
1023 | 108 |
109 ¬f≡t : { c : Level } → ¬ (false {c} ≡ true ) | |
110 ¬f≡t () | |
1025 | 111 |
112 ¬x≡t∧x≡f : { c : Level } → {x : Bool {c}} → ¬ ((x ≡ false) /\ (x ≡ true) ) | |
113 ¬x≡t∧x≡f {_} {true} p = ⊥-elim (¬f≡t (sym (proj₁ p))) | |
114 ¬x≡t∧x≡f {_} {false} p = ⊥-elim (¬f≡t (proj₂ p)) | |
1020 | 115 |
1022 | 116 data _∨_ {c c' : Level } (a : Set c) (b : Set c') : Set (c ⊔ c') where |
1021
8a78ddfdaa24
... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1020
diff
changeset
|
117 case1 : a → a ∨ b |
8a78ddfdaa24
... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1020
diff
changeset
|
118 case2 : b → a ∨ b |
8a78ddfdaa24
... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1020
diff
changeset
|
119 |
1024 | 120 -- |
121 -- m : b → a determins a subset of a as an image | |
122 -- b and m-image in a has one to one correspondence with an equalizer (x : b) → (y : a) ≡ m x. | |
123 -- so tchar m mono and ker (tchar m mono) is Iso. | |
124 -- Finding (x : b) from (y : a) is non constructive. Assuming LEM of image. | |
125 -- | |
126 data image {c : Level} {a b : Set c} (m : Hom Sets b a) : a → Set c where | |
127 isImage : (x : b ) → image m (m x) | |
1022 | 128 |
129 topos : {c : Level } → ({ c : Level} → (b : Set c) → b ∨ (¬ b)) → Topos (Sets {c}) sets | |
1021
8a78ddfdaa24
... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
1020
diff
changeset
|
130 topos {c} lem = record { |
1023 | 131 Ω = Bool |
132 ; ⊤ = λ _ → true | |
999 | 133 ; Ker = tker |
1020 | 134 ; char = λ m mono → tchar m mono |
999 | 135 ; isTopos = record { |
1018 | 136 char-uniqueness = λ {a} {b} {h} m mono → extensionality Sets ( λ x → uniq h m mono x ) |
137 ; ker-m = imequ | |
999 | 138 } |
139 } where | |
1023 | 140 -- |
1024 | 141 -- In Sets, equalizers exist as |
1022 | 142 -- data sequ {c : Level} (A B : Set c) ( f g : A → B ) : Set c where |
143 -- elem : (x : A ) → (eq : f x ≡ g x) → sequ A B f g | |
1024 | 144 -- m have to be isomorphic to ker (char m). |
145 -- | |
1022 | 146 -- i ○ b |
147 -- ker (char m) ----→ b -----------→ 1 | |
148 -- | ←---- | | | |
1024 | 149 -- | j |m | ⊤ char m : a → Ω = {true,false} |
150 -- | e ↓ char m ↓ if y : a ≡ m (∃ x : b) → true ( data char ) | |
151 -- +-------------→ a -----------→ Ω else false | |
1022 | 152 -- h |
153 -- | |
1023 | 154 tker : {a : Obj Sets} (h : Hom Sets a Bool) → Equalizer Sets h (Sets [ (λ _ → true ) o CCC.○ sets a ]) |
999 | 155 tker {a} h = record { |
1023 | 156 equalizer-c = sequ a Bool h (λ _ → true ) |
1018 | 157 ; equalizer = λ e → equ e |
158 ; isEqualizer = SetsIsEqualizer _ _ _ _ } | |
1023 | 159 tchar : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) → a → Bool {c} |
1024 | 160 tchar {a} {b} m mono y with lem (image m y ) |
1023 | 161 ... | case1 t = true |
162 ... | case2 f = false | |
1025 | 163 tcharImg : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) → (y : a) → tchar m mono y ≡ true → image m y |
164 tcharImg m mono y eq with lem (image m y) | |
165 ... | case1 t = t | |
166 tchar¬Img : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) (y : a) → tchar m mono y ≡ false → ¬ image m y | |
167 tchar¬Img m mono y eq im with lem (image m y) | |
168 ... | case2 n = n im | |
169 img-x : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → {y : a} → image m y → b | |
170 img-x m {.(m x)} (isImage x) = x | |
1026 | 171 m-img-x : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → {y : a} → (t : image m y ) → m (img-x m t) ≡ y |
172 m-img-x m (isImage x) = refl | |
1025 | 173 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ ) |
174 img-cong : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y y' : a) → y ≡ y' → (s : image m y ) (t : image m y') → s ≅ t | |
175 img-cong {a} {b} m mono .(m x) .(m x₁) eq (isImage x) (isImage x₁) | |
176 with cong (λ k → k ! ) ( Mono.isMono mono {One} (λ _ → x) (λ _ → x₁ ) ( extensionality Sets ( λ _ → eq )) ) | |
177 ... | refl = HE.refl | |
178 img-x-cong : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y y' : a) → y ≡ y' → (s : image m y ) →( t : image m y') → img-x m s ≡ img-x m t | |
179 img-x-cong {a} {b} m mono .(m x) .(m x₁) eq (isImage x) (isImage x₁) | |
180 with cong (λ k → k ! ) ( Mono.isMono mono {One} (λ _ → x) (λ _ → x₁ ) ( extensionality Sets ( λ _ → eq )) ) | |
181 ... | refl = refl | |
182 img-x-cong0 : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y : a) → (s t : image m y ) → img-x m s ≡ img-x m t | |
183 img-x-cong0 m mono y s t = img-x-cong m mono y y refl s t | |
1023 | 184 isol : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → IsoL Sets m (λ (e : sequ a Bool (tchar m mono) (λ _ → true )) → equ e ) |
1025 | 185 isol {a} {b} m mono = record { iso-L = record { ≅→ = b→s ; ≅← = b←s ; |
1026 | 186 iso→ = extensionality Sets ( λ x → iso1 x ) |
1025 | 187 ; iso← = extensionality Sets ( λ x → iso2 x) } ; iso≈L = {!!} } where |
1023 | 188 b→s : Hom Sets b (sequ a Bool (tchar m mono) (λ _ → true)) |
1025 | 189 b→s x with tchar m mono (m x) | inspect (tchar m mono ) (m x) |
190 ... | true | record { eq = eq1 } = elem (m x) eq1 | |
191 b→s x | false | record { eq = eq1 } with tchar¬Img m mono (m x) eq1 | |
192 ... | t = ⊥-elim (t (isImage x)) | |
1023 | 193 b←s : Hom Sets (sequ a Bool (tchar m mono) (λ _ → true)) b |
1025 | 194 b←s (elem y eq) with tchar m mono y | inspect (tchar m mono ) y |
195 ... | true | record { eq = eq1 } = img-x m (tcharImg m mono y eq1 ) | |
1027 | 196 bs1 : (y : a) → (eq1 : tchar m mono y ≡ true ) → b←s ( elem y eq1 ) ≡ img-x m (tcharImg m mono y eq1 ) |
197 bs1 y eq1 = ? | |
1026 | 198 iso1 : (x : b) → b←s ( b→s x ) ≡ x |
1025 | 199 iso1 x with tchar m mono (m x) | inspect (tchar m mono ) (m x) |
1026 | 200 ... | true | record { eq = eq1 } with tcharImg m mono (m x) eq1 | inspect ( tcharImg m mono (m x) ) eq1 |
201 ... | t | record { eq = eq2 } = begin | |
1027 | 202 b←s ( elem (m x) eq1 ) ≡⟨ bs1 (m x) eq1 ⟩ |
1026 | 203 img-x m (tcharImg m mono (m x) eq1 ) ≡⟨ cong (λ k → img-x m k ) eq2 ⟩ |
204 img-x m t ≡⟨ img-x-cong0 m mono (m x ) t (isImage x) ⟩ | |
205 img-x m (isImage x) ≡⟨⟩ | |
206 x ∎ where open ≡-Reasoning | |
1025 | 207 iso1 x | false | record { eq = eq1 } = ⊥-elim ( tchar¬Img m mono (m x) eq1 (isImage x)) |
208 iso2 : (x : sequ a Bool (tchar m mono) (λ _ → true) ) → (Sets [ b→s o b←s ]) x ≡ id1 Sets (sequ a Bool (tchar m mono) (λ _ → true)) x | |
209 iso2 (elem y eq) = {!!} | |
1023 | 210 imequ : {a b : Obj Sets} (m : Hom Sets b a) (mono : Mono Sets m) → IsEqualizer Sets m (tchar m mono) (Sets [ (λ _ → true ) o CCC.○ sets a ]) |
1022 | 211 imequ {a} {b} m mono = equalizerIso _ _ (tker (tchar m mono)) m (isol m mono) |
1023 | 212 uniq : {a : Obj (Sets {c})} {b : Obj Sets} (h : Hom Sets a Bool) (m : Hom Sets b a) (mono : Mono Sets m) (y : a) → |
213 tchar (Equalizer.equalizer (tker h)) (record { isMono = λ f g → monic (tker h) }) y ≡ h y | |
1025 | 214 uniq {a} {b} h m mono y with h y | inspect h y | lem (image (Equalizer.equalizer (tker h)) y ) | inspect (tchar (Equalizer.equalizer (tker h)) (record { isMono = λ f g → monic (tker h) })) y |
215 ... | true | record { eq = eqhy } | case1 x | record { eq = eq1 } = eq1 | |
216 ... | true | record { eq = eqhy } | case2 x | record { eq = eq1 } = ⊥-elim (x (isImage (elem y eqhy))) | |
217 ... | false | record { eq = eqhy } | case1 (isImage (elem x eq)) | record { eq = eq1 } = ⊥-elim ( ¬x≡t∧x≡f record {fst = eqhy ; snd = eq }) | |
218 ... | false | record { eq = eqhy } | case2 x | record { eq = eq1 } = eq1 | |
1010 | 219 |
999 | 220 |
221 open import graph | |
222 module ccc-from-graph {c₁ c₂ : Level } (G : Graph {c₁} {c₂}) where | |
223 | |
224 open import Relation.Binary.PropositionalEquality renaming ( cong to ≡-cong ) hiding ( [_] ) | |
225 open Graph | |
226 | |
227 V = vertex G | |
228 E : V → V → Set c₂ | |
229 E = edge G | |
230 | |
231 data Objs : Set c₁ where | |
232 atom : V → Objs | |
233 ⊤ : Objs | |
234 _∧_ : Objs → Objs → Objs | |
235 _<=_ : Objs → Objs → Objs | |
236 | |
237 data Arrows : (b c : Objs ) → Set (c₁ ⊔ c₂) | |
238 data Arrow : Objs → Objs → Set (c₁ ⊔ c₂) where --- case i | |
239 arrow : {a b : V} → E a b → Arrow (atom a) (atom b) | |
240 π : {a b : Objs } → Arrow ( a ∧ b ) a | |
241 π' : {a b : Objs } → Arrow ( a ∧ b ) b | |
242 ε : {a b : Objs } → Arrow ((a <= b) ∧ b ) a | |
243 _* : {a b c : Objs } → Arrows (c ∧ b ) a → Arrow c ( a <= b ) --- case v | |
244 | |
245 data Arrows where | |
246 id : ( a : Objs ) → Arrows a a --- case i | |
247 ○ : ( a : Objs ) → Arrows a ⊤ --- case i | |
248 <_,_> : {a b c : Objs } → Arrows c a → Arrows c b → Arrows c (a ∧ b) -- case iii | |
249 iv : {b c d : Objs } ( f : Arrow d c ) ( g : Arrows b d ) → Arrows b c -- cas iv | |
250 | |
251 _・_ : {a b c : Objs } (f : Arrows b c ) → (g : Arrows a b) → Arrows a c | |
252 id a ・ g = g | |
253 ○ a ・ g = ○ _ | |
254 < f , g > ・ h = < f ・ h , g ・ h > | |
255 iv f g ・ h = iv f ( g ・ h ) | |
256 | |
257 | |
258 identityL : {A B : Objs} {f : Arrows A B} → (id B ・ f) ≡ f | |
259 identityL = refl | |
260 | |
261 identityR : {A B : Objs} {f : Arrows A B} → (f ・ id A) ≡ f | |
262 identityR {a} {a} {id a} = refl | |
263 identityR {a} {⊤} {○ a} = refl | |
264 identityR {a} {_} {< f , f₁ >} = cong₂ (λ j k → < j , k > ) identityR identityR | |
265 identityR {a} {b} {iv f g} = cong (λ k → iv f k ) identityR | |
266 | |
267 assoc≡ : {a b c d : Objs} (f : Arrows c d) (g : Arrows b c) (h : Arrows a b) → | |
268 (f ・ (g ・ h)) ≡ ((f ・ g) ・ h) | |
269 assoc≡ (id a) g h = refl | |
270 assoc≡ (○ a) g h = refl | |
271 assoc≡ < f , f₁ > g h = cong₂ (λ j k → < j , k > ) (assoc≡ f g h) (assoc≡ f₁ g h) | |
272 assoc≡ (iv f f1) g h = cong (λ k → iv f k ) ( assoc≡ f1 g h ) | |
273 | |
274 -- positive intutionistic calculus | |
275 PL : Category c₁ (c₁ ⊔ c₂) (c₁ ⊔ c₂) | |
276 PL = record { | |
277 Obj = Objs; | |
278 Hom = λ a b → Arrows a b ; | |
279 _o_ = λ{a} {b} {c} x y → x ・ y ; | |
280 _≈_ = λ x y → x ≡ y ; | |
281 Id = λ{a} → id a ; | |
282 isCategory = record { | |
283 isEquivalence = record {refl = refl ; trans = trans ; sym = sym} ; | |
284 identityL = λ {a b f} → identityL {a} {b} {f} ; | |
285 identityR = λ {a b f} → identityR {a} {b} {f} ; | |
286 o-resp-≈ = λ {a b c f g h i} → o-resp-≈ {a} {b} {c} {f} {g} {h} {i} ; | |
287 associative = λ{a b c d f g h } → assoc≡ f g h | |
288 } | |
289 } where | |
290 o-resp-≈ : {A B C : Objs} {f g : Arrows A B} {h i : Arrows B C} → | |
291 f ≡ g → h ≡ i → (h ・ f) ≡ (i ・ g) | |
292 o-resp-≈ refl refl = refl | |
293 -------- | |
294 -- | |
295 -- Functor from Positive Logic to Sets | |
296 -- | |
297 | |
298 -- open import Category.Sets | |
299 -- postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Relation.Binary.PropositionalEquality.Extensionalit y c₂ c₂ | |
300 | |
301 open import Data.List | |
302 | |
303 C = graphtocat.Chain G | |
304 | |
305 tr : {a b : vertex G} → edge G a b → ((y : vertex G) → C y a) → (y : vertex G) → C y b | |
306 tr f x y = graphtocat.next f (x y) | |
307 | |
308 fobj : ( a : Objs ) → Set (c₁ ⊔ c₂) | |
309 fobj (atom x) = ( y : vertex G ) → C y x | |
310 fobj ⊤ = One | |
311 fobj (a ∧ b) = ( fobj a /\ fobj b) | |
312 fobj (a <= b) = fobj b → fobj a | |
313 | |
314 fmap : { a b : Objs } → Hom PL a b → fobj a → fobj b | |
315 amap : { a b : Objs } → Arrow a b → fobj a → fobj b | |
316 amap (arrow x) y = tr x y -- tr x | |
317 amap π ( x , y ) = x | |
318 amap π' ( x , y ) = y | |
319 amap ε (f , x ) = f x | |
320 amap (f *) x = λ y → fmap f ( x , y ) | |
321 fmap (id a) x = x | |
1020 | 322 fmap (○ a) x = ! |
999 | 323 fmap < f , g > x = ( fmap f x , fmap g x ) |
324 fmap (iv x f) a = amap x ( fmap f a ) | |
325 | |
326 -- CS is a map from Positive logic to Sets | |
327 -- Sets is CCC, so we have a cartesian closed category generated by a graph | |
328 -- as a sub category of Sets | |
329 | |
330 CS : Functor PL (Sets {c₁ ⊔ c₂}) | |
331 FObj CS a = fobj a | |
332 FMap CS {a} {b} f = fmap {a} {b} f | |
333 isFunctor CS = isf where | |
334 _+_ = Category._o_ PL | |
335 ++idR = IsCategory.identityR ( Category.isCategory PL ) | |
336 distr : {a b c : Obj PL} { f : Hom PL a b } { g : Hom PL b c } → (z : fobj a ) → fmap (g + f) z ≡ fmap g (fmap f z) | |
337 distr {a} {a₁} {a₁} {f} {id a₁} z = refl | |
338 distr {a} {a₁} {⊤} {f} {○ a₁} z = refl | |
339 distr {a} {b} {c ∧ d} {f} {< g , g₁ >} z = cong₂ (λ j k → j , k ) (distr {a} {b} {c} {f} {g} z) (distr {a} {b} {d} {f} {g₁} z) | |
340 distr {a} {b} {c} {f} {iv {_} {_} {d} x g} z = adistr (distr {a} {b} {d} {f} {g} z) x where | |
341 adistr : fmap (g + f) z ≡ fmap g (fmap f z) → | |
342 ( x : Arrow d c ) → fmap ( iv x (g + f) ) z ≡ fmap ( iv x g ) (fmap f z ) | |
343 adistr eq x = cong ( λ k → amap x k ) eq | |
344 isf : IsFunctor PL Sets fobj fmap | |
345 IsFunctor.identity isf = extensionality Sets ( λ x → refl ) | |
346 IsFunctor.≈-cong isf refl = refl | |
347 IsFunctor.distr isf {a} {b} {c} {g} {f} = extensionality Sets ( λ z → distr {a} {b} {c} {g} {f} z ) | |
348 | |
349 |