Skip to content

Commit

Permalink
remove levelOf calls
Browse files Browse the repository at this point in the history
c.f. #10
  • Loading branch information
HuStmpHrrr committed Jul 13, 2019
1 parent cd60a79 commit f40c009
Show file tree
Hide file tree
Showing 7 changed files with 15 additions and 13 deletions.
6 changes: 3 additions & 3 deletions Categories/Adjoint.agda
Original file line number Diff line number Diff line change
Expand Up @@ -3,15 +3,15 @@ module Categories.Adjoint where

-- Adjoints

open import Level using (Level; _⊔_)
open import Level using (Level; _⊔_; levelOfTerm)

open import Data.Product using (_,_; _×_)
open import Function using () renaming (_∘_ to _∙_)
open import Function.Inverse using (Inverse)
open import Relation.Binary using (Rel; IsEquivalence; Setoid)

-- be explicit in imports to 'see' where the information comes from
open import Categories.Category using (Category; levelOf)
open import Categories.Category using (Category)
open import Categories.Category.Product using (Product; _⁂_)
open import Categories.Category.Instance.Setoids
open import Categories.Functor using (Functor; _∘F_) renaming (id to idF)
Expand All @@ -27,7 +27,7 @@ private
o ℓ e : Level
C D : Category o ℓ e

record Adjoint (L : Functor C D) (R : Functor D C) : Set (levelOf C ⊔ levelOf D) where
record Adjoint (L : Functor C D) (R : Functor D C) : Set (levelOfTerm C ⊔ levelOfTerm D) where
private
module C = Category C
module D = Category D
Expand Down
4 changes: 0 additions & 4 deletions Categories/Category.agda
Original file line number Diff line number Diff line change
Expand Up @@ -20,7 +20,3 @@ _[_≈_] = Category._≈_

_[_∘_] : {o ℓ e} (C : Category o ℓ e) {X Y Z} (f : C [ Y , Z ]) (g : C [ X , Y ]) C [ X , Z ]
_[_∘_] = Category._∘_

-- Currently defined here for easier level polymorphism, but should move to main library
levelOf : {o ℓ e} Category o ℓ e Level
levelOf {o} {ℓ} {e} _ = o ⊔ ℓ ⊔ e
4 changes: 2 additions & 2 deletions Categories/Category/Cartesian.agda
Original file line number Diff line number Diff line change
Expand Up @@ -30,7 +30,7 @@ private
A B C D X Y Z : Obj
f f′ g g′ h i : A ⇒ B

record BinaryProducts : Set (levelOf 𝒞) where
record BinaryProducts : Set (levelOfTerm 𝒞) where

infixr 5 _×_
infix 8 _⁂_
Expand Down Expand Up @@ -238,7 +238,7 @@ record BinaryProducts : Set (levelOf 𝒞) where
_×- = appˡ -×-

-- Cartesian monoidal category
record Cartesian : Set (levelOf 𝒞) where
record Cartesian : Set (levelOfTerm 𝒞) where
field
terminal : Terminal
products : BinaryProducts
Expand Down
2 changes: 1 addition & 1 deletion Categories/Category/CartesianClosed.agda
Original file line number Diff line number Diff line change
Expand Up @@ -28,7 +28,7 @@ private

-- Cartesian closed category
-- is a category with all products and exponentials
record CartesianClosed : Set (levelOf 𝒞) where
record CartesianClosed : Set (levelOfTerm 𝒞) where
infixl 7 _^_

field
Expand Down
4 changes: 3 additions & 1 deletion Categories/Category/Monoidal/Braided.agda
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,8 @@ open import Categories.Category.Monoidal

module Categories.Category.Monoidal.Braided {o ℓ e} {C : Category o ℓ e} (M : Monoidal C) where

open import Level

open import Data.Product using (Σ; _,_)

open import Categories.Functor.Bifunctor
Expand All @@ -20,7 +22,7 @@ private
-- braided monoidal category
-- it has a braiding natural isomorphism has two hexagon identities.
-- these two identities are directly expressed in the morphism level.
record Braided : Set (levelOf C) where
record Braided : Set (levelOfTerm C) where
open Monoidal M public

field
Expand Down
4 changes: 3 additions & 1 deletion Categories/Category/Monoidal/Symmetric.agda
Original file line number Diff line number Diff line change
Expand Up @@ -5,6 +5,8 @@ open import Categories.Category.Monoidal

module Categories.Category.Monoidal.Symmetric {o ℓ e} {C : Category o ℓ e} (M : Monoidal C) where

open import Level

open import Data.Product using (Σ; _,_)

open import Categories.Functor.Bifunctor
Expand All @@ -21,7 +23,7 @@ private

-- symmetric monoidal category
-- commutative braided monoidal category
record Symmetric : Set (levelOf C) where
record Symmetric : Set (levelOfTerm C) where
field
braided : Braided

Expand Down
4 changes: 3 additions & 1 deletion Categories/Category/Monoidal/Traced.agda
Original file line number Diff line number Diff line change
Expand Up @@ -7,6 +7,8 @@ module Categories.Category.Monoidal.Traced {o ℓ e} {C : Category o ℓ e} (M :

open Category C

open import Level

open import Data.Product using (_,_)

open import Categories.Category.Monoidal.Symmetric M
Expand Down Expand Up @@ -37,7 +39,7 @@ private
--
-- note that the definition in this library is significantly easier than the previous one because
-- we adopt a simpler definition of monoidal category to begin with.
record Traced : Set (levelOf C) where
record Traced : Set (levelOfTerm C) where
field
symmetric : Symmetric

Expand Down

0 comments on commit f40c009

Please sign in to comment.