@@ -47,7 +47,7 @@ open import Relation.Binary.PropositionalEquality.Core
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
- using (Dec; does; yes; no; _because_; ¬?; decidable-stable)
50
+ using (Dec; does; yes; no; _because_; ¬?; decidable-stable; dec-true )
51
51
open import Relation.Unary
52
52
using (Decidable; Pred; Universal; ∁; _∩_; _⟨×⟩_) renaming (_⊆_ to _⋐_)
53
53
open import Relation.Unary.Properties using (∁?)
@@ -461,6 +461,11 @@ all-head-dropWhile P? (x ∷ xs) with P? x
461
461
... | yes px = all-head-dropWhile P? xs
462
462
... | no ¬px = just ¬px
463
463
464
+ all⇒dropWhile≡[] : (P? : Decidable P) → All P xs → dropWhile P? xs ≡ []
465
+ all⇒dropWhile≡[] P? [] = refl
466
+ all⇒dropWhile≡[] P? (px ∷ pxs) rewrite dec-true (P? _) px
467
+ = all⇒dropWhile≡[] P? pxs
468
+
464
469
take⁺ : ∀ n → All P xs → All P (take n xs)
465
470
take⁺ zero pxs = []
466
471
take⁺ (suc n) [] = []
@@ -484,6 +489,11 @@ all-takeWhile P? (x ∷ xs) with P? x
484
489
... | yes px = px ∷ all-takeWhile P? xs
485
490
... | no ¬px = []
486
491
492
+ all⇒takeWhile≗id : (P? : Decidable P) → All P xs → takeWhile P? xs ≡ xs
493
+ all⇒takeWhile≗id P? [] = refl
494
+ all⇒takeWhile≗id P? (px ∷ pxs) rewrite dec-true (P? _) px
495
+ = cong (_ ∷_) (all⇒takeWhile≗id P? pxs)
496
+
487
497
------------------------------------------------------------------------
488
498
-- applyUpTo
489
499
0 commit comments