Mercurial > hg > Members > kono > Proof > category
annotate em-category.agda @ 377:2dfa2d59268c
associative in twocat dead end?
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 11 Mar 2016 18:49:58 +0900 |
parents | 93cf0a6c21fe |
children | d3cd28a71b3f |
rev | line source |
---|---|
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
1 -- -- -- -- -- -- -- -- |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
2 -- Monad to Eilenberg-Moore Category |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
3 -- defines U^T and F^T as a resolution of Monad |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
4 -- checks Adjointness |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
5 -- |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
6 -- Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
7 -- -- -- -- -- -- -- -- |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
8 |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
9 -- Monad |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
10 -- Category A |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
11 -- A = Category |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
12 -- Functor T : A → A |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
13 --T(a) = t(a) |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
14 --T(f) = tf(f) |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
15 |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
16 open import Category -- https://github.com/konn/category-agda |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
17 open import Level |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
18 --open import Category.HomReasoning |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
19 open import HomReasoning |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
20 open import cat-utility |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
21 open import Category.Cat |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
22 |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
23 module em-category { c₁ c₂ ℓ : Level} { A : Category c₁ c₂ ℓ } |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
24 { T : Functor A A } |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
25 { η : NTrans A A identityFunctor T } |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
26 { μ : NTrans A A (T ○ T) T } |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
27 { M : Monad A T η μ } where |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
28 |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
29 -- |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
30 -- Hom in Eilenberg-Moore Category |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
31 -- |
101 | 32 open Functor |
33 open NTrans | |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
34 |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
35 record IsAlgebra {a : Obj A} { phi : Hom A (FObj T a) a } : Set (c₁ ⊔ c₂ ⊔ ℓ) where |
101 | 36 field |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
37 identity : A [ A [ phi o TMap η a ] ≈ id1 A a ] |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
38 eval : A [ A [ phi o TMap μ a ] ≈ A [ phi o FMap T phi ] ] |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
39 |
113 | 40 record EMObj : Set (c₁ ⊔ c₂ ⊔ ℓ) where |
41 field | |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
42 a : Obj A |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
43 phi : Hom A (FObj T a) a |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
44 isAlgebra : IsAlgebra {a} {phi} |
105 | 45 obj : Obj A |
46 obj = a | |
113 | 47 φ : Hom A (FObj T a) a |
48 φ = phi | |
108 | 49 open EMObj |
104 | 50 |
113 | 51 record Eilenberg-Moore-Hom (a : EMObj ) (b : EMObj ) : Set (c₁ ⊔ c₂ ⊔ ℓ) where |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
52 field |
113 | 53 EMap : Hom A (obj a) (obj b) |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
54 homomorphism : A [ A [ (φ b) o FMap T EMap ] ≈ A [ EMap o (φ a) ] ] |
104 | 55 open Eilenberg-Moore-Hom |
105 | 56 |
300 | 57 EMHom : (a : EMObj ) (b : EMObj ) → Set (c₁ ⊔ c₂ ⊔ ℓ) |
58 EMHom = λ a b → Eilenberg-Moore-Hom a b | |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
59 |
113 | 60 Lemma-EM1 : {x : Obj A} {φ : Hom A (FObj T x) x} (a : EMObj ) |
300 | 61 → A [ A [ φ o FMap T (id1 A x) ] ≈ A [ (id1 A x) o φ ] ] |
108 | 62 Lemma-EM1 {x} {φ} a = let open ≈-Reasoning (A) in |
101 | 63 begin |
64 φ o FMap T (id1 A x) | |
65 ≈⟨ cdr ( IsFunctor.identity (isFunctor T) ) ⟩ | |
66 φ o (id1 A (FObj T x)) | |
67 ≈⟨ idR ⟩ | |
68 φ | |
69 ≈⟨ sym idL ⟩ | |
70 (id1 A x) o φ | |
71 ∎ | |
72 | |
300 | 73 EM-id : { a : EMObj } → EMHom a a |
113 | 74 EM-id {a} = record { EMap = id1 A (obj a); |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
75 homomorphism = Lemma-EM1 {obj a} {phi a} a } |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
76 |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
77 open import Relation.Binary.Core |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
78 |
113 | 79 Lemma-EM2 : (a : EMObj ) |
80 (b : EMObj ) | |
81 (c : EMObj ) | |
82 (g : EMHom b c) | |
83 (f : EMHom a b) | |
300 | 84 → A [ A [ φ c o FMap T (A [ (EMap g) o (EMap f) ] ) ] |
113 | 85 ≈ A [ (A [ (EMap g) o (EMap f) ]) o φ a ] ] |
86 Lemma-EM2 a b c g f = let | |
103 | 87 open ≈-Reasoning (A) in |
101 | 88 begin |
113 | 89 φ c o FMap T ((EMap g) o (EMap f)) |
106 | 90 ≈⟨ cdr (distr T) ⟩ |
113 | 91 φ c o ( FMap T (EMap g) o FMap T (EMap f) ) |
106 | 92 ≈⟨ assoc ⟩ |
113 | 93 ( φ c o FMap T (EMap g)) o FMap T (EMap f) |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
94 ≈⟨ car (homomorphism (g)) ⟩ |
113 | 95 ( EMap g o φ b) o FMap T (EMap f) |
106 | 96 ≈⟨ sym assoc ⟩ |
113 | 97 EMap g o (φ b o FMap T (EMap f) ) |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
98 ≈⟨ cdr (homomorphism (f)) ⟩ |
113 | 99 EMap g o (EMap f o φ a) |
107
da14b7f03ff8
no yellow. ready to define category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
106
diff
changeset
|
100 ≈⟨ assoc ⟩ |
113 | 101 ( (EMap g) o (EMap f) ) o φ a |
101 | 102 ∎ |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
103 |
300 | 104 _∙_ : {a b c : EMObj } → EMHom b c → EMHom a b → EMHom a c |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
105 _∙_ {a} {b} {c} g f = record { EMap = A [ EMap g o EMap f ] ; homomorphism = Lemma-EM2 a b c g f } |
111
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
106 |
300 | 107 _≗_ : {a : EMObj } {b : EMObj } (f g : EMHom a b ) → Set ℓ |
113 | 108 _≗_ f g = A [ EMap f ≈ EMap g ] |
108 | 109 |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
110 -- |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
111 -- cannot use as identityL = EMidL |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
112 -- |
300 | 113 EMidL : {C D : EMObj} → {f : EMHom C D} → (EM-id ∙ f) ≗ f |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
114 EMidL {C} {D} {f} = let open ≈-Reasoning (A) in idL {obj D} |
300 | 115 EMidR : {C D : EMObj} → {f : EMHom C D} → (f ∙ EM-id) ≗ f |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
116 EMidR {C} {_} {_} = let open ≈-Reasoning (A) in idR {obj C} |
300 | 117 EMo-resp : {a b c : EMObj} → {f g : EMHom a b } → {h i : EMHom b c } → |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
118 f ≗ g → h ≗ i → (h ∙ f) ≗ (i ∙ g) |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
119 EMo-resp {a} {b} {c} {f} {g} {h} {i} = ( IsCategory.o-resp-≈ (Category.isCategory A) ) |
300 | 120 EMassoc : {a b c d : EMObj} → {f : EMHom c d } → {g : EMHom b c } → {h : EMHom a b } → |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
121 (f ∙ (g ∙ h)) ≗ ((f ∙ g) ∙ h) |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
122 EMassoc {_} {_} {_} {_} {f} {g} {h} = ( IsCategory.associative (Category.isCategory A) ) |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
123 |
111
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
124 isEilenberg-MooreCategory : IsCategory EMObj EMHom _≗_ _∙_ EM-id |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
125 isEilenberg-MooreCategory = record { isEquivalence = isEquivalence |
114
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
126 ; identityL = IsCategory.identityL (Category.isCategory A) |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
127 ; identityR = IsCategory.identityR (Category.isCategory A) |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
128 ; o-resp-≈ = IsCategory.o-resp-≈ (Category.isCategory A) |
2032c438b6a6
EM Category constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
113
diff
changeset
|
129 ; associative = IsCategory.associative (Category.isCategory A) |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
130 } |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
131 where |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
132 open ≈-Reasoning (A) |
300 | 133 isEquivalence : {a : EMObj } {b : EMObj } → |
113 | 134 IsEquivalence {_} {_} {EMHom a b } _≗_ |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
135 isEquivalence {C} {D} = -- this is the same function as A's equivalence but has different types |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
136 record { refl = refl-hom |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
137 ; sym = sym |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
138 ; trans = trans-hom |
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
139 } |
111
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
140 Eilenberg-MooreCategory : Category (c₁ ⊔ c₂ ⊔ ℓ) (c₁ ⊔ c₂ ⊔ ℓ) ℓ |
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
141 Eilenberg-MooreCategory = |
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
142 record { Obj = EMObj |
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
143 ; Hom = EMHom |
112
5f8d6d1aece4
constructed but some yellow remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
111
diff
changeset
|
144 ; _o_ = _∙_ |
111
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
145 ; _≈_ = _≗_ |
112
5f8d6d1aece4
constructed but some yellow remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
111
diff
changeset
|
146 ; Id = EM-id |
111
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
147 ; isCategory = isEilenberg-MooreCategory |
c670d0e6b1e2
Category._o_ /= Category.Category.Id
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
110
diff
changeset
|
148 } |
100
59dec035602c
Eilenberg-Moore Category start
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
149 |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
150 |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
151 -- Resolution |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
152 -- T = U^T U^F |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
153 -- ε^t |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
154 -- η^T |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
155 |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
156 U^T : Functor Eilenberg-MooreCategory A |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
157 U^T = record { |
300 | 158 FObj = λ a → obj a |
159 ; FMap = λ f → EMap f | |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
160 ; isFunctor = record |
300 | 161 { ≈-cong = λ eq → eq |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
162 ; identity = refl-hom |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
163 ; distr = refl-hom |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
164 } |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
165 } where open ≈-Reasoning (A) |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
166 |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
167 open Monad |
300 | 168 Lemma-EM4 : (x : Obj A ) → A [ A [ TMap μ x o TMap η (FObj T x) ] ≈ id1 A (FObj T x) ] |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
169 Lemma-EM4 x = let open ≈-Reasoning (A) in |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
170 begin |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
171 TMap μ x o TMap η (FObj T x) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
172 ≈⟨ IsMonad.unity1 (isMonad M) ⟩ |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
173 id1 A (FObj T x) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
174 ∎ |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
175 |
300 | 176 Lemma-EM5 : (x : Obj A ) → A [ A [ ( TMap μ x) o TMap μ (FObj T x) ] ≈ A [ ( TMap μ x) o FMap T ( TMap μ x) ] ] |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
177 Lemma-EM5 x = let open ≈-Reasoning (A) in |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
178 begin |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
179 ( TMap μ x) o TMap μ (FObj T x) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
180 ≈⟨ IsMonad.assoc (isMonad M) ⟩ |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
181 ( TMap μ x) o FMap T ( TMap μ x) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
182 ∎ |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
183 |
300 | 184 ftobj : Obj A → EMObj |
185 ftobj = λ x → record { a = FObj T x ; phi = TMap μ x; | |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
186 isAlgebra = record { |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
187 identity = Lemma-EM4 x; |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
188 eval = Lemma-EM5 x |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
189 } } |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
190 |
300 | 191 Lemma-EM6 : (a b : Obj A ) → (f : Hom A a b ) → |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
192 A [ A [ (φ (ftobj b)) o FMap T (FMap T f) ] ≈ A [ FMap T f o (φ (ftobj a)) ] ] |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
193 Lemma-EM6 a b f = let open ≈-Reasoning (A) in |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
194 begin |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
195 (φ (ftobj b)) o FMap T (FMap T f) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
196 ≈⟨⟩ |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
197 TMap μ b o FMap T (FMap T f) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
198 ≈⟨ sym (nat μ) ⟩ |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
199 FMap T f o TMap μ a |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
200 ≈⟨⟩ |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
201 FMap T f o (φ (ftobj a)) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
202 ∎ |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
203 |
300 | 204 ftmap : {a b : Obj A} → (Hom A a b) → EMHom (ftobj a) (ftobj b) |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
205 ftmap {a} {b} f = record { EMap = FMap T f; homomorphism = Lemma-EM6 a b f } |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
206 |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
207 F^T : Functor A Eilenberg-MooreCategory |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
208 F^T = record { |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
209 FObj = ftobj |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
210 ; FMap = ftmap |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
211 ; isFunctor = record { |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
212 ≈-cong = ≈-cong |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
213 ; identity = identity |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
214 ; distr = distr1 |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
215 } |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
216 } where |
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
217 ≈-cong : {a b : Obj A} {f g : Hom A a b} → A [ f ≈ g ] → (ftmap f) ≗ (ftmap g) |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
218 ≈-cong = let open ≈-Reasoning (A) in ( fcong T ) |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
219 identity : {a : Obj A} → ftmap (id1 A a) ≗ EM-id {ftobj a} |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
220 identity = IsFunctor.identity ( isFunctor T ) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
221 distr1 : {a b c : Obj A} {f : Hom A a b} {g : Hom A b c} → ftmap (A [ g o f ]) ≗ ( ftmap g ∙ ftmap f ) |
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
222 distr1 = let open ≈-Reasoning (A) in ( distr T ) |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
223 |
117 | 224 -- T = (U^T ○ F^T) |
225 | |
300 | 226 Lemma-EM7 : ∀{a b : Obj A} → (f : Hom A a b ) → A [ FMap T f ≈ FMap (U^T ○ F^T) f ] |
117 | 227 Lemma-EM7 {a} {b} f = let open ≈-Reasoning (A) in |
228 sym ( begin | |
229 FMap (U^T ○ F^T) f | |
230 ≈⟨⟩ | |
231 EMap ( ftmap f ) | |
232 ≈⟨⟩ | |
233 FMap T f | |
234 ∎ ) | |
235 | |
236 Lemma-EM8 : T ≃ (U^T ○ F^T) | |
237 Lemma-EM8 f = Category.Cat.refl (Lemma-EM7 f) | |
238 | |
239 η^T : NTrans A A identityFunctor ( U^T ○ F^T ) | |
300 | 240 η^T = record { TMap = λ x → TMap η x ; isNTrans = isNTrans1 } where |
130 | 241 commute1 : {a b : Obj A} {f : Hom A a b} |
117 | 242 → A [ A [ ( FMap ( U^T ○ F^T ) f ) o ( TMap η a ) ] ≈ A [ (TMap η b ) o f ] ] |
130 | 243 commute1 {a} {b} {f} = let open ≈-Reasoning (A) in |
117 | 244 begin |
245 ( FMap ( U^T ○ F^T ) f ) o ( TMap η a ) | |
246 ≈⟨ sym (resp refl-hom (Lemma-EM7 f)) ⟩ | |
247 FMap T f o TMap η a | |
248 ≈⟨ nat η ⟩ | |
249 TMap η b o f | |
250 ∎ | |
300 | 251 isNTrans1 : IsNTrans A A identityFunctor ( U^T ○ F^T ) (λ a → TMap η a) |
130 | 252 isNTrans1 = record { commute = commute1 } |
117 | 253 |
300 | 254 Lemma-EM9 : (a : EMObj) → A [ A [ (φ a) o FMap T (φ a) ] ≈ A [ (φ a) o (φ (FObj ( F^T ○ U^T ) a)) ] ] |
118 | 255 Lemma-EM9 a = let open ≈-Reasoning (A) in |
256 sym ( begin | |
257 (φ a) o (φ (FObj ( F^T ○ U^T ) a)) | |
258 ≈⟨⟩ | |
259 (φ a) o (TMap μ (obj a)) | |
260 ≈⟨ IsAlgebra.eval (isAlgebra a) ⟩ | |
261 (φ a) o FMap T (φ a) | |
262 ∎ ) | |
263 | |
300 | 264 emap-ε : (a : EMObj ) → EMHom (FObj ( F^T ○ U^T ) a) a |
118 | 265 emap-ε a = record { EMap = φ a ; homomorphism = Lemma-EM9 a } |
266 | |
267 ε^T : NTrans Eilenberg-MooreCategory Eilenberg-MooreCategory ( F^T ○ U^T ) identityFunctor | |
300 | 268 ε^T = record { TMap = λ a → emap-ε a; isNTrans = isNTrans1 } where |
130 | 269 commute1 : {a b : EMObj } {f : EMHom a b} |
118 | 270 → (f ∙ ( emap-ε a ) ) ≗ (( emap-ε b ) ∙( FMap (F^T ○ U^T) f ) ) |
130 | 271 commute1 {a} {b} {f} = let open ≈-Reasoning (A) in |
118 | 272 sym ( begin |
273 EMap (( emap-ε b ) ∙ ( FMap (F^T ○ U^T) f )) | |
274 ≈⟨⟩ | |
275 φ b o FMap T (EMap f) | |
276 ≈⟨ homomorphism f ⟩ | |
277 EMap f o φ a | |
278 ≈⟨⟩ | |
279 EMap (f ∙ ( emap-ε a )) | |
280 ∎ ) | |
300 | 281 isNTrans1 : IsNTrans Eilenberg-MooreCategory Eilenberg-MooreCategory ( F^T ○ U^T ) identityFunctor (λ a → emap-ε a ) |
301 | 282 isNTrans1 = record { commute = λ {a} {b} {f} → commute1 {a} {b} {f} } |
118 | 283 |
284 -- | |
285 -- μ = U^T ε U^F | |
286 -- | |
287 | |
300 | 288 emap-μ : (a : Obj A) → Hom A (FObj ( U^T ○ F^T ) (FObj ( U^T ○ F^T ) a)) (FObj ( U^T ○ F^T ) a) |
118 | 289 emap-μ a = FMap U^T ( TMap ε^T ( FObj F^T a )) |
290 | |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
291 μ^T : NTrans A A (( U^T ○ F^T ) ○ ( U^T ○ F^T )) ( U^T ○ F^T ) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
292 μ^T = record { TMap = emap-μ ; isNTrans = isNTrans1 } where |
130 | 293 commute1 : {a b : Obj A} {f : Hom A a b} |
118 | 294 → A [ A [ (FMap (U^T ○ F^T) f) o ( emap-μ a) ] ≈ A [ ( emap-μ b ) o FMap (U^T ○ F^T) ( FMap (U^T ○ F^T) f) ] ] |
130 | 295 commute1 {a} {b} {f} = let open ≈-Reasoning (A) in |
119 | 296 sym ( begin |
297 ( emap-μ b ) o FMap (U^T ○ F^T) ( FMap (U^T ○ F^T) f) | |
298 ≈⟨⟩ | |
299 (FMap U^T ( TMap ε^T ( FObj F^T b ))) o (FMap (U^T ○ F^T) ( FMap (U^T ○ F^T) f) ) | |
300 ≈⟨⟩ | |
301 (TMap μ b) o (FMap T (FMap T f)) | |
302 ≈⟨ sym (nat μ) ⟩ | |
303 FMap T f o ( TMap μ a ) | |
304 ≈⟨⟩ | |
305 (FMap (U^T ○ F^T) f) o ( emap-μ a) | |
306 ∎ ) | |
118 | 307 isNTrans1 : IsNTrans A A (( U^T ○ F^T ) ○ ( U^T ○ F^T )) ( U^T ○ F^T ) emap-μ |
130 | 308 isNTrans1 = record { commute = commute1 } |
118 | 309 |
300 | 310 Lemma-EM10 : {x : Obj A } → A [ TMap μ^T x ≈ FMap U^T ( TMap ε^T ( FObj F^T x ) ) ] |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
311 Lemma-EM10 {x} = let open ≈-Reasoning (A) in |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
312 sym ( begin |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
313 FMap U^T ( TMap ε^T ( FObj F^T x ) ) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
314 ≈⟨⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
315 emap-μ x |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
316 ≈⟨⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
317 TMap μ^T x |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
318 ∎ ) |
117 | 319 |
300 | 320 Lemma-EM11 : {x : Obj A } → A [ TMap μ x ≈ FMap U^T ( TMap ε^T ( FObj F^T x ) ) ] |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
321 Lemma-EM11 {x} = let open ≈-Reasoning (A) in |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
322 sym ( begin |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
323 FMap U^T ( TMap ε^T ( FObj F^T x ) ) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
324 ≈⟨⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
325 TMap μ x |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
326 ∎ ) |
116
0e37b2cf3c73
F^T and U^T constructed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
115
diff
changeset
|
327 |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
328 Adj^T : Adjunction A Eilenberg-MooreCategory U^T F^T η^T ε^T |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
329 Adj^T = record { |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
330 isAdjunction = record { |
300 | 331 adjoint1 = λ {b} → IsAlgebra.identity (isAlgebra b) ; -- adjoint1 |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
332 adjoint2 = adjoint2 |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
333 } |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
334 } where |
122
f8fbd5ecec97
no yellow on em-category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
120
diff
changeset
|
335 adjoint1 : { b : EMObj } → |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
336 A [ A [ ( FMap U^T ( TMap ε^T b)) o ( TMap η^T ( FObj U^T b )) ] ≈ id1 A (FObj U^T b) ] |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
337 adjoint1 {b} = let open ≈-Reasoning (A) in |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
338 begin |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
339 ( FMap U^T ( TMap ε^T b)) o ( TMap η^T ( FObj U^T b )) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
340 ≈⟨⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
341 φ b o TMap η (obj b) |
122
f8fbd5ecec97
no yellow on em-category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
120
diff
changeset
|
342 ≈⟨ IsAlgebra.identity (isAlgebra b) ⟩ |
f8fbd5ecec97
no yellow on em-category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
120
diff
changeset
|
343 id1 A (a b) |
f8fbd5ecec97
no yellow on em-category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
120
diff
changeset
|
344 ≈⟨⟩ |
f8fbd5ecec97
no yellow on em-category
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
120
diff
changeset
|
345 id1 A (FObj U^T b) |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
346 ∎ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
347 adjoint2 : {a : Obj A} → |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
348 Eilenberg-MooreCategory [ Eilenberg-MooreCategory [ ( TMap ε^T ( FObj F^T a )) o ( FMap F^T ( TMap η^T a )) ] |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
349 ≈ id1 Eilenberg-MooreCategory (FObj F^T a) ] |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
350 adjoint2 {a} = let open ≈-Reasoning (A) in |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
351 begin |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
352 EMap (( TMap ε^T ( FObj F^T a )) ∙ ( FMap F^T ( TMap η^T a )) ) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
353 ≈⟨⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
354 TMap μ a o FMap T ( TMap η a ) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
355 ≈⟨ IsMonad.unity2 (isMonad M) ⟩ |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
356 EMap (id1 Eilenberg-MooreCategory (FObj F^T a)) |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
357 ∎ |
115
17e69b05bc5e
U^T and F^T problem written
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
114
diff
changeset
|
358 |
120
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
359 open MResolution |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
360 |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
361 Resolution^T : MResolution A Eilenberg-MooreCategory T U^T F^T {η^T} {ε^T} {μ^T} Adj^T |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
362 Resolution^T = record { |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
363 T=UF = Lemma-EM8; |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
364 μ=UεF = Lemma-EM11 |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
365 } |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
366 |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
367 |
494f870ad54b
EM Resolution complete
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
119
diff
changeset
|
368 -- end |