Skip to content

Commit c8a11e7

Browse files
jamesmckinnaandreasabel
authored andcommitted
Refactor Data.List.Base.scan* and their properties (#2395)
* refactor `scanr` etc. * restore `inits` etc. but lazier; plus knock-on consequences * more refactoring; plus knock-on consequences * tidy up * refactored into `Base` * ... and `Properties` * fix-up `inits` and `tails` * fix up `import`s * knock-ons * Andreas's suggestions * further, better, refactoring is possible * yet further, better, refactoring is possible: remove explicit equational reasoning! * Update CHANGELOG.md Fix deprecated names * Update Base.agda Fix deprecations * Update Properties.agda Fix deprecations * Update CHANGELOG.md Fix deprecated names
1 parent db84b73 commit c8a11e7

File tree

11 files changed

+245
-49
lines changed

11 files changed

+245
-49
lines changed

CHANGELOG.md

+42
Original file line numberDiff line numberDiff line change
@@ -65,6 +65,30 @@ Deprecated names
6565
isRing* ↦ Algebra.Structures.isRing
6666
```
6767

68+
* In `Data.List.Base`:
69+
```agda
70+
scanr ↦ Data.List.Scans.Base.scanr
71+
scanl ↦ Data.List.Scans.Base.scanl
72+
```
73+
74+
* In `Data.List.Properties`:
75+
```agda
76+
scanr-defn ↦ Data.List.Scans.Properties.scanr-defn
77+
scanl-defn ↦ Data.List.Scans.Properties.scanl-defn
78+
```
79+
80+
* In `Data.List.NonEmpty.Base`:
81+
```agda
82+
inits : List A → List⁺ (List A)
83+
tails : List A → List⁺ (List A)
84+
```
85+
86+
* In `Data.List.NonEmpty.Properties`:
87+
```agda
88+
toList-inits : toList ∘ List⁺.inits ≗ List.inits
89+
toList-tails : toList ∘ List⁺.tails ≗ List.tails
90+
```
91+
6892
* In `Data.Nat.Divisibility.Core`:
6993
```agda
7094
*-pres-∣ ↦ Data.Nat.Divisibility.*-pres-∣
@@ -89,6 +113,12 @@ New modules
89113
Algebra.Module.Bundles.Raw
90114
```
91115

116+
* Refactoring of `Data.List.Base.{scanr|scanl}` and their properties:
117+
```
118+
Data.List.Scans.Base
119+
Data.List.Scans.Properties
120+
```
121+
92122
* Prime factorisation of natural numbers.
93123
```
94124
Data.Nat.Primality.Factorisation
@@ -342,12 +372,24 @@ Additions to existing modules
342372
i*j≢0 : .{{_ : NonZero i}} .{{_ : NonZero j}} → NonZero (i * j)
343373
```
344374

375+
* In `Data.List.Base` redefine `inits` and `tails` in terms of:
376+
```agda
377+
tail∘inits : List A → List (List A)
378+
tail∘tails : List A → List (List A)
379+
```
380+
345381
* In `Data.List.Membership.Setoid.Properties`:
346382
```agda
347383
reverse⁺ : x ∈ xs → x ∈ reverse xs
348384
reverse⁻ : x ∈ reverse xs → x ∈ xs
349385
```
350386

387+
* In `Data.List.NonEmpty.Base`:
388+
```agda
389+
inits : List A → List⁺ (List A)
390+
tails : List A → List⁺ (List A)
391+
```
392+
351393
* In `Data.List.Properties`:
352394
```agda
353395
length-catMaybes : length (catMaybes xs) ≤ length xs

src/Codata/Sized/Colist/Properties.agda

+2-1
Original file line numberDiff line numberDiff line change
@@ -22,6 +22,7 @@ open import Codata.Sized.Stream as Stream using (Stream; _∷_)
2222
open import Data.Vec.Bounded as Vec≤ using (Vec≤)
2323
open import Data.List.Base as List using (List; []; _∷_)
2424
open import Data.List.NonEmpty as List⁺ using (List⁺; _∷_)
25+
import Data.List.Scans.Base as Scans
2526
open import Data.List.Relation.Binary.Equality.Propositional using (≋-refl)
2627
open import Data.Maybe.Base as Maybe using (Maybe; nothing; just)
2728
import Data.Maybe.Properties as Maybe
@@ -283,7 +284,7 @@ fromList-++ [] bs = refl
283284
fromList-++ (a ∷ as) bs = ≡.refl ∷ λ where .force fromList-++ as bs
284285

285286
fromList-scanl : (c : B A B) n as
286-
i ⊢ fromList (List.scanl c n as) ≈ scanl c n (fromList as)
287+
i ⊢ fromList (Scans.scanl c n as) ≈ scanl c n (fromList as)
287288
fromList-scanl c n [] = ≡.refl ∷ λ where .force refl
288289
fromList-scanl c n (a ∷ as) =
289290
≡.refl ∷ λ where .force fromList-scanl c (c n a) as

src/Data/List.agda

+4
Original file line numberDiff line numberDiff line change
@@ -8,10 +8,14 @@
88
-- lists.
99

1010
{-# OPTIONS --cubical-compatible --safe #-}
11+
{-# OPTIONS --warn=noUserWarning #-} -- for deprecated scans
1112

1213
module Data.List where
1314

1415
------------------------------------------------------------------------
1516
-- Types and basic operations
1617

1718
open import Data.List.Base public
19+
hiding (scanr; scanl)
20+
open import Data.List.Scans.Base public
21+
using (scanr; scanl)

src/Data/List/Base.agda

+30-16
Original file line numberDiff line numberDiff line change
@@ -189,13 +189,19 @@ iterate : (A → A) → A → ℕ → List A
189189
iterate f e zero = []
190190
iterate f e (suc n) = e ∷ iterate f (f e) n
191191

192+
tail∘inits : List A List (List A)
193+
tail∘inits [] = []
194+
tail∘inits (x ∷ xs) = [ x ] ∷ map (x ∷_) (tail∘inits xs)
195+
192196
inits : List A List (List A)
193-
inits [] = [] ∷ []
194-
inits (x ∷ xs) = [] ∷ map (x ∷_) (inits xs)
197+
inits xs = [] ∷ tail∘inits xs
198+
199+
tail∘tails : List A List (List A)
200+
tail∘tails [] = []
201+
tail∘tails (_ ∷ xs) = xs ∷ tail∘tails xs
195202

196203
tails : List A List (List A)
197-
tails [] = [] ∷ []
198-
tails (x ∷ xs) = (x ∷ xs) ∷ tails xs
204+
tails xs = xs ∷ tail∘tails xs
199205

200206
insertAt : (xs : List A) Fin (suc (length xs)) A List A
201207
insertAt xs zero v = v ∷ xs
@@ -205,18 +211,6 @@ updateAt : (xs : List A) → Fin (length xs) → (A → A) → List A
205211
updateAt (x ∷ xs) zero f = f x ∷ xs
206212
updateAt (x ∷ xs) (suc i) f = x ∷ updateAt xs i f
207213

208-
-- Scans
209-
210-
scanr : (A B B) B List A List B
211-
scanr f e [] = e ∷ []
212-
scanr f e (x ∷ xs) with scanr f e xs
213-
... | [] = [] -- dead branch
214-
... | y ∷ ys = f x y ∷ y ∷ ys
215-
216-
scanl : (A B A) A List B List A
217-
scanl f e [] = e ∷ []
218-
scanl f e (x ∷ xs) = e ∷ scanl f (f e x) xs
219-
220214
-- Tabulation
221215

222216
applyUpTo : (ℕ A) List A
@@ -571,3 +565,23 @@ _─_ = removeAt
571565
"Warning: _─_ was deprecated in v2.0.
572566
Please use removeAt instead."
573567
#-}
568+
569+
-- Version 2.1
570+
571+
scanr : (A B B) B List A List B
572+
scanr f e [] = e ∷ []
573+
scanr f e (x ∷ xs) with scanr f e xs
574+
... | [] = [] -- dead branch
575+
... | ys@(y ∷ _) = f x y ∷ ys
576+
{-# WARNING_ON_USAGE scanr
577+
"Warning: scanr was deprecated in v2.1.
578+
Please use Data.List.Scans.Base.scanr instead."
579+
#-}
580+
581+
scanl : (A B A) A List B List A
582+
scanl f e [] = e ∷ []
583+
scanl f e (x ∷ xs) = e ∷ scanl f (f e x) xs
584+
{-# WARNING_ON_USAGE scanl
585+
"Warning: scanl was deprecated in v2.1.
586+
Please use Data.List.Scans.Base.scanl instead."
587+
#-}

src/Data/List/Membership/Propositional/Properties.agda

+1-2
Original file line numberDiff line numberDiff line change
@@ -338,8 +338,7 @@ module _ {_•_ : Op₂ A} where
338338
-- inits
339339

340340
[]∈inits : {a} {A : Set a} (as : List A) [] ∈ inits as
341-
[]∈inits [] = here refl
342-
[]∈inits (a ∷ as) = here refl
341+
[]∈inits _ = here refl
343342

344343
------------------------------------------------------------------------
345344
-- Other properties

src/Data/List/NonEmpty/Base.agda

+10
Original file line numberDiff line numberDiff line change
@@ -143,6 +143,16 @@ concatMap f = concat ∘′ map f
143143
ap : List⁺ (A B) List⁺ A List⁺ B
144144
ap fs as = concatMap (λ f map f as) fs
145145

146+
-- Inits
147+
148+
inits : List A List⁺ (List A)
149+
inits xs = [] ∷ List.tail∘inits xs
150+
151+
-- Tails
152+
153+
tails : List A List⁺ (List A)
154+
tails xs = xs ∷ List.tail∘tails xs
155+
146156
-- Reverse
147157

148158
reverse : List⁺ A List⁺ A

src/Data/List/NonEmpty/Properties.agda

+13-1
Original file line numberDiff line numberDiff line change
@@ -18,7 +18,7 @@ open import Data.List.Effectful using () renaming (monad to listMonad)
1818
open import Data.List.NonEmpty.Effectful using () renaming (monad to list⁺Monad)
1919
open import Data.List.NonEmpty
2020
using (List⁺; _∷_; tail; head; toList; _⁺++_; _⁺++⁺_; _++⁺_; length; fromList;
21-
drop+; map; groupSeqs; ungroupSeqs)
21+
drop+; map; inits; tails; groupSeqs; ungroupSeqs)
2222
open import Data.List.NonEmpty.Relation.Unary.All using (All; toList⁺; _∷_)
2323
open import Data.List.Relation.Unary.All using ([]; _∷_) renaming (All to ListAll)
2424
import Data.List.Properties as List
@@ -118,6 +118,18 @@ map-cong f≗g (x ∷ xs) = cong₂ _∷_ (f≗g x) (List.map-cong f≗g xs)
118118
map-∘ : {g : B C} {f : A B} map (g ∘ f) ≗ map g ∘ map f
119119
map-∘ (x ∷ xs) = cong (_ ∷_) (List.map-∘ xs)
120120

121+
------------------------------------------------------------------------
122+
-- inits
123+
124+
toList-inits : toList ∘ inits ≗ List.inits {A = A}
125+
toList-inits _ = refl
126+
127+
------------------------------------------------------------------------
128+
-- tails
129+
130+
toList-tails : toList ∘ tails ≗ List.tails {A = A}
131+
toList-tails _ = refl
132+
121133
------------------------------------------------------------------------
122134
-- groupSeqs
123135

src/Data/List/Properties.agda

+33-28
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@
88
-- equalities than _≡_.
99

1010
{-# OPTIONS --cubical-compatible --safe #-}
11+
{-# OPTIONS --warn=noUserWarning #-} -- for deprecated scans
1112

1213
module Data.List.Properties where
1314

@@ -809,34 +810,6 @@ sum-++ (x ∷ xs) ys = begin
809810
∈⇒∣product {n} {n ∷ ns} (here refl) = divides (product ns) (*-comm n (product ns))
810811
∈⇒∣product {n} {m ∷ ns} (there n∈ns) = ∣n⇒∣m*n m (∈⇒∣product n∈ns)
811812

812-
------------------------------------------------------------------------
813-
-- scanr
814-
815-
scanr-defn : (f : A B B) (e : B)
816-
scanr f e ≗ map (foldr f e) ∘ tails
817-
scanr-defn f e [] = refl
818-
scanr-defn f e (x ∷ []) = refl
819-
scanr-defn f e (x ∷ y∷xs@(_ ∷ _))
820-
with eq scanr-defn f e y∷xs
821-
with z ∷ zs scanr f e y∷xs
822-
= let z≡fy⦇f⦈xs , _ = ∷-injective eq in cong₂ (λ z f x z ∷_) z≡fy⦇f⦈xs eq
823-
824-
------------------------------------------------------------------------
825-
-- scanl
826-
827-
scanl-defn : (f : A B A) (e : A)
828-
scanl f e ≗ map (foldl f e) ∘ inits
829-
scanl-defn f e [] = refl
830-
scanl-defn f e (x ∷ xs) = cong (e ∷_) (begin
831-
scanl f (f e x) xs
832-
≡⟨ scanl-defn f (f e x) xs ⟩
833-
map (foldl f (f e x)) (inits xs)
834-
≡⟨ refl ⟩
835-
map (foldl f e ∘ (x ∷_)) (inits xs)
836-
≡⟨ map-∘ (inits xs) ⟩
837-
map (foldl f e) (map (x ∷_) (inits xs))
838-
∎)
839-
840813
------------------------------------------------------------------------
841814
-- applyUpTo
842815

@@ -1582,3 +1555,35 @@ map-─ = map-removeAt
15821555
"Warning: map-─ was deprecated in v2.0.
15831556
Please use map-removeAt instead."
15841557
#-}
1558+
1559+
-- Version 2.1
1560+
1561+
scanr-defn : (f : A B B) (e : B)
1562+
scanr f e ≗ map (foldr f e) ∘ tails
1563+
scanr-defn f e [] = refl
1564+
scanr-defn f e (x ∷ []) = refl
1565+
scanr-defn f e (x ∷ xs@(_ ∷ _))
1566+
with eq scanr-defn f e xs
1567+
with ys@(_ ∷ _) scanr f e xs
1568+
= cong₂ (λ z f x z ∷_) (∷-injectiveˡ eq) eq
1569+
{-# WARNING_ON_USAGE scanr-defn
1570+
"Warning: scanr-defn was deprecated in v2.1.
1571+
Please use Data.List.Scans.Properties.scanr-defn instead."
1572+
#-}
1573+
1574+
scanl-defn : (f : A B A) (e : A)
1575+
scanl f e ≗ map (foldl f e) ∘ inits
1576+
scanl-defn f e [] = refl
1577+
scanl-defn f e (x ∷ xs) = cong (e ∷_) (begin
1578+
scanl f (f e x) xs
1579+
≡⟨ scanl-defn f (f e x) xs ⟩
1580+
map (foldl f (f e x)) (inits xs)
1581+
≡⟨ refl ⟩
1582+
map (foldl f e ∘ (x ∷_)) (inits xs)
1583+
≡⟨ map-∘ (inits xs) ⟩
1584+
map (foldl f e) (map (x ∷_) (inits xs))
1585+
∎)
1586+
{-# WARNING_ON_USAGE scanl-defn
1587+
"Warning: scanl-defn was deprecated in v2.1.
1588+
Please use Data.List.Scans.Properties.scanl-defn instead."
1589+
#-}

src/Data/List/Scans/Base.agda

+49
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,49 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- List scans: definitions
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.List.Scans.Base where
10+
11+
open import Data.List.Base as List using (List; []; _∷_)
12+
open import Data.List.NonEmpty.Base as List⁺ using (List⁺; _∷_; toList)
13+
open import Function.Base using (_∘_)
14+
open import Level using (Level)
15+
16+
private
17+
variable
18+
a b : Level
19+
A : Set a
20+
B : Set b
21+
22+
23+
------------------------------------------------------------------------
24+
-- Definitions
25+
26+
-- Scanr
27+
28+
module _ (f : A B B) where
29+
30+
scanr⁺ : (e : B) List A List⁺ B
31+
scanr⁺ e [] = e ∷ []
32+
scanr⁺ e (x ∷ xs) = let y ∷ ys = scanr⁺ e xs in f x y ∷ y ∷ ys
33+
34+
scanr : (e : B) List A List B
35+
scanr e = toList ∘ scanr⁺ e
36+
37+
-- Scanl
38+
39+
module _ (f : A B A) where
40+
41+
scanl⁺ : A List B List⁺ A
42+
scanl⁺ e xs = e ∷ go e xs
43+
where
44+
go : A List B List A
45+
go _ [] = []
46+
go e (x ∷ xs) = let fex = f e x in fex ∷ go fex xs
47+
48+
scanl : A List B List A
49+
scanl e = toList ∘ scanl⁺ e

0 commit comments

Comments
 (0)