Skip to content

improve implementation of splitAt, take and drop. #2056

New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Merged
merged 6 commits into from
Sep 17, 2023
Merged
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
5 changes: 5 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -1409,6 +1409,11 @@ Deprecated names

* In `Data.Vec.Properties`:
```
take-distr-zipWith ↦ take-zipWith
take-distr-map ↦ take-map
drop-distr-zipWith ↦ drop-zipWith
drop-distr-map ↦ drop-map

updateAt-id-relative ↦ updateAt-id-local
updateAt-compose-relative ↦ updateAt-∘-local
updateAt-compose ↦ updateAt-∘
Expand Down
38 changes: 17 additions & 21 deletions src/Data/Vec/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -12,11 +12,11 @@ open import Data.Bool.Base using (Bool; true; false; if_then_else_)
open import Data.Nat.Base
open import Data.Fin.Base using (Fin; zero; suc)
open import Data.List.Base as List using (List)
open import Data.Product.Base as Prod using (∃; ∃₂; _×_; _,_)
open import Data.Product.Base as Prod using (∃; ∃₂; _×_; _,_; proj₁; proj₂)
open import Data.These.Base as These using (These; this; that; these)
open import Function.Base using (const; _∘′_; id; _∘_)
open import Level using (Level)
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; cong)
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl; trans; cong)
open import Relation.Nullary.Decidable.Core using (does)
open import Relation.Unary using (Pred; Decidable)

Expand Down Expand Up @@ -259,26 +259,23 @@ allFin _ = tabulate id

splitAt : ∀ m {n} (xs : Vec A (m + n)) →
∃₂ λ (ys : Vec A m) (zs : Vec A n) → xs ≡ ys ++ zs
splitAt zero xs = ([] , xs , refl)
splitAt (suc m) (x ∷ xs) with splitAt m xs
splitAt (suc m) (x ∷ .(ys ++ zs)) | (ys , zs , refl) =
((x ∷ ys) , zs , refl)
splitAt zero xs = [] , xs , refl
splitAt (suc m) (x ∷ xs) =
let ys , zs , eq = splitAt m xs in x ∷ ys , zs , cong (x ∷_) eq

take : ∀ m {n} → Vec A (m + n) → Vec A m
take m xs with splitAt m xs
take m .(ys ++ zs) | (ys , zs , refl) = ys
take m xs = proj₁ (splitAt m xs)

drop : ∀ m {n} → Vec A (m + n) → Vec A n
drop m xs with splitAt m xs
drop m .(ys ++ zs) | (ys , zs , refl) = zs
drop m xs = proj₁ (proj₂ (splitAt m xs))

group : ∀ n k (xs : Vec A (n * k)) →
∃ λ (xss : Vec (Vec A k) n) → xs ≡ concat xss
group zero k [] = ([] , refl)
group (suc n) k xs with splitAt k xs
group (suc n) k .(ys ++ zs) | (ys , zs , refl) with group n k zs
group (suc n) k .(ys ++ concat zss) | (ys , ._ , refl) | (zss , refl) =
((ys ∷ zss) , refl)
group (suc n) k xs =
let ys , zs , eq-split = splitAt k xs in
let zss , eq-group = group n k zs in
(ys ∷ zss) , trans eq-split (cong (ys ++_) eq-group)

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

initLast : ∀ (xs : Vec A (1 + n)) → ∃₂ λ ys y → xs ≡ ys ∷ʳ y
initLast {n = zero} (x ∷ []) = ([] , x , refl)
initLast {n = suc n} (x ∷ xs) with initLast xs
... | (ys , y , refl) = (x ∷ ys , y , refl)
initLast {n = zero} (x ∷ []) = [] , x , refl
initLast {n = suc n} (x ∷ xs) =
let ys , y , eq = initLast xs in
x ∷ ys , y , cong (x ∷_) eq

init : Vec A (1 + n) → Vec A n
init {n = zero} (x ∷ _) = []
init {n = suc n} (x ∷ xs) = x ∷ init xs
init xs = proj₁ (initLast xs)

last : Vec A (1 + n) → A
last {n = zero} (x ∷ _) = x
last {n = suc n} (_ ∷ xs) = last xs
last xs = proj₁ (proj₂ (initLast xs))

------------------------------------------------------------------------
-- Other operations
Expand Down
126 changes: 50 additions & 76 deletions src/Data/Vec/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -71,86 +71,43 @@ private
-- take

unfold-take : ∀ n x (xs : Vec A (n + m)) → take (suc n) (x ∷ xs) ≡ x ∷ take n xs
unfold-take n x xs with splitAt n xs
... | xs , ys , refl = refl

take-distr-zipWith : ∀ (f : A → B → C) →
(xs : Vec A (m + n)) (ys : Vec B (m + n)) →
take m (zipWith f xs ys) ≡ zipWith f (take m xs) (take m ys)
take-distr-zipWith {m = zero} f xs ys = refl
take-distr-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = begin
take (suc m) (zipWith f (x ∷ xs) (y ∷ ys))
≡⟨⟩
take (suc m) (f x y ∷ (zipWith f xs ys))
≡⟨ unfold-take m (f x y) (zipWith f xs ys) ⟩
f x y ∷ take m (zipWith f xs ys)
≡⟨ cong (f x y ∷_) (take-distr-zipWith f xs ys) ⟩
f x y ∷ (zipWith f (take m xs) (take m ys))
≡⟨⟩
zipWith f (x ∷ (take m xs)) (y ∷ (take m ys))
≡˘⟨ cong₂ (zipWith f) (unfold-take m x xs) (unfold-take m y ys) ⟩
zipWith f (take (suc m) (x ∷ xs)) (take (suc m) (y ∷ ys))

take-distr-map : ∀ (f : A → B) (m : ℕ) (xs : Vec A (m + n)) →
take m (map f xs) ≡ map f (take m xs)
take-distr-map f zero xs = refl
take-distr-map f (suc m) (x ∷ xs) = begin
take (suc m) (map f (x ∷ xs)) ≡⟨⟩
take (suc m) (f x ∷ map f xs) ≡⟨ unfold-take m (f x) (map f xs) ⟩
f x ∷ (take m (map f xs)) ≡⟨ cong (f x ∷_) (take-distr-map f m xs) ⟩
f x ∷ (map f (take m xs)) ≡⟨⟩
map f (x ∷ take m xs) ≡˘⟨ cong (map f) (unfold-take m x xs) ⟩
map f (take (suc m) (x ∷ xs)) ∎
unfold-take n x xs = refl

take-zipWith : ∀ (f : A → B → C) →
(xs : Vec A (m + n)) (ys : Vec B (m + n)) →
take m (zipWith f xs ys) ≡ zipWith f (take m xs) (take m ys)
take-zipWith {m = zero} f xs ys = refl
take-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = cong (f x y ∷_) (take-zipWith f xs ys)

take-map : ∀ (f : A → B) (m : ℕ) (xs : Vec A (m + n)) →
take m (map f xs) ≡ map f (take m xs)
take-map f zero xs = refl
take-map f (suc m) (x ∷ xs) = cong (f x ∷_) (take-map f m xs)

------------------------------------------------------------------------
-- drop

unfold-drop : ∀ n x (xs : Vec A (n + m)) →
drop (suc n) (x ∷ xs) ≡ drop n xs
unfold-drop n x xs with splitAt n xs
... | xs , ys , refl = refl

drop-distr-zipWith : (f : A → B → C) →
(x : Vec A (m + n)) (y : Vec B (m + n)) →
drop m (zipWith f x y) ≡ zipWith f (drop m x) (drop m y)
drop-distr-zipWith {m = zero} f xs ys = refl
drop-distr-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = begin
drop (suc m) (zipWith f (x ∷ xs) (y ∷ ys))
≡⟨⟩
drop (suc m) (f x y ∷ (zipWith f xs ys))
≡⟨ unfold-drop m (f x y) (zipWith f xs ys) ⟩
drop m (zipWith f xs ys)
≡⟨ drop-distr-zipWith f xs ys ⟩
zipWith f (drop m xs) (drop m ys)
≡˘⟨ cong₂ (zipWith f) (unfold-drop m x xs) (unfold-drop m y ys) ⟩
zipWith f (drop (suc m) (x ∷ xs)) (drop (suc m) (y ∷ ys))

drop-distr-map : ∀ (f : A → B) (m : ℕ) (x : Vec A (m + n)) →
drop m (map f x) ≡ map f (drop m x)
drop-distr-map f zero x = refl
drop-distr-map f (suc m) (x ∷ xs) = begin
drop (suc m) (map f (x ∷ xs)) ≡⟨⟩
drop (suc m) (f x ∷ map f xs) ≡⟨ unfold-drop m (f x) (map f xs) ⟩
drop m (map f xs) ≡⟨ drop-distr-map f m xs ⟩
map f (drop m xs) ≡˘⟨ cong (map f) (unfold-drop m x xs) ⟩
map f (drop (suc m) (x ∷ xs)) ∎
unfold-drop n x xs = refl

drop-zipWith : (f : A → B → C) →
(xs : Vec A (m + n)) (ys : Vec B (m + n)) →
drop m (zipWith f xs ys) ≡ zipWith f (drop m xs) (drop m ys)
drop-zipWith {m = zero} f xs ys = refl
drop-zipWith {m = suc m} f (x ∷ xs) (y ∷ ys) = drop-zipWith f xs ys

drop-map : ∀ (f : A → B) (m : ℕ) (xs : Vec A (m + n)) →
drop m (map f xs) ≡ map f (drop m xs)
drop-map f zero xs = refl
drop-map f (suc m) (x ∷ xs) = drop-map f m xs

------------------------------------------------------------------------
-- take and drop together

take++drop≡id : ∀ (m : ℕ) (x : Vec A (m + n)) → take m x ++ drop m x ≡ x
take++drop≡id zero x = refl
take++drop≡id (suc m) (x ∷ xs) = begin
take (suc m) (x ∷ xs) ++ drop (suc m) (x ∷ xs)
≡⟨ cong₂ _++_ (unfold-take m x xs) (unfold-drop m x xs) ⟩
(x ∷ take m xs) ++ (drop m xs)
≡⟨⟩
x ∷ (take m xs ++ drop m xs)
≡⟨ cong (x ∷_) (take++drop≡id m xs) ⟩
x ∷ xs
take++drop≡id : ∀ (m : ℕ) (xs : Vec A (m + n)) → take m xs ++ drop m xs ≡ xs
take++drop≡id zero xs = refl
take++drop≡id (suc m) (x ∷ xs) = cong (x ∷_) (take++drop≡id m xs)

------------------------------------------------------------------------
-- truncate
Expand Down Expand Up @@ -215,12 +172,8 @@ lookup⇒[]= (suc i) (_ ∷ xs) p = there (lookup⇒[]= i xs p)

lookup-inject≤-take : ∀ m (m≤m+n : m ≤ m + n) (i : Fin m) (xs : Vec A (m + n)) →
lookup xs (Fin.inject≤ i m≤m+n) ≡ lookup (take m xs) i
lookup-inject≤-take (suc m) m≤m+n zero (x ∷ xs)
rewrite unfold-take m x xs = refl
lookup-inject≤-take (suc (suc m)) (s≤s m≤m+n) (suc zero) (x ∷ y ∷ xs)
rewrite unfold-take (suc m) x (y ∷ xs) | unfold-take m y xs = refl
lookup-inject≤-take (suc (suc m)) (s≤s (s≤s m≤m+n)) (suc (suc i)) (x ∷ y ∷ xs)
rewrite unfold-take (suc m) x (y ∷ xs) | unfold-take m y xs = lookup-inject≤-take m m≤m+n i xs
lookup-inject≤-take (suc m) m≤m+n zero (x ∷ xs) = refl
lookup-inject≤-take (suc m) (s≤s m≤m+n) (suc i) (x ∷ xs) = lookup-inject≤-take m m≤m+n i xs

------------------------------------------------------------------------
-- updateAt (_[_]%=_)
Expand Down Expand Up @@ -1223,3 +1176,24 @@ take-drop-id = take++drop≡id
"Warning: take-drop-id was deprecated in v2.0.
Please use take++drop≡id instead."
#-}

take-distr-zipWith = take-zipWith
{-# WARNING_ON_USAGE take-distr-zipWith
"Warning: take-distr-zipWith was deprecated in v2.0.
Please use take-zipWith instead."
#-}
take-distr-map = take-map
{-# WARNING_ON_USAGE take-distr-map
"Warning: take-distr-map was deprecated in v2.0.
Please use take-map instead."
#-}
drop-distr-zipWith = drop-zipWith
{-# WARNING_ON_USAGE drop-distr-zipWith
"Warning: drop-distr-zipWith was deprecated in v2.0.
Please use tdrop-zipWith instead."
#-}
drop-distr-map = drop-map
{-# WARNING_ON_USAGE drop-distr-map
"Warning: drop-distr-map was deprecated in v2.0.
Please use drop-map instead."
#-}
6 changes: 2 additions & 4 deletions src/Data/Vec/Relation/Unary/All/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -148,13 +148,11 @@ module _ {P : A → Set p} where

drop⁺ : ∀ m {xs} → All P {m + n} xs → All P {n} (drop m xs)
drop⁺ zero pxs = pxs
drop⁺ (suc m) {x ∷ xs} (px ∷ pxs)
rewrite Vecₚ.unfold-drop m x xs = drop⁺ m pxs
drop⁺ (suc m) {x ∷ xs} (px ∷ pxs) = drop⁺ m pxs

take⁺ : ∀ m {xs} → All P {m + n} xs → All P {m} (take m xs)
take⁺ zero pxs = []
take⁺ (suc m) {x ∷ xs} (px ∷ pxs)
rewrite Vecₚ.unfold-take m x xs = px ∷ take⁺ m pxs
take⁺ (suc m) {x ∷ xs} (px ∷ pxs) = px ∷ take⁺ m pxs

------------------------------------------------------------------------
-- toList
Expand Down
6 changes: 2 additions & 4 deletions src/Data/Vec/Relation/Unary/AllPairs/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -70,13 +70,11 @@ module _ {R : Rel A ℓ} where

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

drop⁺ : ∀ {n} m {xs} → AllPairs R {m + n} xs → AllPairs R {n} (drop m xs)
drop⁺ zero pxs = pxs
drop⁺ (suc m) {x ∷ xs} (_ ∷ pxs)
rewrite Vecₚ.unfold-drop m x xs = drop⁺ m pxs
drop⁺ (suc m) (_ ∷ pxs) = drop⁺ m pxs

------------------------------------------------------------------------
-- tabulate
Expand Down