comparison CCCGraph.agda @ 825:8f41ad966eaa

rename discrete
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Fri, 03 May 2019 17:11:33 +0900
parents 878d8643214f
children d1569e80fe0b b8c5f15ee561
comparison
equal deleted inserted replaced
824:878d8643214f 825:8f41ad966eaa
95 *-cong refl = refl 95 *-cong refl = refl
96 96
97 module ccc-from-graph where 97 module ccc-from-graph where
98 98
99 open import Relation.Binary.PropositionalEquality renaming ( cong to ≡-cong ) hiding ( [_] ) 99 open import Relation.Binary.PropositionalEquality renaming ( cong to ≡-cong ) hiding ( [_] )
100 open import discrete 100 open import graph
101 open graphtocat 101 open graphtocat
102 102
103 open Graph 103 open Graph
104 104
105 data Objs (G : Graph {Level.zero} {Level.zero} ) : Set where -- formula 105 data Objs (G : Graph {Level.zero} {Level.zero} ) : Set where -- formula
263 ; identityR = λ {x} {y} {f} → let open ≈-Reasoning (CAT) in idR {cat x} {cat y} {cmap f} 263 ; identityR = λ {x} {y} {f} → let open ≈-Reasoning (CAT) in idR {cat x} {cat y} {cmap f}
264 ; o-resp-≈ = λ {x} {y} {z} {f} {g} {h} {i} → IsCategory.o-resp-≈ ( Category.isCategory CAT) {cat x}{cat y}{cat z} {cmap f} {cmap g} {cmap h} {cmap i} 264 ; o-resp-≈ = λ {x} {y} {z} {f} {g} {h} {i} → IsCategory.o-resp-≈ ( Category.isCategory CAT) {cat x}{cat y}{cat z} {cmap f} {cmap g} {cmap h} {cmap i}
265 ; associative = λ {a} {b} {c} {d} {f} {g} {h} → let open ≈-Reasoning (CAT) in assoc {cat a} {cat b} {cat c} {cat d} {cmap f} {cmap g} {cmap h} 265 ; associative = λ {a} {b} {c} {d} {f} {g} {h} → let open ≈-Reasoning (CAT) in assoc {cat a} {cat b} {cat c} {cat d} {cmap f} {cmap g} {cmap h}
266 }} 266 }}
267 267
268 open import discrete 268 open import graph
269 open Graph 269 open Graph
270 270
271 record GMap {v v' w w' : Level} (x : Graph {v} {v'} ) (y : Graph {w} {w'} ) : Set (suc (v ⊔ w) ⊔ v' ⊔ w' ) where 271 record GMap {v v' : Level} (x y : Graph {v} {v'} ) : Set (suc (v ⊔ v') ) where
272 field 272 field
273 vmap : vertex x → vertex y 273 vmap : vertex x → vertex y
274 emap : {a b : vertex x} → edge x a b → edge y (vmap a) (vmap b) 274 emap : {a b : vertex x} → edge x a b → edge y (vmap a) (vmap b)
275 275
276 open GMap 276 open GMap