Mercurial > hg > Members > kono > Proof > category
annotate SetsCompleteness.agda @ 547:c0078b03201c
on going ...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Tue, 04 Apr 2017 13:51:48 +0900 |
parents | 73a5606fa362 |
children | 03b851adc4fb 97b98b81e990 |
rev | line source |
---|---|
500
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
1 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
2 open import Category -- https://github.com/konn/category-agda |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
3 open import Level |
535 | 4 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
|
5 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
6 module SetsCompleteness where |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
7 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
8 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
9 open import cat-utility |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
10 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
|
11 open import Function |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
12 import Relation.Binary.PropositionalEquality |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
13 -- 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
|
14 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
|
15 |
520 | 16 ≡cong = Relation.Binary.PropositionalEquality.cong |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
17 |
524 | 18 lemma1 : { c₂ : Level } {a b : Obj (Sets { c₂})} {f g : Hom Sets a b} → |
19 Sets [ f ≈ g ] → (x : a ) → f x ≡ g x | |
20 lemma1 refl x = refl | |
503 | 21 |
504 | 22 record Σ {a} (A : Set a) (B : Set a) : Set a where |
503 | 23 constructor _,_ |
24 field | |
25 proj₁ : A | |
504 | 26 proj₂ : B |
503 | 27 |
28 open Σ public | |
500
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
29 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
30 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
31 SetsProduct : { c₂ : Level} → CreateProduct ( Sets { c₂} ) |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
32 SetsProduct { c₂ } = record { |
504 | 33 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
|
34 ; π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
|
35 ; π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
|
36 ; isProduct = λ a b → record { |
503 | 37 _×_ = λ 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
|
38 ; π1fxg=f = refl |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
39 ; π2fxg=g = refl |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
40 ; uniqueness = refl |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
41 ; ×-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
|
42 } |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
43 } where |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
44 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
|
45 → 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
|
46 → 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
|
47 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
|
48 |
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
49 |
508
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
50 record iproduct {a} (I : Set a) ( pi0 : I → Set a ) : Set a where |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
51 field |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
52 pi1 : ( i : I ) → pi0 i |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
53 |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
54 open iproduct |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
55 |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
56 SetsIProduct : { c₂ : Level} → (I : Obj Sets) (ai : I → Obj Sets ) |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
57 → IProduct ( Sets { c₂} ) I |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
58 SetsIProduct I fi = record { |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
59 ai = fi |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
60 ; iprod = iproduct I fi |
3ce21b2a671a
IProduct is written in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
507
diff
changeset
|
61 ; pi = λ i prod → pi1 prod i |
509 | 62 ; isIProduct = record { |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
63 iproduct = iproduct1 |
509 | 64 ; pif=q = pif=q |
65 ; ip-uniqueness = ip-uniqueness | |
66 ; ip-cong = ip-cong | |
67 } | |
68 } where | |
69 iproduct1 : {q : Obj Sets} → ((i : I) → Hom Sets q (fi i)) → Hom Sets q (iproduct I fi) | |
70 iproduct1 {q} qi x = record { pi1 = λ i → (qi i) x } | |
71 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 ] | |
72 pif=q {q} qi {i} = refl | |
73 ip-uniqueness : {q : Obj Sets} {h : Hom Sets q (iproduct I fi)} → Sets [ iproduct1 (λ i → Sets [ (λ prod → pi1 prod i) o h ]) ≈ h ] | |
74 ip-uniqueness = refl | |
75 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 | |
76 ipcx {q} {qi} {qi'} qi=qi x = | |
77 begin | |
78 record { pi1 = λ i → (qi i) x } | |
520 | 79 ≡⟨ ≡cong ( λ QIX → record { pi1 = QIX } ) ( extensionality Sets (λ i → ≡cong ( λ f → f x ) (qi=qi i) )) ⟩ |
509 | 80 record { pi1 = λ i → (qi' i) x } |
81 ∎ where | |
82 open import Relation.Binary.PropositionalEquality | |
83 open ≡-Reasoning | |
84 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' ] | |
85 ip-cong {q} {qi} {qi'} qi=qi = extensionality Sets ( ipcx qi=qi ) | |
86 | |
87 | |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
88 -- |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
89 -- e f |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
90 -- c -------→ a ---------→ b f ( f' |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
91 -- ^ . ---------→ |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
92 -- | . g |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
93 -- |k . |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
94 -- | . h |
514 | 95 --y : d |
509 | 96 |
522
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
97 -- 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
|
98 |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
99 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
|
100 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
|
101 |
532
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
102 equ : { c₂ : Level} {a b : Obj (Sets {c₂}) } { f g : Hom (Sets {c₂}) a b } → ( sequ a b f g ) → a |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
103 equ (elem x eq) = x |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
104 |
533 | 105 fe=ge0 : { c₂ : Level} {a b : Obj (Sets {c₂}) } { f g : Hom (Sets {c₂}) a b } → |
106 (x : sequ a b f g) → (Sets [ f o (λ e → equ e) ]) x ≡ (Sets [ g o (λ e → equ e) ]) x | |
107 fe=ge0 (elem x eq ) = eq | |
108 | |
541 | 109 irr : { c₂ : Level} {d : Set c₂ } { x y : d } ( eq eq' : x ≡ y ) → eq ≡ eq' |
110 irr refl refl = refl | |
111 | |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
112 open sequ |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
113 |
532
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
114 -- equalizer-c = sequ a b f g |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
115 -- ; equalizer = λ e → equ e |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
116 |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
117 SetsIsEqualizer : { c₂ : Level} → (a b : Obj (Sets {c₂}) ) (f g : Hom (Sets {c₂}) a b) → IsEqualizer Sets (λ e → equ e )f g |
d5d7163f2a1d
equalizer does not fit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
531
diff
changeset
|
118 SetsIsEqualizer {c₂} a b f g = record { |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
119 fe=ge = fe=ge |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
120 ; k = k |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
121 ; ek=h = λ {d} {h} {eq} → ek=h {d} {h} {eq} |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
122 ; uniqueness = uniqueness |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
123 } where |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
124 fe=ge : Sets [ Sets [ f o (λ e → equ e ) ] ≈ Sets [ g o (λ e → equ e ) ] ] |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
125 fe=ge = extensionality Sets (fe=ge0 ) |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
126 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) |
520 | 127 k {d} h eq = λ x → elem (h x) ( ≡cong ( λ y → y x ) eq ) |
510
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
128 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 ] |
5eb4b69bf541
equalizer in Sets , uniquness remains
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
509
diff
changeset
|
129 ek=h {d} {h} {eq} = refl |
523 | 130 injection : { c₂ : Level } {a b : Obj (Sets { c₂})} (f : Hom Sets a b) → Set c₂ |
131 injection f = ∀ x y → f x ≡ f y → x ≡ y | |
522
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
132 elm-cong : (x y : sequ a b f g) → equ x ≡ equ y → x ≡ y |
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
133 elm-cong ( elem x eq ) (elem .x eq' ) refl = ≡cong ( λ ee → elem x ee ) ( irr eq eq' ) |
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
134 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)} → |
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
135 Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → (x : d ) → equ (k h fh=gh x) ≡ equ (k' x) |
8fd030f9f572
Equalizer in Sets done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
521
diff
changeset
|
136 lemma5 refl x = refl -- somehow this is not equal to lemma1 |
512 | 137 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)} → |
138 Sets [ Sets [ (λ e → equ e) o k' ] ≈ h ] → Sets [ k h fh=gh ≈ k' ] | |
525 | 139 uniqueness {d} {h} {fh=gh} {k'} ek'=h = extensionality Sets ( λ ( x : d ) → begin |
140 k h fh=gh x | |
141 ≡⟨ elm-cong ( k h fh=gh x) ( k' x ) (lemma5 {d} {h} {fh=gh} {k'} ek'=h x ) ⟩ | |
142 k' x | |
143 ∎ ) where | |
144 open import Relation.Binary.PropositionalEquality | |
145 open ≡-Reasoning | |
146 | |
500
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
147 |
501
61daa68a70c4
Sets completeness failed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
500
diff
changeset
|
148 open Functor |
500
6c993c1fe9de
try to make prodcut and equalizer in Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff
changeset
|
149 |
538 | 150 ---- |
151 -- C is locally small i.e. Hom C i j is a set c₁ | |
152 -- | |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
153 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
|
154 : Set (suc (c₁ ⊔ c₂ ⊔ ℓ )) where |
507 | 155 field |
540 | 156 hom→ : {i j : Obj C } → Hom C i j → I |
157 hom← : {i j : Obj C } → ( f : I ) → Hom C i j | |
158 hom-iso : {i j : Obj C } → { f : Hom C i j } → hom← ( hom→ f ) ≡ f | |
536 | 159 -- ≈-≡ : {a b : Obj C } { x y : Hom C a b } → (x≈y : C [ x ≈ y ] ) → x ≡ y |
507 | 160 |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
161 open Small |
507 | 162 |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
163 ΓObj : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} )) |
538 | 164 (i : Obj C ) → Set c₁ |
165 ΓObj s Γ i = FObj Γ i | |
507 | 166 |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
167 ΓMap : { c₁ c₂ ℓ : Level} { C : Category c₁ c₂ ℓ } { I : Set c₁ } ( s : Small C I ) ( Γ : Functor C ( Sets { c₁} )) |
539 | 168 {i j : Obj C } → ( f : I ) → ΓObj s Γ i → ΓObj s Γ j |
540 | 169 ΓMap s Γ {i} {j} f = FMap Γ ( hom← s f ) |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
170 |
538 | 171 record snat { c₂ } { I OC : Set c₂ } ( sobj : OC → Set c₂ ) |
539 | 172 ( smap : { i j : OC } → (f : I )→ sobj i → sobj j ) : Set c₂ where |
527 | 173 field |
538 | 174 snmap : ( i : OC ) → sobj i |
539 | 175 sncommute : { i j : OC } → ( f : I ) → smap f ( snmap i ) ≡ snmap j |
507 | 176 |
534 | 177 open snat |
501
61daa68a70c4
Sets completeness failed
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
500
diff
changeset
|
178 |
546 | 179 snat1 : { c₂ : Level } { I OC : Set c₂ } ( sobj : OC → Set c₂ ) ( smap : { i j : OC } → (f : I )→ sobj i → sobj j ) |
180 ( snm : ( i : OC ) → sobj i ) → (( snm1 : ( i : OC ) → sobj i ) ( i j : OC ) → (f : I ) → smap f ( snm1 i ) ≡ snm1 j ) → snat sobj smap | |
181 snat1 SO SM snm eq = record { snmap = λ i → snm i ; sncommute = λ {i} {j} f → eq snm i j f } | |
182 | |
541 | 183 ≡cong2 : { c : Level } { A B C : Set c } { a a' : A } { b b' : B } ( f : A → B → C ) |
184 → a ≡ a' | |
185 → b ≡ b' | |
186 → f a b ≡ f a' b' | |
187 ≡cong2 _ refl refl = refl | |
188 | |
543 | 189 snat-cong : { c : Level } { I OC : Set c } ( SObj : OC → Set c ) ( SMap : { i j : OC } → (f : I )→ SObj i → SObj j) |
190 { s t : snat SObj SMap } | |
191 → (( i : OC ) → snmap s i ≡ snmap t i ) | |
192 → ( ( i j : OC ) ( f : I ) → SMap {i} {j} f ( snmap s i ) ≡ snmap s j ) | |
193 → ( ( i j : OC ) ( f : I ) → SMap {i} {j} f ( snmap t i ) ≡ snmap t j ) | |
194 → s ≡ t | |
546 | 195 snat-cong {_} {I} {OC} SO SM {s} {t} eq1 eq2 eq3 = ≡cong2 ( λ x y → snat1 SO SM x y ) |
543 | 196 ( extensionality Sets ( λ i → eq1 i ) ) |
546 | 197 ( extensionality Sets ( λ snm → |
198 ( extensionality Sets ( λ i → | |
199 ( extensionality Sets ( λ j → | |
547 | 200 ( extensionality Sets ( λ f → irr2 snm i j f ( eq2s i j f (eq1 i) snm (eq2 i j f ) ) {!!} )))))))) |
543 | 201 where |
547 | 202 eq2s1 : { i j : OC } { f : I } → {si : SO i} { sj : SO j } |
203 ( snm : ( i : OC ) → SO i ) → ( si ≡ snm i ) → ( sj ≡ snm j ) → SM {i} {j} f si ≡ sj → SM f ( snm i) ≡ snm j | |
204 eq2s1 {i} {j} {f} {si} {sj} snm eq1 eq2 eq3 = begin | |
205 SM f ( snm i) | |
206 ≡⟨ ≡cong (λ x → SM f x ) (sym eq1) ⟩ | |
207 SM f si | |
208 ≡⟨ eq3 ⟩ | |
209 sj | |
210 ≡⟨ eq2 ⟩ | |
211 snm j | |
212 ∎ where | |
213 open import Relation.Binary.PropositionalEquality | |
214 open ≡-Reasoning | |
215 eq2s : ( i j : OC ) ( f : I ) → ( snmap s i ≡ snmap t i ) → | |
216 ( snm : ( i : OC ) → SO i ) → SM {i} {j} f ( snmap s i ) ≡ snmap s j → SM f ( snm i) ≡ snm j | |
217 eq2s i j f eq1 snm eq2 = eq2s1 snm {!!} {!!} eq2 | |
218 irr3 : { i j : OC } { f : I } → {snmapsi : SO i } → {snmapsj : SO j } → | |
219 ( es et : SM f ( snmapsi ) ≡ snmapsj ) → es ≡ et | |
220 irr3 refl refl = refl | |
546 | 221 irr2 : ( snm : ( i : OC ) → SO i ) ( i j : OC ) (f : I ) → |
222 ( es et : SM f (snm i ) ≡ snm j ) → es ≡ et | |
547 | 223 irr2 snm i j f es et = irr3 es et |
541 | 224 |
543 | 225 |
530 | 226 open import HomReasoning |
227 open NTrans | |
533 | 228 |
535 | 229 Cone : { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) ( s : Small C I ) ( Γ : Functor C (Sets {c₁} ) ) |
534 | 230 → NTrans C Sets (K Sets C (snat (ΓObj s Γ) (ΓMap s Γ) ) ) Γ |
535 | 231 Cone C I s Γ = record { |
538 | 232 TMap = λ i → λ sn → snmap sn i |
531 | 233 ; isNTrans = record { commute = comm1 } |
530 | 234 } where |
534 | 235 comm1 : {a b : Obj C} {f : Hom C a b} → |
538 | 236 Sets [ Sets [ FMap Γ f o (λ sn → snmap sn a) ] ≈ |
237 Sets [ (λ sn → (snmap sn b)) o FMap (K Sets C (snat (ΓObj s Γ) (ΓMap s Γ))) f ] ] | |
534 | 238 comm1 {a} {b} {f} = extensionality Sets ( λ sn → begin |
538 | 239 FMap Γ f (snmap sn a ) |
540 | 240 ≡⟨ ≡cong ( λ f → ( FMap Γ f (snmap sn a ))) (sym ( hom-iso s )) ⟩ |
241 FMap Γ ( hom← s ( hom→ s f)) (snmap sn a ) | |
535 | 242 ≡⟨⟩ |
540 | 243 ΓMap s Γ (hom→ s f) (snmap sn a ) |
244 ≡⟨ sncommute sn (hom→ s f) ⟩ | |
538 | 245 snmap sn b |
534 | 246 ∎ ) where |
247 open import Relation.Binary.PropositionalEquality | |
248 open ≡-Reasoning | |
249 | |
530 | 250 |
526
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
251 SetsLimit : { c₁ c₂ ℓ : Level} ( C : Category c₁ c₂ ℓ ) ( I : Set c₁ ) ( small : Small C I ) ( Γ : Functor C (Sets {c₁} ) ) |
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
252 → Limit Sets C Γ |
b035cd3be525
Small Category for Sets Limit
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
525
diff
changeset
|
253 SetsLimit { c₂} C I s Γ = record { |
534 | 254 a0 = snat (ΓObj s Γ) (ΓMap s Γ) |
535 | 255 ; t0 = Cone C I s Γ |
523 | 256 ; isLimit = record { |
530 | 257 limit = limit1 |
539 | 258 ; t0f=t = λ {a t i } → t0f=t {a} {t} {i} |
259 ; limit-uniqueness = λ {a t i } → limit-uniqueness {a} {t} {i} | |
260 } | |
523 | 261 } where |
527 | 262 a0 : Obj Sets |
534 | 263 a0 = snat (ΓObj s Γ) (ΓMap s Γ) |
539 | 264 comm2 : { a : Obj Sets } {x : a } {i j : Obj C} (t : NTrans C Sets (K Sets C a) Γ) (f : I) |
538 | 265 → ΓMap s Γ f (TMap t i x) ≡ TMap t j x |
266 comm2 {a} {x} t f = ≡cong ( λ f → f x ) ( IsNTrans.commute ( isNTrans t ) ) | |
534 | 267 limit1 : (a : Obj Sets) → NTrans C Sets (K Sets C a) Γ → Hom Sets a (snat (ΓObj s Γ) (ΓMap s Γ)) |
538 | 268 limit1 a t = λ x → record { snmap = λ i → ( TMap t i ) x ; |
537 | 269 sncommute = λ f → comm2 t f } |
539 | 270 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 ] |
540 | 271 t0f=t {a} {t} {i} = extensionality Sets ( λ x → begin |
272 ( Sets [ TMap (Cone C I s Γ) i o limit1 a t ]) x | |
273 -- ≡⟨⟩ | |
274 -- snmap ( record { snmap = λ i → ( TMap t i ) x ; sncommute = λ {i j} f → comm2 {a} {x} {i} {j} t f } ) i | |
275 ≡⟨⟩ | |
276 TMap t i x | |
277 ∎ ) where | |
278 open import Relation.Binary.PropositionalEquality | |
279 open ≡-Reasoning | |
539 | 280 limit-uniqueness : {a : Obj Sets} {t : NTrans C Sets (K Sets C a) Γ} {f : Hom Sets a (snat (ΓObj s Γ) (ΓMap s Γ))} → |
281 ({i : Obj C} → Sets [ Sets [ TMap (Cone C I s Γ) i o f ] ≈ TMap t i ]) → Sets [ limit1 a t ≈ f ] | |
540 | 282 limit-uniqueness {a} {t} {f} cif=t = extensionality Sets ( λ x → begin |
283 limit1 a t x | |
284 ≡⟨⟩ | |
285 record { snmap = λ i → ( TMap t i ) x ; sncommute = λ f → comm2 t f } | |
544 | 286 ≡⟨ snat-cong (ΓObj s Γ) (ΓMap s Γ) {!!} {!!} {!!} ⟩ |
541 | 287 record { snmap = λ i → snmap (f x ) i ; sncommute = sncommute (f x ) } |
540 | 288 ≡⟨⟩ |
289 f x | |
290 ∎ ) where | |
291 open import Relation.Binary.PropositionalEquality | |
292 open ≡-Reasoning | |
537 | 293 |
294 | |
295 | |
296 | |
297 | |
298 | |
299 | |
539 | 300 |
301 |