File tree 3 files changed +18
-2
lines changed
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
14
14
open import Algebra.Bundles.Raw
15
15
using (RawMagma; RawMonoid; RawGroup; RawNearSemiring; RawSemiring; RawRing)
16
16
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 ℕ )
18
18
open import Data.Sign.Base as Sign using (Sign)
19
19
open import Level using (0ℓ)
20
20
open import Relation.Binary.Core using (Rel)
@@ -140,6 +140,9 @@ record Negative (i : ℤ) : Set where
140
140
141
141
-- Instances
142
142
143
+ open ℕ public
144
+ using (nonZero)
145
+
143
146
instance
144
147
pos : ∀ {n} → Positive +[1+ n ]
145
148
pos = _
Original file line number Diff line number Diff line change @@ -10,7 +10,9 @@ module Data.Rational.Base where
10
10
11
11
open import Algebra.Bundles.Raw
12
12
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 ℤ )
14
16
open import Data.Nat.GCD
15
17
open import Data.Nat.Coprimality as C
16
18
using (Coprime; Bézout-coprime; coprime-/gcd; coprime?; ¬0-coprimeTo-2+)
@@ -176,6 +178,11 @@ NonPositive p = ℚᵘ.NonPositive (toℚᵘ p)
176
178
NonNegative : Pred ℚ 0ℓ
177
179
NonNegative p = ℚᵘ.NonNegative (toℚᵘ p)
178
180
181
+ -- Instances
182
+
183
+ open ℤ public
184
+ using (nonZero; pos; nonNeg; nonPos0; nonPos; neg)
185
+
179
186
-- Constructors
180
187
181
188
≢-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
12
12
open import Data.Bool.Base using (Bool; true; false; if_then_else_)
13
13
open import Data.Integer.Base as ℤ
14
14
using (ℤ; +_; +0; +[1+_]; -[1+_]; +<+; +≤+)
15
+ hiding (module ℤ )
15
16
open import Data.Nat.Base as ℕ using (ℕ; zero; suc)
16
17
open import Level using (0ℓ)
17
18
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
@@ -149,6 +150,11 @@ NonPositive p = ℤ.NonPositive (↥ p)
149
150
NonNegative : Pred ℚᵘ 0ℓ
150
151
NonNegative p = ℤ.NonNegative (↥ p)
151
152
153
+ -- Instances
154
+
155
+ open ℤ public
156
+ using (nonZero; pos; nonNeg; nonPos0; nonPos; neg)
157
+
152
158
-- Constructors and destructors
153
159
154
160
-- Note: these could be proved more elegantly using the constructors
You can’t perform that action at this time.
0 commit comments