Mercurial > hg > Members > kono > Proof > category
annotate S.agda @ 793:f37f11e1b871
Hom a,b = Hom 1 b^a
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Mon, 22 Apr 2019 02:41:53 +0900 |
parents | 749df4959d19 |
children |
rev | line source |
---|---|
608 | 1 --I'd like to write the Limit in Sets Category using Agda. Assuming local smallness, a functor is a pair of map on Set OC and I, like this. |
2 -- | |
3 -- sobj : OC → Set c₂ | |
4 -- smap : { i j : OC } → (f : I ) → sobj i → sobj j | |
5 -- | |
6 --A cone for the functor is a record with two fields. Using the record, commutativity of the cone and the propertiesy of the Limit | |
7 --are easity shown, except uniquness. The uniquness of the Limit turned out that congruence of the record with two fields. | |
8 -- | |
9 --In the following agda code, I'd like to prove snat-cong lemma. | |
10 | |
11 open import Level | |
12 module S where | |
13 | |
14 open import Relation.Binary.Core | |
15 open import Function | |
16 import Relation.Binary.PropositionalEquality | |
663
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
17 open import Relation.Binary.HeterogeneousEquality using (_≅_;refl) |
608 | 18 |
19 record snat { c₂ } { I OC : Set c₂ } ( sobj : OC → Set c₂ ) | |
20 ( smap : { i j : OC } → (f : I ) → sobj i → sobj j ) : Set c₂ where | |
21 field | |
22 snmap : ( i : OC ) → sobj i | |
23 sncommute : ( i j : OC ) → ( f : I ) → smap f ( snmap i ) ≡ snmap j | |
24 smap0 : { i j : OC } → (f : I ) → sobj i → sobj j | |
25 smap0 {i} {j} f x = smap f x | |
26 | |
27 open snat | |
28 | |
672 | 29 -- snat-cong' : { c : Level } { I OC : Set c } { SObj : OC → Set c } { SMap : { i j : OC } → (f : I )→ SObj i → SObj j } |
30 -- ( s t : snat SObj SMap ) | |
31 -- → ( ( λ i → snmap s i ) ≡ ( λ i → snmap t i ) ) | |
32 -- → ( ( λ i j f → smap0 s f ( snmap s i ) ≡ snmap s j ) ≡ ( ( λ i j f → smap0 t f ( snmap t i ) ≡ snmap t j ) ) ) | |
33 -- → s ≡ t | |
34 -- snat-cong' s t refl refl = {!!} | |
663
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
35 |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
36 snat-cong : {c : Level} |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
37 {I OC : Set c} |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
38 {sobj : OC → Set c} |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
39 {smap : {i j : OC} → (f : I) → sobj i → sobj j} |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
40 → (s t : snat sobj smap) |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
41 → (snmap-≡ : snmap s ≡ snmap t) |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
42 → (sncommute-≅ : sncommute s ≅ sncommute t) |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
43 → s ≡ t |
855e497a9c8f
introducd HeterogeneousEquality
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
608
diff
changeset
|
44 snat-cong _ _ refl refl = refl |
608 | 45 |
46 --This is quite simlar to the answer of | |
47 -- | |
48 -- Equality on dependent record types | |
49 -- https://stackoverflow.com/questions/37488098/equality-on-dependent-record-types | |
50 -- | |
51 --So it should work something like | |
52 -- | |
53 -- snat-cong s t refl refl = refl | |
54 -- | |
55 --but it gives an error like this. | |
56 -- | |
57 -- .sncommute i j f != sncommute t i j f of type | |
58 -- .SMap f (snmap t i) ≡ snmap t j | |
59 -- | |
60 --Is there any help? |