File tree Expand file tree Collapse file tree 3 files changed +18
-2
lines changed Expand file tree Collapse file tree 3 files changed +18
-2
lines changed Original file line number Diff line number Diff line change @@ -14,7 +14,7 @@ module Data.Integer.Base where
1414open import Algebra.Bundles.Raw
1515 using (RawMagma; RawMonoid; RawGroup; RawNearSemiring; RawSemiring; RawRing)
1616open 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 ℕ )
1818open import Data.Sign.Base as Sign using (Sign)
1919open import Level using (0ℓ)
2020open 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+
143146instance
144147 pos : ∀ {n} → Positive +[1+ n ]
145148 pos = _
Original file line number Diff line number Diff line change @@ -10,7 +10,9 @@ module Data.Rational.Base where
1010
1111open import Algebra.Bundles.Raw
1212open 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 ℤ )
1416open import Data.Nat.GCD
1517open 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)
176178NonNegative : Pred ℚ 0ℓ
177179NonNegative 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
Original file line number Diff line number Diff line change @@ -12,6 +12,7 @@ open import Algebra.Bundles.Raw
1212open import Data.Bool.Base using (Bool; true; false; if_then_else_)
1313open import Data.Integer.Base as ℤ
1414 using (ℤ; +_; +0; +[1+_]; -[1+_]; +<+; +≤+)
15+ hiding (module ℤ )
1516open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
1617open import Level using (0ℓ)
1718open import Relation.Nullary.Negation.Core using (¬_; contradiction)
@@ -149,6 +150,11 @@ NonPositive p = ℤ.NonPositive (↥ p)
149150NonNegative : Pred ℚᵘ 0ℓ
150151NonNegative 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
You can’t perform that action at this time.
0 commit comments