[ new ] added util functions for SortedMap (#3254)

Co-authored-by: G. Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
Claudio Etterli 2024-06-11 12:05:48 +02:00 committed by GitHub
parent 109033c7b0
commit 0174618724
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
3 changed files with 25 additions and 0 deletions

View File

@ -145,6 +145,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* Removed need for the runtime value of the implicit argument in `succNotLTEpred`.
* Added utility functions `insertWith`, `insertFromWith` and `fromListWith` for
`SortedMap`.
* Implemented `leftMost` and `rightMost` for `SortedSet`.
* Added `funExt0` and `funExt1`, functions analogous to `funExt` but for functions

View File

@ -80,6 +80,7 @@ troiganto
Wen Kokke
Wind Wong
Zoe Stafford
Claudio Etterli
Apologies to anyone we've missed - let us know and we'll correct it (or just
send a PR with the correction).

View File

@ -30,6 +30,15 @@ export
insert : k -> v -> SortedMap k v -> SortedMap k v
insert k v = M . insert k v . unM
||| Inserts a key value pair into a map and merges duplicated values
||| with the given function.
export
insertWith : (v -> v -> v) -> k -> v -> SortedMap k v -> SortedMap k v
insertWith f k v xs =
case lookup k xs of
Just x => insert k (f v x) xs
Nothing => insert k v xs
public export %inline
insert' : SortedMap k v -> (k, v) -> SortedMap k v
insert' = flip $ uncurry insert
@ -46,6 +55,12 @@ public export %inline
insertFrom' : Foldable f => SortedMap k v -> f (k, v) -> SortedMap k v
insertFrom' = flip insertFrom
||| Inserts any foldable of a key value pair into a map and merges duplicated
||| values with the given function.
export
insertFromWith : Foldable f => (v -> v -> v) -> f (k, v) -> SortedMap k v -> SortedMap k v
insertFromWith f = flip $ foldl $ flip $ uncurry $ insertWith f
export
delete : k -> SortedMap k v -> SortedMap k v
delete k = M . delete k . unM
@ -88,6 +103,12 @@ export
fromList : Ord k => List (k, v) -> SortedMap k v
fromList = insertFrom' empty
||| Converts a list of key-value pairs into a map and merges duplicated
||| values with the given function.
export
fromListWith : Ord k => (v -> v -> v) -> List (k, v) -> SortedMap k v
fromListWith f = flip (insertFromWith f) empty
export
toList : SortedMap k v -> List (k, v)
toList = map unDPair . toList . unM