diff --git a/CHANGELOG.md b/CHANGELOG.md index a243edfd5f..732ff78e7e 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -99,6 +99,15 @@ Deprecated names map-compose ↦ map-∘ ``` +* In `Data.Nat.Base`: the following pattern synonyms and definitions are all + deprecated in favour of direct pattern matching on `Algebra.Definitions.RawMagma._∣ˡ_._,_` + ```agda + pattern less-than-or-equal {k} eq = k , eq + pattern ≤″-offset k = k , refl + pattern <″-offset k = k , refl + s≤″s⁻¹ + ``` + * In `Data.Nat.Divisibility.Core`: ```agda *-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣ @@ -663,12 +672,16 @@ Additions to existing modules * Added new proofs in `Data.Nat.Properties`: ```agda - m≤n+o⇒m∸n≤o : ∀ m n {o} → m ≤ n + o → m ∸ n ≤ o - m″_ _≤″_ : (m n : ℕ) → Set _≤″_ = _∣ˡ_ +-rawMagma -pattern less-than-or-equal {k} proof = k , proof - _<″_ : Rel ℕ 0ℓ m <″ n = suc m ≤″ n @@ -374,18 +372,10 @@ m ≥″ n = n ≤″ m _>″_ : Rel ℕ 0ℓ m >″ n = n <″ m --- Smart constructors of _≤″_ and _<″_ - -pattern ≤″-offset k = less-than-or-equal {k = k} refl -pattern <″-offset k = ≤″-offset k - --- Smart destructors of _<″_ - -s≤″s⁻¹ : ∀ {m n} → suc m ≤″ suc n → m ≤″ n -s≤″s⁻¹ (≤″-offset k) = ≤″-offset k +-- Smart destructor of _<″_ s<″s⁻¹ : ∀ {m n} → suc m <″ suc n → m <″ n -s<″s⁻¹ (<″-offset k) = <″-offset k +s<″s⁻¹ (k , refl) = k , refl -- _≤‴_: this definition is useful for induction with an upper bound. @@ -429,5 +419,26 @@ compare (suc m) (suc n) with compare m n -- Please use the new names as continuing support for the old names is -- not guaranteed. --- Version 2.0 +-- Version 2.1 + +-- Smart constructors of _≤″_ and _<″_ +pattern less-than-or-equal {k} eq = k , eq +{-# WARNING_ON_USAGE less-than-or-equal +"Warning: less-than-or-equal was deprecated in v2.1. Please match directly on proofs of ≤″ using constructor Algebra.Definitions.RawMagma._∣ˡ_._,_ instead. " +#-} +pattern ≤″-offset k = k , refl +{-# WARNING_ON_USAGE ≤″-offset +"Warning: ≤″-offset was deprecated in v2.1. Please match directly on proofs of ≤″ using pattern (_, refl) from Algebra.Definitions.RawMagma._∣ˡ_ instead. " +#-} +pattern <″-offset k = k , refl +{-# WARNING_ON_USAGE <″-offset +"Warning: <″-offset was deprecated in v2.1. Please match directly on proofs of ≤″ using pattern (_, refl) from Algebra.Definitions.RawMagma._∣ˡ_ instead. " +#-} + +-- Smart destructors of _<″_ +s≤″s⁻¹ : ∀ {m n} → suc m ≤″ suc n → m ≤″ n +s≤″s⁻¹ (k , refl) = k , refl +{-# WARNING_ON_USAGE s≤″s⁻¹ +"Warning: s≤″s⁻¹ was deprecated in v2.1. Please match directly on proofs of ≤″ using pattern (_, refl) from Algebra.Definitions.RawMagma._∣ˡ_ instead. " +#-} diff --git a/src/Data/Nat/DivMod.agda b/src/Data/Nat/DivMod.agda index cddb2768f7..86bc8d3168 100644 --- a/src/Data/Nat/DivMod.agda +++ b/src/Data/Nat/DivMod.agda @@ -17,6 +17,7 @@ open import Data.Nat.DivMod.Core open import Data.Nat.Divisibility.Core open import Data.Nat.Induction open import Data.Nat.Properties +open import Data.Product.Base using (_,_) open import Data.Sum.Base using (inj₁; inj₂) open import Function.Base using (_$_; _∘_) open import Relation.Binary.PropositionalEquality.Core @@ -105,7 +106,7 @@ m%n≤m : ∀ m n .{{_ : NonZero n}} → m % n ≤ m m%n≤m m (suc n-1) = a[modₕ]n≤a 0 m n-1 m≤n⇒m%n≡m : m ≤ n → m % suc n ≡ m -m≤n⇒m%n≡m {m = m} m≤n with ≤″-offset k ← ≤⇒≤″ m≤n +m≤n⇒m%n≡m {m = m} m≤n with k , refl ← m≤n⇒∃[o]m+o≡n m≤n = a≤n⇒a[modₕ]n≡a 0 (m + k) m k m″?_ : Decidable _>″_ _>″?_ = flip _<″?_ ≤″-irrelevant : Irrelevant _≤″_ -≤″-irrelevant {m} (less-than-or-equal eq₁) - (less-than-or-equal eq₂) +≤″-irrelevant {m} (_ , eq₁) (_ , eq₂) with refl ← +-cancelˡ-≡ m _ _ (trans eq₁ (sym eq₂)) - = cong less-than-or-equal (≡-irrelevant eq₁ eq₂) + = cong (_ ,_) (≡-irrelevant eq₁ eq₂) <″-irrelevant : Irrelevant _<″_ <″-irrelevant = ≤″-irrelevant @@ -2173,8 +2188,8 @@ _>″?_ = flip _<″?_ ------------------------------------------------------------------------ ≤‴⇒≤″ : ∀{m n} → m ≤‴ n → m ≤″ n -≤‴⇒≤″ {m = m} ≤‴-refl = less-than-or-equal {k = 0} (+-identityʳ m) -≤‴⇒≤″ {m = m} (≤‴-step m≤n) = less-than-or-equal (trans (+-suc m _) (≤″-proof (≤‴⇒≤″ m≤n))) +≤‴⇒≤″ {m = m} ≤‴-refl = 0 , +-identityʳ m +≤‴⇒≤″ {m = m} (≤‴-step m≤n) = _ , trans (+-suc m _) (≤″-proof (≤‴⇒≤″ m≤n)) m≤‴m+k : ∀{m n k} → m + k ≡ n → m ≤‴ n m≤‴m+k {m} {k = zero} refl = subst (λ z → m ≤‴ z) (sym (+-identityʳ m)) (≤‴-refl {m}) @@ -2397,4 +2412,3 @@ open Data.Nat.Base public {-# WARNING_ON_USAGE <-transˡ "Warning: <-transˡ was deprecated in v2.0. Please use <-≤-trans instead. " #-} - diff --git a/src/Data/Nat/WithK.agda b/src/Data/Nat/WithK.agda index 49966a93f0..bb6a97a7aa 100644 --- a/src/Data/Nat/WithK.agda +++ b/src/Data/Nat/WithK.agda @@ -8,8 +8,9 @@ module Data.Nat.WithK where +open import Algebra.Definitions.RawMagma using (_,_) open import Data.Nat.Base open import Relation.Binary.PropositionalEquality.WithK ≤″-erase : ∀ {m n} → m ≤″ n → m ≤″ n -≤″-erase (less-than-or-equal eq) = less-than-or-equal (≡-erase eq) +≤″-erase (_ , eq) = _ , ≡-erase eq diff --git a/src/Data/Vec/Bounded/Base.agda b/src/Data/Vec/Bounded/Base.agda index 7476afb2d4..fa7eb758ed 100644 --- a/src/Data/Vec/Bounded/Base.agda +++ b/src/Data/Vec/Bounded/Base.agda @@ -21,7 +21,7 @@ open import Data.These.Base as These using (These) open import Function.Base using (_∘_; id; _$_) open import Level using (Level) open import Relation.Nullary.Decidable.Core using (recompute) -open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_) +open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_; refl) open import Relation.Binary.PropositionalEquality.Properties using (module ≡-Reasoning) @@ -43,20 +43,27 @@ record Vec≤ (A : Set a) (n : ℕ) : Set a where vec : Vec A length .bound : length ≤ n +-- projection to recompute irrelevant field +isBounded : (as : Vec≤ A n) → Vec≤.length as ≤ n +isBounded as@(_ , m≤n) = recompute (_ ℕ.≤? _) m≤n + ------------------------------------------------------------------------ -- Conversion functions +toVec : (as : Vec≤ A n) → Vec A (Vec≤.length as) +toVec as@(vs , _) = vs + fromVec : Vec A n → Vec≤ A n fromVec v = v , ℕ.≤-refl padRight : A → Vec≤ A n → Vec A n -padRight a (vs , m≤n) - with ≤″-offset k ← recompute (_ ℕ.≤″? _) (ℕ.≤⇒≤″ m≤n) +padRight a as@(vs , m≤n) + with k , refl ← ℕ.m≤n⇒∃[o]m+o≡n m≤n = vs Vec.++ Vec.replicate k a padLeft : A → Vec≤ A n → Vec A n padLeft a record { length = m ; vec = vs ; bound = m≤n } - with ≤″-offset k ← recompute (_ ℕ.≤″? _) (ℕ.≤⇒≤″ m≤n) + with k , refl ← ℕ.m≤n⇒∃[o]m+o≡n m≤n rewrite ℕ.+-comm m k = Vec.replicate k a Vec.++ vs @@ -72,7 +79,7 @@ private padBoth : A → A → Vec≤ A n → Vec A n padBoth aₗ aᵣ record { length = m ; vec = vs ; bound = m≤n } - with ≤″-offset k ← recompute (_ ℕ.≤″? _) (ℕ.≤⇒≤″ m≤n) + with k , refl ← ℕ.m≤n⇒∃[o]m+o≡n m≤n rewrite split m k = Vec.replicate ⌊ k /2⌋ aₗ Vec.++ vs @@ -83,7 +90,7 @@ fromList : (as : List A) → Vec≤ A (List.length as) fromList = fromVec ∘ Vec.fromList toList : Vec≤ A n → List A -toList = Vec.toList ∘ Vec≤.vec +toList = Vec.toList ∘ toVec ------------------------------------------------------------------------ -- Creating new Vec≤ vectors