annotate src/CCCSets.agda @ 1027:5ae0290c34b4

...
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 29 Mar 2021 21:57:12 +0900
parents 7916bde5e57b
children 28569574e3cf
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
1 {-# OPTIONS --allow-unsolved-metas #-}
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
2 module CCCSets where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
4 open import Level
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
5 open import Category
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
6 open import HomReasoning
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
7 open import cat-utility
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
8 open import Data.Product renaming (_×_ to _/\_ ) hiding ( <_,_> )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9 open import Category.Constructions.Product
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
10 open import Relation.Binary.PropositionalEquality hiding ( [_] )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
11 open import CCC
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
12
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
13 open Functor
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
14
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
15 -- ccc-1 : Hom A a 1 ≅ {*}
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
16 -- ccc-2 : Hom A c (a × b) ≅ (Hom A c a ) × ( Hom A c b )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
17 -- ccc-3 : Hom A a (c ^ b) ≅ Hom A (a × b) c
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
18
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
19 open import Category.Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
20
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
21 -- Sets is a CCC
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
22
1018
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
23 open import SetsCompleteness hiding (ki1)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
24
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
25 -- import Axiom.Extensionality.Propositional
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
26 -- postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Axiom.Extensionality.Propositional.Extensionality c₂ c₂
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
27
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
28 data One {c : Level } : Set c where
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
29 ! : One -- () in Haskell ( or any one object set )
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 sets : {c : Level } → CCC (Sets {c})
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 sets = record {
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 1 = One
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
34 ; ○ = λ _ → λ _ → !
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
35 ; _∧_ = _∧_
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 ; <_,_> = <,>
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 ; π = π
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 ; π' = π'
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 ; _<=_ = _<=_
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
40 ; _* = _*
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
41 ; ε = ε
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
42 ; isCCC = isCCC
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
43 } where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
44 1 : Obj Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
45 1 = One
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
46 ○ : (a : Obj Sets ) → Hom Sets a 1
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
47 ○ a = λ _ → !
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
48 _∧_ : Obj Sets → Obj Sets → Obj Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
49 _∧_ a b = a /\ b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
50 <,> : {a b c : Obj Sets } → Hom Sets c a → Hom Sets c b → Hom Sets c ( a ∧ b)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
51 <,> f g = λ x → ( f x , g x )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
52 π : {a b : Obj Sets } → Hom Sets (a ∧ b) a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
53 π {a} {b} = proj₁
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
54 π' : {a b : Obj Sets } → Hom Sets (a ∧ b) b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
55 π' {a} {b} = proj₂
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
56 _<=_ : (a b : Obj Sets ) → Obj Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
57 a <= b = b → a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
58 _* : {a b c : Obj Sets } → Hom Sets (a ∧ b) c → Hom Sets a (c <= b)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
59 f * = λ x → λ y → f ( x , y )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
60 ε : {a b : Obj Sets } → Hom Sets ((a <= b ) ∧ b) a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
61 ε {a} {b} = λ x → ( proj₁ x ) ( proj₂ x )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
62 isCCC : CCC.IsCCC Sets 1 ○ _∧_ <,> π π' _<=_ _* ε
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
63 isCCC = record {
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
64 e2 = e2
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
65 ; e3a = λ {a} {b} {c} {f} {g} → e3a {a} {b} {c} {f} {g}
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
66 ; e3b = λ {a} {b} {c} {f} {g} → e3b {a} {b} {c} {f} {g}
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
67 ; e3c = e3c
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
68 ; π-cong = π-cong
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
69 ; e4a = e4a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
70 ; e4b = e4b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
71 ; *-cong = *-cong
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
72 } where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
73 e2 : {a : Obj Sets} {f : Hom Sets a 1} → Sets [ f ≈ ○ a ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
74 e2 {a} {f} = extensionality Sets ( λ x → e20 x )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
75 where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
76 e20 : (x : a ) → f x ≡ ○ a x
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
77 e20 x with f x
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
78 e20 x | ! = refl
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
79 e3a : {a b c : Obj Sets} {f : Hom Sets c a} {g : Hom Sets c b} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
80 Sets [ ( Sets [ π o ( <,> f g) ] ) ≈ f ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
81 e3a = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
82 e3b : {a b c : Obj Sets} {f : Hom Sets c a} {g : Hom Sets c b} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
83 Sets [ Sets [ π' o ( <,> f g ) ] ≈ g ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
84 e3b = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
85 e3c : {a b c : Obj Sets} {h : Hom Sets c (a ∧ b)} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
86 Sets [ <,> (Sets [ π o h ]) (Sets [ π' o h ]) ≈ h ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
87 e3c = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
88 π-cong : {a b c : Obj Sets} {f f' : Hom Sets c a} {g g' : Hom Sets c b} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
89 Sets [ f ≈ f' ] → Sets [ g ≈ g' ] → Sets [ <,> f g ≈ <,> f' g' ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
90 π-cong refl refl = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
91 e4a : {a b c : Obj Sets} {h : Hom Sets (c ∧ b) a} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
92 Sets [ Sets [ ε o <,> (Sets [ h * o π ]) π' ] ≈ h ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
93 e4a = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
94 e4b : {a b c : Obj Sets} {k : Hom Sets c (a <= b)} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
95 Sets [ (Sets [ ε o <,> (Sets [ k o π ]) π' ]) * ≈ k ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
96 e4b = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
97 *-cong : {a b c : Obj Sets} {f f' : Hom Sets (a ∧ b) c} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
98 Sets [ f ≈ f' ] → Sets [ f * ≈ f' * ]
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
99 *-cong refl = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
100
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
101 open import Relation.Nullary
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
102 open import Data.Empty
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
103 open import equalizer
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
104
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
105 data Bool { c : Level } : Set c where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
106 true : Bool
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
107 false : Bool
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
108
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
109 ¬f≡t : { c : Level } → ¬ (false {c} ≡ true )
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
110 ¬f≡t ()
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
111
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
112 ¬x≡t∧x≡f : { c : Level } → {x : Bool {c}} → ¬ ((x ≡ false) /\ (x ≡ true) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
113 ¬x≡t∧x≡f {_} {true} p = ⊥-elim (¬f≡t (sym (proj₁ p)))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
114 ¬x≡t∧x≡f {_} {false} p = ⊥-elim (¬f≡t (proj₂ p))
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
115
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
116 data _∨_ {c c' : Level } (a : Set c) (b : Set c') : Set (c ⊔ c') where
1021
8a78ddfdaa24 ... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1020
diff changeset
117 case1 : a → a ∨ b
8a78ddfdaa24 ... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1020
diff changeset
118 case2 : b → a ∨ b
8a78ddfdaa24 ... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1020
diff changeset
119
1024
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
120 --
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
121 -- m : b → a determins a subset of a as an image
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
122 -- b and m-image in a has one to one correspondence with an equalizer (x : b) → (y : a) ≡ m x.
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
123 -- so tchar m mono and ker (tchar m mono) is Iso.
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
124 -- Finding (x : b) from (y : a) is non constructive. Assuming LEM of image.
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
125 --
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
126 data image {c : Level} {a b : Set c} (m : Hom Sets b a) : a → Set c where
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
127 isImage : (x : b ) → image m (m x)
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
128
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
129 topos : {c : Level } → ({ c : Level} → (b : Set c) → b ∨ (¬ b)) → Topos (Sets {c}) sets
1021
8a78ddfdaa24 ... use LEM for Topos Sets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1020
diff changeset
130 topos {c} lem = record {
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
131 Ω = Bool
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
132 ; ⊤ = λ _ → true
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
133 ; Ker = tker
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
134 ; char = λ m mono → tchar m mono
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
135 ; isTopos = record {
1018
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
136 char-uniqueness = λ {a} {b} {h} m mono → extensionality Sets ( λ x → uniq h m mono x )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
137 ; ker-m = imequ
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
138 }
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
139 } where
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
140 --
1024
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
141 -- In Sets, equalizers exist as
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
142 -- data sequ {c : Level} (A B : Set c) ( f g : A → B ) : Set c where
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
143 -- elem : (x : A ) → (eq : f x ≡ g x) → sequ A B f g
1024
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
144 -- m have to be isomorphic to ker (char m).
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
145 --
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
146 -- i ○ b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
147 -- ker (char m) ----→ b -----------→ 1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
148 -- | ←---- | |
1024
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
149 -- | j |m | ⊤ char m : a → Ω = {true,false}
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
150 -- | e ↓ char m ↓ if y : a ≡ m (∃ x : b) → true ( data char )
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
151 -- +-------------→ a -----------→ Ω else false
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
152 -- h
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
153 --
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
154 tker : {a : Obj Sets} (h : Hom Sets a Bool) → Equalizer Sets h (Sets [ (λ _ → true ) o CCC.○ sets a ])
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
155 tker {a} h = record {
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
156 equalizer-c = sequ a Bool h (λ _ → true )
1018
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
157 ; equalizer = λ e → equ e
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1016
diff changeset
158 ; isEqualizer = SetsIsEqualizer _ _ _ _ }
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
159 tchar : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) → a → Bool {c}
1024
447bbacacf64 fix comment
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1023
diff changeset
160 tchar {a} {b} m mono y with lem (image m y )
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
161 ... | case1 t = true
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
162 ... | case2 f = false
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
163 tcharImg : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) → (y : a) → tchar m mono y ≡ true → image m y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
164 tcharImg m mono y eq with lem (image m y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
165 ... | case1 t = t
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
166 tchar¬Img : {a b : Obj Sets} (m : Hom Sets b a) → (mono : Mono Sets m ) (y : a) → tchar m mono y ≡ false → ¬ image m y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
167 tchar¬Img m mono y eq im with lem (image m y)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
168 ... | case2 n = n im
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
169 img-x : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → {y : a} → image m y → b
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
170 img-x m {.(m x)} (isImage x) = x
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
171 m-img-x : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → {y : a} → (t : image m y ) → m (img-x m t) ≡ y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
172 m-img-x m (isImage x) = refl
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
173 open import Relation.Binary.HeterogeneousEquality as HE using (_≅_ )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
174 img-cong : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y y' : a) → y ≡ y' → (s : image m y ) (t : image m y') → s ≅ t
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
175 img-cong {a} {b} m mono .(m x) .(m x₁) eq (isImage x) (isImage x₁)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
176 with cong (λ k → k ! ) ( Mono.isMono mono {One} (λ _ → x) (λ _ → x₁ ) ( extensionality Sets ( λ _ → eq )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
177 ... | refl = HE.refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
178 img-x-cong : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y y' : a) → y ≡ y' → (s : image m y ) →( t : image m y') → img-x m s ≡ img-x m t
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
179 img-x-cong {a} {b} m mono .(m x) .(m x₁) eq (isImage x) (isImage x₁)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
180 with cong (λ k → k ! ) ( Mono.isMono mono {One} (λ _ → x) (λ _ → x₁ ) ( extensionality Sets ( λ _ → eq )) )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
181 ... | refl = refl
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
182 img-x-cong0 : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → (y : a) → (s t : image m y ) → img-x m s ≡ img-x m t
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
183 img-x-cong0 m mono y s t = img-x-cong m mono y y refl s t
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
184 isol : {a b : Obj (Sets {c}) } (m : Hom Sets b a) → (mono : Mono Sets m ) → IsoL Sets m (λ (e : sequ a Bool (tchar m mono) (λ _ → true )) → equ e )
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
185 isol {a} {b} m mono = record { iso-L = record { ≅→ = b→s ; ≅← = b←s ;
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
186 iso→ = extensionality Sets ( λ x → iso1 x )
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
187 ; iso← = extensionality Sets ( λ x → iso2 x) } ; iso≈L = {!!} } where
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
188 b→s : Hom Sets b (sequ a Bool (tchar m mono) (λ _ → true))
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
189 b→s x with tchar m mono (m x) | inspect (tchar m mono ) (m x)
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
190 ... | true | record { eq = eq1 } = elem (m x) eq1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
191 b→s x | false | record { eq = eq1 } with tchar¬Img m mono (m x) eq1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
192 ... | t = ⊥-elim (t (isImage x))
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
193 b←s : Hom Sets (sequ a Bool (tchar m mono) (λ _ → true)) b
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
194 b←s (elem y eq) with tchar m mono y | inspect (tchar m mono ) y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
195 ... | true | record { eq = eq1 } = img-x m (tcharImg m mono y eq1 )
1027
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
196 bs1 : (y : a) → (eq1 : tchar m mono y ≡ true ) → b←s ( elem y eq1 ) ≡ img-x m (tcharImg m mono y eq1 )
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
197 bs1 y eq1 = ?
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
198 iso1 : (x : b) → b←s ( b→s x ) ≡ x
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
199 iso1 x with tchar m mono (m x) | inspect (tchar m mono ) (m x)
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
200 ... | true | record { eq = eq1 } with tcharImg m mono (m x) eq1 | inspect ( tcharImg m mono (m x) ) eq1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
201 ... | t | record { eq = eq2 } = begin
1027
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1026
diff changeset
202 b←s ( elem (m x) eq1 ) ≡⟨ bs1 (m x) eq1 ⟩
1026
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
203 img-x m (tcharImg m mono (m x) eq1 ) ≡⟨ cong (λ k → img-x m k ) eq2 ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
204 img-x m t ≡⟨ img-x-cong0 m mono (m x ) t (isImage x) ⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
205 img-x m (isImage x) ≡⟨⟩
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1025
diff changeset
206 x ∎ where open ≡-Reasoning
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
207 iso1 x | false | record { eq = eq1 } = ⊥-elim ( tchar¬Img m mono (m x) eq1 (isImage x))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
208 iso2 : (x : sequ a Bool (tchar m mono) (λ _ → true) ) → (Sets [ b→s o b←s ]) x ≡ id1 Sets (sequ a Bool (tchar m mono) (λ _ → true)) x
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
209 iso2 (elem y eq) = {!!}
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
210 imequ : {a b : Obj Sets} (m : Hom Sets b a) (mono : Mono Sets m) → IsEqualizer Sets m (tchar m mono) (Sets [ (λ _ → true ) o CCC.○ sets a ])
1022
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1021
diff changeset
211 imequ {a} {b} m mono = equalizerIso _ _ (tker (tchar m mono)) m (isol m mono)
1023
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
212 uniq : {a : Obj (Sets {c})} {b : Obj Sets} (h : Hom Sets a Bool) (m : Hom Sets b a) (mono : Mono Sets m) (y : a) →
227e1fe321ea using Bool and LEM
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1022
diff changeset
213 tchar (Equalizer.equalizer (tker h)) (record { isMono = λ f g → monic (tker h) }) y ≡ h y
1025
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
214 uniq {a} {b} h m mono y with h y | inspect h y | lem (image (Equalizer.equalizer (tker h)) y ) | inspect (tchar (Equalizer.equalizer (tker h)) (record { isMono = λ f g → monic (tker h) })) y
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
215 ... | true | record { eq = eqhy } | case1 x | record { eq = eq1 } = eq1
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
216 ... | true | record { eq = eqhy } | case2 x | record { eq = eq1 } = ⊥-elim (x (isImage (elem y eqhy)))
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
217 ... | false | record { eq = eqhy } | case1 (isImage (elem x eq)) | record { eq = eq1 } = ⊥-elim ( ¬x≡t∧x≡f record {fst = eqhy ; snd = eq })
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1024
diff changeset
218 ... | false | record { eq = eqhy } | case2 x | record { eq = eq1 } = eq1
1010
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 999
diff changeset
219
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
220
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
221 open import graph
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
222 module ccc-from-graph {c₁ c₂ : Level } (G : Graph {c₁} {c₂}) where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
223
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
224 open import Relation.Binary.PropositionalEquality renaming ( cong to ≡-cong ) hiding ( [_] )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
225 open Graph
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
226
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
227 V = vertex G
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
228 E : V → V → Set c₂
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
229 E = edge G
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
230
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
231 data Objs : Set c₁ where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
232 atom : V → Objs
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
233 ⊤ : Objs
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
234 _∧_ : Objs → Objs → Objs
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
235 _<=_ : Objs → Objs → Objs
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
236
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
237 data Arrows : (b c : Objs ) → Set (c₁ ⊔ c₂)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
238 data Arrow : Objs → Objs → Set (c₁ ⊔ c₂) where --- case i
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
239 arrow : {a b : V} → E a b → Arrow (atom a) (atom b)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
240 π : {a b : Objs } → Arrow ( a ∧ b ) a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
241 π' : {a b : Objs } → Arrow ( a ∧ b ) b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
242 ε : {a b : Objs } → Arrow ((a <= b) ∧ b ) a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
243 _* : {a b c : Objs } → Arrows (c ∧ b ) a → Arrow c ( a <= b ) --- case v
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
244
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
245 data Arrows where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
246 id : ( a : Objs ) → Arrows a a --- case i
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
247 ○ : ( a : Objs ) → Arrows a ⊤ --- case i
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
248 <_,_> : {a b c : Objs } → Arrows c a → Arrows c b → Arrows c (a ∧ b) -- case iii
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
249 iv : {b c d : Objs } ( f : Arrow d c ) ( g : Arrows b d ) → Arrows b c -- cas iv
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
250
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
251 _・_ : {a b c : Objs } (f : Arrows b c ) → (g : Arrows a b) → Arrows a c
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
252 id a ・ g = g
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
253 ○ a ・ g = ○ _
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
254 < f , g > ・ h = < f ・ h , g ・ h >
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
255 iv f g ・ h = iv f ( g ・ h )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
256
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
257
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
258 identityL : {A B : Objs} {f : Arrows A B} → (id B ・ f) ≡ f
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
259 identityL = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
260
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
261 identityR : {A B : Objs} {f : Arrows A B} → (f ・ id A) ≡ f
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
262 identityR {a} {a} {id a} = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
263 identityR {a} {⊤} {○ a} = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
264 identityR {a} {_} {< f , f₁ >} = cong₂ (λ j k → < j , k > ) identityR identityR
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
265 identityR {a} {b} {iv f g} = cong (λ k → iv f k ) identityR
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
266
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
267 assoc≡ : {a b c d : Objs} (f : Arrows c d) (g : Arrows b c) (h : Arrows a b) →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
268 (f ・ (g ・ h)) ≡ ((f ・ g) ・ h)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
269 assoc≡ (id a) g h = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
270 assoc≡ (○ a) g h = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
271 assoc≡ < f , f₁ > g h = cong₂ (λ j k → < j , k > ) (assoc≡ f g h) (assoc≡ f₁ g h)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
272 assoc≡ (iv f f1) g h = cong (λ k → iv f k ) ( assoc≡ f1 g h )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
273
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
274 -- positive intutionistic calculus
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
275 PL : Category c₁ (c₁ ⊔ c₂) (c₁ ⊔ c₂)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
276 PL = record {
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
277 Obj = Objs;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
278 Hom = λ a b → Arrows a b ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
279 _o_ = λ{a} {b} {c} x y → x ・ y ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
280 _≈_ = λ x y → x ≡ y ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
281 Id = λ{a} → id a ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
282 isCategory = record {
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
283 isEquivalence = record {refl = refl ; trans = trans ; sym = sym} ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
284 identityL = λ {a b f} → identityL {a} {b} {f} ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
285 identityR = λ {a b f} → identityR {a} {b} {f} ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
286 o-resp-≈ = λ {a b c f g h i} → o-resp-≈ {a} {b} {c} {f} {g} {h} {i} ;
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
287 associative = λ{a b c d f g h } → assoc≡ f g h
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
288 }
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
289 } where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
290 o-resp-≈ : {A B C : Objs} {f g : Arrows A B} {h i : Arrows B C} →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
291 f ≡ g → h ≡ i → (h ・ f) ≡ (i ・ g)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
292 o-resp-≈ refl refl = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
293 --------
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
294 --
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
295 -- Functor from Positive Logic to Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
296 --
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
297
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
298 -- open import Category.Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
299 -- postulate extensionality : { c₁ c₂ ℓ : Level} ( A : Category c₁ c₂ ℓ ) → Relation.Binary.PropositionalEquality.Extensionalit y c₂ c₂
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
300
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
301 open import Data.List
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
302
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
303 C = graphtocat.Chain G
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
304
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
305 tr : {a b : vertex G} → edge G a b → ((y : vertex G) → C y a) → (y : vertex G) → C y b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
306 tr f x y = graphtocat.next f (x y)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
307
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
308 fobj : ( a : Objs ) → Set (c₁ ⊔ c₂)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
309 fobj (atom x) = ( y : vertex G ) → C y x
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
310 fobj ⊤ = One
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
311 fobj (a ∧ b) = ( fobj a /\ fobj b)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
312 fobj (a <= b) = fobj b → fobj a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
313
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
314 fmap : { a b : Objs } → Hom PL a b → fobj a → fobj b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
315 amap : { a b : Objs } → Arrow a b → fobj a → fobj b
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
316 amap (arrow x) y = tr x y -- tr x
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
317 amap π ( x , y ) = x
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
318 amap π' ( x , y ) = y
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
319 amap ε (f , x ) = f x
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
320 amap (f *) x = λ y → fmap f ( x , y )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
321 fmap (id a) x = x
1020
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 1019
diff changeset
322 fmap (○ a) x = !
999
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
323 fmap < f , g > x = ( fmap f x , fmap g x )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
324 fmap (iv x f) a = amap x ( fmap f a )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
325
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
326 -- CS is a map from Positive logic to Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
327 -- Sets is CCC, so we have a cartesian closed category generated by a graph
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
328 -- as a sub category of Sets
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
329
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
330 CS : Functor PL (Sets {c₁ ⊔ c₂})
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
331 FObj CS a = fobj a
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
332 FMap CS {a} {b} f = fmap {a} {b} f
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
333 isFunctor CS = isf where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
334 _+_ = Category._o_ PL
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
335 ++idR = IsCategory.identityR ( Category.isCategory PL )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
336 distr : {a b c : Obj PL} { f : Hom PL a b } { g : Hom PL b c } → (z : fobj a ) → fmap (g + f) z ≡ fmap g (fmap f z)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
337 distr {a} {a₁} {a₁} {f} {id a₁} z = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
338 distr {a} {a₁} {⊤} {f} {○ a₁} z = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
339 distr {a} {b} {c ∧ d} {f} {< g , g₁ >} z = cong₂ (λ j k → j , k ) (distr {a} {b} {c} {f} {g} z) (distr {a} {b} {d} {f} {g₁} z)
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
340 distr {a} {b} {c} {f} {iv {_} {_} {d} x g} z = adistr (distr {a} {b} {d} {f} {g} z) x where
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
341 adistr : fmap (g + f) z ≡ fmap g (fmap f z) →
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
342 ( x : Arrow d c ) → fmap ( iv x (g + f) ) z ≡ fmap ( iv x g ) (fmap f z )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
343 adistr eq x = cong ( λ k → amap x k ) eq
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
344 isf : IsFunctor PL Sets fobj fmap
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
345 IsFunctor.identity isf = extensionality Sets ( λ x → refl )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
346 IsFunctor.≈-cong isf refl = refl
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
347 IsFunctor.distr isf {a} {b} {c} {g} {f} = extensionality Sets ( λ z → distr {a} {b} {c} {g} {f} z )
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
348
d89f2c8cf0f4 separate CCCSets
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
349