Skip to content

Commit 16aadff

Browse files
Re-export numeric subclass instances (#2122)
1 parent af593fc commit 16aadff

File tree

3 files changed

+18
-2
lines changed

3 files changed

+18
-2
lines changed

src/Data/Integer/Base.agda

+4-1
Original file line numberDiff line numberDiff line change
@@ -14,7 +14,7 @@ module Data.Integer.Base where
1414
open import Algebra.Bundles.Raw
1515
using (RawMagma; RawMonoid; RawGroup; RawNearSemiring; RawSemiring; RawRing)
1616
open import Data.Bool.Base using (Bool; T; true; false)
17-
open import Data.Nat.Base as ℕ using (ℕ; z≤n; s≤s)
17+
open import Data.Nat.Base as ℕ using (ℕ; z≤n; s≤s) hiding (module )
1818
open import Data.Sign.Base as Sign using (Sign)
1919
open import Level using (0ℓ)
2020
open import Relation.Binary.Core using (Rel)
@@ -140,6 +140,9 @@ record Negative (i : ℤ) : Set where
140140

141141
-- Instances
142142

143+
open public
144+
using (nonZero)
145+
143146
instance
144147
pos : {n} Positive +[1+ n ]
145148
pos = _

src/Data/Rational/Base.agda

+8-1
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,9 @@ module Data.Rational.Base where
1010

1111
open import Algebra.Bundles.Raw
1212
open import Data.Bool.Base using (Bool; true; false; if_then_else_)
13-
open import Data.Integer.Base as ℤ using (ℤ; +_; +0; +[1+_]; -[1+_])
13+
open import Data.Integer.Base as ℤ
14+
using (ℤ; +_; +0; +[1+_]; -[1+_])
15+
hiding (module )
1416
open import Data.Nat.GCD
1517
open import Data.Nat.Coprimality as C
1618
using (Coprime; Bézout-coprime; coprime-/gcd; coprime?; ¬0-coprimeTo-2+)
@@ -176,6 +178,11 @@ NonPositive p = ℚᵘ.NonPositive (toℚᵘ p)
176178
NonNegative : Pred ℚ 0ℓ
177179
NonNegative p = ℚᵘ.NonNegative (toℚᵘ p)
178180

181+
-- Instances
182+
183+
open public
184+
using (nonZero; pos; nonNeg; nonPos0; nonPos; neg)
185+
179186
-- Constructors
180187

181188
≢-nonZero : {p} p ≢ 0ℚ NonZero p

src/Data/Rational/Unnormalised/Base.agda

+6
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ open import Algebra.Bundles.Raw
1212
open import Data.Bool.Base using (Bool; true; false; if_then_else_)
1313
open import Data.Integer.Base as ℤ
1414
using (ℤ; +_; +0; +[1+_]; -[1+_]; +<+; +≤+)
15+
hiding (module )
1516
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1617
open import Level using (0ℓ)
1718
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
@@ -149,6 +150,11 @@ NonPositive p = ℤ.NonPositive (↥ p)
149150
NonNegative : Pred ℚᵘ 0ℓ
150151
NonNegative p = ℤ.NonNegative (↥ p)
151152

153+
-- Instances
154+
155+
open public
156+
using (nonZero; pos; nonNeg; nonPos0; nonPos; neg)
157+
152158
-- Constructors and destructors
153159

154160
-- Note: these could be proved more elegantly using the constructors

0 commit comments

Comments
 (0)