diff --git a/CHANGELOG.md b/CHANGELOG.md index 268e1ba5f..6106d741d 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -71,12 +71,23 @@ * 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 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 * Changes `getNProcessors` to return the number of online processors rather than 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 * The `data` subfolder of an installed or local dependency package is now automatically recognized as a "data" directory by Idris 2. See the diff --git a/libs/contrib/Data/List/HasLength.idr b/libs/base/Data/List/HasLength.idr similarity index 59% rename from libs/contrib/Data/List/HasLength.idr rename to libs/base/Data/List/HasLength.idr index 1ef804efd..be6d647d2 100644 --- a/libs/contrib/Data/List/HasLength.idr +++ b/libs/base/Data/List/HasLength.idr @@ -12,8 +12,8 @@ ||| ||| We would write either of: ||| ```idris example -||| f1 : (n : Nat) -> (0 _ : HasLength xs n) -> P xs -||| f2 : (n : Subset n (HasLength xs)) -> P xs +||| f1 : (n : Nat) -> (0 _ : HasLength n 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) @@ -23,6 +23,7 @@ module Data.List.HasLength import Data.DPair import Data.List +import Data.Nat %default total @@ -31,27 +32,45 @@ import Data.List ||| Ensure that the list's length is the provided natural number public export -data HasLength : List a -> Nat -> Type where - Z : HasLength [] Z - S : HasLength xs n -> HasLength (x :: xs) (S n) +data HasLength : Nat -> List a -> Type where + Z : HasLength Z [] + 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 ||| The length is unique export -hasLengthUnique : HasLength xs m -> HasLength xs n -> m === n +hasLengthUnique : HasLength m xs -> HasLength n xs -> m === n hasLengthUnique Z Z = Refl hasLengthUnique (S p) (S q) = cong S (hasLengthUnique p q) -||| This specification corresponds to the length function export -hasLength : (xs : List a) -> HasLength xs (length xs) -hasLength [] = Z -hasLength (_ :: xs) = S (hasLength xs) +hasLengthAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys) +hasLengthAppend Z ys = ys +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 -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 (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) ||| but relying on n together with an erased HasLength proof instead is O(1) 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 (S n) = S (sucR n) @@ -69,27 +88,27 @@ sucR (S n) = S (sucR n) namespace SubsetView ||| 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. ||| See e.g. listTerminating below for such a function. 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) - 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 ||| 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 ||| 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. - 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) ||| Proof that the view covers all possible cases. 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 (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. ||| See e.g. natTerminating below for (a contrived example of) such a function. 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 - 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. 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 (S n) (S p) = S n p @@ -117,22 +136,24 @@ namespace CurriedView -- /!\ Do NOT re-export these examples -listTerminating : (p : Subset Nat (HasLength xs)) -> HasLength (xs ++ [x]) (S (fst p)) -listTerminating p = case view p of - Z => S Z - S p => S (listTerminating p) +listTerminating : (p : Subset Nat (flip HasLength xs)) -> HasLength (S (fst p)) (xs ++ [x]) +listTerminating p with (view p) + listTerminating (Element 0 Z) | Z = S Z + listTerminating (Element (S (fst y)) (S (snd y))) | (S y) = S (listTerminating y) data P : List Nat -> Type where PNil : P [] PCon : P (map f xs) -> P (x :: xs) covering -notListTerminating : (p : Subset Nat (HasLength xs)) -> P xs -notListTerminating p = case view p of - Z => PNil - S p => PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } p)) +notListTerminating : (p : Subset Nat (flip HasLength xs)) -> P xs +notListTerminating p with (view p) + notListTerminating (Element 0 Z) | Z = PNil + 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 Z => PNil S n p => PCon (natTerminating n (map id p)) + diff --git a/libs/base/base.ipkg b/libs/base/base.ipkg index 757075683..a46a0c126 100644 --- a/libs/base/base.ipkg +++ b/libs/base/base.ipkg @@ -58,6 +58,7 @@ modules = Control.App, Data.SnocList.Quantifiers, Data.SnocList.Operations, Data.List.Elem, + Data.List.HasLength, Data.List.Views, Data.List.Quantifiers, Data.List1, diff --git a/libs/contrib/Data/List/AtIndex.idr b/libs/contrib/Data/List/AtIndex.idr index db145ce13..ee8cf6c7b 100644 --- a/libs/contrib/Data/List/AtIndex.idr +++ b/libs/contrib/Data/List/AtIndex.idr @@ -93,24 +93,24 @@ weakenR Z = Z weakenR (S p) = S (weakenR p) export -weakenL : (p : Subset Nat (HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n) -weakenL m p = case view m of - Z => p - (S m) => S (weakenL m p) +weakenL : (p : Subset Nat (flip HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n) +weakenL m p with (view m) + weakenL (Element 0 Z) p | Z = p + weakenL (Element (S (fst m)) (S (snd m))) p | (S m) = S (weakenL m p) export -strengthenL : (p : Subset Nat (HasLength xs)) -> +strengthenL : (p : Subset Nat (flip HasLength xs)) -> lt n (fst p) === True -> AtIndex x (xs ++ ys) n -> AtIndex x xs n -strengthenL m lt idx = case view m of - S m => case idx of - Z => Z - S idx => S (strengthenL m lt idx) +strengthenL m lt idx with (view m) + strengthenL (Element (S (fst m)) (S (snd m))) lt Z | (S m) = Z + strengthenL (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = S (strengthenL m lt k) export -strengthenR : (p : Subset Nat (HasLength ws)) -> +strengthenR : (p : Subset Nat (flip HasLength ws)) -> lte (fst p) n === True -> AtIndex x (ws ++ xs) n -> AtIndex x xs (minus n (fst p)) -strengthenR m lt idx = case view m of - Z => rewrite minusZeroRight n in idx - S m => case idx of S idx => strengthenR m lt idx +strengthenR m lt idx with (view m) + strengthenR (Element 0 Z) lt idx | Z = rewrite minusZeroRight n in idx + strengthenR (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = strengthenR m lt k + diff --git a/libs/contrib/contrib.ipkg b/libs/contrib/contrib.ipkg index 65d55079f..16fd18d75 100644 --- a/libs/contrib/contrib.ipkg +++ b/libs/contrib/contrib.ipkg @@ -46,7 +46,6 @@ modules = Control.ANSI, Data.List.Reverse, Data.List.Views.Extra, Data.List.Palindrome, - Data.List.HasLength, Data.List.AtIndex, Data.List.Alternating, diff --git a/libs/papers/Data/OpenUnion.idr b/libs/papers/Data/OpenUnion.idr index cc8423faf..d955626f6 100644 --- a/libs/papers/Data/OpenUnion.idr +++ b/libs/papers/Data/OpenUnion.idr @@ -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 ||| came from the left or the right list used in the index. public export -split : Subset Nat (HasLength ss) -> +split : Subset Nat (flip HasLength ss) -> 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) | (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 ||| thing we need to keep around at runtime. 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 diff --git a/src/Core/TT.idr b/src/Core/TT.idr index 6c4e06149..8d78eba36 100644 --- a/src/Core/TT.idr +++ b/src/Core/TT.idr @@ -721,26 +721,31 @@ Show (Var ns) where namespace HasLength + -- TODO: delete in favor of base lib's List.HasLength once next version _after_ v0.6.0 ships. public export data HasLength : Nat -> List a -> Type where Z : HasLength Z [] 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 sucR : HasLength n xs -> HasLength (S n) (xs ++ [x]) sucR Z = S Z 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 hlAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys) hlAppend Z ys = 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 mkHasLength : (xs : List a) -> HasLength (length xs) xs mkHasLength [] = Z mkHasLength (_ :: xs) = S (mkHasLength xs) + -- TODO: Use List.HasLength.take from the base lib instead once next vresion _after_ v0.6.0 ships. export take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs) take Z _ = Z @@ -754,10 +759,12 @@ namespace HasLength succInjective : (0 l, r : Nat) -> S l = S r -> l = r 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 p Z = rewrite plusZeroRightNeutral m in p 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 hlReverse : HasLength m acc -> HasLength m (reverse acc) hlReverse = hlReverseOnto Z