Skip to content

Commit 9897c20

Browse files
authored
[ refactor ] Algebra.Properties.Semiring.Divisibility and friends (#2579)
* refactor: simplify `import`s * fix: remove redundant parentheses * fix: revert `using` directive
1 parent dfa9ed6 commit 9897c20

File tree

5 files changed

+31
-25
lines changed

5 files changed

+31
-25
lines changed

src/Algebra/Properties/Magma/Divisibility.agda

+1-1
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@
66

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

9-
open import Algebra using (Magma)
9+
open import Algebra.Bundles using (Magma)
1010

1111
module Algebra.Properties.Magma.Divisibility {a ℓ} (M : Magma a ℓ) where
1212

src/Algebra/Properties/Monoid/Divisibility.agda

+8-7
Original file line numberDiff line numberDiff line change
@@ -6,16 +6,17 @@
66

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

9-
open import Algebra using (Monoid)
9+
open import Algebra.Bundles using (Monoid)
10+
11+
module Algebra.Properties.Monoid.Divisibility
12+
{a ℓ} (M : Monoid a ℓ) where
13+
1014
open import Data.Product.Base using (_,_)
1115
open import Relation.Binary.Core using (_⇒_)
1216
open import Relation.Binary.Bundles using (Preorder)
1317
open import Relation.Binary.Structures using (IsPreorder; IsEquivalence)
1418
open import Relation.Binary.Definitions using (Reflexive)
1519

16-
module Algebra.Properties.Monoid.Divisibility
17-
{a ℓ} (M : Monoid a ℓ) where
18-
1920
open Monoid M
2021

2122
------------------------------------------------------------------------
@@ -60,9 +61,9 @@ infix 4 ε∣_
6061

6162
∥-isEquivalence : IsEquivalence _∥_
6263
∥-isEquivalence = record
63-
{ refl = λ {x} ∥-refl {x}
64-
; sym = λ {x} {y} ∥-sym {x} {y}
65-
; trans = λ {x} {y} {z} ∥-trans {x} {y} {z}
64+
{ refl = ∥-refl
65+
; sym = ∥-sym
66+
; trans = ∥-trans
6667
}
6768

6869

src/Algebra/Properties/Semigroup/Divisibility.agda

+4-3
Original file line numberDiff line numberDiff line change
@@ -6,13 +6,14 @@
66

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

9-
open import Algebra using (Semigroup)
10-
open import Data.Product.Base using (_,_)
11-
open import Relation.Binary.Definitions using (Transitive)
9+
open import Algebra.Bundles using (Semigroup)
1210

1311
module Algebra.Properties.Semigroup.Divisibility
1412
{a ℓ} (S : Semigroup a ℓ) where
1513

14+
open import Data.Product.Base using (_,_)
15+
open import Relation.Binary.Definitions using (Transitive)
16+
1617
open Semigroup S
1718

1819
------------------------------------------------------------------------

src/Algebra/Properties/Semiring/Divisibility.agda

+3-6
Original file line numberDiff line numberDiff line change
@@ -6,10 +6,7 @@
66

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

9-
open import Algebra using (Semiring)
10-
import Algebra.Properties.Monoid.Divisibility as MonoidDivisibility
11-
open import Data.Product.Base using (_,_)
12-
open import Data.Sum.Base using (_⊎_; inj₁; inj₂)
9+
open import Algebra.Bundles using (Semiring)
1310

1411
module Algebra.Properties.Semiring.Divisibility
1512
{a ℓ} (R : Semiring a ℓ) where
@@ -19,7 +16,7 @@ open Semiring R
1916
------------------------------------------------------------------------
2017
-- Re-exporting divisibility over monoids
2118

22-
open MonoidDivisibility *-monoid public
19+
open import Algebra.Properties.Monoid.Divisibility *-monoid public
2320
renaming (ε∣_ to 1∣_)
2421

2522
------------------------------------------------------------------------
@@ -37,4 +34,4 @@ x∣y∧y≉0⇒x≉0 : ∀ {x y} → x ∣ y → y ≉ 0# → x ≉ 0#
3734
x∣y∧y≉0⇒x≉0 x∣y y≉0 x≈0 = y≉0 (0∣x⇒x≈0 (∣-respˡ x≈0 x∣y))
3835

3936
0∤1 : 0# ≉ 1# 0# ∤ 1#
40-
0∤1 0≉1 (q , q*0≈1) = 0≉1 (trans (sym (zeroʳ q)) q*0≈1)
37+
0∤1 0≉1 0∣1 = 0≉1 (sym (0∣x⇒x≈0 0∣1))
Original file line numberDiff line numberDiff line change
@@ -1,22 +1,29 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- Some theory for CancellativeCommutativeSemiring.
4+
-- Properties of Coprimality and Irreducibility for Semiring.
55
------------------------------------------------------------------------
66

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

9-
open import Algebra using (Semiring)
10-
open import Data.Sum.Base using (reduce)
11-
open import Function.Base using (flip)
12-
open import Relation.Binary.Definitions using (Symmetric)
9+
open import Algebra.Bundles using (Semiring)
1310

1411
module Algebra.Properties.Semiring.Primality
1512
{a ℓ} (R : Semiring a ℓ)
1613
where
1714

15+
open import Data.Sum.Base using (reduce)
16+
open import Function.Base using (flip)
17+
open import Relation.Binary.Definitions using (Symmetric)
18+
1819
open Semiring R renaming (Carrier to A)
1920
open import Algebra.Properties.Semiring.Divisibility R
21+
using (_∣_; ∣-trans; 0∤1)
22+
23+
private
24+
variable
25+
x p : A
26+
2027

2128
------------------------------------------------------------------------
2229
-- Re-export primality definitions
@@ -30,12 +37,12 @@ open import Algebra.Definitions.RawSemiring rawSemiring public
3037
Coprime-sym : Symmetric Coprime
3138
Coprime-sym coprime = flip coprime
3239

33-
∣1⇒Coprime : {x} y x ∣ 1# Coprime x y
34-
∣1⇒Coprime {x} y x∣1 z∣x _ = ∣-trans z∣x x∣1
40+
∣1⇒Coprime : y x ∣ 1# Coprime x y
41+
∣1⇒Coprime _ x∣1 z∣x _ = ∣-trans z∣x x∣1
3542

3643
------------------------------------------------------------------------
3744
-- Properties of Irreducible
3845

39-
Irreducible⇒≉0 : 0# ≉ 1# {p} Irreducible p p ≉ 0#
46+
Irreducible⇒≉0 : 0# ≉ 1# Irreducible p p ≉ 0#
4047
Irreducible⇒≉0 0≉1 (mkIrred _ chooseInvertible) p≈0 =
4148
0∤1 0≉1 (reduce (chooseInvertible (trans p≈0 (sym (zeroˡ 0#)))))

0 commit comments

Comments
 (0)