66
77{-# OPTIONS --cubical-compatible --safe #-}
88
9- open import Relation.Binary.Core using (Rel)
10- open import Relation.Binary.Bundles using (Setoid)
11- open import Relation.Binary.Structures using (IsEquivalence)
12-
139module Algebra.Module.Structures where
1410
1511open import Algebra.Bundles
@@ -23,6 +19,10 @@ open import Algebra.Module.Definitions
2319open import Algebra.Structures using (IsCommutativeMonoid; IsAbelianGroup)
2420open import Data.Product.Base using (_,_; proj₁; proj₂)
2521open import Level using (Level; _⊔_)
22+ open import Relation.Binary.Core using (Rel)
23+ open import Relation.Binary.Bundles using (Setoid)
24+ open import Relation.Binary.Structures using (IsEquivalence)
25+
2626
2727private
2828 variable
@@ -216,6 +216,15 @@ module _ (ring : Ring r ℓr)
216216 ; uniqueˡ-⁻¹ to uniqueˡ‿-ᴹ
217217 ; uniqueʳ-⁻¹ to uniqueʳ‿-ᴹ
218218 )
219+ {-# WARNING_ON_USAGE uniqueˡ‿-ᴹ
220+ "Warning: uniqueˡ‿-ᴹ was deprecated in v2.3.
221+ Please use Algebra.Module.Properties.LeftModule.inverseˡ-uniqueᴹ instead."
222+ #-}
223+ {-# WARNING_ON_USAGE uniqueʳ‿-ᴹ
224+ "Warning: uniqueʳ‿-ᴹ was deprecated in v2.3.
225+ Please use Algebra.Module.Properties.LeftModule.inverseʳ-uniqueᴹ instead."
226+ #-}
227+
219228
220229 record IsRightModule (*ᵣ : Opᵣ R M) : Set (r ⊔ m ⊔ ℓr ⊔ ℓm) where
221230 open Defs ≈ᴹ
@@ -244,6 +253,15 @@ module _ (ring : Ring r ℓr)
244253 ; uniqueˡ-⁻¹ to uniqueˡ‿-ᴹ
245254 ; uniqueʳ-⁻¹ to uniqueʳ‿-ᴹ
246255 )
256+ {-# WARNING_ON_USAGE uniqueˡ‿-ᴹ
257+ "Warning: uniqueˡ‿-ᴹ was deprecated in v2.3.
258+ Please use Algebra.Module.Properties.RightModule.inverseˡ-uniqueᴹ instead."
259+ #-}
260+ {-# WARNING_ON_USAGE uniqueʳ‿-ᴹ
261+ "Warning: uniqueʳ‿-ᴹ was deprecated in v2.3.
262+ Please use Algebra.Module.Properties.RightModule.inverseʳ-uniqueᴹ instead."
263+ #-}
264+
247265
248266module _ (R-ring : Ring r ℓr) (S-ring : Ring s ℓs)
249267 (≈ᴹ : Rel {m} M ℓm) (+ᴹ : Op₂ M) (0ᴹ : M) (-ᴹ : Op₁ M)
0 commit comments