Skip to content

Commit 255f8a8

Browse files
authored
Deprecate Data.List.Relation.Unary.All.Properties.takeWhile⁻ (#2522)
* refactor: generalise type of `takeWhile⁻` * deprecate rather than refactor * refactor: generalise `takeWhile⁺` * refactor: generalise `dropWhile⁻` * roll back to simple deprecation PR
1 parent cd7963b commit 255f8a8

File tree

2 files changed

+18
-10
lines changed

2 files changed

+18
-10
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -88,6 +88,11 @@ Deprecated names
8888
with a more informative type (see below).
8989
```
9090
91+
* In `Data.List.Relation.Unary.All.Properties`:
92+
```agda
93+
takeWhile⁻ ↦ all-takeWhile
94+
```
95+
9196
* In `Data.Vec.Properties`:
9297
```agda
9398
++-assoc _ ↦ ++-assoc-eqFree

src/Data/List/Relation/Unary/All/Properties.agda

+13-10
Original file line numberDiff line numberDiff line change
@@ -43,13 +43,13 @@ open import Relation.Binary.Core using (REL)
4343
open import Relation.Binary.Bundles using (Setoid)
4444
import Relation.Binary.Definitions as B
4545
open import Relation.Binary.PropositionalEquality.Core
46-
using (_≡_; refl; cong; cong₂; _≗_)
46+
using (_≡_; refl; sym; cong; cong₂; _≗_)
4747
open import Relation.Nullary.Reflects using (invert)
4848
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
4949
open import Relation.Nullary.Decidable
5050
using (Dec; does; yes; no; _because_; ¬?; decidable-stable; dec-true)
5151
open import Relation.Unary
52-
using (Decidable; Pred; Universal; ∁; _∩_; _⟨×⟩_) renaming (_⊆_ to _⋐_)
52+
using (Decidable; Pred; ∁; _∩_; _⟨×⟩_) renaming (_⊆_ to _⋐_)
5353
open import Relation.Unary.Properties using (∁?)
5454

5555
private
@@ -444,9 +444,9 @@ drop⁺ (suc n) (px ∷ pxs) = drop⁺ n pxs
444444

445445
dropWhile⁺ : (Q? : Decidable Q) All P xs All P (dropWhile Q? xs)
446446
dropWhile⁺ Q? [] = []
447-
dropWhile⁺ {xs = x ∷ xs} Q? (px ∷ pxs) with does (Q? x)
447+
dropWhile⁺ {xs = x ∷ xs} Q? px∷pxs@(_ ∷ pxs) with does (Q? x)
448448
... | true = dropWhile⁺ Q? pxs
449-
... | false = pxpxs
449+
... | false = pxpxs
450450

451451
dropWhile⁻ : (P? : Decidable P) dropWhile P? xs ≡ [] All P xs
452452
dropWhile⁻ {xs = []} P? eq = []
@@ -477,12 +477,6 @@ takeWhile⁺ {xs = x ∷ xs} Q? (px ∷ pxs) with does (Q? x)
477477
... | true = px ∷ takeWhile⁺ Q? pxs
478478
... | false = []
479479

480-
takeWhile⁻ : (P? : Decidable P) takeWhile P? xs ≡ xs All P xs
481-
takeWhile⁻ {xs = []} P? eq = []
482-
takeWhile⁻ {xs = x ∷ xs} P? eq with P? x
483-
... | yes px = px ∷ takeWhile⁻ P? (List.∷-injectiveʳ eq)
484-
... | no ¬px = case eq of λ ()
485-
486480
all-takeWhile : (P? : Decidable P) xs All P (takeWhile P? xs)
487481
all-takeWhile P? [] = []
488482
all-takeWhile P? (x ∷ xs) with P? x
@@ -765,3 +759,12 @@ map-compose = map-∘
765759
"Warning: map-compose was deprecated in v2.1.
766760
Please use map-∘ instead."
767761
#-}
762+
763+
-- Version 2.2
764+
765+
takeWhile⁻ : (P? : Decidable P) takeWhile P? xs ≡ xs All P xs
766+
takeWhile⁻ {xs = xs} P? eq rewrite sym eq = all-takeWhile P? xs
767+
{-# WARNING_ON_USAGE takeWhile⁻
768+
"Warning: takeWhile⁻ was deprecated in v2.2.
769+
Please use all-takeWhile instead."
770+
#-}

0 commit comments

Comments
 (0)