Skip to content

Commit

Permalink
[ base ] Add update functions to sorted maps
Browse files Browse the repository at this point in the history
  • Loading branch information
buzden committed Sep 23, 2023
1 parent a643e3a commit cf98d3d
Show file tree
Hide file tree
Showing 3 changed files with 28 additions and 0 deletions.
2 changes: 2 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -203,6 +203,8 @@

* Adds `infixOfBy` and `isInfixOfBy` into `Data.List`.

* Adds updating functions to `SortedMap` and `SortedDMap`.

#### System

* Changes `getNProcessors` to return the number of online processors rather than
Expand Down
13 changes: 13 additions & 0 deletions libs/base/Data/SortedMap.idr
Original file line number Diff line number Diff line change
Expand Up @@ -37,6 +37,19 @@ export
delete : k -> SortedMap k v -> SortedMap k v
delete k = M . delete k . unM

export
update : (Maybe v -> Maybe v) -> k -> SortedMap k v -> SortedMap k v
update f k m = case f $ lookup k m of
Just v => insert k v m
Nothing => delete k m

||| Updates existing value, if it is present, and does nothing otherwise
export
updateExisting : (v -> v) -> k -> SortedMap k v -> SortedMap k v
updateExisting f k m = case lookup k m of
Just v => insert k (f v) m
Nothing => m

export
fromList : Ord k => List (k, v) -> SortedMap k v
fromList = flip insertFrom empty
Expand Down
13 changes: 13 additions & 0 deletions libs/base/Data/SortedMap/Dependent.idr
Original file line number Diff line number Diff line change
Expand Up @@ -254,6 +254,19 @@ delete k (M (S n) t) =
Left t' => (M _ t')
Right t' => (M _ t')

export
update : DecEq k => (x : k) -> (Maybe (v x) -> Maybe (v x)) -> SortedDMap k v -> SortedDMap k v
update k f m = case f $ lookupPrecise k m of
Just v => insert k v m
Nothing => delete k m

||| Updates existing value, if it is present, and does nothing otherwise
export
updateExisting : DecEq k => (x : k) -> (v x -> v x) -> SortedDMap k v -> SortedDMap k v
updateExisting k f m = case lookupPrecise k m of
Just v => insert k (f v) m
Nothing => m

export
fromList : Ord k => List (x : k ** v x) -> SortedDMap k v
fromList = foldl (flip (uncurry insert)) empty
Expand Down

0 comments on commit cf98d3d

Please sign in to comment.