Skip to content

Commit e417e4a

Browse files
JacquesCaretteandreasabel
authored andcommitted
Improve implementation of splitAt, take and drop in Data.List. (#2056)
1 parent d95bb96 commit e417e4a

File tree

5 files changed

+76
-105
lines changed

5 files changed

+76
-105
lines changed

CHANGELOG.md

+5
Original file line numberDiff line numberDiff line change
@@ -1409,6 +1409,11 @@ Deprecated names
14091409

14101410
* In `Data.Vec.Properties`:
14111411
```
1412+
take-distr-zipWith ↦ take-zipWith
1413+
take-distr-map ↦ take-map
1414+
drop-distr-zipWith ↦ drop-zipWith
1415+
drop-distr-map ↦ drop-map
1416+
14121417
updateAt-id-relative ↦ updateAt-id-local
14131418
updateAt-compose-relative ↦ updateAt-∘-local
14141419
updateAt-compose ↦ updateAt-∘

src/Data/Vec/Base.agda

+17-21
Original file line numberDiff line numberDiff line change
@@ -12,11 +12,11 @@ open import Data.Bool.Base using (Bool; true; false; if_then_else_)
1212
open import Data.Nat.Base
1313
open import Data.Fin.Base using (Fin; zero; suc)
1414
open import Data.List.Base as List using (List)
15-
open import Data.Product.Base as Prod using (∃; ∃₂; _×_; _,_)
15+
open import Data.Product.Base as Prod using (∃; ∃₂; _×_; _,_; proj₁; proj₂)
1616
open import Data.These.Base as These using (These; this; that; these)
1717
open import Function.Base using (const; _∘′_; id; _∘_)
1818
open import Level using (Level)
19-
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong)
19+
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; trans; cong)
2020
open import Relation.Nullary.Decidable.Core using (does)
2121
open import Relation.Unary using (Pred; Decidable)
2222

@@ -259,26 +259,23 @@ allFin _ = tabulate id
259259

260260
splitAt : m {n} (xs : Vec A (m + n))
261261
∃₂ λ (ys : Vec A m) (zs : Vec A n) xs ≡ ys ++ zs
262-
splitAt zero xs = ([] , xs , refl)
263-
splitAt (suc m) (x ∷ xs) with splitAt m xs
264-
splitAt (suc m) (x ∷ .(ys ++ zs)) | (ys , zs , refl) =
265-
((x ∷ ys) , zs , refl)
262+
splitAt zero xs = [] , xs , refl
263+
splitAt (suc m) (x ∷ xs) =
264+
let ys , zs , eq = splitAt m xs in x ∷ ys , zs , cong (x ∷_) eq
266265

267266
take : m {n} Vec A (m + n) Vec A m
268-
take m xs with splitAt m xs
269-
take m .(ys ++ zs) | (ys , zs , refl) = ys
267+
take m xs = proj₁ (splitAt m xs)
270268

271269
drop : m {n} Vec A (m + n) Vec A n
272-
drop m xs with splitAt m xs
273-
drop m .(ys ++ zs) | (ys , zs , refl) = zs
270+
drop m xs = proj₁ (proj₂ (splitAt m xs))
274271

275272
group : n k (xs : Vec A (n * k))
276273
λ (xss : Vec (Vec A k) n) xs ≡ concat xss
277274
group zero k [] = ([] , refl)
278-
group (suc n) k xs with splitAt k xs
279-
group (suc n) k .(ys ++ zs) | (ys , zs , refl) with group n k zs
280-
group (suc n) k .(ys ++ concat zss) | (ys , ._ , refl) | (zss , refl) =
281-
((ys ∷ zss) , refl)
275+
group (suc n) k xs =
276+
let ys , zs , eq-split = splitAt k xs in
277+
let zss , eq-group = group n k zs in
278+
(ys ∷ zss) , trans eq-split (cong (ys ++_) eq-group)
282279

283280
split : Vec A n Vec A ⌈ n /2⌉ × Vec A ⌊ n /2⌋
284281
split [] = ([] , [])
@@ -338,17 +335,16 @@ xs ʳ++ ys = foldl (Vec _ ∘ (_+ _)) (λ rev x → x ∷ rev) ys xs
338335
-- init and last
339336

340337
initLast : (xs : Vec A (1 + n)) ∃₂ λ ys y xs ≡ ys ∷ʳ y
341-
initLast {n = zero} (x ∷ []) = ([] , x , refl)
342-
initLast {n = suc n} (x ∷ xs) with initLast xs
343-
... | (ys , y , refl) = (x ∷ ys , y , refl)
338+
initLast {n = zero} (x ∷ []) = [] , x , refl
339+
initLast {n = suc n} (x ∷ xs) =
340+
let ys , y , eq = initLast xs in
341+
x ∷ ys , y , cong (x ∷_) eq
344342

345343
init : Vec A (1 + n) Vec A n
346-
init {n = zero} (x ∷ _) = []
347-
init {n = suc n} (x ∷ xs) = x ∷ init xs
344+
init xs = proj₁ (initLast xs)
348345

349346
last : Vec A (1 + n) A
350-
last {n = zero} (x ∷ _) = x
351-
last {n = suc n} (_ ∷ xs) = last xs
347+
last xs = proj₁ (proj₂ (initLast xs))
352348

353349
------------------------------------------------------------------------
354350
-- Other operations

src/Data/Vec/Properties.agda

+50-76
Original file line numberDiff line numberDiff line change
@@ -72,86 +72,43 @@ private
7272
-- take
7373

7474
unfold-take : n x (xs : Vec A (n + m)) take (suc n) (x ∷ xs) ≡ x ∷ take n xs
75-
unfold-take n x xs with splitAt n xs
76-
... | xs , ys , refl = refl
77-
78-
take-distr-zipWith : (f : A B C)
79-
(xs : Vec A (m + n)) (ys : Vec B (m + n))
80-
take m (zipWith f xs ys) ≡ zipWith f (take m xs) (take m ys)
81-
take-distr-zipWith {m = zero} f xs ys = refl
82-
take-distr-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = begin
83-
take (suc m) (zipWith f (x ∷ xs) (y ∷ ys))
84-
≡⟨⟩
85-
take (suc m) (f x y ∷ (zipWith f xs ys))
86-
≡⟨ unfold-take m (f x y) (zipWith f xs ys) ⟩
87-
f x y ∷ take m (zipWith f xs ys)
88-
≡⟨ cong (f x y ∷_) (take-distr-zipWith f xs ys) ⟩
89-
f x y ∷ (zipWith f (take m xs) (take m ys))
90-
≡⟨⟩
91-
zipWith f (x ∷ (take m xs)) (y ∷ (take m ys))
92-
≡˘⟨ cong₂ (zipWith f) (unfold-take m x xs) (unfold-take m y ys) ⟩
93-
zipWith f (take (suc m) (x ∷ xs)) (take (suc m) (y ∷ ys))
94-
95-
96-
take-distr-map : (f : A B) (m : ℕ) (xs : Vec A (m + n))
97-
take m (map f xs) ≡ map f (take m xs)
98-
take-distr-map f zero xs = refl
99-
take-distr-map f (suc m) (x ∷ xs) = begin
100-
take (suc m) (map f (x ∷ xs)) ≡⟨⟩
101-
take (suc m) (f x ∷ map f xs) ≡⟨ unfold-take m (f x) (map f xs) ⟩
102-
f x ∷ (take m (map f xs)) ≡⟨ cong (f x ∷_) (take-distr-map f m xs) ⟩
103-
f x ∷ (map f (take m xs)) ≡⟨⟩
104-
map f (x ∷ take m xs) ≡˘⟨ cong (map f) (unfold-take m x xs) ⟩
105-
map f (take (suc m) (x ∷ xs)) ∎
75+
unfold-take n x xs = refl
76+
77+
take-zipWith : (f : A B C)
78+
(xs : Vec A (m + n)) (ys : Vec B (m + n))
79+
take m (zipWith f xs ys) ≡ zipWith f (take m xs) (take m ys)
80+
take-zipWith {m = zero} f xs ys = refl
81+
take-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = cong (f x y ∷_) (take-zipWith f xs ys)
82+
83+
take-map : (f : A B) (m : ℕ) (xs : Vec A (m + n))
84+
take m (map f xs) ≡ map f (take m xs)
85+
take-map f zero xs = refl
86+
take-map f (suc m) (x ∷ xs) = cong (f x ∷_) (take-map f m xs)
10687

10788
------------------------------------------------------------------------
10889
-- drop
10990

11091
unfold-drop : n x (xs : Vec A (n + m))
11192
drop (suc n) (x ∷ xs) ≡ drop n xs
112-
unfold-drop n x xs with splitAt n xs
113-
... | xs , ys , refl = refl
114-
115-
drop-distr-zipWith : (f : A B C)
116-
(x : Vec A (m + n)) (y : Vec B (m + n))
117-
drop m (zipWith f x y) ≡ zipWith f (drop m x) (drop m y)
118-
drop-distr-zipWith {m = zero} f xs ys = refl
119-
drop-distr-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = begin
120-
drop (suc m) (zipWith f (x ∷ xs) (y ∷ ys))
121-
≡⟨⟩
122-
drop (suc m) (f x y ∷ (zipWith f xs ys))
123-
≡⟨ unfold-drop m (f x y) (zipWith f xs ys) ⟩
124-
drop m (zipWith f xs ys)
125-
≡⟨ drop-distr-zipWith f xs ys ⟩
126-
zipWith f (drop m xs) (drop m ys)
127-
≡˘⟨ cong₂ (zipWith f) (unfold-drop m x xs) (unfold-drop m y ys) ⟩
128-
zipWith f (drop (suc m) (x ∷ xs)) (drop (suc m) (y ∷ ys))
129-
130-
131-
drop-distr-map : (f : A B) (m : ℕ) (x : Vec A (m + n))
132-
drop m (map f x) ≡ map f (drop m x)
133-
drop-distr-map f zero x = refl
134-
drop-distr-map f (suc m) (x ∷ xs) = begin
135-
drop (suc m) (map f (x ∷ xs)) ≡⟨⟩
136-
drop (suc m) (f x ∷ map f xs) ≡⟨ unfold-drop m (f x) (map f xs) ⟩
137-
drop m (map f xs) ≡⟨ drop-distr-map f m xs ⟩
138-
map f (drop m xs) ≡˘⟨ cong (map f) (unfold-drop m x xs) ⟩
139-
map f (drop (suc m) (x ∷ xs)) ∎
93+
unfold-drop n x xs = refl
94+
95+
drop-zipWith : (f : A B C)
96+
(xs : Vec A (m + n)) (ys : Vec B (m + n))
97+
drop m (zipWith f xs ys) ≡ zipWith f (drop m xs) (drop m ys)
98+
drop-zipWith {m = zero} f xs ys = refl
99+
drop-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = drop-zipWith f xs ys
100+
101+
drop-map : (f : A B) (m : ℕ) (xs : Vec A (m + n))
102+
drop m (map f xs) ≡ map f (drop m xs)
103+
drop-map f zero xs = refl
104+
drop-map f (suc m) (x ∷ xs) = drop-map f m xs
140105

141106
------------------------------------------------------------------------
142107
-- take and drop together
143108

144-
take++drop≡id : (m : ℕ) (x : Vec A (m + n)) take m x ++ drop m x ≡ x
145-
take++drop≡id zero x = refl
146-
take++drop≡id (suc m) (x ∷ xs) = begin
147-
take (suc m) (x ∷ xs) ++ drop (suc m) (x ∷ xs)
148-
≡⟨ cong₂ _++_ (unfold-take m x xs) (unfold-drop m x xs) ⟩
149-
(x ∷ take m xs) ++ (drop m xs)
150-
≡⟨⟩
151-
x ∷ (take m xs ++ drop m xs)
152-
≡⟨ cong (x ∷_) (take++drop≡id m xs) ⟩
153-
x ∷ xs
154-
109+
take++drop≡id : (m : ℕ) (xs : Vec A (m + n)) take m xs ++ drop m xs ≡ xs
110+
take++drop≡id zero xs = refl
111+
take++drop≡id (suc m) (x ∷ xs) = cong (x ∷_) (take++drop≡id m xs)
155112

156113
------------------------------------------------------------------------
157114
-- truncate
@@ -216,12 +173,8 @@ lookup⇒[]= (suc i) (_ ∷ xs) p = there (lookup⇒[]= i xs p)
216173

217174
lookup-inject≤-take : m (m≤m+n : m ≤ m + n) (i : Fin m) (xs : Vec A (m + n))
218175
lookup xs (Fin.inject≤ i m≤m+n) ≡ lookup (take m xs) i
219-
lookup-inject≤-take (suc m) m≤m+n zero (x ∷ xs)
220-
rewrite unfold-take m x xs = refl
221-
lookup-inject≤-take (suc (suc m)) (s≤s m≤m+n) (suc zero) (x ∷ y ∷ xs)
222-
rewrite unfold-take (suc m) x (y ∷ xs) | unfold-take m y xs = refl
223-
lookup-inject≤-take (suc (suc m)) (s≤s (s≤s m≤m+n)) (suc (suc i)) (x ∷ y ∷ xs)
224-
rewrite unfold-take (suc m) x (y ∷ xs) | unfold-take m y xs = lookup-inject≤-take m m≤m+n i xs
176+
lookup-inject≤-take (suc m) m≤m+n zero (x ∷ xs) = refl
177+
lookup-inject≤-take (suc m) (s≤s m≤m+n) (suc i) (x ∷ xs) = lookup-inject≤-take m m≤m+n i xs
225178

226179
------------------------------------------------------------------------
227180
-- updateAt (_[_]%=_)
@@ -1298,3 +1251,24 @@ take-drop-id = take++drop≡id
12981251
"Warning: take-drop-id was deprecated in v2.0.
12991252
Please use take++drop≡id instead."
13001253
#-}
1254+
1255+
take-distr-zipWith = take-zipWith
1256+
{-# WARNING_ON_USAGE take-distr-zipWith
1257+
"Warning: take-distr-zipWith was deprecated in v2.0.
1258+
Please use take-zipWith instead."
1259+
#-}
1260+
take-distr-map = take-map
1261+
{-# WARNING_ON_USAGE take-distr-map
1262+
"Warning: take-distr-map was deprecated in v2.0.
1263+
Please use take-map instead."
1264+
#-}
1265+
drop-distr-zipWith = drop-zipWith
1266+
{-# WARNING_ON_USAGE drop-distr-zipWith
1267+
"Warning: drop-distr-zipWith was deprecated in v2.0.
1268+
Please use tdrop-zipWith instead."
1269+
#-}
1270+
drop-distr-map = drop-map
1271+
{-# WARNING_ON_USAGE drop-distr-map
1272+
"Warning: drop-distr-map was deprecated in v2.0.
1273+
Please use drop-map instead."
1274+
#-}

src/Data/Vec/Relation/Unary/All/Properties.agda

+2-4
Original file line numberDiff line numberDiff line change
@@ -148,13 +148,11 @@ module _ {P : A → Set p} where
148148

149149
drop⁺ : m {xs} All P {m + n} xs All P {n} (drop m xs)
150150
drop⁺ zero pxs = pxs
151-
drop⁺ (suc m) {x ∷ xs} (px ∷ pxs)
152-
rewrite Vecₚ.unfold-drop m x xs = drop⁺ m pxs
151+
drop⁺ (suc m) {x ∷ xs} (px ∷ pxs) = drop⁺ m pxs
153152

154153
take⁺ : m {xs} All P {m + n} xs All P {m} (take m xs)
155154
take⁺ zero pxs = []
156-
take⁺ (suc m) {x ∷ xs} (px ∷ pxs)
157-
rewrite Vecₚ.unfold-take m x xs = px ∷ take⁺ m pxs
155+
take⁺ (suc m) {x ∷ xs} (px ∷ pxs) = px ∷ take⁺ m pxs
158156

159157
------------------------------------------------------------------------
160158
-- toList

src/Data/Vec/Relation/Unary/AllPairs/Properties.agda

+2-4
Original file line numberDiff line numberDiff line change
@@ -70,13 +70,11 @@ module _ {R : Rel A ℓ} where
7070

7171
take⁺ : {n} m {xs} AllPairs R {m + n} xs AllPairs R {m} (take m xs)
7272
take⁺ zero pxs = []
73-
take⁺ (suc m) {x ∷ xs} (px ∷ pxs)
74-
rewrite Vecₚ.unfold-take m x xs = Allₚ.take⁺ m px ∷ take⁺ m pxs
73+
take⁺ (suc m) {x ∷ xs} (px ∷ pxs) = Allₚ.take⁺ m px ∷ take⁺ m pxs
7574

7675
drop⁺ : {n} m {xs} AllPairs R {m + n} xs AllPairs R {n} (drop m xs)
7776
drop⁺ zero pxs = pxs
78-
drop⁺ (suc m) {x ∷ xs} (_ ∷ pxs)
79-
rewrite Vecₚ.unfold-drop m x xs = drop⁺ m pxs
77+
drop⁺ (suc m) (_ ∷ pxs) = drop⁺ m pxs
8078

8179
------------------------------------------------------------------------
8280
-- tabulate

0 commit comments

Comments
 (0)