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
1411module 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+
1819open Semiring R renaming (Carrier to A)
1920open 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
3037Coprime-sym : Symmetric Coprime
3138Coprime-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#
4047Irreducible⇒≉0 0≉1 (mkIrred _ chooseInvertible) p≈0 =
4148 0∤1 0≉1 (reduce (chooseInvertible (trans p≈0 (sym (zeroˡ 0#)))))
0 commit comments