Skip to content

Commit

Permalink
add search function to Data.List.Relation.Unary.All (#2428)
Browse files Browse the repository at this point in the history
* adds `search` function

* Nathan's suggestions

* added `CHANGELOG` entry

---------

Co-authored-by: MatthewDaggitt <matthewdaggitt@gmail.com>
  • Loading branch information
jamesmckinna and MatthewDaggitt authored Aug 14, 2024
1 parent a6da449 commit 61e7707
Show file tree
Hide file tree
Showing 2 changed files with 9 additions and 1 deletion.
5 changes: 5 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -32,6 +32,11 @@ New modules
Additions to existing modules
-----------------------------

* In `Data.List.Relation.Unary.All`:
```agda
search : Decidable P → ∀ xs → All (∁ P) xs ⊎ Any P xs
```

* In `Data.List.Relation.Binary.Equality.Setoid`:
```agda
++⁺ʳ : ∀ xs → ys ≋ zs → xs ++ ys ≋ xs ++ zs
Expand Down
5 changes: 4 additions & 1 deletion src/Data/List/Relation/Unary/All.agda
Original file line number Diff line number Diff line change
Expand Up @@ -14,7 +14,7 @@ open import Data.List.Membership.Propositional renaming (_∈_ to _∈ₚ_)
import Data.List.Membership.Setoid as SetoidMembership
open import Data.Product.Base as Product
using (∃; -,_; _×_; _,_; proj₁; proj₂; uncurry)
open import Data.Sum.Base as Sum using (inj₁; inj₂)
open import Data.Sum.Base as Sum using (_⊎_; inj₁; inj₂)
open import Effect.Applicative
open import Effect.Monad
open import Function.Base using (_∘_; _∘′_; id; const)
Expand Down Expand Up @@ -232,6 +232,9 @@ decide p∪q (x ∷ xs) with p∪q x
... | inj₂ qx = inj₂ (here qx)
... | inj₁ px = Sum.map (px ∷_) there (decide p∪q xs)

search : Decidable P xs All (∁ P) xs ⊎ Any P xs
search P? = decide (Sum.swap ∘ Dec.toSum ∘ P?)

------------------------------------------------------------------------
-- DEPRECATED
------------------------------------------------------------------------
Expand Down

0 comments on commit 61e7707

Please sign in to comment.