mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 15:07:37 +03:00
Add list difference to base Data.List module.
This commit is contained in:
parent
08d9e7d82c
commit
ef91cc01c7
@ -206,6 +206,31 @@ public export
|
|||||||
delete : Eq a => a -> List a -> List a
|
delete : Eq a => a -> List a -> List a
|
||||||
delete = deleteBy (==)
|
delete = deleteBy (==)
|
||||||
|
|
||||||
|
||| Delete the first occurrence of each element from the second list in the first
|
||||||
|
||| list where equality is determined by the predicate passed as the first argument.
|
||||||
|
||| @ p A function that returns true when its two arguments should be considered equal.
|
||||||
|
||| @ source The list to delete elements from.
|
||||||
|
||| @ undesirables The list of elements to delete.
|
||||||
|
public export
|
||||||
|
deleteFirstsBy : (p : a -> a -> Bool) -> (source : List a) -> (undesirables : List a) -> List a
|
||||||
|
deleteFirstsBy p = foldl (flip (deleteBy p))
|
||||||
|
|
||||||
|
infix 7 \\
|
||||||
|
|
||||||
|
||| The non-associative list-difference.
|
||||||
|
||| A specialized form of @deleteFirstsBy@ where the predicate is equality under the @Eq@
|
||||||
|
||| interface.
|
||||||
|
||| Deletes the first occurrence of each element of the second list from the first list.
|
||||||
|
|||
|
||||||
|
||| In the following example, the result is `[2, 4]`.
|
||||||
|
||| ```idris example
|
||||||
|
||| [1, 2, 3, 4] // [1, 3]
|
||||||
|
||| ```
|
||||||
|
|||
|
||||||
|
public export
|
||||||
|
(\\) : (Eq a) => List a -> List a -> List a
|
||||||
|
(\\) = foldl (flip delete)
|
||||||
|
|
||||||
||| The unionBy function returns the union of two lists by user-supplied equality predicate.
|
||| The unionBy function returns the union of two lists by user-supplied equality predicate.
|
||||||
public export
|
public export
|
||||||
unionBy : (a -> a -> Bool) -> List a -> List a -> List a
|
unionBy : (a -> a -> Bool) -> List a -> List a -> List a
|
||||||
|
Loading…
Reference in New Issue
Block a user