Skip to content

Commit 05a3ff3

Browse files
mattpolzinjoelberkeley
authored andcommitted
Moving Data.List.HasLength into base (idris-lang#2844)
1 parent 477c583 commit 05a3ff3

File tree

7 files changed

+85
-46
lines changed

7 files changed

+85
-46
lines changed

CHANGELOG.md

+11
Original file line numberDiff line numberDiff line change
@@ -71,12 +71,23 @@
7171
* Add an alias for `HVect` to `All id` in `Data.Vect.Quantifiers.All`. This is the
7272
approach to getting a heterogeneous Vect of elements that is generall preferred by
7373
the community vs. a standalone type as seen in `contrib`.
74+
* Add Data.List.HasLength from the compiler codebase slash contrib library but
75+
adopt the type signature from the compiler codebase and some of the naming
76+
from the contrib library. The type ended up being `HasLength n xs` rather than
77+
`HasLength xs n`.
7478

7579
#### System
7680

7781
* Changes `getNProcessors` to return the number of online processors rather than
7882
the number of configured processors.
7983

84+
85+
#### Contrib
86+
* Remove Data.List.HasLength from contrib library but add it to the base library
87+
with the type signature from the compiler codebase and some of the naming
88+
from the contrib library. The type ended up being `HasLength n xs` rather than
89+
`HasLength xs n`.
90+
8091
### Other Changes
8192
* The `data` subfolder of an installed or local dependency package is now automatically
8293
recognized as a "data" directory by Idris 2. See the

libs/contrib/Data/List/HasLength.idr renamed to libs/base/Data/List/HasLength.idr

+51-30
Original file line numberDiff line numberDiff line change
@@ -12,8 +12,8 @@
1212
|||
1313
||| We would write either of:
1414
||| ```idris example
15-
||| f1 : (n : Nat) -> (0 _ : HasLength xs n) -> P xs
16-
||| f2 : (n : Subset n (HasLength xs)) -> P xs
15+
||| f1 : (n : Nat) -> (0 _ : HasLength n xs) -> P xs
16+
||| f2 : (n : Subset n (flip HasLength xs)) -> P xs
1717
||| ```
1818
|||
1919
||| See `sucR` for an example where the update to the runtime-relevant Nat is O(1)
@@ -23,6 +23,7 @@ module Data.List.HasLength
2323

2424
import Data.DPair
2525
import Data.List
26+
import Data.Nat
2627

2728
%default total
2829

@@ -31,35 +32,53 @@ import Data.List
3132

3233
||| Ensure that the list's length is the provided natural number
3334
public export
34-
data HasLength : List a -> Nat -> Type where
35-
Z : HasLength [] Z
36-
S : HasLength xs n -> HasLength (x :: xs) (S n)
35+
data HasLength : Nat -> List a -> Type where
36+
Z : HasLength Z []
37+
S : HasLength n xs -> HasLength (S n) (x :: xs)
38+
39+
||| This specification corresponds to the length function
40+
export
41+
hasLength : (xs : List a) -> HasLength (length xs) xs
42+
hasLength [] = Z
43+
hasLength (_ :: xs) = S (hasLength xs)
44+
45+
export
46+
take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs)
47+
take Z _ = Z
48+
take (S n) (x :: xs) = S (take n xs)
3749

3850
------------------------------------------------------------------------
3951
-- Properties
4052

4153
||| The length is unique
4254
export
43-
hasLengthUnique : HasLength xs m -> HasLength xs n -> m === n
55+
hasLengthUnique : HasLength m xs -> HasLength n xs -> m === n
4456
hasLengthUnique Z Z = Refl
4557
hasLengthUnique (S p) (S q) = cong S (hasLengthUnique p q)
4658

47-
||| This specification corresponds to the length function
4859
export
49-
hasLength : (xs : List a) -> HasLength xs (length xs)
50-
hasLength [] = Z
51-
hasLength (_ :: xs) = S (hasLength xs)
60+
hasLengthAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys)
61+
hasLengthAppend Z ys = ys
62+
hasLengthAppend (S xs) ys = S (hasLengthAppend xs ys)
63+
64+
hasLengthReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs)
65+
hasLengthReverseOnto p Z = rewrite plusZeroRightNeutral m in p
66+
hasLengthReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hasLengthReverseOnto (S p) q
5267

5368
export
54-
map : (f : a -> b) -> HasLength xs n -> HasLength (map f xs) n
69+
hasLengthReverse : HasLength m acc -> HasLength m (reverse acc)
70+
hasLengthReverse = hasLengthReverseOnto Z
71+
72+
export
73+
map : (f : a -> b) -> HasLength n xs -> HasLength n (map f xs)
5574
map f Z = Z
5675
map f (S n) = S (map f n)
5776

5877
||| @sucR demonstrates that snoc only increases the lenght by one
5978
||| So performing this operation while carrying the list around would cost O(n)
6079
||| but relying on n together with an erased HasLength proof instead is O(1)
6180
export
62-
sucR : HasLength xs n -> HasLength (snoc xs x) (S n)
81+
sucR : HasLength n xs -> HasLength (S n) (snoc xs x)
6382
sucR Z = S Z
6483
sucR (S n) = S (sucR n)
6584

@@ -69,27 +88,27 @@ sucR (S n) = S (sucR n)
6988
namespace SubsetView
7089

7190
||| We provide this view as a convenient way to perform nested pattern-matching
72-
||| on values of type `Subset Nat (HasLength xs)`. Functions using this view will
91+
||| on values of type `Subset Nat (flip HasLength xs)`. Functions using this view will
7392
||| be seen as terminating as long as the index list `xs` is left untouched.
7493
||| See e.g. listTerminating below for such a function.
7594
public export
76-
data View : (xs : List a) -> Subset Nat (HasLength xs) -> Type where
95+
data View : (xs : List a) -> Subset Nat (flip HasLength xs) -> Type where
7796
Z : View [] (Element Z Z)
78-
S : (p : Subset Nat (HasLength xs)) -> View (x :: xs) (Element (S (fst p)) (S (snd p)))
97+
S : (p : Subset Nat (flip HasLength xs)) -> View (x :: xs) (Element (S (fst p)) (S (snd p)))
7998

8099
||| This auxiliary function gets around the limitation of the check ensuring that
81100
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
82-
viewZ : (0 p : HasLength xs Z) -> View xs (Element Z p)
101+
viewZ : (0 p : HasLength Z xs) -> View xs (Element Z p)
83102
viewZ Z = Z
84103

85104
||| This auxiliary function gets around the limitation of the check ensuring that
86105
||| we do not match on runtime-irrelevant data to produce runtime-relevant data.
87-
viewS : (n : Nat) -> (0 p : HasLength xs (S n)) -> View xs (Element (S n) p)
106+
viewS : (n : Nat) -> (0 p : HasLength (S n) xs) -> View xs (Element (S n) p)
88107
viewS n (S p) = S (Element n p)
89108

90109
||| Proof that the view covers all possible cases.
91110
export
92-
view : (p : Subset Nat (HasLength xs)) -> View xs p
111+
view : (p : Subset Nat (flip HasLength xs)) -> View xs p
93112
view (Element Z p) = viewZ p
94113
view (Element (S n) p) = viewS n p
95114

@@ -102,13 +121,13 @@ namespace CurriedView
102121
||| separately from the HasLength proof and the Subset view is not as useful anymore.
103122
||| See e.g. natTerminating below for (a contrived example of) such a function.
104123
public export
105-
data View : (xs : List a) -> (n : Nat) -> HasLength xs n -> Type where
124+
data View : (xs : List a) -> (n : Nat) -> HasLength n xs -> Type where
106125
Z : View [] Z Z
107-
S : (n : Nat) -> (0 p : HasLength xs n) -> View (x :: xs) (S n) (S p)
126+
S : (n : Nat) -> (0 p : HasLength n xs) -> View (x :: xs) (S n) (S p)
108127

109128
||| Proof that the view covers all possible cases.
110129
export
111-
view : (n : Nat) -> (0 p : HasLength xs n) -> View xs n p
130+
view : (n : Nat) -> (0 p : HasLength n xs) -> View xs n p
112131
view Z Z = Z
113132
view (S n) (S p) = S n p
114133

@@ -117,22 +136,24 @@ namespace CurriedView
117136

118137
-- /!\ Do NOT re-export these examples
119138

120-
listTerminating : (p : Subset Nat (HasLength xs)) -> HasLength (xs ++ [x]) (S (fst p))
121-
listTerminating p = case view p of
122-
Z => S Z
123-
S p => S (listTerminating p)
139+
listTerminating : (p : Subset Nat (flip HasLength xs)) -> HasLength (S (fst p)) (xs ++ [x])
140+
listTerminating p with (view p)
141+
listTerminating (Element 0 Z) | Z = S Z
142+
listTerminating (Element (S (fst y)) (S (snd y))) | (S y) = S (listTerminating y)
124143

125144
data P : List Nat -> Type where
126145
PNil : P []
127146
PCon : P (map f xs) -> P (x :: xs)
128147

129148
covering
130-
notListTerminating : (p : Subset Nat (HasLength xs)) -> P xs
131-
notListTerminating p = case view p of
132-
Z => PNil
133-
S p => PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } p))
149+
notListTerminating : (p : Subset Nat (flip HasLength xs)) -> P xs
150+
notListTerminating p with (view p)
151+
notListTerminating (Element 0 Z) | Z = PNil
152+
notListTerminating (Element (S (fst y)) (S (snd y))) | (S y) =
153+
PCon (notListTerminating {xs = map id (tail xs)} ({ snd $= map id } y))
134154

135-
natTerminating : (n : Nat) -> (0 p : HasLength xs n) -> P xs
155+
natTerminating : (n : Nat) -> (0 p : HasLength n xs ) -> P xs
136156
natTerminating n p = case view n p of
137157
Z => PNil
138158
S n p => PCon (natTerminating n (map id p))
159+

libs/base/base.ipkg

+1
Original file line numberDiff line numberDiff line change
@@ -58,6 +58,7 @@ modules = Control.App,
5858
Data.SnocList.Quantifiers,
5959
Data.SnocList.Operations,
6060
Data.List.Elem,
61+
Data.List.HasLength,
6162
Data.List.Views,
6263
Data.List.Quantifiers,
6364
Data.List1,

libs/contrib/Data/List/AtIndex.idr

+13-13
Original file line numberDiff line numberDiff line change
@@ -93,24 +93,24 @@ weakenR Z = Z
9393
weakenR (S p) = S (weakenR p)
9494

9595
export
96-
weakenL : (p : Subset Nat (HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n)
97-
weakenL m p = case view m of
98-
Z => p
99-
(S m) => S (weakenL m p)
96+
weakenL : (p : Subset Nat (flip HasLength ws)) -> AtIndex x xs n -> AtIndex x (ws ++ xs) (fst p + n)
97+
weakenL m p with (view m)
98+
weakenL (Element 0 Z) p | Z = p
99+
weakenL (Element (S (fst m)) (S (snd m))) p | (S m) = S (weakenL m p)
100100

101101
export
102-
strengthenL : (p : Subset Nat (HasLength xs)) ->
102+
strengthenL : (p : Subset Nat (flip HasLength xs)) ->
103103
lt n (fst p) === True ->
104104
AtIndex x (xs ++ ys) n -> AtIndex x xs n
105-
strengthenL m lt idx = case view m of
106-
S m => case idx of
107-
Z => Z
108-
S idx => S (strengthenL m lt idx)
105+
strengthenL m lt idx with (view m)
106+
strengthenL (Element (S (fst m)) (S (snd m))) lt Z | (S m) = Z
107+
strengthenL (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = S (strengthenL m lt k)
109108

110109
export
111-
strengthenR : (p : Subset Nat (HasLength ws)) ->
110+
strengthenR : (p : Subset Nat (flip HasLength ws)) ->
112111
lte (fst p) n === True ->
113112
AtIndex x (ws ++ xs) n -> AtIndex x xs (minus n (fst p))
114-
strengthenR m lt idx = case view m of
115-
Z => rewrite minusZeroRight n in idx
116-
S m => case idx of S idx => strengthenR m lt idx
113+
strengthenR m lt idx with (view m)
114+
strengthenR (Element 0 Z) lt idx | Z = rewrite minusZeroRight n in idx
115+
strengthenR (Element (S (fst m)) (S (snd m))) lt (S k) | (S m) = strengthenR m lt k
116+

libs/contrib/contrib.ipkg

-1
Original file line numberDiff line numberDiff line change
@@ -46,7 +46,6 @@ modules = Control.ANSI,
4646
Data.List.Reverse,
4747
Data.List.Views.Extra,
4848
Data.List.Palindrome,
49-
Data.List.HasLength,
5049
Data.List.AtIndex,
5150
Data.List.Alternating,
5251

libs/papers/Data/OpenUnion.idr

+2-2
Original file line numberDiff line numberDiff line change
@@ -58,7 +58,7 @@ prj = let (Element n p) = isMember t ts in prj' n p
5858
||| By doing a bit of arithmetic we can figure out whether the union's value
5959
||| came from the left or the right list used in the index.
6060
public export
61-
split : Subset Nat (HasLength ss) ->
61+
split : Subset Nat (flip HasLength ss) ->
6262
Union elt (ss ++ ts) -> Either (Union elt ss) (Union elt ts)
6363
split m (Element n p t) with (@@ lt n (fst m))
6464
split m (Element n p t) | (True ** lt)
@@ -92,5 +92,5 @@ weakenR (Element n p t) = Element n (weakenR p) t
9292
||| the number of members introduced. Note that this number is the only
9393
||| thing we need to keep around at runtime.
9494
public export
95-
weakenL : Subset Nat (HasLength ss) -> Union elt ts -> Union elt (ss ++ ts)
95+
weakenL : Subset Nat (flip HasLength ss) -> Union elt ts -> Union elt (ss ++ ts)
9696
weakenL m (Element n p t) = Element (fst m + n) (weakenL m p) t

src/Core/TT.idr

+7
Original file line numberDiff line numberDiff line change
@@ -721,26 +721,31 @@ Show (Var ns) where
721721

722722
namespace HasLength
723723

724+
-- TODO: delete in favor of base lib's List.HasLength once next version _after_ v0.6.0 ships.
724725
public export
725726
data HasLength : Nat -> List a -> Type where
726727
Z : HasLength Z []
727728
S : HasLength n as -> HasLength (S n) (a :: as)
728729

730+
-- TODO: Use List.HasLength.sucR from the base lib instead once next version _after_ v0.6.0 ships.
729731
export
730732
sucR : HasLength n xs -> HasLength (S n) (xs ++ [x])
731733
sucR Z = S Z
732734
sucR (S n) = S (sucR n)
733735

736+
-- TODO: Use List.HasLength.hasLengthAppend from the base lib instead once next version _after_ v0.6.0 ships.
734737
export
735738
hlAppend : HasLength m xs -> HasLength n ys -> HasLength (m + n) (xs ++ ys)
736739
hlAppend Z ys = ys
737740
hlAppend (S xs) ys = S (hlAppend xs ys)
738741

742+
-- TODO: Use List.HasLength.hasLength from the base lib instead once next version _after_ v0.6.0 ships.
739743
export
740744
mkHasLength : (xs : List a) -> HasLength (length xs) xs
741745
mkHasLength [] = Z
742746
mkHasLength (_ :: xs) = S (mkHasLength xs)
743747

748+
-- TODO: Use List.HasLength.take from the base lib instead once next vresion _after_ v0.6.0 ships.
744749
export
745750
take : (n : Nat) -> (xs : Stream a) -> HasLength n (take n xs)
746751
take Z _ = Z
@@ -754,10 +759,12 @@ namespace HasLength
754759
succInjective : (0 l, r : Nat) -> S l = S r -> l = r
755760
succInjective _ _ Refl = Refl
756761

762+
-- TODO: Delete once version _after_ v0.6.0 ships.
757763
hlReverseOnto : HasLength m acc -> HasLength n xs -> HasLength (m + n) (reverseOnto acc xs)
758764
hlReverseOnto p Z = rewrite plusZeroRightNeutral m in p
759765
hlReverseOnto {n = S n} p (S q) = rewrite sym (plusSuccRightSucc m n) in hlReverseOnto (S p) q
760766

767+
-- TODO: Use List.HasLength.hasLengthReverse from the base lib instead once next version _after_ v0.6.0 ships.
761768
export
762769
hlReverse : HasLength m acc -> HasLength m (reverse acc)
763770
hlReverse = hlReverseOnto Z

0 commit comments

Comments
 (0)