[ base ] Add flipped access/update functions for Sorted{Set,Map,DMap} (#3247)

Co-authored-by: G. Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
Denis Buzdalov 2024-06-05 15:59:38 +03:00 committed by GitHub
parent 2c128e216c
commit 1c588f77ec
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
4 changed files with 85 additions and 14 deletions

View File

@ -60,24 +60,24 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
is dropped as soon as possible. This allows you to reuse unique variables and is dropped as soon as possible. This allows you to reuse unique variables and
optimize memory consumption. optimize memory consumption.
* Fix invalid memory read onf strSubStr. * Fix invalid memory read on `strSubStr`.
* Fix memory leaks of IORef. Now that IORef holds values by itself, * Fix memory leaks of `IORef`. Now that `IORef` holds values by itself,
global_IORef_Storage is no longer needed. `global_IORef_Storage` is no longer needed.
* Pattern matching generates simpler code. This reduces malloc/free and memory * Pattern matching generates simpler code. This reduces `malloc`/`free` and memory
consumption. It also makes debugging easier. consumption. It also makes debugging easier.
* Stopped useless string copying in the constructor to save memory. Also, name * Stopped useless string copying in the constructor to save memory. Also, name
generation was stopped for constructors that have tags. generation was stopped for constructors that have tags.
* Special constructors such as Nil and Nothing were eliminated and assigned to * Special constructors such as `Nil` and `Nothing` were eliminated and assigned to
NULL. `NULL`.
* Unbox Bits32,Bits16,Bits8,Int32,Int16,Int8. These types are now packed into * Unbox `Bits32`, `Bits16`, `Bits8`, `Int32`, `Int16`, `Int8`. These types are now packed into
Value*. Now, RefC backend requires at least 32 bits for pointers. Value*. Now, RefC backend requires at least 32 bits for pointers.
16-bit CPUs are no longer supported. And we expect the address returned by 16-bit CPUs are no longer supported. And we expect the address returned by
malloc to be aligned with at least 32 bits. Otherwise it cause a runtime error. `malloc` to be aligned with at least 32 bits. Otherwise it cause a runtime error.
* Rename C function to avoid confliction. But only a part. * Rename C function to avoid confliction. But only a part.
@ -144,6 +144,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
* Added `funExt0` and `funExt1`, functions analogous to `funExt` but for functions * Added `funExt0` and `funExt1`, functions analogous to `funExt` but for functions
with quantities 0 and 1 respectively. with quantities 0 and 1 respectively.
* `SortedSet`, `SortedMap` and `SortedDMap` modules were extended with flipped variants
of functions like `lookup`, `contains`, `update` and `insert`.
* Moved definition of `Data.Vect.nubBy` to the global scope as `nubByImpl` to * Moved definition of `Data.Vect.nubBy` to the global scope as `nubByImpl` to
allow compile time proofs on `nubBy` and `nub`. allow compile time proofs on `nubBy` and `nub`.
@ -160,4 +163,4 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
#### Network #### Network
* Add a missing function parameter (the flag) in the C implementation of idrnet_recv_bytes * Add a missing function parameter (the flag) in the C implementation of `idrnet_recv_bytes`

View File

@ -22,22 +22,38 @@ export
lookup : k -> SortedMap k v -> Maybe v lookup : k -> SortedMap k v -> Maybe v
lookup k = map snd . lookup k . unM lookup k = map snd . lookup k . unM
public export %inline
lookup' : SortedMap k v -> k -> Maybe v
lookup' = flip lookup
export export
insert : k -> v -> SortedMap k v -> SortedMap k v insert : k -> v -> SortedMap k v -> SortedMap k v
insert k v = M . insert k v . unM insert k v = M . insert k v . unM
public export %inline
insert' : SortedMap k v -> (k, v) -> SortedMap k v
insert' = flip $ uncurry insert
export export
singleton : Ord k => k -> v -> SortedMap k v singleton : Ord k => k -> v -> SortedMap k v
singleton = M .: singleton singleton = M .: singleton
export export
insertFrom : Foldable f => f (k, v) -> SortedMap k v -> SortedMap k v insertFrom : Foldable f => f (k, v) -> SortedMap k v -> SortedMap k v
insertFrom = flip $ foldl $ flip $ uncurry insert insertFrom = flip $ foldl insert'
public export %inline
insertFrom' : Foldable f => SortedMap k v -> f (k, v) -> SortedMap k v
insertFrom' = flip insertFrom
export export
delete : k -> SortedMap k v -> SortedMap k v delete : k -> SortedMap k v -> SortedMap k v
delete k = M . delete k . unM delete k = M . delete k . unM
public export %inline
delete' : SortedMap k v -> k -> SortedMap k v
delete' = flip delete
||| Updates or deletes a value based on the decision function ||| Updates or deletes a value based on the decision function
||| |||
||| The decision function takes information about the presence of the value, ||| The decision function takes information about the presence of the value,
@ -51,6 +67,10 @@ update f k m = case f $ lookup k m of
Just v => insert k v m Just v => insert k v m
Nothing => delete k m Nothing => delete k m
public export %inline
update' : SortedMap k v -> (Maybe v -> Maybe v) -> k -> SortedMap k v
update' m f x = update f x m
||| Updates existing value, if it is present, and does nothing otherwise ||| Updates existing value, if it is present, and does nothing otherwise
||| |||
||| The current implementation performs up to two traversals of the original map ||| The current implementation performs up to two traversals of the original map
@ -60,9 +80,13 @@ updateExisting f k m = case lookup k m of
Just v => insert k (f v) m Just v => insert k (f v) m
Nothing => m Nothing => m
public export %inline
updateExisting' : SortedMap k v -> (v -> v) -> k -> SortedMap k v
updateExisting' m f x = updateExisting f x m
export export
fromList : Ord k => List (k, v) -> SortedMap k v fromList : Ord k => List (k, v) -> SortedMap k v
fromList = flip insertFrom empty fromList = insertFrom' empty
export export
toList : SortedMap k v -> List (k, v) toList : SortedMap k v -> List (k, v)

View File

@ -219,6 +219,10 @@ lookup : (x : k) -> SortedDMap k v -> Maybe (y : k ** v y) -- could return also
lookup _ Empty = Nothing lookup _ Empty = Nothing
lookup k (M _ t) = treeLookup k t lookup k (M _ t) = treeLookup k t
public export %inline
lookup' : SortedDMap k v -> (x : k) -> Maybe (y : k ** v y)
lookup' = flip lookup
export export
lookupPrecise : DecEq k => (x : k) -> SortedDMap k v -> Maybe (v x) lookupPrecise : DecEq k => (x : k) -> SortedDMap k v -> Maybe (v x)
lookupPrecise x = lookup x >=> \(y ** v) => lookupPrecise x = lookup x >=> \(y ** v) =>
@ -226,6 +230,10 @@ lookupPrecise x = lookup x >=> \(y ** v) =>
Yes Refl => Just v Yes Refl => Just v
No _ => Nothing No _ => Nothing
public export %inline
lookupPrecise' : DecEq k => SortedDMap k v -> (x : k) -> Maybe (v x)
lookupPrecise' m x = lookupPrecise x m
export export
insert : (x : k) -> v x -> SortedDMap k v -> SortedDMap k v insert : (x : k) -> v x -> SortedDMap k v -> SortedDMap k v
insert k v Empty = M Z (Leaf k v) insert k v Empty = M Z (Leaf k v)
@ -234,13 +242,21 @@ insert k v (M _ t) =
Left t' => (M _ t') Left t' => (M _ t')
Right t' => (M _ t') Right t' => (M _ t')
public export %inline
insert' : SortedDMap k v -> (x : k ** v x) -> SortedDMap k v
insert' m (x ** v) = insert x v m
export export
singleton : Ord k => (x : k) -> v x -> SortedDMap k v singleton : Ord k => (x : k) -> v x -> SortedDMap k v
singleton k v = insert k v empty singleton k v = insert k v empty
export export
insertFrom : Foldable f => f (x : k ** v x) -> SortedDMap k v -> SortedDMap k v insertFrom : Foldable f => f (x : k ** v x) -> SortedDMap k v -> SortedDMap k v
insertFrom = flip $ foldl $ flip $ uncurry insert insertFrom = flip $ foldl insert'
public export %inline
insertFrom' : Foldable f => SortedDMap k v -> f (x : k ** v x) -> SortedDMap k v
insertFrom' = flip insertFrom
export export
delete : k -> SortedDMap k v -> SortedDMap k v delete : k -> SortedDMap k v -> SortedDMap k v
@ -254,6 +270,10 @@ delete k (M (S n) t) =
Left t' => (M _ t') Left t' => (M _ t')
Right t' => (M _ t') Right t' => (M _ t')
public export %inline
delete' : SortedDMap k v -> k -> SortedDMap k v
delete' = flip delete
||| Updates or deletes a value based on the decision function ||| Updates or deletes a value based on the decision function
||| |||
||| The decision function takes information about the presence of the value, ||| The decision function takes information about the presence of the value,
@ -267,6 +287,10 @@ update k f m = case f $ lookupPrecise k m of
Just v => insert k v m Just v => insert k v m
Nothing => delete k m Nothing => delete k m
public export %inline
update' : DecEq k => SortedDMap k v -> (x : k ** Maybe (v x) -> Maybe (v x)) -> SortedDMap k v
update' m (x ** f) = update x f m
||| Updates existing value, if it is present, and does nothing otherwise ||| Updates existing value, if it is present, and does nothing otherwise
||| |||
||| The current implementation performs up to two traversals of the original map ||| The current implementation performs up to two traversals of the original map
@ -276,9 +300,13 @@ updateExisting k f m = case lookupPrecise k m of
Just v => insert k (f v) m Just v => insert k (f v) m
Nothing => m Nothing => m
public export %inline
updateExisting' : DecEq k => SortedDMap k v -> (x : k ** v x -> v x) -> SortedDMap k v
updateExisting' m (x ** f) = updateExisting x f m
export export
fromList : Ord k => List (x : k ** v x) -> SortedDMap k v fromList : Ord k => List (x : k ** v x) -> SortedDMap k v
fromList = foldl (flip (uncurry insert)) empty fromList = foldl insert' empty
export export
toList : SortedDMap k v -> List (x : k ** v x) toList : SortedDMap k v -> List (x : k ** v x)
@ -346,6 +374,10 @@ traverse : Applicative f => ({x : k} -> v x -> f (w x)) -> SortedDMap k v -> f (
traverse _ Empty = pure Empty traverse _ Empty = pure Empty
traverse f (M _ t) = M _ <$> treeTraverse f t traverse f (M _ t) = M _ <$> treeTraverse f t
public export %inline
for : Applicative f => SortedDMap k v -> ({x : k} -> v x -> f (w x)) -> f (SortedDMap k w)
for = flip traverse
||| Merge two maps. When encountering duplicate keys, using a function to combine the values. ||| Merge two maps. When encountering duplicate keys, using a function to combine the values.
||| Uses the ordering of the first map given. ||| Uses the ordering of the first map given.
export export

View File

@ -17,14 +17,26 @@ export
insert : k -> SortedSet k -> SortedSet k insert : k -> SortedSet k -> SortedSet k
insert k (SetWrapper m) = SetWrapper (Data.SortedMap.insert k () m) insert k (SetWrapper m) = SetWrapper (Data.SortedMap.insert k () m)
public export %inline
insert' : SortedSet k -> k -> SortedSet k
insert' = flip insert
export export
delete : k -> SortedSet k -> SortedSet k delete : k -> SortedSet k -> SortedSet k
delete k (SetWrapper m) = SetWrapper (Data.SortedMap.delete k m) delete k (SetWrapper m) = SetWrapper (Data.SortedMap.delete k m)
public export %inline
delete' : SortedSet k -> k -> SortedSet k
delete' = flip delete
export export
contains : k -> SortedSet k -> Bool contains : k -> SortedSet k -> Bool
contains k (SetWrapper m) = isJust (Data.SortedMap.lookup k m) contains k (SetWrapper m) = isJust (Data.SortedMap.lookup k m)
public export %inline
contains' : SortedSet k -> k -> Bool
contains' = flip contains
export export
fromList : Ord k => List k -> SortedSet k fromList : Ord k => List k -> SortedSet k
fromList l = SetWrapper (Data.SortedMap.fromList (map (\i => (i, ())) l)) fromList l = SetWrapper (Data.SortedMap.fromList (map (\i => (i, ())) l))
@ -100,4 +112,4 @@ namespace Dependent
export export
singleton : Ord k => k -> SortedSet k singleton : Ord k => k -> SortedSet k
singleton k = insert k empty singleton = insert' empty