77{-# OPTIONS --cubical-compatible --safe #-}
88
99open import Algebra using (CommutativeMagma)
10- open import Data.Product.Base using (_×_; _,_; map)
1110
1211module Algebra.Properties.CommutativeMagma.Divisibility
1312 {a ℓ} (CM : CommutativeMagma a ℓ)
1413 where
1514
16- open CommutativeMagma CM
15+ open import Data.Product.Base using (_×_; _,_)
1716
18- open import Relation.Binary.Reasoning.Setoid setoid
17+ open CommutativeMagma CM using (magma; _≈_; _∙_; comm)
1918
2019------------------------------------------------------------------------
2120-- Re-export the contents of magmas
@@ -31,8 +30,28 @@ x∣xy x y = y , comm y x
3130xy≈z⇒x∣z : ∀ x y {z} → x ∙ y ≈ z → x ∣ z
3231xy≈z⇒x∣z x y xy≈z = ∣-respʳ xy≈z (x∣xy x y)
3332
34- ∣-factors : ∀ x y → (x ∣ x ∙ y) × (y ∣ x ∙ y)
35- ∣-factors x y = x∣xy x y , x∣yx y x
33+ x|xy∧y|xy : ∀ x y → (x ∣ x ∙ y) × (y ∣ x ∙ y)
34+ x|xy∧y|xy x y = x∣xy x y , x∣yx y x
3635
37- ∣-factors-≈ : ∀ x y {z} → x ∙ y ≈ z → x ∣ z × y ∣ z
38- ∣-factors-≈ x y xy≈z = xy≈z⇒x∣z x y xy≈z , xy≈z⇒y∣z x y xy≈z
36+ xy≈z⇒x|z∧y|z : ∀ x y {z} → x ∙ y ≈ z → x ∣ z × y ∣ z
37+ xy≈z⇒x|z∧y|z x y xy≈z = xy≈z⇒x∣z x y xy≈z , xy≈z⇒y∣z x y xy≈z
38+
39+
40+ ------------------------------------------------------------------------
41+ -- DEPRECATED NAMES
42+ ------------------------------------------------------------------------
43+ -- Please use the new names as continuing support for the old names is
44+ -- not guaranteed.
45+
46+ -- Version 2.2
47+
48+ ∣-factors = x|xy∧y|xy
49+ {-# WARNING_ON_USAGE ∣-factors
50+ "Warning: ∣-factors was deprecated in v2.2.
51+ Please use x|xy∧y|xy instead. "
52+ #-}
53+ ∣-factors-≈ = xy≈z⇒x|z∧y|z
54+ {-# WARNING_ON_USAGE ∣-factors-≈
55+ "Warning: ∣-factors-≈ was deprecated in v2.2.
56+ Please use xy≈z⇒x|z∧y|z instead. "
57+ #-}
0 commit comments