Mercurial > hg > Members > kono > Proof > category
annotate HomReasoning.agda @ 445:00cf84846d81
cont..
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Fri, 14 Oct 2016 19:44:59 +0900 |
parents | 3e44951b9bdb |
children | 55a9b6177ed4 |
rev | line source |
---|---|
56 | 1 module HomReasoning where |
31 | 2 |
3 -- Shinji KONO <kono@ie.u-ryukyu.ac.jp> | |
4 | |
5 open import Category -- https://github.com/konn/category-agda | |
6 open import Level | |
7 open Functor | |
8 | |
9 | |
10 -- F(f) | |
11 -- F(a) ---→ F(b) | |
12 -- | | | |
13 -- |t(a) |t(b) G(f)t(a) = t(b)F(f) | |
14 -- | | | |
15 -- v v | |
16 -- G(a) ---→ G(b) | |
17 -- G(f) | |
18 | |
19 record IsNTrans {c₁ c₂ ℓ c₁′ c₂′ ℓ′ : Level} (D : Category c₁ c₂ ℓ) (C : Category c₁′ c₂′ ℓ′) | |
20 ( F G : Functor D C ) | |
21 (TMap : (A : Obj D) → Hom C (FObj F A) (FObj G A)) | |
22 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁′ ⊔ c₂′ ⊔ ℓ′)) where | |
23 field | |
130 | 24 commute : {a b : Obj D} {f : Hom D a b} |
31 | 25 → C [ C [ ( FMap G f ) o ( TMap a ) ] ≈ C [ (TMap b ) o (FMap F f) ] ] |
26 | |
130 | 27 record NTrans {c₁ c₂ ℓ c₁′ c₂′ ℓ′ : Level} (domain : Category c₁ c₂ ℓ) (codomain : Category c₁′ c₂′ ℓ′) |
28 (F G : Functor domain codomain ) | |
31 | 29 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ ⊔ c₁′ ⊔ c₂′ ⊔ ℓ′)) where |
30 field | |
31 TMap : (A : Obj domain) → Hom codomain (FObj F A) (FObj G A) | |
32 isNTrans : IsNTrans domain codomain F G TMap | |
33 | |
34 | |
35 module ≈-Reasoning {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) where | |
36 open import Relation.Binary.Core | |
37 | |
38 _o_ : {a b c : Obj A } ( x : Hom A a b ) ( y : Hom A c a ) → Hom A c b | |
39 x o y = A [ x o y ] | |
40 | |
41 _≈_ : {a b : Obj A } → Rel (Hom A a b) ℓ | |
42 x ≈ y = A [ x ≈ y ] | |
43 | |
44 infixr 9 _o_ | |
45 infix 4 _≈_ | |
46 | |
47 refl-hom : {a b : Obj A } { x : Hom A a b } → x ≈ x | |
48 refl-hom = IsEquivalence.refl (IsCategory.isEquivalence ( Category.isCategory A )) | |
49 | |
50 trans-hom : {a b : Obj A } { x y z : Hom A a b } → | |
51 x ≈ y → y ≈ z → x ≈ z | |
52 trans-hom b c = ( IsEquivalence.trans (IsCategory.isEquivalence ( Category.isCategory A ))) b c | |
53 | |
54 -- some short cuts | |
55 | |
56 car : {a b c : Obj A } {x y : Hom A a b } { f : Hom A c a } → | |
57 x ≈ y → ( x o f ) ≈ ( y o f ) | |
58 car {f} eq = ( IsCategory.o-resp-≈ ( Category.isCategory A )) ( refl-hom ) eq | |
59 | |
60 cdr : {a b c : Obj A } {x y : Hom A a b } { f : Hom A b c } → | |
61 x ≈ y → f o x ≈ f o y | |
62 cdr {f} eq = ( IsCategory.o-resp-≈ ( Category.isCategory A )) eq (refl-hom ) | |
63 | |
64 id : (a : Obj A ) → Hom A a a | |
65 id a = (Id {_} {_} {_} {A} a) | |
66 | |
67 idL : {a b : Obj A } { f : Hom A b a } → id a o f ≈ f | |
68 idL = IsCategory.identityL (Category.isCategory A) | |
69 | |
70 idR : {a b : Obj A } { f : Hom A a b } → f o id a ≈ f | |
71 idR = IsCategory.identityR (Category.isCategory A) | |
72 | |
73 sym : {a b : Obj A } { f g : Hom A a b } → f ≈ g → g ≈ f | |
74 sym = IsEquivalence.sym (IsCategory.isEquivalence (Category.isCategory A)) | |
75 | |
299
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
76 -- working on another cateogry |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
77 idL1 : { c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) {a b : Obj A } { f : Hom A b a } → A [ A [ Id {_} {_} {_} {A} a o f ] ≈ f ] |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
78 idL1 A = IsCategory.identityL (Category.isCategory A) |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
79 |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
80 idR1 : { c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) {a b : Obj A } { f : Hom A a b } → A [ A [ f o Id {_} {_} {_} {A} a ] ≈ f ] |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
81 idR1 A = IsCategory.identityR (Category.isCategory A) |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
82 |
67 | 83 -- How to prove this? |
75 | 84 ≡-≈ : {a b : Obj A } { x y : Hom A a b } → (x≈y : x ≡ y ) → x ≈ y |
85 ≡-≈ refl = refl-hom | |
86 | |
87 -- ≈-≡ : {a b : Obj A } { x y : Hom A a b } → (x≈y : x ≈ y ) → x ≡ y | |
88 -- ≈-≡ x≈y = irr x≈y | |
89 ≡-cong : { c₁′ c₂′ ℓ′ : Level} {B : Category c₁′ c₂′ ℓ′} {x y : Obj B } { a b : Hom B x y } {x' y' : Obj A } → | |
420
3e44951b9bdb
refl in free-monoid trouble
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
299
diff
changeset
|
90 (f : Hom B x y → Hom A x' y' ) → a ≡ b → f a ≈ f b |
3e44951b9bdb
refl in free-monoid trouble
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
299
diff
changeset
|
91 ≡-cong f refl = ≡-≈ refl |
75 | 92 |
67 | 93 -- cong-≈ : { c₁′ c₂′ ℓ′ : Level} {B : Category c₁′ c₂′ ℓ′} {x y : Obj B } { a b : Hom B x y } {x' y' : Obj A } → |
94 -- B [ a ≈ b ] → (f : Hom B x y → Hom A x' y' ) → f a ≈ f b | |
75 | 95 -- cong-≈ eq f = {!!} |
67 | 96 |
69
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
97 assoc : {a b c d : Obj A } {f : Hom A c d} {g : Hom A b c} {h : Hom A a b} |
299
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
98 → f o ( g o h ) ≈ ( f o g ) o h |
31 | 99 assoc = IsCategory.associative (Category.isCategory A) |
100 | |
299
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
101 -- working on another cateogry |
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
102 assoc1 : { c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) {a b c d : Obj A } {f : Hom A c d} {g : Hom A b c} {h : Hom A a b} |
69
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
103 → A [ A [ f o ( A [ g o h ] ) ] ≈ A [ ( A [ f o g ] ) o h ] ] |
299
8c72f5284bc8
remove module parameter from yoneda functor
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
253
diff
changeset
|
104 assoc1 A = IsCategory.associative (Category.isCategory A) |
69
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
105 |
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
106 distr : { c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ} |
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
107 { c₁′ c₂′ ℓ′ : Level} {D : Category c₁′ c₂′ ℓ′} (T : Functor D A) → {a b c : Obj D} {g : Hom D b c} { f : Hom D a b } |
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
108 → A [ FMap T ( D [ g o f ] ) ≈ A [ FMap T g o FMap T f ] ] |
75 | 109 distr T = IsFunctor.distr ( isFunctor T ) |
110 | |
111 resp : {a b c : Obj A} {f g : Hom A a b} {h i : Hom A b c} → f ≈ g → h ≈ i → (h o f) ≈ (i o g) | |
112 resp = IsCategory.o-resp-≈ (Category.isCategory A) | |
113 | |
114 fcong : { c₁ c₂ ℓ : Level} {C : Category c₁ c₂ ℓ} | |
115 { c₁′ c₂′ ℓ′ : Level} {D : Category c₁′ c₂′ ℓ′} {a b : Obj C} {f g : Hom C a b} → (T : Functor C D) → C [ f ≈ g ] → D [ FMap T f ≈ FMap T g ] | |
116 fcong T = IsFunctor.≈-cong (isFunctor T) | |
31 | 117 |
118 open NTrans | |
69
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
119 nat : { c₁ c₂ ℓ : Level} {A : Category c₁ c₂ ℓ} { c₁′ c₂′ ℓ′ : Level} {D : Category c₁′ c₂′ ℓ′} |
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
120 {a b : Obj D} {f : Hom D a b} {F G : Functor D A } |
31 | 121 → (η : NTrans D A F G ) |
69
84a150c980ce
generalized distr and assco1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
67
diff
changeset
|
122 → A [ A [ FMap G f o TMap η a ] ≈ A [ TMap η b o FMap F f ] ] |
130 | 123 nat η = IsNTrans.commute ( isNTrans η ) |
31 | 124 |
152 | 125 infixr 2 _∎ |
31 | 126 infixr 2 _≈⟨_⟩_ _≈⟨⟩_ |
168 | 127 infixr 2 _≈↑⟨_⟩_ |
31 | 128 infix 1 begin_ |
129 | |
130 ------ If we have this, for example, as an axiom of a category, we can use ≡-Reasoning directly | |
131 -- ≈-to-≡ : {a b : Obj A } { x y : Hom A a b } → A [ x ≈ y ] → x ≡ y | |
132 -- ≈-to-≡ refl-hom = refl | |
133 | |
134 data _IsRelatedTo_ { a b : Obj A } ( x y : Hom A a b ) : | |
135 Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where | |
136 relTo : (x≈y : x ≈ y ) → x IsRelatedTo y | |
137 | |
138 begin_ : { a b : Obj A } { x y : Hom A a b } → | |
139 x IsRelatedTo y → x ≈ y | |
140 begin relTo x≈y = x≈y | |
141 | |
142 _≈⟨_⟩_ : { a b : Obj A } ( x : Hom A a b ) → { y z : Hom A a b } → | |
143 x ≈ y → y IsRelatedTo z → x IsRelatedTo z | |
144 _ ≈⟨ x≈y ⟩ relTo y≈z = relTo (trans-hom x≈y y≈z) | |
145 | |
168 | 146 _≈↑⟨_⟩_ : { a b : Obj A } ( x : Hom A a b ) → { y z : Hom A a b } → |
147 y ≈ x → y IsRelatedTo z → x IsRelatedTo z | |
148 _ ≈↑⟨ y≈x ⟩ relTo y≈z = relTo (trans-hom ( sym y≈x ) y≈z) | |
149 | |
31 | 150 _≈⟨⟩_ : { a b : Obj A } ( x : Hom A a b ) → { y : Hom A a b } → x IsRelatedTo y → x IsRelatedTo y |
151 _ ≈⟨⟩ x∼y = x∼y | |
152 | |
153 _∎ : { a b : Obj A } ( x : Hom A a b ) → x IsRelatedTo x | |
154 _∎ _ = relTo refl-hom | |
155 | |
56 | 156 -- an example |
157 | |
158 Lemma61 : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) → | |
159 { a : Obj A } ( b : Obj A ) → | |
160 ( f : Hom A a b ) | |
161 → A [ A [ (Id {_} {_} {_} {A} b) o f ] ≈ f ] | |
162 Lemma61 c b g = -- IsCategory.identityL (Category.isCategory c) | |
163 let open ≈-Reasoning (c) in | |
164 begin | |
165 c [ Id {_} {_} {_} {c} b o g ] | |
166 ≈⟨ IsCategory.identityL (Category.isCategory c) ⟩ | |
167 g | |
168 ∎ | |
253 | 169 |
170 Lemma62 : {c₁ c₂ ℓ : Level} (A : Category c₁ c₂ ℓ) → | |
171 { a b : Obj A } → | |
172 ( f g : Hom A a b ) | |
173 → A [ A [ (Id {_} {_} {_} {A} b) o f ] ≈ A [ (Id {_} {_} {_} {A} b) o g ] ] | |
174 → A [ g ≈ f ] | |
175 Lemma62 A {a} {b} f g 1g=1f = let open ≈-Reasoning A in | |
176 begin | |
177 g | |
178 ≈↑⟨ idL ⟩ | |
179 id b o g | |
180 ≈↑⟨ 1g=1f ⟩ | |
181 id b o f | |
182 ≈⟨ idL ⟩ | |
183 f | |
184 ∎ |