Skip to content
Merged
Show file tree
Hide file tree
Changes from 4 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
28 changes: 28 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -43,3 +43,31 @@ Additions to existing modules
∸-suc : m ≤ n → suc n ∸ m ≡ suc (n ∸ m)
```

* In `Data.Vec.Properties`:
```agda
take-updateAt : (f : A → A) {m n : ℕ} (xs : Vec A (m + n)) (i : Fin m) →
updateAt (take m xs) i f ≡ take m (updateAt xs (inject≤ i (m≤m+n m n)) f)

truncate-zipWith : (f : A → B → C) (m≤n : m ≤ n) (xs : Vec A n) (ys : Vec B n) →
truncate m≤n (zipWith f xs ys) ≡ zipWith f (truncate m≤n xs) (truncate m≤n ys)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Ditto. re: alignment


truncate-zipWith-truncate : truncate o≤m (zipWith f (truncate m≤n xs) ys) ≡
zipWith f (truncate o≤n xs) (truncate o≤m ys)

zipWith-truncate : zipWith f (truncate p≤p+q xs) (truncate p≤p+q ys) ≡
truncate p≤p+q (zipWith f xs ys)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I don't understand why you've changed to p and q here, rather than m and n?
Also, we 'know' that p ≤ p + q, so the (presumably) mnemonic name must be carrying other information? If not, simply use m≤n?


zipWith-truncate₁ : zipWith f (truncate o≤o+m+n xs) (truncate (o≤o+m) ys) ≡
truncate (o≤o+m) (zipWith f (truncate (o+m≤o+m+n) xs) ys)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Similarly here, the names are getting very unwieldy... I'll look at the code to see what might be going on...


truncate-updateAt : (f : A → A) (m≤n : m ≤ n) (xs : Vec A n) (i : Fin m) →
updateAt (truncate m≤n xs) i f ≡ truncate m≤n (updateAt xs (inject≤ i m≤n) f)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Alignment.
Position of f binding


updateAt-truncate : updateAt (truncate p≤p+q xs) i f ≡ truncate p≤p+q (updateAt xs i′ f)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

What is i′ here? It's out of scope, so the reader won't know what it refers to...


truncate++drop≡id : (xs : Vec A (m + n)) → truncate (m≤m+n m n) xs ++ drop m xs ≡ xs
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I'm in two minds whether this combination of take≡truncate and take++drop≡id meets the Fairbairn threshold?


truncate-map : (f : A → B) (m : ℕ) (m≤n : m ≤ n) (xs : Vec A n) →
map f (truncate m≤n xs) ≡ truncate m≤n (map f xs)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Should this be truncate-map or map-truncate? probably the former, but it's a bit 'odd' wrt our L-R convention for names derived from function compositions...
Also, why is m explicit here? It appears neither in the statement, nor the proof, IIUC.


```
89 changes: 86 additions & 3 deletions src/Data/Vec/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -11,14 +11,14 @@ module Data.Vec.Properties where
open import Algebra.Definitions
open import Data.Bool.Base using (true; false)
open import Data.Fin.Base as Fin
using (Fin; zero; suc; toℕ; fromℕ<; _↑ˡ_; _↑ʳ_)
using (Fin; zero; suc; toℕ; fromℕ<; _↑ˡ_; _↑ʳ_; inject≤)
open import Data.List.Base as List using (List)
import Data.List.Properties as List
open import Data.Nat.Base
using (ℕ; zero; suc; _+_; _≤_; _<_; s≤s; pred; s<s⁻¹; _≥_; s≤s⁻¹; z≤n)
using (ℕ; zero; suc; _+_; _≤_; _<_; s≤s; pred; s<s⁻¹; _≥_; s≤s⁻¹; z≤n; _∸_)
open import Data.Nat.Properties
using (+-assoc; m≤n⇒m≤1+n; m≤m+n; ≤-refl; ≤-trans; ≤-irrelevant; ≤⇒≤″
; suc-injective; +-comm; +-suc; +-identityʳ)
; suc-injective; +-comm; +-suc; +-identityʳ; m≤n⇒∃[o]m+o≡n)
open import Data.Product.Base as Product
using (_×_; _,_; proj₁; proj₂; <_,_>; uncurry)
open import Data.Sum.Base using ([_,_]′)
Expand Down Expand Up @@ -104,6 +104,11 @@ take-map : ∀ (f : A → B) (m : ℕ) (xs : Vec A (m + n)) →
take-map f zero xs = refl
take-map f (suc m) (x ∷ xs) = cong (f x ∷_) (take-map f m xs)

take-updateAt : (f : A → A) {m n : ℕ} (xs : Vec A (m + n)) (i : Fin m) →
updateAt (take m xs) i f ≡ take m (updateAt xs (inject≤ i (m≤m+n m n)) f)
take-updateAt f (_ ∷ _ ) zero = refl
take-updateAt f (x ∷ xs) (suc i) = cong (x ∷_) (take-updateAt f xs i)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Hmmm... alignment is correct here, so why not in CHANGELOG?

Suggested change
take-updateAt : (f : A A) {m n : ℕ} (xs : Vec A (m + n)) (i : Fin m) →
updateAt (take m xs) i f ≡ take m (updateAt xs (inject≤ i (m≤m+n m n)) f)
take-updateAt f (_ ∷ _ ) zero = refl
take-updateAt f (x ∷ xs) (suc i) = cong (x ∷_) (take-updateAt f xs i)
take-updateAt : (xs : Vec A (m + n)) (i : Fin m) (f : A A) →
updateAt (take m xs) i f ≡ take m (updateAt xs (inject≤ i (m≤m+n m n)) f)
take-updateAt (_ ∷ _ ) zero f = refl
take-updateAt (x ∷ xs) (suc i) f = cong (x ∷_) (take-updateAt xs i f)


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

Expand Down Expand Up @@ -154,6 +159,70 @@ take≡truncate : ∀ m (xs : Vec A (m + n)) →
take≡truncate zero _ = refl
take≡truncate (suc m) (x ∷ xs) = cong (x ∷_) (take≡truncate m xs)

truncate-zipWith : (f : A → B → C) (m≤n : m ≤ n) (xs : Vec A n) (ys : Vec B n) →
truncate m≤n (zipWith f xs ys) ≡ zipWith f (truncate m≤n xs) (truncate m≤n ys)
truncate-zipWith f z≤n xs ys = refl
truncate-zipWith f (s≤s m≤n) (x ∷ xs) (y ∷ ys) =
cong (f x y ∷_) (truncate-zipWith f m≤n xs ys)

truncate-zipWith-truncate : ∀ (f : A → B → C) (o≤m : o ≤ m) (m≤n : m ≤ n)
(xs : Vec A n) (ys : Vec B m) →
let o≤n = ≤-trans o≤m m≤n in
truncate o≤m (zipWith f (truncate m≤n xs) ys) ≡
zipWith f (truncate o≤n xs) (truncate o≤m ys)
truncate-zipWith-truncate f o≤m m≤n xs ys =
let o≤n = ≤-trans o≤m m≤n in
trans (truncate-zipWith f o≤m (truncate m≤n xs) ys)
(cong (λ zs → zipWith f zs (truncate o≤m ys)) ((sym (truncate-trans o≤m m≤n xs))))
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

A repeated let binding between type and definition is a UX/cognitive annoyance, for which the only really satisfactory solution is to use an anonymous module to share the definition... but here o≤n doesn't appear in the definition, so can be dropped in favour of

Suggested change
truncate-zipWith-truncate : (f : A B C) (o≤m : o ≤ m) (m≤n : m ≤ n)
(xs : Vec A n) (ys : Vec B m) →
let o≤n = ≤-trans o≤m m≤n in
truncate o≤m (zipWith f (truncate m≤n xs) ys) ≡
zipWith f (truncate o≤n xs) (truncate o≤m ys)
truncate-zipWith-truncate f o≤m m≤n xs ys =
let o≤n = ≤-trans o≤m m≤n in
trans (truncate-zipWith f o≤m (truncate m≤n xs) ys)
(cong (λ zs zipWith f zs (truncate o≤m ys)) ((sym (truncate-trans o≤m m≤n xs))))
truncate-zipWith-truncate : (f : A B C) (o≤m : o ≤ m) (m≤n : m ≤ n)
(xs : Vec A n) (ys : Vec B m) →
truncate o≤m (zipWith f (truncate m≤n xs) ys) ≡
zipWith f (truncate (≤-trans o≤m m≤n) xs) (truncate o≤m ys)
truncate-zipWith-truncate f o≤m m≤n xs ys =
trans (truncate-zipWith f o≤m (truncate m≤n xs) ys)
(cong (λ zs zipWith f zs (truncate o≤m ys)) ((sym (truncate-trans o≤m m≤n xs))))

plus:

  • fix up alignment
  • if at all possible, rename and reorder so that m ≤ n ≤o in terms of the (implicit) binder ordering... much less cognitive strain on the reader


zipWith-truncate : ∀ (f : A → B → C) {p q : ℕ} (xs : Vec A (p + q)) (ys : Vec B (p + q)) →
let p≤p+q = m≤m+n p q in
zipWith f (truncate p≤p+q xs) (truncate p≤p+q ys) ≡ truncate p≤p+q (zipWith f xs ys)
zipWith-truncate f {p} {q} xs ys =
let p≤p+q = m≤m+n p q in
begin
zipWith f (truncate p≤p+q xs) (truncate p≤p+q ys) ≡⟨ cong₂ (zipWith f) (take≡truncate p xs) (take≡truncate p ys) ⟨
zipWith f (take p xs) (take p ys) ≡⟨ take-zipWith f xs ys ⟨
take p (zipWith f xs ys) ≡⟨ take≡truncate p (zipWith f xs ys) ⟩
truncate p≤p+q (zipWith f xs ys) ∎ where open ≡-Reasoning
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Assuming this meets the Fairbairn threshold, it's still very hard to parse, but AFAICT, this is a specialised instance of truncate-zipWith above... or am I missing something?
Suggest: deletion


module _ (f : A → B → C) {o m n : ℕ} (xs : Vec A (o + m + n)) (ys : Vec B (o + m)) where
private
o≤o+m = m≤m+n o m
o+m≤o+m+n = m≤m+n (o + m) n
o≤o+m+n = ≤-trans (m≤m+n o m) (m≤m+n (o + m) n)

zipWith-truncate₁ : zipWith f (truncate o≤o+m+n xs) (truncate (o≤o+m) ys) ≡
truncate (o≤o+m) (zipWith f (truncate (o+m≤o+m+n) xs) ys)
zipWith-truncate₁ =
trans (cong (λ zs → zipWith f zs (truncate (o≤o+m) ys)) (truncate-trans (o≤o+m) (o+m≤o+m+n) xs))
(zipWith-truncate f (truncate (o+m≤o+m+n) xs) ys)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Hmmm. Not quite sure what's being gained here over truncate-zipWith-truncate above, given that

o≤o+m+n = ≤-trans o≤o+m o+m≤o+m+n

Fairbairn threshold?


truncate-updateAt : (f : A → A) (m≤n : m ≤ n) (xs : Vec A n) (i : Fin m) →
updateAt (truncate m≤n xs) i f ≡ truncate m≤n (updateAt xs (inject≤ i m≤n) f)
truncate-updateAt f (s≤s _ ) (_ ∷ _ ) zero = refl
truncate-updateAt f (s≤s m≤n) (x ∷ xs) (suc i) = cong (x ∷_) (truncate-updateAt f m≤n xs i)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Move f binding?


module _ (f : A → A) {p q : ℕ} (xs : Vec A (p + q)) (i : Fin p) where
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
module _ (f : A A) {p q : ℕ} (xs : Vec A (p + q)) (i : Fin p) where
module _ (xs : Vec A (m + n)) (i : Fin m) (f : A A) where

private
p≤p+q : p ≤ p + q
p≤p+q = m≤m+n p q
i′ : Fin (p + q)
i′ = inject≤ i p≤p+q

updateAt-truncate : updateAt (truncate p≤p+q xs) i f ≡ truncate p≤p+q (updateAt xs i′ f)
updateAt-truncate = begin
updateAt (truncate p≤p+q xs) i f ≡⟨ cong (λ l → updateAt l i f) (sym (take≡truncate p xs)) ⟩
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Avoid explicit sym

Suggested change
updateAt (truncate p≤p+q xs) i f ≡⟨ cong (λ l updateAt l i f) (sym (take≡truncate p xs)) ⟩
updateAt (truncate p≤p+q xs) i f ≡⟨ cong (λ l updateAt l i f) (take≡truncate p xs)

updateAt (take p xs) i f ≡⟨ take-updateAt f xs i ⟩
take p (updateAt xs i′ f) ≡⟨ take≡truncate p (updateAt xs i′ f) ⟩
truncate p≤p+q (updateAt xs i′ f) ∎ where open ≡-Reasoning
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
truncate p≤p+q (updateAt xs i′ f) ∎ where open ≡-Reasoning
truncate p≤p+q (updateAt xs i′ f) ∎
where open ≡-Reasoning

See style-guide for layout conventions wrt equational proofs.


truncate++drop≡id : (xs : Vec A (m + n)) → truncate (m≤m+n m n) xs ++ drop m xs ≡ xs
truncate++drop≡id {m = m} {n} xs = begin
truncate (m≤m+n m n) xs ++ drop m xs ≡⟨ cong (_++ drop m xs) (take≡truncate m xs) ⟨
take m xs ++ drop m xs ≡⟨ take++drop≡id m xs ⟩
xs ∎ where open ≡-Reasoning
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Ditto.


------------------------------------------------------------------------
-- pad

Expand Down Expand Up @@ -472,6 +541,20 @@ toList-map : ∀ (f : A → B) (xs : Vec A n) →
toList-map f [] = refl
toList-map f (x ∷ xs) = cong (f x List.∷_) (toList-map f xs)

truncate-map : (f : A → B) (m : ℕ) (m≤n : m ≤ n) (xs : Vec A n) →
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Suggested change
truncate-map : (f : A B) (m : ℕ) (m≤n : m ≤ n) (xs : Vec A n) →
truncate-map : (f : A B) (m≤n : m ≤ n) (xs : Vec A n) →

map f (truncate m≤n xs) ≡ truncate m≤n (map f xs)
truncate-map {n = n} f m m≤n xs =
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

... with corresponding

Suggested change
truncate-map {n = n} f m m≤n xs =
truncate-map {m = m} {n = n} f m≤n xs =

begin
map f (truncate m≤n xs) ≡⟨ cong (map f) (truncate≡take m≤n xs eq) ⟩
map f (take m (cast eq xs)) ≡⟨ take-map f m _ ⟨
take m (map f (cast eq xs)) ≡⟨ cong (take m) (map-cast f eq xs) ⟩
take m (cast eq (map f xs)) ≡⟨ truncate≡take m≤n (map f xs) eq ⟨
truncate m≤n (map f xs) ∎
where
.eq : n ≡ m + (n ∸ m)
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

No need to mark eq as irrelevant; it's the proof which is ignored by cast etc.

Suggested change
.eq : n ≡ m + (n ∸ m)
eq : n ≡ m + (n ∸ m)

eq = sym (proj₂ (m≤n⇒∃[o]m+o≡n m≤n))
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

In fact, Data.Nat.Properties.guarded-∸≗∸ is (more like) what you need here... and could then be inlined?

Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

See also padRight-drop′ and padRight-take′ that I've added to #2787 for comparison.

open ≡-Reasoning

------------------------------------------------------------------------
-- _++_

Expand Down