Skip to content

Commit 67afe7c

Browse files
jamesmckinnaandreasabel
authored andcommittedJul 10, 2024
remove final references to Category.* (#2214)
1 parent 56ed8fe commit 67afe7c

File tree

3 files changed

+10
-10
lines changed

3 files changed

+10
-10
lines changed
 

‎src/Algebra/Lattice/Properties/BooleanAlgebra/Expression.agda

+6-6
Original file line numberDiff line numberDiff line change
@@ -20,8 +20,8 @@ open import Data.Fin.Base using (Fin)
2020
open import Data.Nat.Base
2121
open import Data.Product.Base using (_,_; proj₁; proj₂)
2222
open import Data.Vec.Base as Vec using (Vec)
23-
import Data.Vec.Effectful as VecCat
24-
import Function.Identity.Effectful as IdCat
23+
import Data.Vec.Effectful as Vec
24+
import Function.Identity.Effectful as Identity
2525
open import Data.Vec.Properties using (lookup-map)
2626
open import Data.Vec.Relation.Binary.Pointwise.Extensional as PW
2727
using (Pointwise; ext)
@@ -165,18 +165,18 @@ lift n = record
165165
}
166166
}
167167
where
168-
open RawApplicative VecCat.applicative
168+
open RawApplicative Vec.applicative
169169
using (pure; zipWith) renaming (_<$>_ to map)
170170

171171
⟦_⟧Id : {n} Expr n Vec Carrier n Carrier
172-
⟦_⟧Id = Semantics.⟦_⟧ IdCat.applicative
172+
⟦_⟧Id = Semantics.⟦_⟧ Identity.applicative
173173

174174
⟦_⟧Vec : {m n} Expr n Vec (Vec Carrier m) n Vec Carrier m
175-
⟦_⟧Vec = Semantics.⟦_⟧ VecCat.applicative
175+
⟦_⟧Vec = Semantics.⟦_⟧ Vec.applicative
176176

177177
open module R {n} (i : Fin n) =
178178
Reflection setoid var
179179
(λ e ρ Vec.lookup (⟦ e ⟧Vec ρ) i)
180180
(λ e ρ ⟦ e ⟧Id (Vec.map (flip Vec.lookup i) ρ))
181181
(λ e ρ sym $ reflexive $
182-
Naturality.natural (VecCat.lookup-morphism i) e ρ)
182+
Naturality.natural (Vec.lookup-morphism i) e ρ)

‎src/Relation/Binary/Construct/Closure/ReflexiveTransitive.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -130,11 +130,11 @@ _>>=_ : ∀ {i t u} {I : Set i} {T : Rel I t} {U : Rel I u} {j k} →
130130
m >>= f = (f ⋆) m
131131

132132
-- Note that the monad-like structure above is not an indexed monad
133-
-- (as defined in Category.Monad.Indexed). If it were, then _>>=_
133+
-- (as defined in Effect.Monad.Indexed). If it were, then _>>=_
134134
-- would have a type similar to
135135
--
136136
-- ∀ {I} {T U : Rel I t} {i j k} →
137137
-- Star T i j → (T i j → Star U j k) → Star U i k.
138138
-- ^^^^^
139139
-- Note, however, that there is no scope for applying T to any indices
140-
-- in the definition used in Category.Monad.Indexed.
140+
-- in the definition used in Effect.Monad.Indexed.

‎src/Text/Pretty.agda

+2-2
Original file line numberDiff line numberDiff line change
@@ -23,8 +23,8 @@ open import Data.String.Base using (String; fromList; replicate)
2323
open import Function.Base using (_∘_; _∘′_; _$_)
2424

2525
open import Effect.Monad using (RawMonad)
26-
import Data.List.Effectful as Cat
27-
open RawMonad (Cat.monad {Level.zero})
26+
import Data.List.Effectful as List
27+
open RawMonad (List.monad {Level.zero})
2828

2929
import Data.Nat.Properties as ℕₚ
3030
open import Data.List.Extrema.Core ℕₚ.≤-totalOrder using (⊓ᴸ)

0 commit comments

Comments
 (0)
Please sign in to comment.