Skip to content

Commit 0e3105d

Browse files
gallaisMatthewDaggitt
authored andcommitted
[ re #1993 ] Simplifying the dependency graph (#1995)
1 parent ec2f58c commit 0e3105d

File tree

107 files changed

+243
-268
lines changed

Some content is hidden

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

107 files changed

+243
-268
lines changed

src/Algebra/Consequences/Setoid.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,7 @@ open Setoid S renaming (Carrier to A)
1515
open import Algebra.Core
1616
open import Algebra.Definitions _≈_
1717
open import Data.Sum.Base using (inj₁; inj₂)
18-
open import Data.Product using (_,_)
18+
open import Data.Product.Base using (_,_)
1919
open import Function.Base using (_$_)
2020
import Function.Definitions as FunDefs
2121
import Relation.Binary.Consequences as Bin

src/Algebra/Construct/LexProduct/Base.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -8,10 +8,10 @@
88

99
open import Algebra.Core using (Op₂)
1010
open import Data.Bool.Base using (true; false)
11-
open import Data.Product using (_×_; _,_)
11+
open import Data.Product.Base using (_×_; _,_)
1212
open import Relation.Binary.Core using (Rel)
1313
open import Relation.Binary.Definitions using (Decidable)
14-
open import Relation.Nullary.Decidable using (does; yes; no)
14+
open import Relation.Nullary.Decidable.Core using (does; yes; no)
1515

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

src/Algebra/Construct/NaturalChoice/MinMaxOp.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Algebra.Core
1111
open import Algebra.Bundles
1212
open import Algebra.Construct.NaturalChoice.Base
1313
open import Data.Sum.Base as Sum using (inj₁; inj₂; [_,_])
14-
open import Data.Product using (_,_)
14+
open import Data.Product.Base using (_,_)
1515
open import Function.Base using (id; _∘_; flip)
1616
open import Relation.Binary
1717
open import Relation.Binary.Consequences

src/Algebra/Construct/NaturalChoice/MinOp.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ open import Algebra.Core
1111
open import Algebra.Bundles
1212
open import Algebra.Construct.NaturalChoice.Base
1313
open import Data.Sum.Base as Sum using (inj₁; inj₂; [_,_])
14-
open import Data.Product using (_,_)
14+
open import Data.Product.Base using (_,_)
1515
open import Function.Base using (id; _∘_)
1616
open import Relation.Binary
1717
open import Relation.Binary.Consequences

src/Algebra/Definitions.agda

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -15,17 +15,17 @@
1515

1616
{-# OPTIONS --cubical-compatible --safe #-}
1717

18-
open import Relation.Binary.Core
19-
open import Relation.Nullary.Negation using (¬_)
18+
open import Relation.Binary.Core using (Rel; _Preserves_⟶_; _Preserves₂_⟶_⟶_)
19+
open import Relation.Nullary.Negation.Core using (¬_)
2020

2121
module Algebra.Definitions
2222
{a ℓ} {A : Set a} -- The underlying set
2323
(_≈_ : Rel A ℓ) -- The underlying equality
2424
where
2525

26-
open import Algebra.Core
27-
open import Data.Product
28-
open import Data.Sum.Base
26+
open import Algebra.Core using (Op₁; Op₂)
27+
open import Data.Product.Base using (_×_; ∃-syntax)
28+
open import Data.Sum.Base using (_⊎_)
2929

3030
------------------------------------------------------------------------
3131
-- Properties of operations

src/Algebra/Definitions/RawMagma.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@
1111
{-# OPTIONS --cubical-compatible --safe #-}
1212

1313
open import Algebra.Bundles using (RawMagma)
14-
open import Data.Product using (_×_; ∃)
14+
open import Data.Product.Base using (_×_; ∃)
1515
open import Level using (_⊔_)
1616
open import Relation.Binary.Core
1717
open import Relation.Nullary.Negation using (¬_)

src/Algebra/Lattice/Properties/BooleanAlgebra.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -23,9 +23,9 @@ open import Algebra.Bundles
2323
open import Algebra.Lattice.Structures _≈_
2424
open import Relation.Binary.Reasoning.Setoid setoid
2525
open import Relation.Binary
26-
open import Function.Base
26+
open import Function.Base using (id; _$_; _⟨_⟩_)
2727
open import Function.Bundles using (_⇔_; module Equivalence)
28-
open import Data.Product using (_,_)
28+
open import Data.Product.Base using (_,_)
2929

3030
------------------------------------------------------------------------
3131
-- Export properties from distributive lattices

src/Algebra/Lattice/Properties/Lattice.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -11,7 +11,7 @@ import Algebra.Lattice.Properties.Semilattice as SemilatticeProperties
1111
open import Relation.Binary
1212
import Relation.Binary.Lattice as R
1313
open import Function.Base
14-
open import Data.Product using (_,_; swap)
14+
open import Data.Product.Base using (_,_; swap)
1515

1616
module Algebra.Lattice.Properties.Lattice
1717
{l₁ l₂} (L : Lattice l₁ l₂) where

src/Algebra/Lattice/Properties/Semilattice.agda

Lines changed: 2 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -6,11 +6,8 @@
66

77
{-# OPTIONS --cubical-compatible --safe #-}
88

9-
open import Algebra.Lattice
10-
open import Algebra.Structures
11-
open import Function
12-
open import Data.Product
13-
open import Relation.Binary
9+
open import Algebra.Lattice.Bundles using (Semilattice)
10+
open import Relation.Binary.Bundles using (Poset)
1411
import Relation.Binary.Lattice as B
1512
import Relation.Binary.Properties.Poset as PosetProperties
1613

src/Algebra/Lattice/Structures.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@
1414
{-# OPTIONS --cubical-compatible --safe #-}
1515

1616
open import Algebra.Core
17-
open import Data.Product using (proj₁; proj₂)
17+
open import Data.Product.Base using (proj₁; proj₂)
1818
open import Level using (_⊔_)
1919
open import Relation.Binary using (Rel; Setoid; IsEquivalence)
2020

src/Algebra/Properties/CommutativeSemigroup.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ open CommutativeSemigroup CS
1616

1717
open import Algebra.Definitions _≈_
1818
open import Relation.Binary.Reasoning.Setoid setoid
19-
open import Data.Product
19+
open import Data.Product.Base using (_,_)
2020

2121
------------------------------------------------------------------------------
2222
-- Re-export the contents of semigroup

src/Algebra/Properties/Group.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -13,8 +13,8 @@ module Algebra.Properties.Group {g₁ g₂} (G : Group g₁ g₂) where
1313
open Group G
1414
open import Algebra.Definitions _≈_
1515
open import Relation.Binary.Reasoning.Setoid setoid
16-
open import Function
17-
open import Data.Product
16+
open import Function.Base using (_$_; _⟨_⟩_)
17+
open import Data.Product.Base using (_,_; proj₂)
1818

1919
ε⁻¹≈ε : ε ⁻¹ ≈ ε
2020
ε⁻¹≈ε = begin

src/Algebra/Properties/Ring.agda

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,6 @@ import Algebra.Properties.AbelianGroup as AbelianGroupProperties
1616
open import Function.Base using (_$_)
1717
open import Relation.Binary.Reasoning.Setoid setoid
1818
open import Algebra.Definitions _≈_
19-
open import Data.Product
2019

2120
------------------------------------------------------------------------
2221
-- Export properties of abelian groups

src/Algebra/Properties/Semigroup.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -12,7 +12,7 @@ module Algebra.Properties.Semigroup {a ℓ} (S : Semigroup a ℓ) where
1212

1313
open Semigroup S
1414
open import Algebra.Definitions _≈_
15-
open import Data.Product
15+
open import Data.Product.Base using (_,_)
1616

1717
x∙yz≈xy∙z : x y z x ∙ (y ∙ z) ≈ (x ∙ y) ∙ z
1818
x∙yz≈xy∙z x y z = sym (assoc x y z)

src/Algebra/Structures.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -22,7 +22,7 @@ module Algebra.Structures
2222
open import Algebra.Core
2323
open import Algebra.Definitions _≈_
2424
import Algebra.Consequences.Setoid as Consequences
25-
open import Data.Product using (_,_; proj₁; proj₂)
25+
open import Data.Product.Base using (_,_; proj₁; proj₂)
2626
open import Level using (_⊔_)
2727

2828
------------------------------------------------------------------------

src/Algebra/Structures/Biased.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@
99

1010
open import Algebra.Core
1111
open import Algebra.Consequences.Setoid
12-
open import Data.Product using (_,_; proj₁; proj₂)
12+
open import Data.Product.Base using (_,_; proj₁; proj₂)
1313
open import Level using (_⊔_)
1414
open import Relation.Binary using (Rel; Setoid; IsEquivalence)
1515

src/Codata/Sized/Cowriter.agda

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -16,9 +16,9 @@ open import Codata.Sized.Delay using (Delay; later; now)
1616
open import Codata.Sized.Stream as Stream using (Stream; _∷_)
1717
open import Data.Unit.Base
1818
open import Data.List.Base using (List; []; _∷_)
19-
open import Data.List.NonEmpty using (List⁺; _∷_)
19+
open import Data.List.NonEmpty.Base using (List⁺; _∷_)
2020
open import Data.Nat.Base as Nat using (ℕ; zero; suc)
21-
open import Data.Product as Prod using (_×_; _,_)
21+
open import Data.Product.Base as Prod using (_×_; _,_)
2222
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
2323
open import Data.Vec.Base using (Vec; []; _∷_)
2424
open import Data.Vec.Bounded.Base as Vec≤ using (Vec≤; _,_)

src/Codata/Sized/Stream.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -15,10 +15,10 @@ open import Data.Nat.Base
1515
open import Data.List.Base using (List; []; _∷_)
1616
open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_; _∷⁺_)
1717
open import Data.Vec.Base using (Vec; []; _∷_)
18-
open import Data.Product as P hiding (map)
19-
open import Function.Base
18+
open import Data.Product.Base as P using (Σ; _×_; _,_; <_,_>; proj₁; proj₂)
19+
open import Function.Base using (id; _∘_)
2020
open import Level using (Level)
21-
open import Relation.Binary.PropositionalEquality using (_≡_; refl)
21+
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
2222

2323
private
2424
variable

src/Data/Bool/Properties.agda

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -13,18 +13,18 @@ open import Algebra.Lattice.Bundles
1313
import Algebra.Lattice.Properties.BooleanAlgebra as BooleanAlgebraProperties
1414
open import Data.Bool.Base
1515
open import Data.Empty
16-
open import Data.Product
17-
open import Data.Sum.Base
18-
open import Function.Base
16+
open import Data.Product.Base using (_×_; _,_; proj₁; proj₂)
17+
open import Data.Sum.Base using (_⊎_; inj₁; inj₂; [_,_])
18+
open import Function.Base using (_⟨_⟩_; const; id)
1919
open import Function.Equality using (_⟨$⟩_)
2020
open import Function.Equivalence
2121
using (_⇔_; equivalence; module Equivalence)
2222
open import Induction.WellFounded using (WellFounded; Acc; acc)
2323
open import Level using (Level; 0ℓ)
2424
open import Relation.Binary hiding (_⇔_)
2525
open import Relation.Binary.PropositionalEquality hiding ([_])
26-
open import Relation.Nullary using (ofʸ; ofⁿ; does; proof; yes; no)
27-
open import Relation.Nullary.Decidable using (True)
26+
open import Relation.Nullary.Reflects using (ofʸ; ofⁿ)
27+
open import Relation.Nullary.Decidable.Core using (True; does; proof; yes; no)
2828
import Relation.Unary as U
2929

3030
open import Algebra.Definitions {A = Bool} _≡_

src/Data/Digit.agda

Lines changed: 7 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -9,21 +9,20 @@
99
module Data.Digit where
1010

1111
open import Data.Nat.Base
12-
open import Data.Nat.Properties
13-
open import Data.Nat.Solver
12+
open import Data.Nat.Properties using (_≤?_; _<?_; ≤⇒≤′; module ≤-Reasoning; m≤m+n)
13+
open import Data.Nat.Solver using (module +-*-Solver)
1414
open import Data.Fin.Base as Fin using (Fin; zero; suc; toℕ)
1515
open import Data.Bool.Base using (Bool; true; false)
16-
open import Data.Char using (Char)
16+
open import Data.Char.Base using (Char)
1717
open import Data.List.Base
18-
open import Data.Product
18+
open import Data.Product.Base using (∃; _,_)
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.Decidable using (does)
23-
open import Relation.Nullary.Decidable
24-
open import Relation.Binary using (Decidable)
22+
open import Relation.Nullary.Decidable using (True; does; toWitness)
23+
open import Relation.Binary.Definitions using (Decidable)
2524
open import Relation.Binary.PropositionalEquality as P using (_≡_; refl)
26-
open import Function
25+
open import Function.Base using (_$_)
2726

2827
------------------------------------------------------------------------
2928
-- Digits

src/Data/Fin/Base.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ module Data.Fin.Base where
1414
open import Data.Bool.Base using (Bool; true; false; T; not)
1515
open import Data.Empty using (⊥-elim)
1616
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; z≤n; s≤s; z<s; s<s; _^_)
17-
open import Data.Product as Product using (_×_; _,_; proj₁; proj₂)
17+
open import Data.Product.Base 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ℓ)

src/Data/Fin/Properties.agda

Lines changed: 5 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -18,11 +18,13 @@ open import Data.Bool.Base using (Bool; true; false; not; _∧_; _∨_)
1818
open import Data.Empty using (⊥; ⊥-elim)
1919
open import Data.Fin.Base
2020
open import Data.Fin.Patterns
21-
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; s≤s; z≤n; z<s; s<s; _∸_; _^_)
21+
open import Data.Nat.Base as ℕ
22+
using (ℕ; zero; suc; s≤s; z≤n; z<s; s<s; _∸_; _^_)
2223
import Data.Nat.Properties as ℕₚ
2324
open import Data.Nat.Solver
2425
open import Data.Unit using (⊤; tt)
25-
open import Data.Product using (Σ-syntax; ∃; ∃₂; ∄; _×_; _,_; map; proj₁; proj₂; uncurry; <_,_>)
26+
open import Data.Product.Base as Prod
27+
using (∃; ∃₂; _×_; _,_; map; proj₁; proj₂; uncurry; <_,_>)
2628
open import Data.Product.Properties using (,-injective)
2729
open import Data.Product.Algebra using (×-cong)
2830
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂; [_,_]; [_,_]′)
@@ -619,7 +621,7 @@ splitAt-≥ (suc m) (suc i) (s≤s i≥m) = cong (Sum.map suc id) (splitAt-≥ m
619621
remQuot-combine : {n k} (i : Fin n) j remQuot k (combine i j) ≡ (i , j)
620622
remQuot-combine {suc n} {k} zero j rewrite splitAt-↑ˡ k j (n ℕ.* k) = refl
621623
remQuot-combine {suc n} {k} (suc i) j rewrite splitAt-↑ʳ k (n ℕ.* k) (combine i j) =
622-
cong (Data.Product.map₁ suc) (remQuot-combine i j)
624+
cong (Prod.map₁ suc) (remQuot-combine i j)
623625

624626
combine-remQuot : {n} k (i : Fin (n ℕ.* k)) uncurry combine (remQuot {n} k i) ≡ i
625627
combine-remQuot {suc n} k i with splitAt k i in eq
@@ -1158,4 +1160,3 @@ Please use <⇒<′ instead."
11581160
"Warning: <′⇒≺ was deprecated in v2.0.
11591161
Please use <′⇒< instead."
11601162
#-}
1161-

src/Data/List/Base.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -17,7 +17,7 @@ open import Data.Bool.Base as Bool
1717
open import Data.Fin.Base using (Fin; zero; suc)
1818
open import Data.Maybe.Base as Maybe using (Maybe; nothing; just; maybe′)
1919
open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _+_; _*_ ; _≤_ ; s≤s)
20-
open import Data.Product as Prod using (_×_; _,_; map₁; map₂′)
20+
open import Data.Product.Base as Prod using (_×_; _,_; map₁; map₂′)
2121
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)

src/Data/List/Fresh.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -16,7 +16,7 @@ module Data.List.Fresh where
1616
open import Level using (Level; _⊔_)
1717
open import Data.Bool.Base using (true; false)
1818
open import Data.Unit.Polymorphic.Base using (⊤)
19-
open import Data.Product using (∃; _×_; _,_; -,_; proj₁; proj₂)
19+
open import Data.Product.Base using (∃; _×_; _,_; -,_; proj₁; proj₂)
2020
open import Data.List.Relation.Unary.All using (All; []; _∷_)
2121
open import Data.List.Relation.Unary.AllPairs using (AllPairs; []; _∷_)
2222
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)

src/Data/List/Fresh/NonEmpty.agda

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -11,9 +11,9 @@ module Data.List.Fresh.NonEmpty where
1111
open import Level using (Level; _⊔_)
1212
open import Data.List.Fresh as List# using (List#; []; cons; fresh)
1313
open import Data.Maybe.Base using (Maybe; nothing; just)
14-
open import Data.Nat using (ℕ; suc)
15-
open import Data.Product using (_×_; _,_)
16-
open import Relation.Binary as B using (Rel)
14+
open import Data.Nat.Base using (ℕ; suc)
15+
open import Data.Product.Base using (_×_; _,_)
16+
open import Relation.Binary.Core using (Rel)
1717

1818
private
1919
variable

src/Data/List/Kleene/Base.agda

Lines changed: 6 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -8,11 +8,12 @@
88

99
module Data.List.Kleene.Base where
1010

11-
open import Data.Product as Product using (_×_; _,_; map₂; map₁; proj₁; proj₂)
12-
open import Data.Nat as ℕ using (ℕ; suc; zero)
13-
open import Data.Maybe as Maybe using (Maybe; just; nothing)
14-
open import Data.Sum as Sum using (_⊎_; inj₁; inj₂)
15-
open import Level as Level using (Level)
11+
open import Data.Product.Base as Product
12+
using (_×_; _,_; map₂; map₁; proj₁; proj₂)
13+
open import Data.Nat.Base as ℕ using (ℕ; suc; zero)
14+
open import Data.Maybe.Base as Maybe using (Maybe; just; nothing)
15+
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
16+
open import Level using (Level)
1617

1718
open import Algebra.Core using (Op₂)
1819
open import Function.Base

src/Data/List/Membership/Setoid.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Function.Base using (_∘_; id; flip)
1414
open import Data.List.Base as List using (List; []; _∷_; length; lookup)
1515
open import Data.List.Relation.Unary.Any
1616
using (Any; index; map; here; there)
17-
open import Data.Product as Prod using (∃; _×_; _,_)
17+
open import Data.Product.Base as Prod using (∃; _×_; _,_)
1818
open import Relation.Unary using (Pred)
1919
open import Relation.Nullary.Negation using (¬_)
2020

src/Data/List/NonEmpty/Base.agda

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ open import Data.Bool.Properties using (T?)
1414
open import Data.List.Base as List using (List; []; _∷_)
1515
open import Data.Maybe.Base using (Maybe ; nothing; just)
1616
open import Data.Nat.Base as ℕ
17-
open import Data.Product as Prod using (∃; _×_; proj₁; proj₂; _,_; -,_)
17+
open import Data.Product.Base as Prod using (∃; _×_; proj₁; proj₂; _,_; -,_)
1818
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
1919
open import Data.These.Base as These using (These; this; that; these)
2020
open import Data.Vec.Base as Vec using (Vec; []; _∷_)

src/Data/List/Properties.agda

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -25,7 +25,8 @@ open import Data.Maybe.Base using (Maybe; just; nothing)
2525
open import Data.Nat.Base
2626
open import Data.Nat.Divisibility
2727
open import Data.Nat.Properties
28-
open import Data.Product as Prod hiding (map; zip)
28+
open import Data.Product.Base as Prod
29+
using (_×_; _,_; uncurry; uncurry′; proj₁; proj₂; <_,_>)
2930
import Data.Product.Relation.Unary.All as Prod using (All)
3031
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
3132
open import Data.These.Base as These using (These; this; that; these)

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

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@ module Data.List.Relation.Binary.Lex.Core where
1010

1111
open import Data.Empty using (⊥; ⊥-elim)
1212
open import Data.Unit.Base using (⊤; tt)
13-
open import Data.Product using (_×_; _,_; proj₁; proj₂; uncurry)
13+
open import Data.Product.Base 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; _⊔_)

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

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -8,9 +8,9 @@
88

99
module Data.List.Relation.Binary.Pointwise.Base where
1010

11-
open import Data.Product using (_×_; <_,_>)
11+
open import Data.Product.Base using (_×_; <_,_>)
1212
open import Data.List.Base using (List; []; _∷_)
13-
open import Level
13+
open import Level using (Level; _⊔_)
1414
open import Relation.Binary.Core using (REL; _⇒_)
1515

1616
private

0 commit comments

Comments
 (0)