mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-23 22:22:07 +03:00
[ base ] Add some more properties, functions and interface implementations (#2361)
This commit is contained in:
parent
f38feaa48d
commit
e8d3d788c1
@ -142,16 +142,6 @@ dropWhile : (p : a -> Bool) -> (xs : List a) -> List a
|
||||
dropWhile p [] = []
|
||||
dropWhile p (x::xs) = if p x then dropWhile p xs else x::xs
|
||||
|
||||
||| Applied to a predicate and a list, returns the list of those elements that
|
||||
||| satisfy the predicate.
|
||||
public export
|
||||
filter : (p : a -> Bool) -> List a -> List a
|
||||
filter p [] = []
|
||||
filter p (x :: xs)
|
||||
= if p x
|
||||
then x :: filter p xs
|
||||
else filter p xs
|
||||
|
||||
||| Find the first element of the list that satisfies the predicate.
|
||||
public export
|
||||
find : (p : a -> Bool) -> (xs : List a) -> Maybe a
|
||||
@ -697,16 +687,6 @@ export
|
||||
intercalate : (sep : List a) -> (xss : List (List a)) -> List a
|
||||
intercalate sep xss = concat $ intersperse sep xss
|
||||
|
||||
||| Apply a partial function to the elements of a list, keeping the ones at which
|
||||
||| it is defined.
|
||||
public export
|
||||
mapMaybe : (a -> Maybe b) -> List a -> List b
|
||||
mapMaybe f [] = []
|
||||
mapMaybe f (x::xs) =
|
||||
case f x of
|
||||
Nothing => mapMaybe f xs
|
||||
Just j => j :: mapMaybe f xs
|
||||
|
||||
||| Extract all of the values contained in a List of Maybes
|
||||
public export
|
||||
catMaybes : List (Maybe a) -> List a
|
||||
@ -1046,3 +1026,36 @@ export
|
||||
foldlAppend : (f : acc -> a -> acc) -> (init : acc) -> (xs : List a) -> (ys : List a) -> foldl f init (xs ++ ys) = foldl f (foldl f init xs) ys
|
||||
foldlAppend f init [] ys = Refl
|
||||
foldlAppend f init (x::xs) ys = rewrite foldlAppend f (f init x) xs ys in Refl
|
||||
|
||||
export
|
||||
filterAppend : (f : a -> Bool) -> (xs, ys : List a) -> filter f (xs ++ ys) = filter f xs ++ filter f ys
|
||||
filterAppend f [] ys = Refl
|
||||
filterAppend f (x::xs) ys with (f x)
|
||||
_ | False = rewrite filterAppend f xs ys in Refl
|
||||
_ | True = rewrite filterAppend f xs ys in Refl
|
||||
|
||||
export
|
||||
mapMaybeFusion : (g : b -> Maybe c) -> (f : a -> Maybe b) -> (xs : List a) -> mapMaybe g (mapMaybe f xs) = mapMaybe (f >=> g) xs
|
||||
mapMaybeFusion g f [] = Refl
|
||||
mapMaybeFusion g f (x::xs) with (f x)
|
||||
_ | Nothing = mapMaybeFusion g f xs
|
||||
_ | (Just y) with (g y)
|
||||
_ | Nothing = mapMaybeFusion g f xs
|
||||
_ | (Just z) = rewrite mapMaybeFusion g f xs in Refl
|
||||
|
||||
export
|
||||
mapMaybeAppend : (f : a -> Maybe b) -> (xs, ys : List a) -> mapMaybe f (xs ++ ys) = mapMaybe f xs ++ mapMaybe f ys
|
||||
mapMaybeAppend f [] ys = Refl
|
||||
mapMaybeAppend f (x::xs) ys with (f x)
|
||||
_ | Nothing = rewrite mapMaybeAppend f xs ys in Refl
|
||||
_ | (Just y) = rewrite mapMaybeAppend f xs ys in Refl
|
||||
|
||||
export
|
||||
mapFusion : (g : b -> c) -> (f : a -> b) -> (xs : List a) -> map g (map f xs) = map (g . f) xs
|
||||
mapFusion g f [] = Refl
|
||||
mapFusion g f (x::xs) = rewrite mapFusion g f xs in Refl
|
||||
|
||||
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
|
||||
|
@ -162,27 +162,6 @@ public export
|
||||
Ord a => Ord (List1 a) where
|
||||
compare xs ys = compare (forget xs) (forget ys)
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Properties
|
||||
|
||||
export
|
||||
consInjective : (x ::: xs) === (y ::: ys) -> (x === y, xs === ys)
|
||||
consInjective Refl = (Refl, Refl)
|
||||
|
||||
export
|
||||
{x : a} -> Injective (x :::) where
|
||||
injective Refl = Refl
|
||||
|
||||
export
|
||||
{ys : List a} -> Injective (::: ys) where
|
||||
injective Refl = Refl
|
||||
|
||||
||| Proof that the length of a List1 is the same as the length
|
||||
||| of the List it represents.
|
||||
export
|
||||
listLength : (xs : List1 a) -> length xs = length (forget xs)
|
||||
listLength (head ::: tail) = Refl
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Zippable
|
||||
|
||||
@ -223,6 +202,16 @@ Zippable List1 where
|
||||
(bs, cs, ds) = unzipWith3' xs in
|
||||
(b :: bs, c :: cs, d :: ds)
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Uninhabited
|
||||
|
||||
export
|
||||
Uninhabited a => Uninhabited (List1 a) where
|
||||
uninhabited (hd ::: _) = uninhabited hd
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Filtering
|
||||
|
||||
public export %inline
|
||||
filter : (a -> Bool) -> List1 a -> Maybe $ List1 a
|
||||
filter f = fromList . filter f . forget
|
||||
|
69
libs/base/Data/List1/Properties.idr
Normal file
69
libs/base/Data/List1/Properties.idr
Normal file
@ -0,0 +1,69 @@
|
||||
||| This module contains stuff that may use functions from `Data.List`.
|
||||
||| This separation is needed because `Data.List` uses `List1` type inside it,
|
||||
||| thus the core of `List1` must not import `Data.List`.
|
||||
module Data.List1.Properties
|
||||
|
||||
import Control.Function
|
||||
|
||||
import public Data.Maybe
|
||||
import Data.List
|
||||
import Data.List1
|
||||
|
||||
%default total
|
||||
|
||||
export
|
||||
consInjective : (x ::: xs) === (y ::: ys) -> (x === y, xs === ys)
|
||||
consInjective Refl = (Refl, Refl)
|
||||
|
||||
export
|
||||
{x : a} -> Injective (x :::) where
|
||||
injective Refl = Refl
|
||||
|
||||
export
|
||||
{ys : List a} -> Injective (::: ys) where
|
||||
injective Refl = Refl
|
||||
|
||||
||| Proof that the length of a List1 is the same as the length
|
||||
||| of the List it represents.
|
||||
export
|
||||
listLength : (xs : List1 a) -> length xs = length (forget xs)
|
||||
listLength (head ::: tail) = Refl
|
||||
|
||||
------------------------------------------------------------------------
|
||||
-- Properties of append involving usual `List`'s append
|
||||
|
||||
export
|
||||
appendlNilRightNeutral : (l : List1 a) -> l `appendl` [] = l
|
||||
appendlNilRightNeutral (_:::xs) = rewrite appendNilRightNeutral xs in Refl
|
||||
|
||||
export
|
||||
lappendNilLeftNeutral : (l : List1 a) -> [] `lappend` l = l
|
||||
lappendNilLeftNeutral l = Refl
|
||||
|
||||
export
|
||||
appendAssociative : (l, c, r : List1 a) -> l ++ (c ++ r) = (l ++ c) ++ r
|
||||
appendAssociative (l:::ls) (c:::cs) (r:::rs) = rewrite appendAssociative ls (c::cs) (r::rs) in Refl
|
||||
|
||||
export
|
||||
toListAppendl : (xs : List1 a) -> (ys : List a) -> toList (xs `appendl` ys) = forget xs ++ ys
|
||||
toListAppendl (x:::xs) ys = Refl
|
||||
|
||||
export
|
||||
toListLappend : (xs : List a) -> (ys : List1 a) -> toList (xs `lappend` ys) = xs ++ forget ys
|
||||
toListLappend [] ys = Refl
|
||||
toListLappend (x::xs) ys = Refl
|
||||
|
||||
export
|
||||
toListAppend : (xs, ys : List1 a) -> toList (xs ++ ys) = toList xs ++ toList ys
|
||||
toListAppend (x:::xs) (y:::ys) = Refl
|
||||
|
||||
export
|
||||
fromListAppend : (xs, ys : List a) -> fromList (xs ++ ys) = (fromList xs <+> fromList ys) @{Deep}
|
||||
fromListAppend [] ys with (fromList ys)
|
||||
_ | Nothing = Refl
|
||||
_ | (Just _) = Refl
|
||||
fromListAppend (x::xs) ys with (fromList ys) proof prf
|
||||
fromListAppend (x::xs) [] | Nothing = rewrite appendNilRightNeutral xs in Refl
|
||||
fromListAppend (x::xs) (y::ys) | (Just $ l:::ls) = do
|
||||
let (prfL, prfLs) = consInjective $ injective prf
|
||||
rewrite prfL; rewrite prfLs; Refl
|
@ -69,3 +69,18 @@ public export
|
||||
filter : (a -> Bool) -> Maybe a -> Maybe a
|
||||
filter _ Nothing = Nothing
|
||||
filter f (Just x) = toMaybe (f x) x
|
||||
|
||||
namespace Semigroup
|
||||
|
||||
public export
|
||||
[Deep] Semigroup a => Semigroup (Maybe a) where
|
||||
Nothing <+> Nothing = Nothing
|
||||
Just l <+> Nothing = Just l
|
||||
Nothing <+> Just r = Just r
|
||||
Just l <+> Just r = Just $ l <+> r
|
||||
|
||||
namespace Monoid
|
||||
|
||||
public export
|
||||
[Deep] Semigroup a => Monoid (Maybe a) using Semigroup.Deep where
|
||||
neutral = Nothing
|
||||
|
@ -21,22 +21,6 @@ public export
|
||||
asList : SnocList type -> List type
|
||||
asList = (reverse . cast)
|
||||
|
||||
public export
|
||||
Eq a => Eq (SnocList a) where
|
||||
(==) Lin Lin = True
|
||||
(==) (sx :< x) (sy :< y) = x == y && sx == sy
|
||||
(==) _ _ = False
|
||||
|
||||
public export
|
||||
Ord a => Ord (SnocList a) where
|
||||
compare Lin Lin = EQ
|
||||
compare Lin (sx :< x) = LT
|
||||
compare (sx :< x) Lin = GT
|
||||
compare (sx :< x) (sy :< y)
|
||||
= case compare sx sy of
|
||||
EQ => compare x y
|
||||
c => c
|
||||
|
||||
||| True iff input is Lin
|
||||
public export
|
||||
isLin : SnocList a -> Bool
|
||||
@ -49,16 +33,6 @@ isSnoc : SnocList a -> Bool
|
||||
isSnoc Lin = False
|
||||
isSnoc (sx :< x) = True
|
||||
|
||||
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
|
||||
|
||||
export
|
||||
Show a => Show (SnocList a) where
|
||||
show xs = concat ("[< " :: intersperse ", " (show' [] xs) ++ ["]"])
|
||||
@ -225,6 +199,11 @@ cons x sx = [< x] ++ sx
|
||||
|
||||
--- Folds ---
|
||||
|
||||
export
|
||||
foldAppend : (f : acc -> a -> acc) -> (init : acc) -> (sx, sy : SnocList a) -> foldl f init (sx ++ sy) = foldl f (foldl f init sx) sy
|
||||
foldAppend f init sx [<] = Refl
|
||||
foldAppend f init sx (sy :< x) = rewrite foldAppend f init sx sy in Refl
|
||||
|
||||
export
|
||||
snocFoldlAsListFoldl : (f : acc -> a -> acc) -> (init : acc) -> (xs : SnocList a) -> foldl f init xs = foldl f init (toList xs)
|
||||
snocFoldlAsListFoldl f init [<] = Refl
|
||||
@ -233,3 +212,122 @@ snocFoldlAsListFoldl f init (sx :< x) = do
|
||||
rewrite snocFoldlAsListFoldl f init sx
|
||||
rewrite foldlAppend f init (cast sx) [x]
|
||||
Refl
|
||||
|
||||
--- Filtering ---
|
||||
|
||||
export
|
||||
filterAppend : (f : a -> Bool) -> (sx, sy : SnocList a) -> filter f (sx ++ sy) = filter f sx ++ filter f sy
|
||||
filterAppend f sx [<] = Refl
|
||||
filterAppend f sx (sy :< x) with (f x)
|
||||
_ | False = filterAppend f sx sy
|
||||
_ | True = rewrite filterAppend f sx sy in Refl
|
||||
|
||||
export
|
||||
toListFilter : (f : a -> Bool) -> (sx : SnocList a) -> toList (filter f sx) = filter f (toList sx)
|
||||
toListFilter f [<] = Refl
|
||||
toListFilter f (sx :< x) = do
|
||||
rewrite chipsAsListAppend sx [x]
|
||||
rewrite filterAppend f (cast sx) [x]
|
||||
rewrite filterStepLemma
|
||||
rewrite toListFilter f sx
|
||||
Refl
|
||||
where
|
||||
filterStepLemma : toList (filter f (sx :< x)) = toList (filter f sx) ++ filter f [x]
|
||||
filterStepLemma with (f x)
|
||||
_ | False = rewrite appendNilRightNeutral $ toList $ filter f sx in Refl
|
||||
_ | True = rewrite chipsAsListAppend (filter f sx) [x] in Refl
|
||||
|
||||
export
|
||||
filterCast : (f : a -> Bool) -> (xs : List a) -> filter f (cast {to=SnocList a} xs) = cast (filter f xs)
|
||||
filterCast f [] = Refl
|
||||
filterCast f (x::xs) = do
|
||||
rewrite fishAsSnocAppend [<x] xs
|
||||
rewrite filterAppend f [<x] (cast xs)
|
||||
rewrite filterStepLemma
|
||||
rewrite filterCast f xs
|
||||
Refl
|
||||
where
|
||||
filterStepLemma : cast (filter f (x::xs)) = filter f [<x] ++ cast (filter f xs)
|
||||
filterStepLemma with (f x)
|
||||
_ | False = rewrite appendLinLeftNeutral $ [<] <>< filter f xs in Refl
|
||||
_ | True = rewrite fishAsSnocAppend [<x] (filter f xs) in Refl
|
||||
|
||||
--- Functor map ---
|
||||
|
||||
export
|
||||
mapFusion : (g : b -> c) -> (f : a -> b) -> (sx : SnocList a) -> map g (map f sx) = map (g . f) sx
|
||||
mapFusion g f [<] = Refl
|
||||
mapFusion g f (sx :< x) = rewrite mapFusion g f sx in Refl
|
||||
|
||||
export
|
||||
mapAppend : (f : a -> b) -> (sx, sy : SnocList a) -> map f (sx ++ sy) = map f sx ++ map f sy
|
||||
mapAppend f sx [<] = Refl
|
||||
mapAppend f sx (sy :< x) = rewrite mapAppend f sx sy in Refl
|
||||
|
||||
export
|
||||
toListMap : (f : a -> b) -> (sx : SnocList a) -> toList (map f sx) = map f (toList sx)
|
||||
toListMap f [<] = Refl
|
||||
toListMap f (sx :< x) = do
|
||||
rewrite chipsAsListAppend (map f sx) [f x]
|
||||
rewrite chipsAsListAppend sx [x]
|
||||
rewrite mapAppend f (toList sx) [x]
|
||||
rewrite toListMap f sx
|
||||
Refl
|
||||
|
||||
export
|
||||
mapCast : (f : a -> b) -> (xs : List a) -> map f (cast {to=SnocList a} xs) = cast (map f xs)
|
||||
mapCast f [] = Refl
|
||||
mapCast f (x::xs) = do
|
||||
rewrite fishAsSnocAppend [<f x] (map f xs)
|
||||
rewrite fishAsSnocAppend [<x] xs
|
||||
rewrite mapAppend f [<x] (cast xs)
|
||||
rewrite mapCast f xs
|
||||
Refl
|
||||
|
||||
--- mapMaybe ---
|
||||
|
||||
export
|
||||
mapMaybeFusion : (g : b -> Maybe c) -> (f : a -> Maybe b) -> (sx : SnocList a) -> mapMaybe g (mapMaybe f sx) = mapMaybe (f >=> g) sx
|
||||
mapMaybeFusion g f [<] = Refl
|
||||
mapMaybeFusion g f (sx :< x) with (f x)
|
||||
_ | Nothing = mapMaybeFusion g f sx
|
||||
_ | (Just y) with (g y)
|
||||
_ | Nothing = mapMaybeFusion g f sx
|
||||
_ | (Just z) = rewrite mapMaybeFusion g f sx in Refl
|
||||
|
||||
export
|
||||
mapMaybeAppend : (f : a -> Maybe b) -> (sx, sy : SnocList a) -> mapMaybe f (sx ++ sy) = mapMaybe f sx ++ mapMaybe f sy
|
||||
mapMaybeAppend f sx [<] = Refl
|
||||
mapMaybeAppend f sx (sy :< x) with (f x)
|
||||
_ | Nothing = mapMaybeAppend f sx sy
|
||||
_ | (Just y) = rewrite mapMaybeAppend f sx sy in Refl
|
||||
|
||||
export
|
||||
toListMapMaybe : (f : a -> Maybe b) -> (sx : SnocList a) -> toList (mapMaybe f sx) = mapMaybe f (toList sx)
|
||||
toListMapMaybe f [<] = Refl
|
||||
toListMapMaybe f (sx :< x) = do
|
||||
rewrite chipsAsListAppend sx [x]
|
||||
rewrite mapMaybeAppend f (toList sx) [x]
|
||||
rewrite mapMaybeStepLemma
|
||||
rewrite toListMapMaybe f sx
|
||||
Refl
|
||||
where
|
||||
mapMaybeStepLemma : toList (mapMaybe f (sx :< x)) = toList (mapMaybe f sx) ++ mapMaybe f [x]
|
||||
mapMaybeStepLemma with (f x)
|
||||
_ | Nothing = rewrite appendNilRightNeutral $ toList $ mapMaybe f sx in Refl
|
||||
_ | (Just y) = rewrite chipsAsListAppend (mapMaybe f sx) [y] in Refl
|
||||
|
||||
export
|
||||
mapMaybeCast : (f : a -> Maybe b) -> (xs : List a) -> mapMaybe f (cast {to=SnocList a} xs) = cast (mapMaybe f xs)
|
||||
mapMaybeCast f [] = Refl
|
||||
mapMaybeCast f (x::xs) = do
|
||||
rewrite fishAsSnocAppend [<x] xs
|
||||
rewrite mapMaybeAppend f [<x] (cast xs)
|
||||
rewrite mapMaybeStepLemma
|
||||
rewrite mapMaybeCast f xs
|
||||
Refl
|
||||
where
|
||||
mapMaybeStepLemma : cast (mapMaybe f (x::xs)) = mapMaybe f [<x] ++ cast (mapMaybe f xs)
|
||||
mapMaybeStepLemma with (f x)
|
||||
_ | Nothing = rewrite appendLinLeftNeutral $ [<] <>< mapMaybe f xs in Refl
|
||||
_ | (Just y) = rewrite fishAsSnocAppend [<y] (mapMaybe f xs) in Refl
|
||||
|
@ -6,6 +6,7 @@ import Data.Either
|
||||
import Data.Nat
|
||||
import Data.List
|
||||
import Data.List1
|
||||
import Data.List1.Properties
|
||||
|
||||
import public Decidable.Equality.Core as Decidable.Equality
|
||||
|
||||
|
@ -60,6 +60,7 @@ modules = Control.App,
|
||||
Data.List.Views,
|
||||
Data.List.Quantifiers,
|
||||
Data.List1,
|
||||
Data.List1.Properties,
|
||||
Data.Maybe,
|
||||
Data.Morphisms,
|
||||
Data.Nat,
|
||||
|
@ -1,7 +1,7 @@
|
||||
||| Contains:
|
||||
|||
|
||||
||| 1. Tail recursive versions of the list processing functions from
|
||||
||| Data.List.
|
||||
||| List.
|
||||
|||
|
||||
||| 2. Extensional equality proofs that these variants are
|
||||
||| (extensionally) equivalent to their non-tail-recursive
|
||||
@ -37,7 +37,7 @@ lengthAccSucc [] _ = Refl
|
||||
lengthAccSucc (_::xs) n = rewrite lengthAccSucc xs (S n) in cong S Refl
|
||||
|
||||
export
|
||||
length_ext : (xs : List a) -> List.length xs = Data.List.TailRec.length xs
|
||||
length_ext : (xs : List a) -> List.length xs = List.TailRec.length xs
|
||||
length_ext [] = Refl
|
||||
length_ext (_::xs) = rewrite length_ext xs in sym $ lengthAccSucc xs Z
|
||||
|
||||
@ -53,15 +53,15 @@ take n xs = take_aux n xs []
|
||||
export
|
||||
take_ext :
|
||||
(n : Nat) -> (xs : List a) ->
|
||||
Data.List.take n xs = Data.List.TailRec.take n xs
|
||||
List.take n xs = List.TailRec.take n xs
|
||||
take_ext n xs = Calc $
|
||||
|~ Data.List.take n xs
|
||||
~~ reverse [] ++ (Data.List.take n xs) ...( Refl )
|
||||
~~ reverseOnto (Data.List.take n xs) [] ...( sym (revOnto (Data.List.take n xs) []) )
|
||||
|~ List.take n xs
|
||||
~~ reverse [] ++ (List.take n xs) ...( Refl )
|
||||
~~ reverseOnto (List.take n xs) [] ...( sym (revOnto (List.take n xs) []) )
|
||||
~~ take_aux n xs [] ...( sym (lemma n xs []) )
|
||||
where
|
||||
lemma : (n : Nat) -> (xs, acc : List a) ->
|
||||
take_aux n xs acc = reverseOnto (Data.List.take n xs) acc
|
||||
take_aux n xs acc = reverseOnto (List.take n xs) acc
|
||||
lemma Z xs acc = Refl
|
||||
lemma (S n) [] acc = Refl
|
||||
lemma (S n) (x::xs) acc = lemma n xs (x :: acc)
|
||||
@ -71,7 +71,7 @@ span_aux : (a -> Bool) -> List a -> List a -> (List a, List a)
|
||||
span_aux p [] acc = (reverseOnto [] acc, [])
|
||||
span_aux p (x::xs) acc =
|
||||
if p x then
|
||||
Data.List.TailRec.span_aux p xs (x :: acc)
|
||||
List.TailRec.span_aux p xs (x :: acc)
|
||||
else
|
||||
(reverseOnto [] acc, x::xs)
|
||||
|
||||
@ -83,26 +83,26 @@ coe : (f : (i : a) -> Type) -> i = i' -> f i -> f i'
|
||||
coe f Refl x = x
|
||||
|
||||
span_aux_ext : (p : a -> Bool) -> (xs, acc : List a) ->
|
||||
(reverseOnto (fst $ Data.List.span p xs) acc, snd $ Data.List.span p xs)
|
||||
(reverseOnto (fst $ List.span p xs) acc, snd $ List.span p xs)
|
||||
=
|
||||
span_aux p xs acc
|
||||
span_aux_ext p [] acc = Refl
|
||||
-- This is disgusting. Please teach me a better way.
|
||||
span_aux_ext p (x::xs) acc with (@@(p x), @@(Data.List.span p xs))
|
||||
span_aux_ext p (x::xs) acc with (@@(p x), @@(List.span p xs))
|
||||
span_aux_ext p (x::xs) acc | ((True ** px_tru), ((pre, rest)**dl_pf)) =
|
||||
rewrite px_tru in
|
||||
rewrite dl_pf in
|
||||
let u = span_aux_ext p xs (x::acc) in
|
||||
coe (\u => (reverseOnto (x :: fst u) acc, snd u) =
|
||||
Data.List.TailRec.span_aux p xs (x :: acc)) dl_pf u
|
||||
List.TailRec.span_aux p xs (x :: acc)) dl_pf u
|
||||
span_aux_ext p (x::xs) acc | ((False**px_fls), ((pre,rest)**dl_pf)) =
|
||||
rewrite px_fls in
|
||||
Refl
|
||||
|
||||
export
|
||||
span_ext : (p : a -> Bool) -> (xs : List a) ->
|
||||
Data.List.span p xs = Data.List.TailRec.span p xs
|
||||
span_ext p xs with (@@(Data.List.span p xs))
|
||||
List.span p xs = List.TailRec.span p xs
|
||||
span_ext p xs with (@@(List.span p xs))
|
||||
span_ext p xs | ((pre, rest) ** pf) =
|
||||
rewrite pf in
|
||||
let u = span_aux_ext p xs [] in
|
||||
@ -110,16 +110,16 @@ span_ext p xs with (@@(Data.List.span p xs))
|
||||
|
||||
export
|
||||
break : (a -> Bool) -> List a -> (List a, List a)
|
||||
break p xs = Data.List.TailRec.span (not . p) xs
|
||||
break p xs = List.TailRec.span (not . p) xs
|
||||
|
||||
export
|
||||
break_ext : (p : a -> Bool) -> (xs : List a) ->
|
||||
Data.List.break p xs = Data.List.TailRec.break p xs
|
||||
List.break p xs = List.TailRec.break p xs
|
||||
break_ext p xs = span_ext (not . p) xs
|
||||
|
||||
splitOnto : List (List a) -> (a -> Bool) -> List a -> List1 (List a)
|
||||
splitOnto acc p xs =
|
||||
case Data.List.break p xs of
|
||||
case List.break p xs of
|
||||
(chunk, [] ) => reverseOnto (chunk ::: []) acc
|
||||
(chunk, (c::rest)) => splitOnto (chunk::acc) p $ assert_smaller xs rest
|
||||
|
||||
@ -128,9 +128,9 @@ split : (a -> Bool) -> List a -> List1 (List a)
|
||||
split p xs = splitOnto [] p xs
|
||||
|
||||
splitOnto_ext : (acc : List (List a)) -> (p : a -> Bool) -> (xs : List a) ->
|
||||
reverseOnto (Data.List.split p xs) acc
|
||||
= Data.List.TailRec.splitOnto acc p xs
|
||||
splitOnto_ext acc p xs with (@@(Data.List.break p xs))
|
||||
reverseOnto (List.split p xs) acc
|
||||
= List.TailRec.splitOnto acc p xs
|
||||
splitOnto_ext acc p xs with (@@(List.break p xs))
|
||||
splitOnto_ext acc p xs | ((chunk, [] )**pf) =
|
||||
rewrite pf in
|
||||
Refl
|
||||
@ -141,7 +141,7 @@ splitOnto_ext acc p xs with (@@(Data.List.break p xs))
|
||||
|
||||
export
|
||||
split_ext : (p : a -> Bool) -> (xs : List a) ->
|
||||
Data.List.split p xs = Data.List.TailRec.split p xs
|
||||
List.split p xs = List.TailRec.split p xs
|
||||
split_ext p xs = splitOnto_ext [] p xs
|
||||
|
||||
|
||||
@ -155,11 +155,11 @@ splitAt : (n : Nat) -> (xs : List a) -> (List a, List a)
|
||||
splitAt n xs = splitAtOnto [] n xs
|
||||
|
||||
splitAtOnto_ext : (acc : List a) -> (n : Nat) -> (xs : List a) ->
|
||||
(reverseOnto (fst $ Data.List.splitAt n xs) acc, snd $ Data.List.splitAt n xs)
|
||||
(reverseOnto (fst $ List.splitAt n xs) acc, snd $ List.splitAt n xs)
|
||||
= splitAtOnto acc n xs
|
||||
splitAtOnto_ext acc Z xs = Refl
|
||||
splitAtOnto_ext acc (S n) [] = Refl
|
||||
splitAtOnto_ext acc (S n) (x::xs) with (@@(Data.List.splitAt n xs))
|
||||
splitAtOnto_ext acc (S n) (x::xs) with (@@(List.splitAt n xs))
|
||||
splitAtOnto_ext acc (S n) (x::xs) | ((tk, dr)**pf) =
|
||||
rewrite pf in
|
||||
let u = splitAtOnto_ext (x::acc) n xs in
|
||||
@ -168,9 +168,9 @@ splitAtOnto_ext acc (S n) (x::xs) with (@@(Data.List.splitAt n xs))
|
||||
|
||||
export
|
||||
splitAt_ext : (n : Nat) -> (xs : List a) ->
|
||||
Data.List.splitAt n xs =
|
||||
Data.List.TailRec.splitAt n xs
|
||||
splitAt_ext n xs with (@@(Data.List.splitAt n xs))
|
||||
List.splitAt n xs =
|
||||
List.TailRec.splitAt n xs
|
||||
splitAt_ext n xs with (@@(List.splitAt n xs))
|
||||
splitAt_ext n xs | ((tk, dr)**pf) =
|
||||
rewrite pf in
|
||||
rewrite sym $ splitAtOnto_ext [] n xs in
|
||||
@ -190,11 +190,11 @@ partition : (a -> Bool) -> List a -> (List a, List a)
|
||||
partition p xs = partitionOnto [] [] p xs
|
||||
|
||||
partitionOnto_ext : (lfts, rgts : List a) -> (p : a -> Bool) -> (xs : List a) ->
|
||||
(reverseOnto (fst $ Data.List.partition p xs) lfts
|
||||
,reverseOnto (snd $ Data.List.partition p xs) rgts)
|
||||
= Data.List.TailRec.partitionOnto lfts rgts p xs
|
||||
(reverseOnto (fst $ List.partition p xs) lfts
|
||||
,reverseOnto (snd $ List.partition p xs) rgts)
|
||||
= List.TailRec.partitionOnto lfts rgts p xs
|
||||
partitionOnto_ext lfts rgts p [] = Refl
|
||||
partitionOnto_ext lfts rgts p (x::xs) with (@@(p x), @@(Data.List.partition p xs))
|
||||
partitionOnto_ext lfts rgts p (x::xs) with (@@(p x), @@(List.partition p xs))
|
||||
partitionOnto_ext lfts rgts p (x::xs) | ((True **px_tru), ((dl_l, dl_r)**dl_pf))
|
||||
= rewrite px_tru in
|
||||
rewrite dl_pf in
|
||||
@ -233,8 +233,8 @@ intersperse sep (y::ys) = y :: mergeReplicate_aux sep ys []
|
||||
|
||||
export
|
||||
intersperse_ext : (sep : a) -> (xs : List a) ->
|
||||
Data.List.intersperse sep xs =
|
||||
Data.List.TailRec.intersperse sep xs
|
||||
List.intersperse sep xs =
|
||||
List.TailRec.intersperse sep xs
|
||||
intersperse_ext sep [] = Refl
|
||||
intersperse_ext sep (y::ys) = cong (y::) (sym $ mergeReplicate_ext sep ys [])
|
||||
|
||||
@ -250,7 +250,7 @@ mapMaybe : (a -> Maybe b) -> List a -> List b
|
||||
mapMaybe f xs = mapMaybeOnto [] f xs
|
||||
|
||||
mapMaybeOnto_ext : (acc : List b) -> (f : a -> Maybe b) -> (xs : List a) ->
|
||||
reverseOnto (Data.List.mapMaybe f xs) acc
|
||||
reverseOnto (List.mapMaybe f xs) acc
|
||||
=
|
||||
mapMaybeOnto acc f xs
|
||||
mapMaybeOnto_ext acc f [] = Refl
|
||||
@ -260,7 +260,7 @@ mapMaybeOnto_ext acc f (x::xs) with (f x)
|
||||
|
||||
export
|
||||
mapMaybe_ext : (f : a -> Maybe b) -> (xs : List a) ->
|
||||
Data.List.mapMaybe f xs = Data.List.TailRec.mapMaybe f xs
|
||||
List.mapMaybe f xs = List.TailRec.mapMaybe f xs
|
||||
mapMaybe_ext f xs = mapMaybeOnto_ext [] f xs
|
||||
|
||||
export
|
||||
@ -269,12 +269,12 @@ sorted [ ] = True
|
||||
sorted [x] = True
|
||||
sorted (x :: xs@(y :: ys)) = case (x <= y) of
|
||||
False => False
|
||||
True => Data.List.TailRec.sorted xs
|
||||
True => List.TailRec.sorted xs
|
||||
|
||||
export
|
||||
covering
|
||||
sorted_ext : Ord a => (xs : List a) ->
|
||||
Data.List.sorted xs = Data.List.TailRec.sorted xs
|
||||
List.sorted xs = List.TailRec.sorted xs
|
||||
sorted_ext [] = Refl
|
||||
sorted_ext [x] = Refl
|
||||
sorted_ext (x :: y :: ys) with (x <= y)
|
||||
@ -317,18 +317,18 @@ mergeBy order left right = mergeByOnto [] order left right
|
||||
export
|
||||
covering
|
||||
mergeBy_ext : (order : a -> a -> Ordering) -> (left, right : List a) ->
|
||||
Data.List.mergeBy order left right =
|
||||
Data.List.TailRec.mergeBy order left right
|
||||
List.mergeBy order left right =
|
||||
List.TailRec.mergeBy order left right
|
||||
mergeBy_ext order left right = mergeByOnto_ext [] order left right
|
||||
|
||||
export
|
||||
merge : Ord a => List a -> List a -> List a
|
||||
merge = Data.List.TailRec.mergeBy compare
|
||||
merge = List.TailRec.mergeBy compare
|
||||
|
||||
export
|
||||
covering
|
||||
merge_ext : Ord a => (left, right : List a) ->
|
||||
Data.List.merge left right = Data.List.TailRec.merge left right
|
||||
List.merge left right = List.TailRec.merge left right
|
||||
merge_ext left right = mergeBy_ext compare left right
|
||||
|
||||
|
||||
@ -346,30 +346,30 @@ sortBy : (cmp : a -> a -> Ordering) -> (xs : List a) -> List a
|
||||
sortBy cmp [] = []
|
||||
sortBy cmp [x] = [x]
|
||||
sortBy cmp zs = let (xs, ys) = sortBy_split zs in
|
||||
Data.List.TailRec.mergeBy cmp
|
||||
(Data.List.TailRec.sortBy cmp (assert_smaller zs xs))
|
||||
(Data.List.TailRec.sortBy cmp (assert_smaller zs ys))
|
||||
List.TailRec.mergeBy cmp
|
||||
(List.TailRec.sortBy cmp (assert_smaller zs xs))
|
||||
(List.TailRec.sortBy cmp (assert_smaller zs ys))
|
||||
|
||||
|
||||
{- Can't really finish this proof because Data.List doesn't export the definition of sortBy. -}
|
||||
{-
|
||||
export
|
||||
sortBy_ext : (cmp : a -> a -> Ordering) -> (xs : List a) ->
|
||||
Data.List.sortBy cmp xs = Data.List.TailRec.sortBy cmp xs
|
||||
List.sortBy cmp xs = List.TailRec.sortBy cmp xs
|
||||
sortBy_ext cmp [] = Refl
|
||||
sortBy_ext cmp [x] = Refl
|
||||
sortBy_ext cmp zs'@(z::zs) =
|
||||
Calc $
|
||||
|~ Data.List.sortBy cmp (z::zs)
|
||||
|~ List.sortBy cmp (z::zs)
|
||||
~~ (let (xs, ys) = sortBy_split zs' in
|
||||
Data.List.mergeBy cmp
|
||||
(Data.List.sortBy cmp xs)
|
||||
(Data.List.sortBy cmp ys))
|
||||
List.mergeBy cmp
|
||||
(List.sortBy cmp xs)
|
||||
(List.sortBy cmp ys))
|
||||
...( ?help0 )
|
||||
~~
|
||||
let (xs, ys) = sortBy_split (z::zs) in
|
||||
Data.List.TailRec.mergeBy cmp
|
||||
(Data.List.TailRec.sortBy cmp xs)
|
||||
(Data.List.TailRec.sortBy cmp ys)
|
||||
List.TailRec.mergeBy cmp
|
||||
(List.TailRec.sortBy cmp xs)
|
||||
(List.TailRec.sortBy cmp ys)
|
||||
...( ?help1 )
|
||||
-}
|
||||
|
@ -86,7 +86,7 @@ filter filePred dirPred (MkTree files dirs) = MkTree files' dirs' where
|
||||
files' = filter filePred files
|
||||
|
||||
dirs' : List (SubTree root)
|
||||
dirs' = flip mapMaybe dirs $ \ (dname ** iot) => do
|
||||
dirs' = flip List.mapMaybe dirs $ \ (dname ** iot) => do
|
||||
guard (dirPred dname)
|
||||
pure (dname ** map (assert_total (filter filePred dirPred)) iot)
|
||||
|
||||
|
@ -180,19 +180,3 @@ data SnocList a =
|
||||
(:<) (SnocList a) a
|
||||
|
||||
%name SnocList sx, sy, sz
|
||||
|
||||
|
||||
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
|
||||
|
@ -393,6 +393,7 @@ Ord a => Ord (List a) where
|
||||
c => c
|
||||
|
||||
namespace List
|
||||
|
||||
||| Concatenate one list with another.
|
||||
public export
|
||||
(++) : (xs, ys : List a) -> List a
|
||||
@ -405,6 +406,25 @@ namespace List
|
||||
length [] = Z
|
||||
length (x :: xs) = S (length xs)
|
||||
|
||||
||| Applied to a predicate and a list, returns the list of those elements that
|
||||
||| satisfy the predicate.
|
||||
public export
|
||||
filter : (p : a -> Bool) -> List a -> List a
|
||||
filter p [] = []
|
||||
filter p (x :: xs)
|
||||
= if p x
|
||||
then x :: filter p xs
|
||||
else filter p xs
|
||||
|
||||
||| Apply a partial function to the elements of a list, keeping the ones at which it is defined.
|
||||
public export
|
||||
mapMaybe : (a -> Maybe b) -> List a -> List b
|
||||
mapMaybe f [] = []
|
||||
mapMaybe f (x::xs) =
|
||||
case f x of
|
||||
Nothing => mapMaybe f xs
|
||||
Just j => j :: mapMaybe f xs
|
||||
|
||||
||| Reverse the second list, prepending its elements to the first.
|
||||
public export
|
||||
reverseOnto : List a -> List a -> List a
|
||||
@ -497,6 +517,64 @@ 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
|
||||
(==) (sx :< x) (sy :< y) = x == y && sx == sy
|
||||
(==) _ _ = False
|
||||
|
||||
public export
|
||||
Ord a => Ord (SnocList a) where
|
||||
compare Lin Lin = EQ
|
||||
compare Lin (sx :< x) = LT
|
||||
compare (sx :< x) Lin = GT
|
||||
compare (sx :< x) (sy :< y)
|
||||
= case compare sx sy of
|
||||
EQ => compare x y
|
||||
c => c
|
||||
|
||||
-- This works quickly because when string-concat builds the result, it allocates
|
||||
-- enough room in advance so there's only one allocation, rather than lots!
|
||||
--
|
||||
|
@ -594,7 +594,7 @@ inRange low high t = matches (takeUntil (greater high) t)
|
||||
||| Finds the values matching the exact interval input
|
||||
export
|
||||
exactRange : MeasureRM a => FilePos -> FilePos -> PosMap a -> List a
|
||||
exactRange low high t = flip mapMaybe (inRange low high t) $ \ a =>
|
||||
exactRange low high t = flip List.mapMaybe (inRange low high t) $ \ a =>
|
||||
do let (MkRange rng _) = measureRM a
|
||||
guard (rng == (low, high))
|
||||
pure a
|
||||
|
@ -87,7 +87,7 @@ filter filePred dirPred (MkTree files dirs) = MkTree files' dirs' where
|
||||
files' = filter filePred files
|
||||
|
||||
dirs' : List (SubTree root)
|
||||
dirs' = flip mapMaybe dirs $ \ (dname ** iot) => do
|
||||
dirs' = flip List.mapMaybe dirs $ \ (dname ** iot) => do
|
||||
guard (dirPred dname)
|
||||
pure (dname ** map (assert_total (filter filePred dirPred)) iot)
|
||||
|
||||
|
@ -198,7 +198,7 @@ findUniqueBindableNames fc arg env used t
|
||||
let ctxt = gamma defs
|
||||
ns <- map catMaybes $ for assoc $ \ (n, _) => do
|
||||
ns <- lookupCtxtName (UN (Basic n)) ctxt
|
||||
let ns = flip mapMaybe ns $ \(n, _, d) =>
|
||||
let ns = flip List.mapMaybe ns $ \(n, _, d) =>
|
||||
case definition d of
|
||||
-- do not warn about holes: `?a` is not actually
|
||||
-- getting shadowed as it will not become a
|
||||
|
@ -1,2 +1,2 @@
|
||||
:exec printLn (filter even [S Z,2,3,4,5,6])
|
||||
:exec printLn (Main.filter even [S Z,2,3,4,5,6])
|
||||
:q
|
||||
|
@ -20,6 +20,20 @@ and:
|
||||
List ?a
|
||||
Mismatch between: Nat and List ?a.
|
||||
|
||||
Error:12:18--12:19
|
||||
08 | length [] = Z
|
||||
09 | length (x :: xs) = S (length xs)
|
||||
10 |
|
||||
11 | wrong : Nat -> Nat
|
||||
12 | wrong x = length x
|
||||
^
|
||||
|
||||
If Prelude.SnocList.length: When unifying:
|
||||
Nat
|
||||
and:
|
||||
SnocList ?a
|
||||
Mismatch between: Nat and SnocList ?a.
|
||||
|
||||
Error:12:18--12:19
|
||||
08 | length [] = Z
|
||||
09 | length (x :: xs) = S (length xs)
|
||||
|
@ -2,6 +2,9 @@
|
||||
LOG 0: Name: Prelude.Types.List.++
|
||||
LOG 0: Type: (%pi Rig0 Implicit (Just a) %type (%pi RigW Explicit (Just xs) (Prelude.Basics.List a) (%pi RigW Explicit (Just ys) (Prelude.Basics.List a) (Prelude.Basics.List a))))
|
||||
LOG 0: Pretty Type: (xs : List a) -> (ys : List a) -> List a
|
||||
LOG 0: Name: Prelude.Types.SnocList.++
|
||||
LOG 0: Type: (%pi Rig0 Implicit (Just a) %type (%pi RigW Explicit (Just sx) (Prelude.Basics.SnocList a) (%pi RigW Explicit (Just sy) (Prelude.Basics.SnocList a) (Prelude.Basics.SnocList a))))
|
||||
LOG 0: Pretty Type: (sx : SnocList a) -> (sy : SnocList a) -> SnocList a
|
||||
LOG 0: Name: Prelude.Types.String.++
|
||||
LOG 0: Type: (%pi RigW Explicit (Just x) String (%pi RigW Explicit (Just y) String String))
|
||||
LOG 0: Pretty Type: (x : String) -> (y : String) -> String
|
||||
|
@ -2,6 +2,8 @@ module WithProof
|
||||
|
||||
%default total
|
||||
|
||||
%hide List.filter
|
||||
|
||||
filter : (p : a -> Bool) -> (xs : List a) -> List a
|
||||
filter p [] = []
|
||||
filter p (x :: xs) with (p x)
|
||||
|
@ -1,3 +1,5 @@
|
||||
%hide List.filter
|
||||
|
||||
filter : (p : a -> Bool) -> List a -> List a
|
||||
filter p [] = []
|
||||
filter p (x :: xs) with (p x)
|
||||
|
@ -1,2 +1,2 @@
|
||||
:exec printLn (filter even [S Z,2,3,4,5,6])
|
||||
:exec printLn (Main.filter even [S Z,2,3,4,5,6])
|
||||
:q
|
||||
|
Loading…
Reference in New Issue
Block a user