mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 04:33:45 +03:00
[ performance ] More stack safety in the Prelude (#2704)
This commit is contained in:
parent
b1f2eab6d6
commit
57c589ca80
@ -445,6 +445,10 @@ public export
|
||||
replaceOn : Eq a => (e : a) -> (b : a) -> (l : List a) -> List a
|
||||
replaceOn e = replaceWhen (== e)
|
||||
|
||||
replicateTR : List a -> Nat -> a -> List a
|
||||
replicateTR as Z _ = as
|
||||
replicateTR as (S n) x = replicateTR (x :: as) n x
|
||||
|
||||
||| Construct a list with `n` copies of `x`.
|
||||
|||
|
||||
||| @ n how many copies
|
||||
@ -454,6 +458,10 @@ replicate : (n : Nat) -> (x : a) -> List a
|
||||
replicate Z _ = []
|
||||
replicate (S n) x = x :: replicate n x
|
||||
|
||||
-- Data.List.replicateTRIsReplicate proves these are equivalent.
|
||||
%transform "tailRecReplicate" List.replicate = List.replicateTR Nil
|
||||
|
||||
|
||||
||| Compute the intersect of two lists by user-supplied equality predicate.
|
||||
export
|
||||
intersectBy : (a -> a -> Bool) -> List a -> List a -> List a
|
||||
@ -1066,3 +1074,49 @@ export
|
||||
mapAppend : (f : a -> b) -> (xs, ys : List a) -> map f (xs ++ ys) = map f xs ++ map f ys
|
||||
mapAppend f [] ys = Refl
|
||||
mapAppend f (x::xs) ys = rewrite mapAppend f xs ys in Refl
|
||||
|
||||
0 mapTRIsMap : (f : a -> b) -> (as : List a) -> mapTR f as === map f as
|
||||
mapTRIsMap f = lemma Lin
|
||||
where lemma : (sb : SnocList b)
|
||||
-> (as : List a)
|
||||
-> mapAppend sb f as === (sb <>> map f as)
|
||||
lemma sb [] = Refl
|
||||
lemma sb (x :: xs) = lemma (sb :< f x) xs
|
||||
|
||||
|
||||
0 mapMaybeTRIsMapMaybe : (f : a -> Maybe b)
|
||||
-> (as : List a)
|
||||
-> mapMaybeTR f as === mapMaybe f as
|
||||
mapMaybeTRIsMapMaybe f = lemma Lin
|
||||
where lemma : (sb : SnocList b)
|
||||
-> (as : List a)
|
||||
-> mapMaybeAppend sb f as === (sb <>> mapMaybe f as)
|
||||
lemma sb [] = Refl
|
||||
lemma sb (x :: xs) with (f x)
|
||||
lemma sb (x :: xs) | Nothing = lemma sb xs
|
||||
lemma sb (x :: xs) | Just v = lemma (sb :< v) xs
|
||||
|
||||
0 filterTRIsFilter : (f : a -> Bool)
|
||||
-> (as : List a)
|
||||
-> filterTR f as === filter f as
|
||||
filterTRIsFilter f = lemma Lin
|
||||
|
||||
where lemma : (sa : SnocList a)
|
||||
-> (as : List a)
|
||||
-> filterAppend sa f as === (sa <>> filter f as)
|
||||
lemma sa [] = Refl
|
||||
lemma sa (x :: xs) with (f x)
|
||||
lemma sa (x :: xs) | False = lemma sa xs
|
||||
lemma sa (x :: xs) | True = lemma (sa :< x) xs
|
||||
|
||||
0 replicateTRIsReplicate : (n : Nat) -> (x : a) -> replicateTR [] n x === replicate n x
|
||||
replicateTRIsReplicate n x = trans (lemma [] n) (appendNilRightNeutral _)
|
||||
where lemma1 : (as : List a) -> (m : Nat) -> (x :: replicate m x) ++ as === replicate m x ++ (x :: as)
|
||||
lemma1 as 0 = Refl
|
||||
lemma1 as (S k) = cong (x ::) (lemma1 as k)
|
||||
|
||||
lemma : (as : List a) -> (m : Nat) -> replicateTR as m x === replicate m x ++ as
|
||||
lemma as 0 = Refl
|
||||
lemma as (S k) =
|
||||
let prf := lemma (x :: as) k
|
||||
in trans prf (sym $ lemma1 as k)
|
||||
|
@ -62,9 +62,26 @@ Show a => Show (SnocList a) where
|
||||
show' acc (xs :< x) = show' (show x :: acc) xs
|
||||
|
||||
public export
|
||||
mapImpl : (a -> b) -> SnocList a -> SnocList b
|
||||
mapImpl f Lin = Lin
|
||||
mapImpl f (sx :< x) = (mapImpl f sx) :< (f x)
|
||||
|
||||
-- Utility for implementing `mapTR`
|
||||
mapTR' : List b -> (a -> b) -> SnocList a -> SnocList b
|
||||
mapTR' xs f (sx :< x) = mapTR' (f x :: xs) f sx
|
||||
mapTR' xs f Lin = Lin <>< xs
|
||||
|
||||
-- Tail recursive version of `map`. This is automatically used
|
||||
-- at runtime due to a `transform` rule.
|
||||
mapTR : (a -> b) -> SnocList a -> SnocList b
|
||||
mapTR = mapTR' []
|
||||
|
||||
-- mapTRIsMap proves these are equivalent.
|
||||
%transform "tailRecMapSnocList" SnocList.mapImpl = SnocList.mapTR
|
||||
|
||||
public export %inline
|
||||
Functor SnocList where
|
||||
map f Lin = Lin
|
||||
map f (sx :< x) = (map f sx) :< (f x)
|
||||
map = mapImpl
|
||||
|
||||
public export
|
||||
Semigroup (SnocList a) where
|
||||
@ -357,3 +374,75 @@ mapMaybeCast f (x::xs) = do
|
||||
mapMaybeStepLemma with (f x)
|
||||
_ | Nothing = rewrite appendLinLeftNeutral $ [<] <>< mapMaybe f xs in Refl
|
||||
_ | (Just y) = rewrite fishAsSnocAppend [<y] (mapMaybe f xs) in Refl
|
||||
|
||||
0 mapTRIsMap : (f : a -> b) -> (sa : SnocList a) -> mapTR f sa === map f sa
|
||||
mapTRIsMap f = lemma []
|
||||
where lemma : (bs : List b)
|
||||
-> (sa : SnocList a)
|
||||
-> mapTR' bs f sa === (map f sa <>< bs)
|
||||
lemma bs Lin = Refl
|
||||
lemma bs (sx :< x) = lemma (f x :: bs) sx
|
||||
|
||||
|
||||
0 mapMaybeTRIsMapMaybe : (f : a -> Maybe b)
|
||||
-> (sa : SnocList a)
|
||||
-> mapMaybeTR f sa === mapMaybe f sa
|
||||
mapMaybeTRIsMapMaybe f = lemma []
|
||||
where lemma : (bs : List b)
|
||||
-> (sa : SnocList a)
|
||||
-> mapMaybeAppend bs f sa === (mapMaybe f sa <>< bs)
|
||||
lemma bs Lin = Refl
|
||||
lemma bs (sx :< x) with (f x)
|
||||
lemma bs (sx :< x) | Nothing = lemma bs sx
|
||||
lemma bs (sx :< x) | Just v = lemma (v :: bs) sx
|
||||
|
||||
0 filterTRIsFilter : (f : a -> Bool)
|
||||
-> (sa : SnocList a)
|
||||
-> filterTR f sa === filter f sa
|
||||
filterTRIsFilter f = lemma []
|
||||
where lemma : (as : List a)
|
||||
-> (sa : SnocList a)
|
||||
-> filterAppend as f sa === (filter f sa <>< as)
|
||||
lemma as Lin = Refl
|
||||
lemma as (sx :< x) with (f x)
|
||||
lemma as (sx :< x) | False = lemma as sx
|
||||
lemma as (sx :< x) | True = lemma (x :: as) sx
|
||||
|
||||
-- SnocList `reverse` applied to `reverseOnto` is equivalent to swapping the
|
||||
-- arguments of `reverseOnto`.
|
||||
reverseReverseOnto : (l, r : SnocList a) ->
|
||||
reverse (reverseOnto l r) = reverseOnto r l
|
||||
reverseReverseOnto _ Lin = Refl
|
||||
reverseReverseOnto l (sx :< x) = reverseReverseOnto (l :< x) sx
|
||||
|
||||
||| SnocList `reverse` applied twice yields the identity function.
|
||||
export
|
||||
reverseInvolutive : (sx : SnocList a) -> reverse (reverse sx) = sx
|
||||
reverseInvolutive = reverseReverseOnto Lin
|
||||
|
||||
-- Appending `x` to `l` and then reversing the result onto `r` is the same as
|
||||
-- using (::) with `x` and the result of reversing `l` onto `r`.
|
||||
snocReverse : (x : a) -> (l, r : SnocList a) ->
|
||||
reverseOnto r l :< x = reverseOnto r (reverseOnto [<x] (reverse l))
|
||||
snocReverse _ Lin _ = Refl
|
||||
snocReverse x (sy :< y) r
|
||||
= rewrite snocReverse x sy (r :< y) in
|
||||
rewrite cong (reverseOnto r . reverse) $ snocReverse x sy [<y] in
|
||||
rewrite reverseInvolutive (reverseOnto [<x] (reverse sy) :< y) in
|
||||
Refl
|
||||
|
||||
-- Proof that it is safe to lift a (::) out of the first `tailRecAppend`
|
||||
-- argument.
|
||||
snocTailRecAppend : (x : a) -> (l, r : SnocList a) ->
|
||||
tailRecAppend l (r :< x) = (tailRecAppend l r) :< x
|
||||
snocTailRecAppend x l r =
|
||||
rewrite snocReverse x (reverse r) l in
|
||||
rewrite reverseInvolutive r in
|
||||
Refl
|
||||
|
||||
-- Proof that `(++)` and `tailRecAppend` do the same thing, so the %transform
|
||||
-- directive is safe.
|
||||
tailRecAppendIsAppend : (sx, sy : SnocList a) -> tailRecAppend sx sy = sx ++ sy
|
||||
tailRecAppendIsAppend sx Lin = Refl
|
||||
tailRecAppendIsAppend sx (sy :< y) =
|
||||
trans (snocTailRecAppend y sx sy) (cong (:< y) $ tailRecAppendIsAppend sx sy)
|
||||
|
@ -392,6 +392,119 @@ Ord a => Ord (List a) where
|
||||
EQ => compare xs ys
|
||||
c => c
|
||||
|
||||
namespace SnocList
|
||||
|
||||
infixl 7 <><
|
||||
infixr 6 <>>
|
||||
|
||||
||| 'fish': Action of lists on snoc-lists
|
||||
public export
|
||||
(<><) : SnocList a -> List a -> SnocList a
|
||||
sx <>< [] = sx
|
||||
sx <>< (x :: xs) = sx :< x <>< xs
|
||||
|
||||
||| 'chips': Action of snoc-lists on lists
|
||||
public export
|
||||
(<>>) : SnocList a -> List a -> List a
|
||||
Lin <>> xs = xs
|
||||
(sx :< x) <>> xs = sx <>> x :: xs
|
||||
|
||||
public export
|
||||
(++) : (sx, sy : SnocList a) -> SnocList a
|
||||
(++) sx Lin = sx
|
||||
(++) sx (sy :< y) = (sx ++ sy) :< y
|
||||
|
||||
public export
|
||||
length : SnocList a -> Nat
|
||||
length Lin = Z
|
||||
length (sx :< x) = S $ length sx
|
||||
|
||||
||| Filters a snoc-list according to a simple classifying function
|
||||
public export
|
||||
filter : (a -> Bool) -> SnocList a -> SnocList a
|
||||
filter f [<] = [<]
|
||||
filter f (xs:<x) = let rest = filter f xs in if f x then rest :< x else rest
|
||||
|
||||
||| Apply a partial function to the elements of a list, keeping the ones at which
|
||||
||| it is defined.
|
||||
public export
|
||||
mapMaybe : (a -> Maybe b) -> SnocList a -> SnocList b
|
||||
mapMaybe f [<] = [<]
|
||||
mapMaybe f (sx :< x) = case f x of
|
||||
Nothing => mapMaybe f sx
|
||||
Just j => mapMaybe f sx :< j
|
||||
|
||||
||| Reverse the second snoclist, prepending its elements to the first.
|
||||
public export
|
||||
reverseOnto : SnocList a -> SnocList a -> SnocList a
|
||||
reverseOnto acc Lin = acc
|
||||
reverseOnto acc (sx :< x) = reverseOnto (acc :< x) sx
|
||||
|
||||
||| Reverses the given list.
|
||||
public export
|
||||
reverse : SnocList a -> SnocList a
|
||||
reverse = reverseOnto Lin
|
||||
|
||||
||| Tail-recursive append. Uses of (++) are automatically transformed to
|
||||
||| this. The only reason this is exported is that the proof of equivalence
|
||||
||| lives in a different module.
|
||||
public export
|
||||
tailRecAppend : (sx, sy : SnocList a) -> SnocList a
|
||||
tailRecAppend sx sy = reverseOnto sx (reverse sy)
|
||||
|
||||
-- Always use tailRecAppend at runtime. Data.SnocList.tailRecAppendIsAppend
|
||||
-- proves these are equivalent.
|
||||
%transform "tailRecAppendSnocList" SnocList.(++) = SnocList.tailRecAppend
|
||||
|
||||
||| Returns the first argument plus the length of the second.
|
||||
public export
|
||||
lengthPlus : Nat -> SnocList a -> Nat
|
||||
lengthPlus n Lin = n
|
||||
lengthPlus n (sx :< _) = lengthPlus (S n) sx
|
||||
|
||||
||| `length` implementation that uses tail recursion. Exported so
|
||||
||| lengthTRIsLength can see it.
|
||||
public export
|
||||
lengthTR : SnocList a -> Nat
|
||||
lengthTR = lengthPlus Z
|
||||
|
||||
-- Data.List.lengthTRIsLength proves these are equivalent.
|
||||
%transform "tailRecLengthSnocList" SnocList.length = SnocList.lengthTR
|
||||
|
||||
||| Utility for implementing `mapMaybeTR`
|
||||
public export
|
||||
mapMaybeAppend : List b -> (a -> Maybe b) -> SnocList a -> SnocList b
|
||||
mapMaybeAppend xs f (sx :< x) = case f x of
|
||||
Just v => mapMaybeAppend (v :: xs) f sx
|
||||
Nothing => mapMaybeAppend xs f sx
|
||||
mapMaybeAppend xs f Lin = Lin <>< xs
|
||||
|
||||
||| Tail recursive version of `mapMaybe`. This is automatically used
|
||||
||| at runtime due to a `transform` rule.
|
||||
public export %inline
|
||||
mapMaybeTR : (a -> Maybe b) -> SnocList a -> SnocList b
|
||||
mapMaybeTR = mapMaybeAppend []
|
||||
|
||||
-- Data.List.mapMaybeTRIsMapMaybe proves these are equivalent.
|
||||
%transform "tailRecMapMaybeSnocList" SnocList.mapMaybe = SnocList.mapMaybeTR
|
||||
|
||||
||| Utility for implementing `filterTR`
|
||||
public export
|
||||
filterAppend : List a -> (a -> Bool) -> SnocList a -> SnocList a
|
||||
filterAppend xs f (sx :< x) = case f x of
|
||||
True => filterAppend (x :: xs) f sx
|
||||
False => filterAppend xs f sx
|
||||
filterAppend xs f Lin = Lin <>< xs
|
||||
|
||||
||| Tail recursive version of `filter`. This is automatically used
|
||||
||| at runtime due to a `transform` rule.
|
||||
public export %inline
|
||||
filterTR : (a -> Bool) -> SnocList a -> SnocList a
|
||||
filterTR = filterAppend []
|
||||
|
||||
-- Data.List.listTRIsList proves these are equivalent.
|
||||
%transform "tailRecFilterSnocList" SnocList.filter = SnocList.filterTR
|
||||
|
||||
namespace List
|
||||
|
||||
||| Concatenate one list with another.
|
||||
@ -445,7 +558,7 @@ namespace List
|
||||
|
||||
-- Always use tailRecAppend at runtime. Data.List.tailRecAppendIsAppend
|
||||
-- proves these are equivalent.
|
||||
%transform "tailRecAppend" (++) = tailRecAppend
|
||||
%transform "tailRecAppend" List.(++) = List.tailRecAppend
|
||||
|
||||
||| Returns the first argument plus the length of the second.
|
||||
public export
|
||||
@ -460,12 +573,65 @@ namespace List
|
||||
lengthTR = lengthPlus Z
|
||||
|
||||
-- Data.List.lengthTRIsLength proves these are equivalent.
|
||||
%transform "tailRecLength" length = lengthTR
|
||||
%transform "tailRecLength" List.length = List.lengthTR
|
||||
|
||||
public export
|
||||
public export
|
||||
mapImpl : (a -> b) -> List a -> List b
|
||||
mapImpl f [] = []
|
||||
mapImpl f (x :: xs) = f x :: mapImpl f xs
|
||||
|
||||
||| Utility for implementing `mapTR`
|
||||
public export
|
||||
mapAppend : SnocList b -> (a -> b) -> List a -> List b
|
||||
mapAppend sx f (x :: xs) = mapAppend (sx :< f x) f xs
|
||||
mapAppend sx f [] = sx <>> []
|
||||
|
||||
||| Tail recursive version of `map`. This is automatically used
|
||||
||| at runtime due to a `transform` rule.
|
||||
public export %inline
|
||||
mapTR : (a -> b) -> List a -> List b
|
||||
mapTR = mapAppend Lin
|
||||
|
||||
-- Data.List.mapTRIsMap proves these are equivalent.
|
||||
%transform "tailRecMap" mapImpl = List.mapTR
|
||||
|
||||
||| Utility for implementing `mapMaybeTR`
|
||||
public export
|
||||
mapMaybeAppend : SnocList b -> (a -> Maybe b) -> List a -> List b
|
||||
mapMaybeAppend sx f (x :: xs) = case f x of
|
||||
Just v => mapMaybeAppend (sx :< v) f xs
|
||||
Nothing => mapMaybeAppend sx f xs
|
||||
mapMaybeAppend sx f [] = sx <>> []
|
||||
|
||||
||| Tail recursive version of `mapMaybe`. This is automatically used
|
||||
||| at runtime due to a `transform` rule.
|
||||
public export %inline
|
||||
mapMaybeTR : (a -> Maybe b) -> List a -> List b
|
||||
mapMaybeTR = mapMaybeAppend Lin
|
||||
|
||||
-- Data.List.mapMaybeTRIsMapMaybe proves these are equivalent.
|
||||
%transform "tailRecMapMaybe" List.mapMaybe = List.mapMaybeTR
|
||||
|
||||
||| Utility for implementing `filterTR`
|
||||
public export
|
||||
filterAppend : SnocList a -> (a -> Bool) -> List a -> List a
|
||||
filterAppend sx f (x :: xs) = case f x of
|
||||
True => filterAppend (sx :< x) f xs
|
||||
False => filterAppend sx f xs
|
||||
filterAppend sx f [] = sx <>> []
|
||||
|
||||
||| Tail recursive version of `filter`. This is automatically used
|
||||
||| at runtime due to a `transform` rule.
|
||||
public export %inline
|
||||
filterTR : (a -> Bool) -> List a -> List a
|
||||
filterTR = filterAppend Lin
|
||||
|
||||
-- Data.List.listTRIsList proves these are equivalent.
|
||||
%transform "tailRecFilter" List.filter = List.filterTR
|
||||
|
||||
public export %inline
|
||||
Functor List where
|
||||
map f [] = []
|
||||
map f (x :: xs) = f x :: map f xs
|
||||
map = mapImpl
|
||||
|
||||
public export
|
||||
Semigroup (List a) where
|
||||
@ -519,48 +685,6 @@ Traversable List where
|
||||
traverse f [] = pure []
|
||||
traverse f (x::xs) = [| f x :: traverse f xs |]
|
||||
|
||||
namespace SnocList
|
||||
|
||||
infixl 7 <><
|
||||
infixr 6 <>>
|
||||
|
||||
||| 'fish': Action of lists on snoc-lists
|
||||
public export
|
||||
(<><) : SnocList a -> List a -> SnocList a
|
||||
sx <>< [] = sx
|
||||
sx <>< (x :: xs) = sx :< x <>< xs
|
||||
|
||||
||| 'chips': Action of snoc-lists on lists
|
||||
public export
|
||||
(<>>) : SnocList a -> List a -> List a
|
||||
Lin <>> xs = xs
|
||||
(sx :< x) <>> xs = sx <>> x :: xs
|
||||
|
||||
public export
|
||||
(++) : (sx, sy : SnocList a) -> SnocList a
|
||||
(++) sx Lin = sx
|
||||
(++) sx (sy :< y) = (sx ++ sy) :< y
|
||||
|
||||
public export
|
||||
length : SnocList a -> Nat
|
||||
length Lin = Z
|
||||
length (sx :< x) = S $ length sx
|
||||
|
||||
||| Filters a snoc-list according to a simple classifying function
|
||||
public export
|
||||
filter : (a -> Bool) -> SnocList a -> SnocList a
|
||||
filter f [<] = [<]
|
||||
filter f (xs:<x) = let rest = filter f xs in if f x then rest :< x else rest
|
||||
|
||||
||| Apply a partial function to the elements of a list, keeping the ones at which
|
||||
||| it is defined.
|
||||
public export
|
||||
mapMaybe : (a -> Maybe b) -> SnocList a -> SnocList b
|
||||
mapMaybe f [<] = [<]
|
||||
mapMaybe f (sx :< x) = case f x of
|
||||
Nothing => mapMaybe f sx
|
||||
Just j => mapMaybe f sx :< j
|
||||
|
||||
public export
|
||||
Eq a => Eq (SnocList a) where
|
||||
(==) Lin Lin = True
|
||||
|
@ -72,7 +72,7 @@ bindEnv loc (b :: env) tm
|
||||
Explicit
|
||||
(binderType b)) tm)
|
||||
|
||||
revOnto : (xs, vs : _) -> reverseOnto xs vs = reverse vs ++ xs
|
||||
revOnto : (xs, vs : List a) -> reverseOnto xs vs = reverse vs ++ xs
|
||||
revOnto xs [] = Refl
|
||||
revOnto xs (v :: vs)
|
||||
= rewrite revOnto (v :: xs) vs in
|
||||
|
@ -48,7 +48,7 @@
|
||||
0000cf(:return (:ok " n : Nat
|
||||
------------------------------
|
||||
prf : [2, 1, 0] = ?a" ((7 3 ((:decor :type))) (49 1 ((:decor :data))) (52 1 ((:decor :data))) (55 1 ((:decor :data))) (58 1 ((:decor :keyword))))) 2)
|
||||
0000f6(:return (:ok " xs : List Nat
|
||||
0000fa(:return (:ok " xs : List Nat
|
||||
------------------------------
|
||||
prf2 : map S xs = ?lgjgk" ((8 4 ((:decor :type))) (13 3 ((:decor :type))) (55 3 ((:decor :function))) (59 1 ((:decor :data))) (61 2 ((:decor :bound))) (64 1 ((:decor :keyword))))) 3)
|
||||
prf2 : mapImpl S xs = ?lgjgk" ((8 4 ((:decor :type))) (13 3 ((:decor :type))) (55 7 ((:decor :function))) (63 1 ((:decor :data))) (65 2 ((:decor :bound))) (68 1 ((:decor :keyword))))) 3)
|
||||
Alas the file is done, aborting
|
||||
|
@ -1 +1,8 @@
|
||||
50001
|
||||
50000
|
||||
50000
|
||||
50000
|
||||
50001
|
||||
50000
|
||||
50000
|
||||
50000
|
||||
|
@ -9,14 +9,22 @@ So far, only Prelude.List.++ is tested.
|
||||
|
||||
module Main
|
||||
|
||||
-- Until replicate is tail recursive, we roll our own for this test.
|
||||
import Data.List
|
||||
import Data.SnocList
|
||||
|
||||
replicateOnto : List a -> Nat -> a -> List a
|
||||
replicateOnto acc Z x = acc
|
||||
replicateOnto acc (S n) x = replicateOnto (x :: acc) n x
|
||||
values : List Nat
|
||||
values = replicate 50000 1
|
||||
|
||||
tailRecReplicate : Nat -> a -> List a
|
||||
tailRecReplicate = replicateOnto []
|
||||
seulav : SnocList Nat
|
||||
seulav = Lin <>< values
|
||||
|
||||
main : IO ()
|
||||
main = putStrLn $ show $ length $ tailRecReplicate 50000 () ++ [()]
|
||||
main = do
|
||||
printLn $ length $ values ++ [1]
|
||||
printLn $ length $ map Just values
|
||||
printLn $ length $ mapMaybe Just values
|
||||
printLn $ length $ filter (const True) values
|
||||
printLn $ length $ seulav ++ [<1]
|
||||
printLn $ length $ map Just seulav
|
||||
printLn $ length $ mapMaybe Just seulav
|
||||
printLn $ length $ filter (const True) seulav
|
||||
|
Loading…
Reference in New Issue
Block a user