annotate SetsCompleteness.agda @ 605:af321e38ecee

another snat-cong approach
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Thu, 08 Jun 2017 12:53:54 +0900
parents 75112154faf0
children 92eb707498c7
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
1 open import Category -- https://github.com/konn/category-agda
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
2 open import Level
535
5d7f8921bac0 on going ....
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 534
diff changeset
3 open import Category.Sets renaming ( _o_ to _*_ )
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
4
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
5 module SetsCompleteness where
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
6
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
7 open import cat-utility
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
8 open import Relation.Binary.Core
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9 open import Function
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
10 import Relation.Binary.PropositionalEquality
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
11 -- Extensionality a b = {A : Set a} {B : A → Set b} {f g : (x : A) → B x} → (∀ x → f x ≡ g x) → f ≡ g → ( λ x → f x ≡ λ x → g x )
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
12 postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Relation.Binary.PropositionalEquality.Extensionality c₂ c₂
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
13
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
14 ≡cong = Relation.Binary.PropositionalEquality.cong
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
15
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
16 lemma1 : { c₂ : Level } {a b : Obj (Sets { c₂})} {f g : Hom Sets a b} →
524
d6739779b4ac on going ..
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 523
diff changeset
17 Sets [ f ≈ g ] → (x : a ) → f x ≡ g x
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
18 lemma1 refl x = refl
503
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
19
504
b489f27317fb on going...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
20 record Σ {a} (A : Set a) (B : Set a) : Set a where
503
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
21 constructor _,_
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
22 field
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
23 proj₁ : A
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
24 proj₂ : B
503
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
25
bd33096c189b on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 501
diff changeset
26 open Σ public
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
27
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
28
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29 SetsProduct : { c₂ : Level} → CreateProduct ( Sets { c₂} )
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
30 SetsProduct { c₂ } = record {
504
b489f27317fb on going...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 503
diff changeset
31 product = λ a b → Σ a b
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 ; π1 = λ a b → λ ab → (proj₁ ab)
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 ; π2 = λ a b → λ ab → (proj₂ ab)
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
34 ; isProduct = λ a b → record {
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
35 _×_ = λ f g x → record { proj₁ = f x ; proj₂ = g x } -- ( f x , g x )
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 ; π1fxg=f = refl
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 ; π2fxg=g = refl
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 ; uniqueness = refl
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 ; ×-cong = λ {c} {f} {f'} {g} {g'} f=f g=g → prod-cong a b f=f g=g
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
40 }
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
41 } where
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
42 prod-cong : ( a b : Obj (Sets {c₂}) ) {c : Obj (Sets {c₂}) } {f f' : Hom Sets c a } {g g' : Hom Sets c b }
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
43 → Sets [ f ≈ f' ] → Sets [ g ≈ g' ]
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
44 → Sets [ (λ x → f x , g x) ≈ (λ x → f' x , g' x) ]
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
45 prod-cong a b {c} {f} {.f} {g} {.g} refl refl = refl
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
46
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
47
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
48 record iproduct {a} (I : Set a) ( pi0 : I → Set a ) : Set a where
508
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
49 field
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
50 pi1 : ( i : I ) → pi0 i
508
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
51
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
52 open iproduct
574
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 573
diff changeset
53
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
54 SetsIProduct : { c₂ : Level} → (I : Obj Sets) (ai : I → Obj Sets )
508
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
55 → IProduct ( Sets { c₂} ) I
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
56 SetsIProduct I fi = record {
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
57 ai = fi
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
58 ; iprod = iproduct I fi
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
59 ; pi = λ i prod → pi1 prod i
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
60 ; isIProduct = record {
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
61 iproduct = iproduct1
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
62 ; pif=q = pif=q
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
63 ; ip-uniqueness = ip-uniqueness
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
64 ; ip-cong = ip-cong
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
65 }
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
66 } where
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
67 iproduct1 : {q : Obj Sets} → ((i : I) → Hom Sets q (fi i)) → Hom Sets q (iproduct I fi)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
68 iproduct1 {q} qi x = record { pi1 = λ i → (qi i) x }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
69 pif=q : {q : Obj Sets} (qi : (i : I) → Hom Sets q (fi i)) {i : I} → Sets [ Sets [ (λ prod → pi1 prod i) o iproduct1 qi ] ≈ qi i ]
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
70 pif=q {q} qi {i} = refl
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
71 ip-uniqueness : {q : Obj Sets} {h : Hom Sets q (iproduct I fi)} → Sets [ iproduct1 (λ i → Sets [ (λ prod → pi1 prod i) o h ]) ≈ h ]
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
72 ip-uniqueness = refl
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
73 ipcx : {q : Obj Sets} {qi qi' : (i : I) → Hom Sets q (fi i)} → ((i : I) → Sets [ qi i ≈ qi' i ]) → (x : q) → iproduct1 qi x ≡ iproduct1 qi' x
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
74 ipcx {q} {qi} {qi'} qi=qi x =
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
75 begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
76 record { pi1 = λ i → (qi i) x }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
77 ≡⟨ ≡cong ( λ QIX → record { pi1 = QIX } ) ( extensionality Sets (λ i → ≡cong ( λ f → f x ) (qi=qi i) )) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
78 record { pi1 = λ i → (qi' i) x }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
79 ∎ where
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
80 open import Relation.Binary.PropositionalEquality
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
81 open ≡-Reasoning
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
82 ip-cong : {q : Obj Sets} {qi qi' : (i : I) → Hom Sets q (fi i)} → ((i : I) → Sets [ qi i ≈ qi' i ]) → Sets [ iproduct1 qi ≈ iproduct1 qi' ]
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
83 ip-cong {q} {qi} {qi'} qi=qi = extensionality Sets ( ipcx qi=qi )
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
84
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
85
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
86 --
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
87 -- e f
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
88 -- c -------→ a ---------→ b f ( f'
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
89 -- ^ . ---------→
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
90 -- | . g
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
91 -- |k .
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
92 -- | . h
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
93 --y : d
509
3e82fb1ce170 IProduct in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 508
diff changeset
94
522
8fd030f9f572 Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
95 -- cf. https://github.com/danr/Agda-projects/blob/master/Category-Theory/Equalizer.agda
508
3ce21b2a671a IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 507
diff changeset
96
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
97 data sequ {c : Level} (A B : Set c) ( f g : A → B ) : Set c where
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
98 elem : (x : A ) → (eq : f x ≡ g x) → sequ A B f g
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
99
532
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
100 equ : { c₂ : Level} {a b : Obj (Sets {c₂}) } { f g : Hom (Sets {c₂}) a b } → ( sequ a b f g ) → a
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
101 equ (elem x eq) = x
532
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
102
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
103 fe=ge0 : { c₂ : Level} {a b : Obj (Sets {c₂}) } { f g : Hom (Sets {c₂}) a b } →
533
c3dcea3a92a7 use sequ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
104 (x : sequ a b f g) → (Sets [ f o (λ e → equ e) ]) x ≡ (Sets [ g o (λ e → equ e) ]) x
c3dcea3a92a7 use sequ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
105 fe=ge0 (elem x eq ) = eq
c3dcea3a92a7 use sequ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 532
diff changeset
106
541
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
107 irr : { c₂ : Level} {d : Set c₂ } { x y : d } ( eq eq' : x ≡ y ) → eq ≡ eq'
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
108 irr refl refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 540
diff changeset
109
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
110 open sequ
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
111
532
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
112 -- equalizer-c = sequ a b f g
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
113 -- ; equalizer = λ e → equ e
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
114
d5d7163f2a1d equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 531
diff changeset
115 SetsIsEqualizer : { c₂ : Level} → (a b : Obj (Sets {c₂}) ) (f g : Hom (Sets {c₂}) a b) → IsEqualizer Sets (λ e → equ e )f g
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
116 SetsIsEqualizer {c₂} a b f g = record {
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
117 fe=ge = fe=ge
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
118 ; k = k
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
119 ; ek=h = λ {d} {h} {eq} → ek=h {d} {h} {eq}
510
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
120 ; uniqueness = uniqueness
5eb4b69bf541 equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 509
diff changeset
121 } where
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
122 fe=ge : Sets [ Sets [ f o (λ e → equ e ) ] ≈ Sets [ g o (λ e → equ e ) ] ]
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
123 fe=ge = extensionality Sets (fe=ge0 )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
124 k : {d : Obj Sets} (h : Hom Sets d a) → Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ] → Hom Sets d (sequ a b f g)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
125 k {d} h eq = λ x → elem (h x) ( ≡cong ( λ y → y x ) eq )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
126 ek=h : {d : Obj Sets} {h : Hom Sets d a} {eq : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} → Sets [ Sets [ (λ e → equ e ) o k h eq ] ≈ h ]
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
127 ek=h {d} {h} {eq} = refl
523
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
128 injection : { c₂ : Level } {a b : Obj (Sets { c₂})} (f : Hom Sets a b) → Set c₂
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 522
diff changeset
129 injection f = ∀ x y → f x ≡ f y → x ≡ y
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
130 elm-cong : (x y : sequ a b f g) → equ x ≡ equ y → x ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
131 elm-cong ( elem x eq ) (elem .x eq' ) refl = ≡cong ( λ ee → elem x ee ) ( irr eq eq' )
522
8fd030f9f572 Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 521
diff changeset
132 lemma5 : {d : Obj Sets} {h : Hom Sets d a} {fh=gh : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} {k' : Hom Sets d (sequ a b f g)} →
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
133 Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → (x : d ) → equ (k h fh=gh x) ≡ equ (k' x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
134 lemma5 refl x = refl -- somehow this is not equal to lemma1
512
f19dab948e39 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
135 uniqueness : {d : Obj Sets} {h : Hom Sets d a} {fh=gh : Sets [ Sets [ f o h ] ≈ Sets [ g o h ] ]} {k' : Hom Sets d (sequ a b f g)} →
f19dab948e39 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 511
diff changeset
136 Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → Sets [ k h fh=gh ≈ k' ]
525
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
137 uniqueness {d} {h} {fh=gh} {k'} ek'=h = extensionality Sets ( λ ( x : d ) → begin
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
138 k h fh=gh x
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
139 ≡⟨ elm-cong ( k h fh=gh x) ( k' x ) (lemma5 {d} {h} {fh=gh} {k'} ek'=h x ) ⟩
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
140 k' x
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
141 ∎ ) where
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
142 open import Relation.Binary.PropositionalEquality
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
143 open ≡-Reasoning
cb35d6b25559 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 524
diff changeset
144
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
145
501
61daa68a70c4 Sets completeness failed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 500
diff changeset
146 open Functor
500
6c993c1fe9de try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
147
538
d22c93dca806 locally small
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
148 ----
d22c93dca806 locally small
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
149 -- C is locally small i.e. Hom C i j is a set c₁
d22c93dca806 locally small
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
150 --
526
b035cd3be525 Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
151 record Small { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ )
b035cd3be525 Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 525
diff changeset
152 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
153 field
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
154 hom→ : {i j : Obj C } → Hom C i j → I
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
155 hom← : {i j : Obj C } → ( f : I ) → Hom C i j
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
156 hom-iso : {i j : Obj C } → { f : Hom C i j } → hom← ( hom→ f ) ≡ f
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
157
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
158 open Small
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
159
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
160 ΓObj : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} ))
538
d22c93dca806 locally small
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
161 (i : Obj C ) →  Set c₁
d22c93dca806 locally small
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 537
diff changeset
162 ΓObj s Γ i = FObj Γ i
507
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 506
diff changeset
163
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
164 ΓMap : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} ))
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
165 {i j : Obj C } →  ( f : I ) → ΓObj s Γ i → ΓObj s Γ j
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
166 ΓMap s Γ {i} {j} f = FMap Γ ( hom← s f )
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
167
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
168 record snat { c₂ } { I OC : Set c₂ } ( sobj : OC → Set c₂ )
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
169 ( smap : { i j : OC } → (f : I ) → sobj i → sobj j ) : Set c₂ where
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
170 field
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
171 snmap : ( i : OC ) → sobj i
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
172 sncommute : { i j : OC } → ( f : I ) → smap f ( snmap i ) ≡ snmap j
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
173 smap0 : { i j : OC } → (f : I ) → sobj i → sobj j
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
174 smap0 {i} {j} f x = smap f x
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
175
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
176 open snat
600
3e2ef72d8d2f Set Completeness unfinished
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 599
diff changeset
177
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
178 ≡cong2 : { c c' : Level } { A B : Set c } { C : Set c' } { a a' : A } { b b' : B } ( f : A → B → C )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
179 → a ≡ a'
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
180 → b ≡ b'
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
181 → f a b ≡ f a' b'
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
182 ≡cong2 _ refl refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
183
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
184 snat-cong : { c : Level } { I OC : Set c } { SObj : OC → Set c } { SMap : { i j : OC } → (f : I )→ SObj i → SObj j }
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
185 ( s t : snat SObj SMap )
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
186 → ( ( λ i → snmap s i ) ≡ ( λ i → snmap t i) )
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
187 → ( ( λ {i} {j} f → smap0 s f ( snmap s i ) ≡ snmap s j ) ≡ ( λ {i} {j} f → smap0 t f ( snmap t i ) ≡ snmap t j ) )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
188 → s ≡ t
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
189 snat-cong {_} {I} {OC} {SO} {SM} s t eq1 eq2 = {!!}
590
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
190
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
191 open import HomReasoning
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
192 open NTrans
590
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
193
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
194 Cone : { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) ( s : Small C I ) ( Γ : Functor C (Sets {c₁} ) )
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
195 → NTrans C Sets (K Sets C (snat (ΓObj s Γ) (ΓMap s Γ) ) ) Γ
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
196 Cone C I s Γ = record {
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
197 TMap = λ i → λ sn → snmap sn i
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
198 ; isNTrans = record { commute = comm1 }
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
199 } where
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
200 comm1 : {a b : Obj C} {f : Hom C a b} →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
201 Sets [ Sets [ FMap Γ f o (λ sn → snmap sn a) ] ≈
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
202 Sets [ (λ sn → (snmap sn b)) o FMap (K Sets C (snat (ΓObj s Γ) (ΓMap s Γ))) f ] ]
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
203 comm1 {a} {b} {f} = extensionality Sets ( λ sn → begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
204 FMap Γ f (snmap sn a )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
205 ≡⟨ ≡cong ( λ f → ( FMap Γ f (snmap sn a ))) (sym ( hom-iso s )) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
206 FMap Γ ( hom← s ( hom→ s f)) (snmap sn a )
596
9367813d3f61 lemma-equ retry
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 591
diff changeset
207 ≡⟨⟩
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
208 ΓMap s Γ (hom→ s f) (snmap sn a )
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
209 ≡⟨ sncommute sn (hom→ s f) ⟩
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
210 snmap sn b
596
9367813d3f61 lemma-equ retry
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 591
diff changeset
211 ∎ ) where
590
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
212 open import Relation.Binary.PropositionalEquality
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
213 open ≡-Reasoning
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
214
2c5d8c3c9086 on going ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 589
diff changeset
215
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
216 SetsLimit : { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) ( small : Small C I ) ( Γ : Functor C (Sets {c₁} ) )
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
217 → Limit Sets C Γ
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
218 SetsLimit { c₂} C I s Γ = record {
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
219 a0 = snat (ΓObj s Γ) (ΓMap s Γ)
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
220 ; t0 = Cone C I s Γ
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
221 ; isLimit = record {
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
222 limit = limit1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
223 ; t0f=t = λ {a t i } → t0f=t {a} {t} {i}
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
224 ; limit-uniqueness = λ {a t i } → limit-uniqueness {a} {t} {i}
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
225 }
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
226 } where
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
227 comm2 : { a : Obj Sets } {x : a } {i j : Obj C} (t : NTrans C Sets (K Sets C a) Γ) (f : I)
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
228 → ΓMap s Γ f (TMap t i x) ≡ TMap t j x
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
229 comm2 {a} {x} t f = ≡cong ( λ f → f x ) ( IsNTrans.commute ( isNTrans t ) )
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
230 limit1 : (a : Obj Sets) → NTrans C Sets (K Sets C a) Γ → Hom Sets a (snat (ΓObj s Γ) (ΓMap s Γ))
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
231 limit1 a t = λ x → record { snmap = λ i → ( TMap t i ) x ;
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
232 sncommute = λ f → comm2 t f }
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
233 t0f=t : {a : Obj Sets} {t : NTrans C Sets (K Sets C a) Γ} {i : Obj C} → Sets [ Sets [ TMap (Cone C I s Γ) i o limit1 a t ] ≈ TMap t i ]
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
234 t0f=t {a} {t} {i} = extensionality Sets ( λ x → begin
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
235 ( Sets [ TMap (Cone C I s Γ) i o limit1 a t ]) x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
236 ≡⟨⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
237 TMap t i x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
238 ∎ ) where
562
14483d9d604c dead end again ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
239 open import Relation.Binary.PropositionalEquality
14483d9d604c dead end again ...
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 561
diff changeset
240 open ≡-Reasoning
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
241 limit-uniqueness : {a : Obj Sets} {t : NTrans C Sets (K Sets C a) Γ} {f : Hom Sets a (snat (ΓObj s Γ) (ΓMap s Γ))} →
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
242 ({i : Obj C} → Sets [ Sets [ TMap (Cone C I s Γ) i o f ] ≈ TMap t i ]) → Sets [ limit1 a t ≈ f ]
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
243 limit-uniqueness {a} {t} {f} cif=t = extensionality Sets ( λ x → begin
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
244 limit1 a t x
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
245 ≡⟨⟩
605
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
246 record { snmap = λ i → ( TMap t i ) x ; sncommute = λ f → comm2 t f }
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
247 ≡⟨ snat-cong (limit1 a t x) (f x ) ( extensionality Sets ( λ i → eq1 x i ))
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
248 (
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
249 {!!}
af321e38ecee another snat-cong approach
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 604
diff changeset
250 )⟩
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
251 record { snmap = λ i → snmap (f x ) i ; sncommute = sncommute (f x ) }
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
252 ≡⟨⟩
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
253 f x
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
254 ∎ ) where
598
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
255 open import Relation.Binary.PropositionalEquality
2e3459a9a69f try two field again
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 597
diff changeset
256 open ≡-Reasoning
604
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
257 eq1 : (x : a ) (i : Obj C) → TMap t i x ≡ snmap (f x) i
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 603
diff changeset
258 eq1 x i = sym ( ≡cong ( λ f → f x ) cif=t )