mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 04:33:45 +03:00
Moving Data.List.HasLength into base (#2844)
This commit is contained in:
parent
60a3b80c19
commit
24ac56de88
11
CHANGELOG.md
11
CHANGELOG.md
@ -71,12 +71,23 @@
|
|||||||
* Add an alias for `HVect` to `All id` in `Data.Vect.Quantifiers.All`. This is the
|
* Add an alias for `HVect` to `All id` in `Data.Vect.Quantifiers.All`. This is the
|
||||||
approach to getting a heterogeneous Vect of elements that is generall preferred by
|
approach to getting a heterogeneous Vect of elements that is generall preferred by
|
||||||
the community vs. a standalone type as seen in `contrib`.
|
the community vs. a standalone type as seen in `contrib`.
|
||||||
|
* Add Data.List.HasLength from the compiler codebase slash contrib library but
|
||||||
|
adopt the type signature from the compiler codebase and some of the naming
|
||||||
|
from the contrib library. The type ended up being `HasLength n xs` rather than
|
||||||
|
`HasLength xs n`.
|
||||||
|
|
||||||
#### System
|
#### System
|
||||||
|
|
||||||
* Changes `getNProcessors` to return the number of online processors rather than
|
* Changes `getNProcessors` to return the number of online processors rather than
|
||||||
the number of configured processors.
|
the number of configured processors.
|
||||||
|
|
||||||
|
|
||||||
|
#### Contrib
|
||||||
|
* Remove Data.List.HasLength from contrib library but add it to the base library
|
||||||
|
with the type signature from the compiler codebase and some of the naming
|
||||||
|
from the contrib library. The type ended up being `HasLength n xs` rather than
|
||||||
|
`HasLength xs n`.
|
||||||
|
|
||||||
### Other Changes
|
### Other Changes
|
||||||
* The `data` subfolder of an installed or local dependency package is now automatically
|
* The `data` subfolder of an installed or local dependency package is now automatically
|
||||||
recognized as a "data" directory by Idris 2. See the
|
recognized as a "data" directory by Idris 2. See the
|
||||||
|
@ -12,8 +12,8 @@
|
|||||||
|||
|
|||
|
||||||
||| We would write either of:
|
||| We would write either of:
|
||||||
||| ```idris example
|
||| ```idris example
|
||||||
||| f1 : (n : Nat) -> (0 _ : HasLength xs n) -> P xs
|
||| f1 : (n : Nat) -> (0 _ : HasLength n xs) -> P xs
|
||||||
||| f2 : (n : Subset n (HasLength xs)) -> P xs
|
||| f2 : (n : Subset n (flip HasLength xs)) -> P xs
|
||||||
||| ```
|
||| ```
|
||||||
|||
|
|||
|
||||||
||| See `sucR` for an example where the update to the runtime-relevant Nat is O(1)
|
||| See `sucR` for an example where the update to the runtime-relevant Nat is O(1)
|
||||||
@ -23,6 +23,7 @@ module Data.List.HasLength
|
|||||||
|
|
||||||
import Data.DPair
|
import Data.DPair
|
||||||
import Data.List
|
import Data.List
|
||||||
|
import Data.Nat
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
|
||||||
@ -31,27 +32,45 @@ import Data.List
|
|||||||
|
|
||||||
||| Ensure that the list's length is the provided natural number
|
||| Ensure that the list's length is the provided natural number
|
||||||
public export
|
public export
|
||||||
data HasLength : List a -> Nat -> Type where
|
data HasLength : Nat -> List a -> Type where
|
||||||
Z : HasLength [] Z
|
Z : HasLength Z []
|
||||||
S : HasLength xs n -> HasLength (x :: xs) (S n)
|
S : HasLength n xs -> HasLength (S n) (x :: xs)
|
||||||
|
|
||||||
|
||| This specification corresponds to the length function
|
||||||
|
export
|
||||||
|
hasLength : (xs : List a) -> HasLength (length xs) xs
|
||||||
|
hasLength [] = Z
|
||||||
|
hasLength (_ :: xs) = S (hasLength xs)
|
||||||
|
|
||||||
|
export
|
||||||
|
take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs)
|
||||||
|
take Z _ = Z
|
||||||
|
take (S n) (x :: xs) = S (take n xs)
|
||||||
|
|
||||||
------------------------------------------------------------------------
|
------------------------------------------------------------------------
|
||||||
-- Properties
|
-- Properties
|
||||||
|
|
||||||
||| The length is unique
|
||| The length is unique
|
||||||
export
|
export
|
||||||
hasLengthUnique : HasLength xs m -> HasLength xs n -> m === n
|
hasLengthUnique : HasLength m xs -> HasLength n xs -> m === n
|
||||||
hasLengthUnique Z Z = Refl
|
hasLengthUnique Z Z = Refl
|
||||||
hasLengthUnique (S p) (S q) = cong S (hasLengthUnique p q)
|
hasLengthUnique (S p) (S q) = cong S (hasLengthUnique p q)
|
||||||
|
|
||||||
||| This specification corresponds to the length function
|
|
||||||
export
|
export
|
||||||
hasLength : (xs : List a) -> HasLength xs (length xs)
|
hasLengthAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys)
|
||||||
hasLength [] = Z
|
hasLengthAppend Z ys = ys
|
||||||
hasLength (_ :: xs) = S (hasLength xs)
|
hasLengthAppend (S xs) ys = S (hasLengthAppend xs ys)
|
||||||
|
|
||||||
|
hasLengthReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs)
|
||||||
|
hasLengthReverseOnto p Z = rewrite plusZeroRightNeutral m in p
|
||||||
|
hasLengthReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hasLengthReverseOnto (S p) q
|
||||||
|
|
||||||
export
|
export
|
||||||
map : (f : a -> b) -> HasLength xs n -> HasLength (map f xs) n
|
hasLengthReverse : HasLength m acc -> HasLength m (reverse acc)
|
||||||
|
hasLengthReverse = hasLengthReverseOnto Z
|
||||||
|
|
||||||
|
export
|
||||||
|
map : (f : a -> b) -> HasLength n xs -> HasLength n (map f xs)
|
||||||
map f Z = Z
|
map f Z = Z
|
||||||
map f (S n) = S (map f n)
|
map f (S n) = S (map f n)
|
||||||
|
|
||||||
@ -59,7 +78,7 @@ map f (S n) = S (map f n)
|
|||||||
||| So performing this operation while carrying the list around would cost O(n)
|
||| So performing this operation while carrying the list around would cost O(n)
|
||||||
||| but relying on n together with an erased HasLength proof instead is O(1)
|
||| but relying on n together with an erased HasLength proof instead is O(1)
|
||||||
export
|
export
|
||||||
sucR : HasLength xs n -> HasLength (snoc xs x) (S n)
|
sucR : HasLength n xs -> HasLength (S n) (snoc xs x)
|
||||||
sucR Z = S Z
|
sucR Z = S Z
|
||||||
sucR (S n) = S (sucR n)
|
sucR (S n) = S (sucR n)
|
||||||
|
|
||||||
@ -69,27 +88,27 @@ sucR (S n) = S (sucR n)
|
|||||||
namespace SubsetView
|
namespace SubsetView
|
||||||
|
|
||||||
||| We provide this view as a convenient way to perform nested pattern-matching
|
||| We provide this view as a convenient way to perform nested pattern-matching
|
||||||
||| on values of type `Subset Nat (HasLength xs)`. Functions using this view will
|
||| on values of type `Subset Nat (flip HasLength xs)`. Functions using this view will
|
||||||
||| be seen as terminating as long as the index list `xs` is left untouched.
|
||| be seen as terminating as long as the index list `xs` is left untouched.
|
||||||
||| See e.g. listTerminating below for such a function.
|
||| See e.g. listTerminating below for such a function.
|
||||||
public export
|
public export
|
||||||
data View : (xs : List a) -> Subset Nat (HasLength xs) -> Type where
|
data View : (xs : List a) -> Subset Nat (flip HasLength xs) -> Type where
|
||||||
Z : View [] (Element Z Z)
|
Z : View [] (Element Z Z)
|
||||||
S : (p : Subset Nat (HasLength xs)) -> View (x :: xs) (Element (S (fst p)) (S (snd p)))
|
S : (p : Subset Nat (flip HasLength xs)) -> View (x :: xs) (Element (S (fst p)) (S (snd p)))
|
||||||
|
|
||||||
||| This auxiliary function gets around the limitation of the check ensuring that
|
||| This auxiliary function gets around the limitation of the check ensuring that
|
||||||
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
|
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
|
||||||
viewZ : (0 p : HasLength xs Z) -> View xs (Element Z p)
|
viewZ : (0 p : HasLength Z xs) -> View xs (Element Z p)
|
||||||
viewZ Z = Z
|
viewZ Z = Z
|
||||||
|
|
||||||
||| This auxiliary function gets around the limitation of the check ensuring that
|
||| This auxiliary function gets around the limitation of the check ensuring that
|
||||||
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
|
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
|
||||||
viewS : (n : Nat) -> (0 p : HasLength xs (S n)) -> View xs (Element (S n) p)
|
viewS : (n : Nat) -> (0 p : HasLength (S n) xs) -> View xs (Element (S n) p)
|
||||||
viewS n (S p) = S (Element n p)
|
viewS n (S p) = S (Element n p)
|
||||||
|
|
||||||
||| Proof that the view covers all possible cases.
|
||| Proof that the view covers all possible cases.
|
||||||
export
|
export
|
||||||
view : (p : Subset Nat (HasLength xs)) -> View xs p
|
view : (p : Subset Nat (flip HasLength xs)) -> View xs p
|
||||||
view (Element Z p) = viewZ p
|
view (Element Z p) = viewZ p
|
||||||
view (Element (S n) p) = viewS n p
|
view (Element (S n) p) = viewS n p
|
||||||
|
|
||||||
@ -102,13 +121,13 @@ namespace CurriedView
|
|||||||
||| separately from the HasLength proof and the Subset view is not as useful anymore.
|
||| separately from the HasLength proof and the Subset view is not as useful anymore.
|
||||||
||| See e.g. natTerminating below for (a contrived example of) such a function.
|
||| See e.g. natTerminating below for (a contrived example of) such a function.
|
||||||
public export
|
public export
|
||||||
data View : (xs : List a) -> (n : Nat) -> HasLength xs n -> Type where
|
data View : (xs : List a) -> (n : Nat) -> HasLength n xs -> Type where
|
||||||
Z : View [] Z Z
|
Z : View [] Z Z
|
||||||
S : (n : Nat) -> (0 p : HasLength xs n) -> View (x :: xs) (S n) (S p)
|
S : (n : Nat) -> (0 p : HasLength n xs) -> View (x :: xs) (S n) (S p)
|
||||||
|
|
||||||
||| Proof that the view covers all possible cases.
|
||| Proof that the view covers all possible cases.
|
||||||
export
|
export
|
||||||
view : (n : Nat) -> (0 p : HasLength xs n) -> View xs n p
|
view : (n : Nat) -> (0 p : HasLength n xs) -> View xs n p
|
||||||
view Z Z = Z
|
view Z Z = Z
|
||||||
view (S n) (S p) = S n p
|
view (S n) (S p) = S n p
|
||||||
|
|
||||||
@ -117,22 +136,24 @@ namespace CurriedView
|
|||||||
|
|
||||||
-- /!\ Do NOT re-export these examples
|
-- /!\ Do NOT re-export these examples
|
||||||
|
|
||||||
listTerminating : (p : Subset Nat (HasLength xs)) -> HasLength (xs ++ [x]) (S (fst p))
|
listTerminating : (p : Subset Nat (flip HasLength xs)) -> HasLength (S (fst p)) (xs ++ [x])
|
||||||
listTerminating p = case view p of
|
listTerminating p with (view p)
|
||||||
Z => S Z
|
listTerminating (Element 0 Z) | Z = S Z
|
||||||
S p => S (listTerminating p)
|
listTerminating (Element (S (fst y)) (S (snd y))) | (S y) = S (listTerminating y)
|
||||||
|
|
||||||
data P : List Nat -> Type where
|
data P : List Nat -> Type where
|
||||||
PNil : P []
|
PNil : P []
|
||||||
PCon : P (map f xs) -> P (x :: xs)
|
PCon : P (map f xs) -> P (x :: xs)
|
||||||
|
|
||||||
covering
|
covering
|
||||||
notListTerminating : (p : Subset Nat (HasLength xs)) -> P xs
|
notListTerminating : (p : Subset Nat (flip HasLength xs)) -> P xs
|
||||||
notListTerminating p = case view p of
|
notListTerminating p with (view p)
|
||||||
Z => PNil
|
notListTerminating (Element 0 Z) | Z = PNil
|
||||||
S p => PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } p))
|
notListTerminating (Element (S (fst y)) (S (snd y))) | (S y) =
|
||||||
|
PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } y))
|
||||||
|
|
||||||
natTerminating : (n : Nat) -> (0 p : HasLength xs n) -> P xs
|
natTerminating : (n : Nat) -> (0 p : HasLength n xs ) -> P xs
|
||||||
natTerminating n p = case view n p of
|
natTerminating n p = case view n p of
|
||||||
Z => PNil
|
Z => PNil
|
||||||
S n p => PCon (natTerminating n (map id p))
|
S n p => PCon (natTerminating n (map id p))
|
||||||
|
|
@ -58,6 +58,7 @@ modules = Control.App,
|
|||||||
Data.SnocList.Quantifiers,
|
Data.SnocList.Quantifiers,
|
||||||
Data.SnocList.Operations,
|
Data.SnocList.Operations,
|
||||||
Data.List.Elem,
|
Data.List.Elem,
|
||||||
|
Data.List.HasLength,
|
||||||
Data.List.Views,
|
Data.List.Views,
|
||||||
Data.List.Quantifiers,
|
Data.List.Quantifiers,
|
||||||
Data.List1,
|
Data.List1,
|
||||||
|
@ -93,24 +93,24 @@ weakenR Z = Z
|
|||||||
weakenR (S p) = S (weakenR p)
|
weakenR (S p) = S (weakenR p)
|
||||||
|
|
||||||
export
|
export
|
||||||
weakenL : (p : Subset Nat (HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n)
|
weakenL : (p : Subset Nat (flip HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n)
|
||||||
weakenL m p = case view m of
|
weakenL m p with (view m)
|
||||||
Z => p
|
weakenL (Element 0 Z) p | Z = p
|
||||||
(S m) => S (weakenL m p)
|
weakenL (Element (S (fst m)) (S (snd m))) p | (S m) = S (weakenL m p)
|
||||||
|
|
||||||
export
|
export
|
||||||
strengthenL : (p : Subset Nat (HasLength xs)) ->
|
strengthenL : (p : Subset Nat (flip HasLength xs)) ->
|
||||||
lt n (fst p) === True ->
|
lt n (fst p) === True ->
|
||||||
AtIndex x (xs ++ ys) n -> AtIndex x xs n
|
AtIndex x (xs ++ ys) n -> AtIndex x xs n
|
||||||
strengthenL m lt idx = case view m of
|
strengthenL m lt idx with (view m)
|
||||||
S m => case idx of
|
strengthenL (Element (S (fst m)) (S (snd m))) lt Z | (S m) = Z
|
||||||
Z => Z
|
strengthenL (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = S (strengthenL m lt k)
|
||||||
S idx => S (strengthenL m lt idx)
|
|
||||||
|
|
||||||
export
|
export
|
||||||
strengthenR : (p : Subset Nat (HasLength ws)) ->
|
strengthenR : (p : Subset Nat (flip HasLength ws)) ->
|
||||||
lte (fst p) n === True ->
|
lte (fst p) n === True ->
|
||||||
AtIndex x (ws ++ xs) n -> AtIndex x xs (minus n (fst p))
|
AtIndex x (ws ++ xs) n -> AtIndex x xs (minus n (fst p))
|
||||||
strengthenR m lt idx = case view m of
|
strengthenR m lt idx with (view m)
|
||||||
Z => rewrite minusZeroRight n in idx
|
strengthenR (Element 0 Z) lt idx | Z = rewrite minusZeroRight n in idx
|
||||||
S m => case idx of S idx => strengthenR m lt idx
|
strengthenR (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = strengthenR m lt k
|
||||||
|
|
||||||
|
@ -46,7 +46,6 @@ modules = Control.ANSI,
|
|||||||
Data.List.Reverse,
|
Data.List.Reverse,
|
||||||
Data.List.Views.Extra,
|
Data.List.Views.Extra,
|
||||||
Data.List.Palindrome,
|
Data.List.Palindrome,
|
||||||
Data.List.HasLength,
|
|
||||||
Data.List.AtIndex,
|
Data.List.AtIndex,
|
||||||
Data.List.Alternating,
|
Data.List.Alternating,
|
||||||
|
|
||||||
|
@ -58,7 +58,7 @@ prj = let (Element n p) = isMember t ts in prj' n p
|
|||||||
||| By doing a bit of arithmetic we can figure out whether the union's value
|
||| By doing a bit of arithmetic we can figure out whether the union's value
|
||||||
||| came from the left or the right list used in the index.
|
||| came from the left or the right list used in the index.
|
||||||
public export
|
public export
|
||||||
split : Subset Nat (HasLength ss) ->
|
split : Subset Nat (flip HasLength ss) ->
|
||||||
Union elt (ss ++ ts) -> Either (Union elt ss) (Union elt ts)
|
Union elt (ss ++ ts) -> Either (Union elt ss) (Union elt ts)
|
||||||
split m (Element n p t) with (@@ lt n (fst m))
|
split m (Element n p t) with (@@ lt n (fst m))
|
||||||
split m (Element n p t) | (True ** lt)
|
split m (Element n p t) | (True ** lt)
|
||||||
@ -92,5 +92,5 @@ weakenR (Element n p t) = Element n (weakenR p) t
|
|||||||
||| the number of members introduced. Note that this number is the only
|
||| the number of members introduced. Note that this number is the only
|
||||||
||| thing we need to keep around at runtime.
|
||| thing we need to keep around at runtime.
|
||||||
public export
|
public export
|
||||||
weakenL : Subset Nat (HasLength ss) -> Union elt ts -> Union elt (ss ++ ts)
|
weakenL : Subset Nat (flip HasLength ss) -> Union elt ts -> Union elt (ss ++ ts)
|
||||||
weakenL m (Element n p t) = Element (fst m + n) (weakenL m p) t
|
weakenL m (Element n p t) = Element (fst m + n) (weakenL m p) t
|
||||||
|
@ -721,26 +721,31 @@ Show (Var ns) where
|
|||||||
|
|
||||||
namespace HasLength
|
namespace HasLength
|
||||||
|
|
||||||
|
-- TODO: delete in favor of base lib's List.HasLength once next version _after_ v0.6.0 ships.
|
||||||
public export
|
public export
|
||||||
data HasLength : Nat -> List a -> Type where
|
data HasLength : Nat -> List a -> Type where
|
||||||
Z : HasLength Z []
|
Z : HasLength Z []
|
||||||
S : HasLength n as -> HasLength (S n) (a :: as)
|
S : HasLength n as -> HasLength (S n) (a :: as)
|
||||||
|
|
||||||
|
-- TODO: Use List.HasLength.sucR from the base lib instead once next version _after_ v0.6.0 ships.
|
||||||
export
|
export
|
||||||
sucR : HasLength n xs -> HasLength (S n) (xs ++ [x])
|
sucR : HasLength n xs -> HasLength (S n) (xs ++ [x])
|
||||||
sucR Z = S Z
|
sucR Z = S Z
|
||||||
sucR (S n) = S (sucR n)
|
sucR (S n) = S (sucR n)
|
||||||
|
|
||||||
|
-- TODO: Use List.HasLength.hasLengthAppend from the base lib instead once next version _after_ v0.6.0 ships.
|
||||||
export
|
export
|
||||||
hlAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys)
|
hlAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys)
|
||||||
hlAppend Z ys = ys
|
hlAppend Z ys = ys
|
||||||
hlAppend (S xs) ys = S (hlAppend xs ys)
|
hlAppend (S xs) ys = S (hlAppend xs ys)
|
||||||
|
|
||||||
|
-- TODO: Use List.HasLength.hasLength from the base lib instead once next version _after_ v0.6.0 ships.
|
||||||
export
|
export
|
||||||
mkHasLength : (xs : List a) -> HasLength (length xs) xs
|
mkHasLength : (xs : List a) -> HasLength (length xs) xs
|
||||||
mkHasLength [] = Z
|
mkHasLength [] = Z
|
||||||
mkHasLength (_ :: xs) = S (mkHasLength xs)
|
mkHasLength (_ :: xs) = S (mkHasLength xs)
|
||||||
|
|
||||||
|
-- TODO: Use List.HasLength.take from the base lib instead once next vresion _after_ v0.6.0 ships.
|
||||||
export
|
export
|
||||||
take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs)
|
take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs)
|
||||||
take Z _ = Z
|
take Z _ = Z
|
||||||
@ -754,10 +759,12 @@ namespace HasLength
|
|||||||
succInjective : (0 l, r : Nat) -> S l = S r -> l = r
|
succInjective : (0 l, r : Nat) -> S l = S r -> l = r
|
||||||
succInjective _ _ Refl = Refl
|
succInjective _ _ Refl = Refl
|
||||||
|
|
||||||
|
-- TODO: Delete once version _after_ v0.6.0 ships.
|
||||||
hlReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs)
|
hlReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs)
|
||||||
hlReverseOnto p Z = rewrite plusZeroRightNeutral m in p
|
hlReverseOnto p Z = rewrite plusZeroRightNeutral m in p
|
||||||
hlReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hlReverseOnto (S p) q
|
hlReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hlReverseOnto (S p) q
|
||||||
|
|
||||||
|
-- TODO: Use List.HasLength.hasLengthReverse from the base lib instead once next version _after_ v0.6.0 ships.
|
||||||
export
|
export
|
||||||
hlReverse : HasLength m acc -> HasLength m (reverse acc)
|
hlReverse : HasLength m acc -> HasLength m (reverse acc)
|
||||||
hlReverse = hlReverseOnto Z
|
hlReverse = hlReverseOnto Z
|
||||||
|
Loading…
Reference in New Issue
Block a user