@@ -20,7 +20,8 @@ open import Data.Nat.Base as ℕ using (ℕ; zero; suc; _+_; _*_ ; _≤_ ; s≤s
2020open import Data.Product.Base as Prod using (_×_; _,_; map₁; map₂′)
2121open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
2222open import Data.These.Base as These using (These; this; that; these)
23- open import Function.Base using (id; _∘_ ; _∘′_; _∘₂_; const; flip)
23+ open import Function.Base
24+ using (id; _∘_ ; _∘′_; _∘₂_; _$_; const; flip)
2425open import Level using (Level)
2526open import Relation.Nullary.Decidable.Core using (does; ¬?)
2627open import Relation.Unary using (Pred; Decidable)
@@ -395,6 +396,26 @@ deduplicateᵇ : (A → A → Bool) → List A → List A
395396deduplicateᵇ r [] = []
396397deduplicateᵇ r (x ∷ xs) = x ∷ filterᵇ (not ∘ r x) (deduplicateᵇ r xs)
397398
399+ -- Finds the first element satisfying the boolean predicate
400+ findᵇ : (A → Bool) → List A → Maybe A
401+ findᵇ p [] = nothing
402+ findᵇ p (x ∷ xs) = if p x then just x else findᵇ p xs
403+
404+ -- Finds the index of the first element satisfying the boolean predicate
405+ findIndexᵇ : (A → Bool) → (xs : List A) → Maybe $ Fin (length xs)
406+ findIndexᵇ p [] = nothing
407+ findIndexᵇ p (x ∷ xs) = if p x
408+ then just zero
409+ else Maybe.map suc (findIndexᵇ p xs)
410+
411+ -- Finds indices of all the elements satisfying the boolean predicate
412+ findIndicesᵇ : (A → Bool) → (xs : List A) → List $ Fin (length xs)
413+ findIndicesᵇ p [] = []
414+ findIndicesᵇ p (x ∷ xs) = if p x
415+ then zero ∷ indices
416+ else indices
417+ where indices = map suc (findIndicesᵇ p xs)
418+
398419-- Equivalent functions that use a decidable predicate instead of a
399420-- boolean function.
400421
@@ -436,6 +457,15 @@ derun R? = derunᵇ (does ∘₂ R?)
436457deduplicate : ∀ {R : Rel A p} → B.Decidable R → List A → List A
437458deduplicate R? = deduplicateᵇ (does ∘₂ R?)
438459
460+ find : ∀ {P : Pred A p} → Decidable P → List A → Maybe A
461+ find P? = findᵇ (does ∘ P?)
462+
463+ findIndex : ∀ {P : Pred A p} → Decidable P → (xs : List A) → Maybe $ Fin (length xs)
464+ findIndex P? = findIndexᵇ (does ∘ P?)
465+
466+ findIndices : ∀ {P : Pred A p} → Decidable P → (xs : List A) → List $ Fin (length xs)
467+ findIndices P? = findIndicesᵇ (does ∘ P?)
468+
439469------------------------------------------------------------------------
440470-- Actions on single elements
441471
0 commit comments