Skip to content

Commit 7602ca3

Browse files
Stage 2 of refactoring of Relation.Nullary (#1869)
* Deprecate Relation.Nullary.(Sum/Product/Implication) * Some grepped refining of imports * Added missing deprecation warnings
1 parent 51a138b commit 7602ca3

File tree

142 files changed

+305
-355
lines changed

Some content is hidden

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

142 files changed

+305
-355
lines changed

CHANGELOG.md

+10
Original file line numberDiff line numberDiff line change
@@ -534,6 +534,16 @@ Non-backwards compatible changes
534534
535535
* Backwards compatibility has been maintained, as `Relation.Nullary` still re-exports these publicly.
536536
537+
* The modules:
538+
```
539+
Relation.Nullary.Product
540+
Relation.Nullary.Sum
541+
Relation.Nullary.Implication
542+
```
543+
have been deprecated and their contents moved to `Relation.Nullary.(Negation/Reflects/Decidable)`
544+
however all their contents is re-exported by `Relation.Nullary` which is the easiest way to access
545+
it now.
546+
537547
* In order to facilitate this reorganisation the following breaking moves have occured:
538548
- `¬?` has been moved from `Relation.Nullary.Negation.Core` to `Relation.Nullary.Decidable.Core`
539549
- `¬-reflects` has been moved from `Relation.Nullary.Negation.Core` to `Relation.Nullary.Reflects`.

README/Axiom.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -31,7 +31,7 @@ private variable ℓ : Level
3131
-- The types for which `P` or `¬P` holds is called `Dec P` in the
3232
-- standard library (short for `Decidable`).
3333

34-
open import Relation.Nullary using (Dec)
34+
open import Relation.Nullary.Decidable using (Dec)
3535

3636
-- The type of the proof of saying that excluded middle holds for
3737
-- all types at universe level ℓ is therefore:

README/Data/List/Relation/Binary/Pointwise.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@ open import Data.Nat using (ℕ; _<_; s≤s; z≤n)
1010
open import Data.List using (List; []; _∷_; length)
1111
open import Relation.Binary.PropositionalEquality
1212
using (_≡_; refl; sym; cong; setoid)
13-
open import Relation.Nullary using (¬_)
13+
open import Relation.Nullary.Negation using (¬_)
1414

1515
------------------------------------------------------------------------
1616
-- Pointwise

README/Design/Decidability.agda

+10-8
Original file line numberDiff line numberDiff line change
@@ -8,13 +8,6 @@
88

99
module README.Design.Decidability where
1010

11-
-- Reflects and Dec are defined in Relation.Nullary, and operations on them can
12-
-- be found in Relation.Nullary.Reflects and Relation.Nullary.Decidable.
13-
14-
open import Relation.Nullary as Nullary
15-
open import Relation.Nullary.Reflects
16-
open import Relation.Nullary.Decidable
17-
1811
open import Data.Bool
1912
open import Data.List
2013
open import Data.List.Properties using (∷-injective)
@@ -25,7 +18,11 @@ open import Data.Unit
2518
open import Function
2619
open import Relation.Binary.PropositionalEquality
2720
open import Relation.Nary
28-
open import Relation.Nullary.Product
21+
22+
------------------------------------------------------------------------
23+
-- Reflects
24+
25+
open import Relation.Nullary.Reflects
2926

3027
infix 4 _≟₀_ _≟₁_ _≟₂_
3128

@@ -43,6 +40,11 @@ ex₂ : (b : Bool) → Reflects (T b) b
4340
ex₂ false = ofⁿ id
4441
ex₂ true = ofʸ tt
4542

43+
------------------------------------------------------------------------
44+
-- Dec
45+
46+
open import Relation.Nullary.Decidable
47+
4648
-- A proof of `Dec P` is a proof of `Reflects P b` for some `b`.
4749
-- `Dec P` is declared as a record, with fields:
4850
-- does : Bool

README/Text/Regex.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open import Data.Bool using (true; false)
1212
open import Data.List using (_∷_; [])
1313
open import Data.String
1414
open import Function.Base using () renaming (_$′_ to _$_)
15-
open import Relation.Nullary using (yes)
15+
open import Relation.Nullary.Decidable using (yes)
1616
open import Relation.Nullary.Decidable using (True; False; from-yes)
1717

1818
-- Our library available via the Text.Regex module is safe but it works on

src/Algebra/Apartness/Structures.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,7 @@ open import Algebra.Definitions _≈_ using (Invertible)
2222
open import Algebra.Structures _≈_ using (IsCommutativeRing)
2323
open import Relation.Binary.Structures using (IsEquivalence; IsApartnessRelation)
2424
open import Relation.Binary.Definitions using (Tight)
25-
open import Relation.Nullary using (¬_)
25+
open import Relation.Nullary.Negation using (¬_)
2626
import Relation.Binary.Properties.ApartnessRelation as AR
2727

2828

src/Algebra/Bundles/Raw.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ module Algebra.Bundles.Raw where
1111
open import Algebra.Core
1212
open import Relation.Binary.Core using (Rel)
1313
open import Level using (suc; _⊔_)
14-
open import Relation.Nullary using (¬_)
14+
open import Relation.Nullary.Negation using (¬_)
1515

1616
------------------------------------------------------------------------
1717
-- Raw bundles with 1 binary operation

src/Algebra/Construct/LexProduct/Base.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Data.Bool.Base using (true; false)
1111
open import Data.Product using (_×_; _,_)
1212
open import Relation.Binary.Core using (Rel)
1313
open import Relation.Binary.Definitions using (Decidable)
14-
open import Relation.Nullary using (does; yes; no)
14+
open import Relation.Nullary.Decidable using (does; yes; no)
1515

1616
module Algebra.Construct.LexProduct.Base
1717
{a b ℓ} {A : Set a} {B : Set b}

src/Algebra/Construct/LexProduct/Inner.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ open import Data.Product using (_×_; _,_; swap; map; uncurry′)
1212
open import Function.Base using (_∘_)
1313
open import Level using (Level; _⊔_)
1414
open import Relation.Binary.Definitions using (Decidable)
15-
open import Relation.Nullary using (does; yes; no)
15+
open import Relation.Nullary.Decidable using (does; yes; no)
1616
open import Relation.Nullary.Negation
1717
using (contradiction; contradiction₂)
1818
import Relation.Binary.Reasoning.Setoid as SetoidReasoning

src/Algebra/Definitions.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@
1616
{-# OPTIONS --without-K --safe #-}
1717

1818
open import Relation.Binary.Core
19-
open import Relation.Nullary using (¬_)
19+
open import Relation.Nullary.Negation using (¬_)
2020

2121
module Algebra.Definitions
2222
{a ℓ} {A : Set a} -- The underlying set

src/Algebra/Definitions/RawMagma.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Algebra.Bundles using (RawMagma)
1414
open import Data.Product using (_×_; ∃)
1515
open import Level using (_⊔_)
1616
open import Relation.Binary.Core
17-
open import Relation.Nullary using (¬_)
17+
open import Relation.Nullary.Negation using (¬_)
1818

1919
module Algebra.Definitions.RawMagma
2020
{a ℓ} (M : RawMagma a ℓ)

src/Algebra/Properties/CancellativeCommutativeSemiring.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@ open import Algebra using (CancellativeCommutativeSemiring)
1010
open import Algebra.Definitions using (AlmostRightCancellative)
1111
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
1212
open import Relation.Binary using (Decidable)
13-
open import Relation.Nullary using (yes; no)
13+
open import Relation.Nullary.Decidable using (yes; no)
1414
open import Relation.Nullary.Negation using (contradiction)
1515

1616
module Algebra.Properties.CancellativeCommutativeSemiring

src/Algebra/Solver/CommutativeMonoid.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -28,7 +28,7 @@ import Relation.Nullary.Decidable as Dec
2828
import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise
2929

3030
open import Relation.Binary.PropositionalEquality as P using (_≡_; decSetoid)
31-
open import Relation.Nullary using (Dec)
31+
open import Relation.Nullary.Decidable using (Dec)
3232

3333
open CommutativeMonoid M
3434
open EqReasoning setoid

src/Algebra/Solver/IdempotentCommutativeMonoid.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ import Relation.Nullary.Decidable as Dec
2626
import Data.Vec.Relation.Binary.Pointwise.Inductive as Pointwise
2727

2828
open import Relation.Binary.PropositionalEquality as P using (_≡_; decSetoid)
29-
open import Relation.Nullary using (Dec)
29+
open import Relation.Nullary.Decidable using (Dec)
3030

3131
module Algebra.Solver.IdempotentCommutativeMonoid
3232
{m₁ m₂} (M : IdempotentCommutativeMonoid m₁ m₂) where

src/Algebra/Solver/Ring.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -40,7 +40,7 @@ open _-Raw-AlmostCommutative⟶_ morphism renaming (⟦_⟧ to ⟦_⟧′)
4040
open import Algebra.Properties.Semiring.Exp semiring
4141

4242
open import Relation.Binary
43-
open import Relation.Nullary using (yes; no)
43+
open import Relation.Nullary.Decidable using (yes; no)
4444
open import Relation.Binary.Reasoning.Setoid setoid
4545
import Relation.Binary.PropositionalEquality as PropEq
4646
import Relation.Binary.Reflection as Reflection

src/Data/Digit.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -19,7 +19,7 @@ open import Data.Product
1919
open import Data.Vec.Base as Vec using (Vec; _∷_; [])
2020
open import Data.Nat.DivMod
2121
open import Data.Nat.Induction
22-
open import Relation.Nullary using (does)
22+
open import Relation.Nullary.Decidable using (does)
2323
open import Relation.Nullary.Decidable
2424
open import Relation.Binary using (Decidable)
2525
open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)

src/Data/Fin/Base.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ open import Data.Product as Product using (_×_; _,_; proj₁; proj₂)
1818
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]′)
1919
open import Function.Base using (id; _∘_; _on_; flip)
2020
open import Level using (0ℓ)
21-
open import Relation.Nullary using (yes; no)
21+
open import Relation.Nullary.Decidable using (yes; no)
2222
open import Relation.Nullary.Negation using (contradiction)
2323
open import Relation.Nullary.Decidable.Core using (True; toWitness)
2424
open import Relation.Binary.Core

src/Data/Fin/Induction.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -29,7 +29,7 @@ import Relation.Binary.Construct.NonStrictToStrict as ToStrict
2929
import Relation.Binary.Construct.On as On
3030
open import Relation.Binary.Definitions using (Tri; tri<; tri≈; tri>)
3131
open import Relation.Binary.PropositionalEquality
32-
open import Relation.Nullary using (yes; no)
32+
open import Relation.Nullary.Decidable using (yes; no)
3333
open import Relation.Nullary.Negation using (contradiction)
3434
open import Relation.Unary using (Pred)
3535

src/Data/Fin/Properties.agda

+3-6
Original file line numberDiff line numberDiff line change
@@ -37,13 +37,10 @@ open import Level using (Level)
3737
open import Relation.Binary as B hiding (Decidable; _⇔_)
3838
open import Relation.Binary.PropositionalEquality as P
3939
using (_≡_; _≢_; refl; sym; trans; cong; cong₂; subst; _≗_; module ≡-Reasoning)
40-
open import Relation.Nullary.Decidable as Dec using (map′)
41-
open import Relation.Nullary.Reflects
42-
open import Relation.Nullary.Negation using (contradiction)
4340
open import Relation.Nullary
44-
using (Reflects; ofʸ; ofⁿ; Dec; _because_; does; proof; yes; no; ¬_)
45-
open import Relation.Nullary.Product using (_×-dec_)
46-
open import Relation.Nullary.Sum using (_⊎-dec_)
41+
using (Reflects; ofʸ; ofⁿ; Dec; _because_; does; proof; yes; no; ¬_; _×-dec_; _⊎-dec_; contradiction)
42+
open import Relation.Nullary.Reflects
43+
open import Relation.Nullary.Decidable as Dec using (map′)
4744
open import Relation.Unary as U
4845
using (U; Pred; Decidable; _⊆_; Satisfiable; Universal)
4946
open import Relation.Unary.Properties using (U?)

src/Data/Fin/Show.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Data.Nat as ℕ using (ℕ; _≤?_; _<?_)
1414
import Data.Nat.Show as ℕ
1515
open import Data.String.Base using (String)
1616
open import Function.Base
17-
open import Relation.Nullary using (yes; no)
17+
open import Relation.Nullary.Decidable using (yes; no)
1818
open import Relation.Nullary.Decidable using (True)
1919

2020
show : {n} Fin n String

src/Data/Fin/Subset/Properties.agda

+1-3
Original file line numberDiff line numberDiff line change
@@ -32,10 +32,8 @@ open import Function.Bundles using (_⇔_; mk⇔)
3232
open import Level using (Level)
3333
open import Relation.Binary as B hiding (Decidable; _⇔_)
3434
open import Relation.Binary.PropositionalEquality
35-
open import Relation.Nullary using (Dec; yes; no)
36-
import Relation.Nullary.Decidable as Dec
35+
open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _⊎-dec_)
3736
open import Relation.Nullary.Negation using (contradiction)
38-
open import Relation.Nullary.Sum using (_⊎-dec_)
3937
open import Relation.Unary using (Pred; Decidable; Satisfiable)
4038

4139
private

src/Data/Integer/Base.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -20,7 +20,7 @@ open import Level using (0ℓ)
2020
open import Relation.Binary.Core using (Rel)
2121
open import Relation.Binary.PropositionalEquality.Core
2222
using (_≡_; _≢_; refl)
23-
open import Relation.Nullary using (¬_)
23+
open import Relation.Nullary.Negation using (¬_)
2424
open import Relation.Nullary.Negation using (contradiction)
2525
open import Relation.Unary using (Pred)
2626

src/Data/Integer/Divisibility/Signed.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -24,7 +24,7 @@ open import Level
2424
open import Relation.Binary
2525
open import Relation.Binary.PropositionalEquality
2626
import Relation.Binary.Reasoning.Preorder as PreorderReasoning
27-
open import Relation.Nullary using (yes; no)
27+
open import Relation.Nullary.Decidable using (yes; no)
2828
import Relation.Nullary.Decidable as DEC
2929

3030
------------------------------------------------------------------------

src/Data/List/Base.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,7 @@ open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
2222
open import Data.These.Base as These using (These; this; that; these)
2323
open import Function.Base using (id; _∘_ ; _∘′_; _∘₂_; const; flip)
2424
open import Level using (Level)
25-
open import Relation.Nullary using (does)
25+
open import Relation.Nullary.Decidable using (does)
2626
open import Relation.Nullary.Decidable using (¬?)
2727
open import Relation.Unary using (Pred; Decidable)
2828
open import Relation.Unary.Properties using (∁?)

src/Data/List/Fresh/Relation/Unary/All.agda

+1-3
Original file line numberDiff line numberDiff line change
@@ -11,9 +11,7 @@ module Data.List.Fresh.Relation.Unary.All where
1111
open import Level using (Level; _⊔_; Lift)
1212
open import Data.Product using (_×_; _,_; proj₁; uncurry)
1313
open import Data.Sum.Base as Sum using (inj₁; inj₂)
14-
open import Relation.Nullary
15-
import Relation.Nullary.Decidable as Dec
16-
open import Relation.Nullary.Product using (_×-dec_)
14+
open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _×-dec_)
1715
open import Relation.Unary as U
1816
open import Relation.Binary as B using (Rel)
1917

src/Data/List/Fresh/Relation/Unary/Any.agda

+2-3
Original file line numberDiff line numberDiff line change
@@ -13,9 +13,8 @@ open import Data.Empty
1313
open import Data.Product using (∃; _,_; -,_)
1414
open import Data.Sum.Base using (_⊎_; [_,_]′; inj₁; inj₂)
1515
open import Function.Bundles using (_⇔_; mk⇔)
16-
open import Relation.Nullary
17-
import Relation.Nullary.Decidable as Dec
18-
open import Relation.Nullary.Sum using (_⊎-dec_)
16+
open import Relation.Nullary.Negation using (¬_)
17+
open import Relation.Nullary.Decidable as Dec using (Dec; yes; no; _⊎-dec_)
1918
open import Relation.Unary as U
2019
open import Relation.Binary as B using (Rel)
2120

src/Data/List/Membership/Setoid.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ open import Data.List.Relation.Unary.Any
1616
using (Any; index; map; here; there)
1717
open import Data.Product as Prod using (∃; _×_; _,_)
1818
open import Relation.Unary using (Pred)
19-
open import Relation.Nullary using (¬_)
19+
open import Relation.Nullary.Negation using (¬_)
2020

2121
open Setoid S renaming (Carrier to A)
2222

src/Data/List/NonEmpty/Base.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -23,7 +23,7 @@ open import Relation.Binary.PropositionalEquality.Core
2323
using (_≡_; _≢_; refl)
2424
open import Relation.Unary using (Pred; Decidable; U; ∅)
2525
open import Relation.Unary.Properties using (U?; ∅?)
26-
open import Relation.Nullary using (does)
26+
open import Relation.Nullary.Decidable using (does)
2727

2828
private
2929
variable

src/Data/List/Properties.agda

+2-4
Original file line numberDiff line numberDiff line change
@@ -36,10 +36,8 @@ import Relation.Binary.Reasoning.Setoid as EqR
3636
open import Relation.Binary.PropositionalEquality as P hiding ([_])
3737
open import Relation.Binary as B using (Rel)
3838
open import Relation.Nullary.Reflects using (invert)
39-
open import Relation.Nullary using (¬_; Dec; does; _because_; yes; no)
40-
open import Relation.Nullary.Negation using (contradiction)
41-
open import Relation.Nullary.Decidable as Decidable using (isYes; map′; ⌊_⌋; ¬?)
42-
open import Relation.Nullary.Product using (_×-dec_)
39+
open import Relation.Nullary using (¬_; Dec; does; _because_; yes; no; contradiction)
40+
open import Relation.Nullary.Decidable as Decidable using (isYes; map′; ⌊_⌋; ¬?; _×-dec_)
4341
open import Relation.Unary using (Pred; Decidable; ∁)
4442
open import Relation.Unary.Properties using (∁?)
4543

src/Data/List/Relation/Binary/Disjoint/Setoid.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Relation.Binary
1111
module Data.List.Relation.Binary.Disjoint.Setoid {c ℓ} (S : Setoid c ℓ) where
1212

1313
open import Level using (_⊔_)
14-
open import Relation.Nullary using (¬_)
14+
open import Relation.Nullary.Negation using (¬_)
1515
open import Function.Base using (_∘_)
1616
open import Data.List.Base using (List; []; [_]; _∷_)
1717
open import Data.List.Relation.Unary.Any using (here; there)

src/Data/List/Relation/Binary/Disjoint/Setoid/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ open import Data.List.Relation.Unary.Any.Properties using (++⁻)
1717
open import Data.Product using (_,_)
1818
open import Data.Sum.Base using (inj₁; inj₂)
1919
open import Relation.Binary
20-
open import Relation.Nullary using (¬_)
20+
open import Relation.Nullary.Negation using (¬_)
2121

2222
------------------------------------------------------------------------
2323
-- Relational properties

src/Data/List/Relation/Binary/Infix/Heterogeneous/Properties.agda

+2-4
Original file line numberDiff line numberDiff line change
@@ -17,10 +17,8 @@ import Data.Nat.Properties as ℕₚ
1717
open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_]′)
1818
open import Function.Base using (case_of_; _$′_)
1919

20-
open import Relation.Nullary using (¬_; yes; no; does)
21-
open import Relation.Nullary.Decidable using (map′)
22-
open import Relation.Nullary.Negation using (contradiction)
23-
open import Relation.Nullary.Sum using (_⊎-dec_)
20+
open import Relation.Nullary.Decidable using (yes; no; does; map′; _⊎-dec_)
21+
open import Relation.Nullary.Negation using (¬_; contradiction)
2422
open import Relation.Unary as U using (Pred)
2523
open import Relation.Binary using (REL; _⇒_; Decidable; Trans; Antisym)
2624
open import Relation.Binary.PropositionalEquality using (_≢_; refl; cong)

src/Data/List/Relation/Binary/Lex.agda

+3-4
Original file line numberDiff line numberDiff line change
@@ -16,10 +16,9 @@ open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_])
1616
open import Function.Base using (_∘_; flip; id)
1717
open import Function.Bundles using (_⇔_; mk⇔)
1818
open import Level using (_⊔_)
19-
open import Relation.Nullary using (Dec; yes; no; ¬_)
20-
import Relation.Nullary.Decidable as Dec
21-
open import Relation.Nullary.Product using (_×-dec_)
22-
open import Relation.Nullary.Sum using (_⊎-dec_)
19+
open import Relation.Nullary.Negation using (¬_)
20+
open import Relation.Nullary.Decidable as Dec
21+
using (Dec; yes; no; _×-dec_; _⊎-dec_)
2322
open import Relation.Binary hiding (_⇔_)
2423
open import Data.List.Relation.Binary.Pointwise.Base
2524
using (Pointwise; []; _∷_; head; tail)

src/Data/List/Relation/Binary/Lex/Core.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Data.Product using (_×_; _,_; proj₁; proj₂; uncurry)
1414
open import Data.List.Base using (List; []; _∷_)
1515
open import Function.Base using (_∘_; flip; id)
1616
open import Level using (Level; _⊔_)
17-
open import Relation.Nullary using (¬_)
17+
open import Relation.Nullary.Negation using (¬_)
1818
open import Relation.Binary.Core using (Rel)
1919
open import Data.List.Relation.Binary.Pointwise.Base
2020
using (Pointwise; []; _∷_; head; tail)

src/Data/List/Relation/Binary/Permutation/Setoid/Properties.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -38,7 +38,7 @@ open import Relation.Unary using (Pred; Decidable)
3838
open import Relation.Binary.Properties.Setoid S using (≉-resp₂)
3939
open import Relation.Binary.PropositionalEquality as ≡
4040
using (_≡_ ; refl; sym; cong; cong₂; subst; _≢_; inspect)
41-
open import Relation.Nullary using (yes; no; does)
41+
open import Relation.Nullary.Decidable using (yes; no; does)
4242
open import Relation.Nullary.Negation using (contradiction)
4343

4444
private

src/Data/List/Relation/Binary/Pointwise.agda

-2
Original file line numberDiff line numberDiff line change
@@ -23,9 +23,7 @@ open import Data.Nat.Base using (ℕ; zero; suc)
2323
open import Data.Nat.Properties
2424
open import Level
2525
open import Relation.Nullary hiding (Irrelevant)
26-
open import Relation.Nullary.Negation.Core using (contradiction)
2726
import Relation.Nullary.Decidable as Dec using (map′)
28-
open import Relation.Nullary.Product using (_×-dec_)
2927
open import Relation.Unary as U using (Pred)
3028
open import Relation.Binary renaming (Rel to Rel₂)
3129
open import Relation.Binary.PropositionalEquality as P using (_≡_)

0 commit comments

Comments
 (0)