Skip to content

Commit 4c5e885

Browse files
committed
Tidy up: s/Set/Type + flags + levels
1 parent 2d7cb5c commit 4c5e885

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

43 files changed

+84
-105
lines changed

Class/Applicative.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Applicative where
33

44
open import Class.Applicative.Core public

Class/Applicative/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Applicative.Core where
33

44
open import Class.Prelude

Class/Applicative/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Applicative.Instances where
33

44
open import Class.Prelude

Class/Bifunctor.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Bifunctor where
33

44
open import Class.Prelude hiding (A; B; C)

Class/CommutativeMonoid.agda

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,4 @@
1-
{-# OPTIONS --safe --without-K #-}
2-
1+
{-# OPTIONS --cubical-compatible #-}
32
module Class.CommutativeMonoid where
43

54
open import Class.CommutativeMonoid.Core public

Class/CommutativeMonoid/Core.agda

Lines changed: 5 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -1,12 +1,12 @@
1-
{-# OPTIONS --safe --without-K #-}
2-
1+
{-# OPTIONS --cubical-compatible #-}
32
module Class.CommutativeMonoid.Core where
43

5-
import Algebra as Alg
64
open import Class.Prelude
75
open import Class.Semigroup
86
open import Class.Monoid
97

8+
import Algebra as Alg
9+
1010
record CommutativeMonoid c ℓ Carrier : Type (lsuc (c ⊔ ℓ)) where
1111
infix 4 _≈_
1212
field
@@ -19,7 +19,8 @@ module Conversion {c ℓ} where
1919
toBundle : {Carrier} CommutativeMonoid c ℓ Carrier Alg.CommutativeMonoid c ℓ
2020
toBundle c = record { CommutativeMonoid c }
2121

22-
fromBundle : (m : Alg.CommutativeMonoid c ℓ) CommutativeMonoid c ℓ (Alg.CommutativeMonoid.Carrier m)
22+
fromBundle : (m : Alg.CommutativeMonoid c ℓ)
23+
CommutativeMonoid c ℓ (Alg.CommutativeMonoid.Carrier m)
2324
fromBundle c = record
2425
{ Alg.CommutativeMonoid c using (_≈_; isCommutativeMonoid)
2526
; semigroup = λ where ._◇_ Alg.CommutativeMonoid._∙_ c

Class/CommutativeMonoid/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --safe --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22

33
module Class.CommutativeMonoid.Instances where
44

Class/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Core where
33

44
open import Class.Prelude

Class/Default.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2,7 +2,7 @@
22
-- Types with a default value.
33
------------------------------------------------------------------------
44

5-
{-# OPTIONS --without-K #-}
5+
{-# OPTIONS --cubical-compatible #-}
66
module Class.Default where
77

88
open import Class.Prelude

Class/Foldable.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Foldable where
33

44
open import Class.Foldable.Core public

Class/Foldable/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Foldable.Core where
33

44
open import Class.Prelude

Class/Foldable/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Foldable.Instances where
33

44
open import Class.Prelude

Class/Functor.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Functor where
33

44
open import Class.Functor.Core public

Class/Functor/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Functor.Core where
33

44
open import Class.Prelude

Class/Functor/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Functor.Instances where
33

44
open import Class.Prelude

Class/HasAdd.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --safe --cubical-compatible #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.HasAdd where
33

44
open import Class.HasAdd.Core public

Class/HasAdd/Core.agda

Lines changed: 5 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -1,8 +1,10 @@
1-
{-# OPTIONS --safe --cubical-compatible #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.HasAdd.Core where
33

4-
record HasAdd (A : Set) : Set where
4+
open import Class.Prelude
5+
6+
record HasAdd (A : Type ℓ) : Type ℓ where
57
infixl 6 _+_
68
field _+_ : A A A
79

8-
open HasAdd ... public
10+
open HasAdd ⦃...⦄ public

Class/HasAdd/Instance.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --safe --cubical-compatible #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.HasAdd.Instance where
33

44
open import Class.HasAdd.Core

Class/HasOrder.agda

Lines changed: 0 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,3 @@
1-
{-# OPTIONS --safe #-}
2-
31
module Class.HasOrder where
42

53
open import Class.HasOrder.Core public

Class/HasOrder/Core.agda

Lines changed: 25 additions & 28 deletions
Original file line numberDiff line numberDiff line change
@@ -1,30 +1,28 @@
1-
{-# OPTIONS --safe #-}
2-
1+
{-# OPTIONS --without-K #-}
32
module Class.HasOrder.Core where
43

4+
open import Class.Prelude
55
open import Class.Decidable
6-
open import Data.Empty
7-
open import Data.Product
8-
open import Data.Sum
9-
open import Function
10-
open import Level
11-
open import Relation.Binary
12-
open import Relation.Binary using () renaming (Decidable to Decidable²)
13-
open import Relation.Binary.PropositionalEquality
14-
open import Relation.Nullary
6+
open import Function.Bundles using (module Equivalence; mk⇔; _⇔_)
7+
open import Relation.Binary using
8+
( IsPreorder; IsPartialOrder; IsEquivalence; Total; IsTotalOrder
9+
; IsStrictTotalOrder; IsStrictPartialOrder
10+
; Irreflexive; Antisymmetric; Asymmetric; Transitive
11+
)
1512

1613
open Equivalence
1714

18-
module _ {a} {A : Set a} where
19-
module _ {_≈_ : Rel A a} where
15+
module _ {A : Type ℓ} where
16+
module _ {_≈_ : Rel A ℓ′} {ℓ″ ℓ‴} (let 𝐿 = lsuc ℓ ⊔ ℓ′ ⊔ lsuc ℓ″ ⊔ lsuc ℓ‴) where
2017

21-
record HasPreorder : Set (suc a) where
18+
record HasPreorder : Type 𝐿 where
2219
infix 4 _≤_ _<_ _≥_ _>_
2320
field
24-
_≤_ _<_ : Rel A a
21+
_≤_ : Rel A ℓ″
22+
_<_ : Rel A ℓ‴
2523
≤-isPreorder : IsPreorder _≈_ _≤_
2624
<-irrefl : Irreflexive _≈_ _<_
27-
≤⇔<∨≈ : {x y} x ≤ y ⇔ (x < y ⊎ x ≈ y)
25+
≤⇔<∨≈ : {x y : A} x ≤ y ⇔ (x < y ⊎ x ≈ y)
2826

2927
_≥_ = flip _≤_
3028
_>_ = flip _<_
@@ -33,10 +31,10 @@ module _ {a} {A : Set a} where
3331
using ()
3432
renaming (isEquivalence to ≈-isEquivalence; refl to ≤-refl; trans to ≤-trans)
3533

36-
_≤?_ : ⦃ _≤_ ⁇² ⦄ Decidable _≤_
34+
_≤?_ : ⦃ _≤_ ⁇² ⦄ Decidable² _≤_
3735
_≤?_ = dec²
3836

39-
_<?_ : ⦃ _<_ ⁇² ⦄ Decidable _<_
37+
_<?_ : ⦃ _<_ ⁇² ⦄ Decidable² _<_
4038
_<?_ = dec²
4139

4240
infix 4 _<?_ _≤?_
@@ -55,12 +53,12 @@ module _ {a} {A : Set a} where
5553

5654
open HasPreorder ⦃...⦄
5755

58-
record HasDecPreorder : Set (suc a) where
56+
record HasDecPreorder : Type 𝐿 where
5957
field ⦃ hasPreorder ⦄ : HasPreorder
6058
⦃ dec-≤ ⦄ : _≤_ ⁇²
6159
⦃ dec-< ⦄ : _<_ ⁇²
6260

63-
record HasPartialOrder : Set (suc a) where
61+
record HasPartialOrder : Type 𝐿 where
6462
field
6563
⦃ hasPreorder ⦄ : HasPreorder
6664
≤-antisym : Antisymmetric _≈_ _≤_
@@ -90,13 +88,13 @@ module _ {a} {A : Set a} where
9088

9189
open HasPartialOrder ⦃...⦄
9290

93-
record HasDecPartialOrder : Set (suc a) where
91+
record HasDecPartialOrder : Type 𝐿 where
9492
field
9593
⦃ hasPartialOrder ⦄ : HasPartialOrder
9694
⦃ dec-≤ ⦄ : _≤_ ⁇²
9795
⦃ dec-< ⦄ : _<_ ⁇²
9896

99-
record HasTotalOrder : Set (suc a) where
97+
record HasTotalOrder : Type 𝐿 where
10098
field
10199
⦃ hasPartialOrder ⦄ : HasPartialOrder
102100
≤-total : Total _≤_
@@ -106,15 +104,15 @@ module _ {a} {A : Set a} where
106104

107105
open IsEquivalence ≈-isEquivalence renaming (sym to ≈-sym)
108106

109-
≮⇒≥ : Decidable _≈_ {x y} ¬ (x < y) y ≤ x
107+
≮⇒≥ : Decidable² _≈_ {x y} ¬ (x < y) y ≤ x
110108
≮⇒≥ _≈?_ {x} {y} x≮y with x ≈? y | ≤-total y x
111109
... | yes x≈y | _ = IsPreorder.reflexive ≤-isPreorder (≈-sym x≈y)
112110
... | _ | inj₁ y≤x = y≤x
113111
... | no x≉y | inj₂ x≤y = contradiction (≤∧≉⇒< (x≤y , x≉y)) x≮y
114112

115113
open HasTotalOrder ⦃...⦄
116114

117-
record HasDecTotalOrder : Set (suc a) where
115+
record HasDecTotalOrder : Type 𝐿 where
118116
field
119117
⦃ hasTotalOrder ⦄ : HasTotalOrder
120118
⦃ dec-≤ ⦄ : _≤_ ⁇²
@@ -133,9 +131,9 @@ open HasDecPartialOrder ⦃...⦄ public hiding (hasPartialOrder)
133131
open HasTotalOrder ⦃...⦄ public hiding (hasPartialOrder)
134132
open HasDecTotalOrder ⦃...⦄ public hiding (hasTotalOrder)
135133

136-
module _ {a} {A : Set a} {_≈_ : Rel A a} where
134+
module _ {A : Type ℓ} {_≈_ : Rel A ℓ′} where
137135

138-
module _ {_≤_ : Rel A a} where
136+
module _ {_≤_ : Rel A ℓ″} where
139137
import Relation.Binary.Construct.NonStrictToStrict _≈_ _≤_ as SNS
140138

141139
module _ (≤-isPreorder : IsPreorder _≈_ _≤_)
@@ -158,7 +156,7 @@ module _ {a} {A : Set a} {_≈_ : Rel A a} where
158156
; ≤-antisym = antsym
159157
}
160158

161-
module _ {_<_ : Rel A a} where
159+
module _ {_<_ : Rel A ℓ″} where
162160

163161
import Relation.Binary.Construct.StrictToNonStrict _≈_ _<_ as SNS
164162

@@ -181,7 +179,6 @@ module _ {a} {A : Set a} {_≈_ : Rel A a} where
181179
; ≤-antisym = SNS.isPartialOrder <-isStrictPartialOrder .IsPartialOrder.antisym
182180
}
183181

184-
185182
module _ (<-isStrictTotalOrder : IsStrictTotalOrder _≈_ _<_) where
186183

187184
private spo = IsStrictTotalOrder.isStrictPartialOrder <-isStrictTotalOrder

Class/HasOrder/Instance.agda

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,4 @@
1-
{-# OPTIONS --safe #-}
2-
1+
{-# OPTIONS --without-K #-}
32
module Class.HasOrder.Instance where
43

54
open import Class.DecEq

Class/Monad.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monad where
33

44
open import Class.Monad.Core public

Class/Monad/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monad.Core where
33

44
open import Class.Prelude

Class/Monad/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monad.Instances where
33

44
open import Class.Prelude

Class/Monoid.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monoid where
33

44
open import Class.Monoid.Core public

Class/Monoid/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monoid.Core where
33

44
open import Class.Prelude

Class/Monoid/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Monoid.Instances where
33

44
open import Class.Prelude

Class/Prelude.agda

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -1,12 +1,12 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Prelude where
33

44
open import Agda.Primitive public
55
using () renaming (Set to Type; Setω to Typeω)
66
open import Level public
77
using (Level; _⊔_) renaming (suc to lsuc)
8-
open import Function public
9-
using (id; _∘_; _∋_; _$_; const; constᵣ; flip; it)
8+
open import Function.Base public
9+
using (id; _∘_; _∋_; _$_; const; constᵣ; flip; it; case_of_)
1010

1111
open import Data.Empty public
1212
using (⊥; ⊥-elim)
@@ -44,7 +44,7 @@ open import Data.These public
4444
using (These; this; that; these)
4545

4646
open import Relation.Nullary public
47-
using (¬_; Dec; yes; no)
47+
using (¬_; Dec; yes; no; contradiction)
4848
open import Relation.Nullary.Decidable public
4949
using (⌊_⌋; dec-yes; isYes)
5050
open import Relation.Unary public
@@ -54,7 +54,7 @@ open import Relation.Binary public
5454
using (REL; Rel; DecidableEquality)
5555
renaming (Decidable to Decidable²)
5656
module _ {a b c} where
57-
3REL : (A : Set a) (B : Set b) (C : Set c) (ℓ : Level) Type _
57+
3REL : (A : Type a) (B : Type b) (C : Type c) (ℓ : Level) Type _
5858
3REL A B C ℓ = A B C Type ℓ
5959

6060
Decidable³ : {A B C ℓ} 3REL A B C ℓ Type _
@@ -66,7 +66,7 @@ open import Reflection public
6666
using (TC; Arg; Abs)
6767

6868
variable
69-
ℓ ℓ′ ℓ″ : Level
69+
ℓ ℓ′ ℓ″ ℓ‴ : Level
7070
A B C : Type ℓ
7171

7272
module Alg (_~_ : Rel A ℓ) where

Class/Semigroup.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Semigroup where
33

44
open import Class.Semigroup.Core public

Class/Semigroup/Core.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Semigroup.Core where
33

44
open import Class.Prelude

Class/Semigroup/Derive.agda

Lines changed: 0 additions & 10 deletions
This file was deleted.

Class/Semigroup/Instances.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,4 +1,4 @@
1-
{-# OPTIONS --without-K #-}
1+
{-# OPTIONS --cubical-compatible #-}
22
module Class.Semigroup.Instances where
33

44
open import Class.Prelude

0 commit comments

Comments
 (0)