@@ -43,13 +43,13 @@ open import Relation.Binary.Core using (REL)
43
43
open import Relation.Binary.Bundles using (Setoid)
44
44
import Relation.Binary.Definitions as B
45
45
open import Relation.Binary.PropositionalEquality.Core
46
- using (_≡_; refl; cong; cong₂; _≗_)
46
+ using (_≡_; refl; sym; cong; cong₂; _≗_)
47
47
open import Relation.Nullary.Reflects using (invert)
48
48
open import Relation.Nullary.Negation.Core using (¬_; contradiction)
49
49
open import Relation.Nullary.Decidable
50
50
using (Dec; does; yes; no; _because_; ¬?; decidable-stable; dec-true)
51
51
open import Relation.Unary
52
- using (Decidable; Pred; Universal; ∁; _∩_; _⟨×⟩_) renaming (_⊆_ to _⋐_)
52
+ using (Decidable; Pred; ∁; _∩_; _⟨×⟩_) renaming (_⊆_ to _⋐_)
53
53
open import Relation.Unary.Properties using (∁?)
54
54
55
55
private
@@ -444,9 +444,9 @@ drop⁺ (suc n) (px ∷ pxs) = drop⁺ n pxs
444
444
445
445
dropWhile⁺ : (Q? : Decidable Q) → All P xs → All P (dropWhile Q? xs)
446
446
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)
448
448
... | true = dropWhile⁺ Q? pxs
449
- ... | false = px ∷ pxs
449
+ ... | false = px∷ pxs
450
450
451
451
dropWhile⁻ : (P? : Decidable P) → dropWhile P? xs ≡ [] → All P xs
452
452
dropWhile⁻ {xs = []} P? eq = []
@@ -477,12 +477,6 @@ takeWhile⁺ {xs = x ∷ xs} Q? (px ∷ pxs) with does (Q? x)
477
477
... | true = px ∷ takeWhile⁺ Q? pxs
478
478
... | false = []
479
479
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
-
486
480
all-takeWhile : (P? : Decidable P) → ∀ xs → All P (takeWhile P? xs)
487
481
all-takeWhile P? [] = []
488
482
all-takeWhile P? (x ∷ xs) with P? x
@@ -765,3 +759,12 @@ map-compose = map-∘
765
759
"Warning: map-compose was deprecated in v2.1.
766
760
Please use map-∘ instead."
767
761
#-}
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