How can I implement a `rotate` function on Vec by using `splitAt`? - agda

Question
I'm trying to implement a rotate function on Vec, which moves every element n positions to the left, looping around. I could implement that function by using splitAt. Here is a sketch:
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Fin
open import Data.Vec
open import Relation.Nullary.Decidable
open import Relation.Binary.PropositionalEquality
rotateLeft : {A : Set} -> {w : ℕ} -> {w≢0 : False (w ≟ 0)} -> ℕ -> Vec A w -> Vec A w
rotateLeft {A} {w} n vec =
let parts = splitAt (toℕ (n mod w)) {n = ?} vec
in ?
The problem is that splitAt requires two inputs, m and n, such that the size of the vector is m + n. Since the size of the vector here is w, I need to find a k such that k + toℕ (n mod w) = w. I couldn't find any standard function handy for that. What is the best way to proceed?
Some possibilities?
Perhaps it would be helpful if k = n mod w gave me a proof that k < w, that way I could try implementing a function diff : ∀ {k w} -> k < w -> ∃ (λ a : Nat) -> a + k = w. Another possibility would be to just receive a and b as inputs, rather than the bits to shift and size of the vector, but I'm not sure that is the best interface.
Update
I've implemented the following:
add-diff : (a : ℕ) -> (b : Fin (suc a)) -> toℕ b + (a ℕ-ℕ b) ≡ a
add-diff zero zero = refl
add-diff zero (suc ())
add-diff (suc a) zero = refl
add-diff (suc a) (suc b) = cong suc (aaa a b)
Now I just need a proof that ∀ {n m} -> n mod m < m.

Here's what I came up with.
open import Data.Vec
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Fin hiding (_+_)
open import Data.Product
open import Relation.Binary.PropositionalEquality
open import Data.Nat.Properties using (+-comm)
difference : ∀ m (n : Fin m) → ∃ λ o → m ≡ toℕ n + o
difference zero ()
difference (suc m) zero = suc m , refl
difference (suc m) (suc n) with difference m n
difference (suc m) (suc n) | o , p1 = o , cong suc p1
rotate-help : ∀ {A : Set} {m} (n : Fin m) → Vec A m → Vec A m
rotate-help {A} {m = m} n vec with difference m n
... | o , p rewrite p with splitAt (toℕ n) vec
... | xs , ys , _ = subst (Vec A) (+-comm o (toℕ n)) (ys ++ xs)
rotate : ∀ {A : Set} {m} (n : ℕ) → Vec A m → Vec A m
rotate {m = zero} n v = v
rotate {m = suc m} n v = rotate-help (n mod suc m) v

After talking with adamse on IRC, I've came up with this:
open import Data.Fin hiding (_+_)
open import Data.Vec
open import Data.Nat
open import Data.Nat.Properties
open import Data.Nat.DivMod
open import Data.Empty
open import Data.Product
open import Relation.Binary.PropositionalEquality
open import Relation.Nullary.Decidable
diff : {a : ℕ} → {b : Fin a} → ∃ λ c → toℕ b + c ≡ a
diff {zero} {()}
diff {suc a} {zero} = suc a , refl
diff {suc a} {suc b} with diff {a} {b}
... | c , prf = c , cong suc prf
rotateLeft : {A : Set} → {w : ℕ} → {w≢0 : False (w ≟ 0)} → ℕ → Vec A w → Vec A w
rotateLeft {A} {w} {w≢0} n v =
let m = _mod_ n w {w≢0}
d = diff {w} {m}
d₁ = proj₁ d
d₂ = proj₂ d
d₃ = subst (λ x → x ≡ w) (+-comm (toℕ (n mod w)) d₁) d₂
v₁ = subst (λ x → Vec A x) (sym d₂) v
sp = splitAt {A = A} (toℕ m) {n = d₁} v₁
xs = proj₁ (proj₂ sp)
ys = proj₁ sp
in subst (λ x → Vec A x) d₃ (xs ++ ys)
Which is nowhere as elegant as his implementation (partly because I'm still learning Agda's syntax so I opt to just use functions), but works. Now I should return a more refined type, I believe. (Can't thank him enough!)

For your last question to just prove k < w, since k = toℕ (n mod w), you can use bounded from Data.Fin.Properties:
bounded : ∀ {n} (i : Fin n) → toℕ i ℕ< n

Related

Why won't the following Agda code typecheck?

I'm new to Agda and am puzzled by this one.
open import Data.Vec
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Fin hiding (_+_ ; splitAt)
open import Data.Product
open import Relation.Binary.PropositionalEquality
difference : ∀ m (n : Fin m) → ∃ λ o → m ≡ toℕ n + o
difference zero ()
difference (suc m) zero = suc m , refl
difference (suc m) (suc n) with difference m n
difference (suc m) (suc n) | o , p1 = o , cong suc p1
takeFin : ∀ {A : Set} {m : ℕ} (n : Fin m) → Vec A m → Vec A (toℕ n)
takeFin {A} {m = m} n vec with difference m n
... | o , p rewrite p with splitAt (toℕ n) vec
... | xs , _ , _ = xs
The takeFin function gives the error message:
m != lhs of type ℕ
when checking that the type
{m : ℕ} (n : Fin m) (o : ℕ) (p : m ≡ toℕ n + o) (lhs : ℕ) →
lhs ≡ toℕ n + o → {A : Set} (vec : Vec A lhs) → Vec A (toℕ n)
of the generated with function is well-formed
but the following functions do compile
takeFin' : ∀ {A : Set} {m : ℕ} (n : Fin m) → Vec A m → Vec A m
takeFin' {A} {m = m} n a vec with difference m n
... | o , p rewrite p with splitAt (toℕ n) vec
... | xs , ys , _ = xs ++ ys
takeFin'' : ∀ {A : Set} {m : ℕ} (n : Fin m) → A → Vec A m → Vec A (toℕ n)
takeFin'' {A} {m = m} n a vec = replicate a
Can anyone help me out?
As new Agda users tend to do, you did complicate matters a lot more than you needed to. What you intend to prove can actually be done in a much simpler way, as follows:
open import Data.Vec
open import Data.Fin
takeFin : ∀ {a} {A : Set a} {m} {n : Fin m} → Vec A m → Vec A (toℕ n)
takeFin {n = zero} (x ∷ v) = []
takeFin {n = suc _} (x ∷ v) = x ∷ takeFin v
You should always try to write simple inductive proofs rather than using unnecessary intermediate lemmas.
As to why your version does not typecheck (it's not compilation, it's type checking) the reason lies in your rewrite call which is made on an element of m ≡ toℕ n + o while your goal is of type Vec A (toℕ n) and does not contain any occurrence of m. What you want to do instead is to transform the type of vec in your context, while rewrite only acts over the goal. Here is how I would make it work:
takeFin : ∀ {A : Set} {m} {n : Fin m} → Vec A m → Vec A (toℕ n)
takeFin {m = m} {n} vec with difference m n
... | _ , p = proj₁ (splitAt (toℕ n) (subst (Vec _) p vec))
It works but as you can see it is far less elegant (and it also requires the difference function that you defined) and, more importantly, it uses subst which is often discouraged.
As a side note, and mostly for fun, it's possible to make the function a bit more concise and elegant (but less understandable) as follows:
open import Function
takeFin : ∀ {A : Set} {m} {n : Fin m} → Vec A m → Vec A (toℕ n)
takeFin {n = n} = proj₁ ∘ (splitAt (toℕ n)) ∘ (subst (Vec _) (proj₂ (difference _ n)))
This version, while a lot more complicated to read, shows how powerful Agda is in inferring the values of parameters, as only n is explicitly given.

Problems with a conductive proof

I'm trying to understand coinduction (I'm reading Sangiorgi's book) using Agda. I already managed to prove some simple equalities between streams, but I'm stuck trying to prove that all natural numbers (values of type ℕ) are in the stream allℕ --- function allℕisℕ. Any tip on how should I proceed with this?
open import Coinduction
open import Data.Nat
module Simple where
data Stream (A : Set) : Set where
_∷_ : A → ∞ (Stream A) → Stream A
infix 4 _∈_
data _∈_ {A : Set} : A → Stream A → Set where
here : ∀ {x xs} → x ∈ x ∷ xs
there : ∀ {x y xs} → (x ∈ ♭ xs) → x ∈ y ∷ xs
enum : ℕ → Stream ℕ
enum n = n ∷ (♯ enum (suc n))
allℕ : Stream ℕ
allℕ = enum 0
allℕisℕ : ∀ (n : ℕ) → n ∈ allℕ
allℕisℕ n = ?
Just sharing the complete solution...
open import Coinduction
open import Data.Nat
module Simple where
data Stream (A : Set) : Set where
_∷_ : A → ∞ (Stream A) → Stream A
infix 4 _∈_
data _∈_ {A : Set} : A → Stream A → Set where
here : ∀ {x xs} → x ∈ x ∷ xs
there : ∀ {x y xs} → (x ∈ ♭ xs) → x ∈ y ∷ xs
enum : ℕ → Stream ℕ
enum n = n ∷ (♯ enum (suc n))
allℕ : Stream ℕ
allℕ = enum 0
∈-suc : ∀ {n m : ℕ} → n ∈ enum m → suc n ∈ enum (suc m)
∈-suc here = here
∈-suc (there p) = there (∈-suc p)
allℕisℕ : ∀ (n : ℕ) → n ∈ allℕ
allℕisℕ zero = here
allℕisℕ (suc n) = there (∈-suc (allℕisℕ n))

How to prove unfold-reverse for Vec?

The Agda standard library has a few properties on how reverse and _++_ work on List. Trying to transfer these proofs to Vec appears to be non-trivial (disregarding universes):
open import Data.Nat
open import Data.Vec
open import Relation.Binary.HeterogeneousEquality
unfold-reverse : {A : Set} → (x : A) → {n : ℕ} → (xs : Vec A n) →
reverse (x ∷ xs) ≅ reverse xs ++ [ x ]
TL;DR: How to prove unfold-reverse?
The rest of this question outlines approaches to doing so and explains what problems surface.
The type of this property is very similar to the List counter part in Data.List.Properties. The proof involves a helper which roughly translates to:
open import Function
helper : ∀ {n m} → (xs : Vec A n) → (ys : Vec A m) →
foldl (Vec A ∘ (flip _+_ n)) (flip _∷_) xs ys ≅ reverse ys ++ xs
Trying to insert this helper in unfold-reverse fails, because the left hand reverse is a foldl application with Vec A ∘ suc as first argument whereas the left hand side of helper has a foldl application with Vec A ∘ (flip _+_ 1) as first argument. Even though suc ≗ flip _+_ 1 is readily available from Data.Nat.Properties.Simple, it cannot be used here as cong would need a non-pointwise equality here and we don't have extensionality without further assumptions.
Removing the flip from flip _+_ n in helper yields a type error, so that is no option either.
Any other ideas?
The Data.Vec.Properties module contains this function:
foldl-cong : ∀ {a b} {A : Set a}
{B₁ : ℕ → Set b}
{f₁ : ∀ {n} → B₁ n → A → B₁ (suc n)} {e₁}
{B₂ : ℕ → Set b}
{f₂ : ∀ {n} → B₂ n → A → B₂ (suc n)} {e₂} →
(∀ {n x} {y₁ : B₁ n} {y₂ : B₂ n} →
y₁ ≅ y₂ → f₁ y₁ x ≅ f₂ y₂ x) →
e₁ ≅ e₂ →
∀ {n} (xs : Vec A n) →
foldl B₁ f₁ e₁ xs ≅ foldl B₂ f₂ e₂ xs
foldl-cong _ e₁=e₂ [] = e₁=e₂
foldl-cong {B₁ = B₁} f₁=f₂ e₁=e₂ (x ∷ xs) =
foldl-cong {B₁ = B₁ ∘ suc} f₁=f₂ (f₁=f₂ e₁=e₂) xs
Here is more or less elaborated solution:
unfold-reverse : {A : Set} → (x : A) → {n : ℕ} → (xs : Vec A n) →
reverse (x ∷ xs) ≅ reverse xs ++ (x ∷ [])
unfold-reverse x xs = begin
foldl (Vec _ ∘ _+_ 1) (flip _∷_) (x ∷ []) xs
≅⟨ (foldl-cong
{B₁ = Vec _ ∘ _+_ 1}
{f₁ = flip _∷_}
{e₁ = x ∷ []}
{B₂ = Vec _ ∘ flip _+_ 1}
{f₂ = flip _∷_}
{e₂ = x ∷ []}
(λ {n} {a} {as₁} {as₂} as₁≅as₂ -> {!!})
refl
xs) ⟩
foldl (Vec _ ∘ flip _+_ 1) (flip _∷_) (x ∷ []) xs
≅⟨ helper (x ∷ []) xs ⟩
reverse xs ++ x ∷ []
∎
Note, that only B₁ and B₂ are distinct in the arguments of the foldl-cong function. After simplifying context in the hole we have
Goal: a ∷ as₁ ≅ a ∷ as₂
————————————————————————————————————————————————————————————
as₁≅as₂ : as₁ ≅ as₂
as₂ : Vec A (n + 1)
as₁ : Vec A (1 + n)
a : A
n : ℕ
A : Set
So we need to prove, that at each recursive call adding an element to an accumulator of type Vec A (n + 1) is equal to adding an element to an accumulator of type Vec A (1 + n), and then results of two foldls are equal. The proof itself is simple. Here is the whole code:
open import Function
open import Relation.Binary.HeterogeneousEquality
open import Data.Nat
open import Data.Vec
open import Data.Nat.Properties.Simple
open import Data.Vec.Properties
open ≅-Reasoning
postulate
helper : ∀ {n m} {A : Set} (xs : Vec A n) (ys : Vec A m)
-> foldl (Vec A ∘ flip _+_ n) (flip _∷_) xs ys ≅ reverse ys ++ xs
cong' : ∀ {α β γ} {I : Set α} {i j : I}
-> (A : I -> Set β) {B : {k : I} -> A k -> Set γ} {x : A i} {y : A j}
-> i ≅ j
-> (f : {k : I} -> (x : A k) -> B x)
-> x ≅ y
-> f x ≅ f y
cong' _ refl _ refl = refl
unfold-reverse : {A : Set} → (x : A) → {n : ℕ} → (xs : Vec A n) →
reverse (x ∷ xs) ≅ reverse xs ++ (x ∷ [])
unfold-reverse x xs = begin
foldl (Vec _ ∘ _+_ 1) (flip _∷_) (x ∷ []) xs
≅⟨ (foldl-cong
{B₁ = Vec _ ∘ _+_ 1}
{f₁ = flip _∷_}
{e₁ = x ∷ []}
{B₂ = Vec _ ∘ flip _+_ 1}
{f₂ = flip _∷_}
{e₂ = x ∷ []}
(λ {n} {a} {as₁} {as₂} as₁≅as₂ -> begin
a ∷ as₁
≅⟨ cong' (Vec _) (sym (≡-to-≅ (+-comm n 1))) (_∷_ a) as₁≅as₂ ⟩
a ∷ as₂
∎)
refl
xs) ⟩
foldl (Vec _ ∘ flip _+_ 1) (flip _∷_) (x ∷ []) xs
≅⟨ helper (x ∷ []) xs ⟩
reverse xs ++ x ∷ []
∎

Lexicographic ordering of pairs/lists in Agda using the standard library

The Agda standard library contains some modules Relation.Binary.*.(Non)StrictLex (currently only for Product and List). We can use these modules to easily construct an instance of, for example, IsStrictTotalOrder for pairs of natural numbers (i.e. ℕ × ℕ).
open import Data.Nat as ℕ using (ℕ; _<_)
open import Data.Nat.Properties as ℕ
open import Relation.Binary using (module StrictTotalOrder; IsStrictTotalOrder)
open import Relation.Binary.PropositionalEquality using (_≡_)
open import Relation.Binary.Product.StrictLex using (×-Lex; _×-isStrictTotalOrder_)
open import Relation.Binary.Product.Pointwise using (_×-Rel_)
ℕ-isSTO : IsStrictTotalOrder _≡_ _<_
ℕ-isSTO = StrictTotalOrder.isStrictTotalOrder ℕ.strictTotalOrder
ℕ×ℕ-isSTO : IsStrictTotalOrder (_≡_ ×-Rel _≡_) (×-Lex _≡_ _<_ _<_)
ℕ×ℕ-isSTO = ℕ-isSTO ×-isStrictTotalOrder ℕ-isSTO
This creates an instance using the pointwise equality _≡_ ×-Rel _≡_. In the case of propositional equality, this should be equivalent to using just propositional equality.
Is there an easy way of converting the instance above to an instance of type IsStrictTotalOrder _≡_ (×-Lex _≡_ _<_ _<_), using normal propositional equality?
The kit required isn't too hard to assemble:
open import Data.Product
open import Function using (_∘_; case_of_)
open import Relation.Binary
_⇔₂_ : ∀ {a ℓ₁ ℓ₂} {A : Set a} → Rel A ℓ₁ → Rel A ℓ₂ → Set _
_≈_ ⇔₂ _≈′_ = (∀ {x y} → x ≈ y → x ≈′ y) × (∀ {x y} → x ≈′ y → x ≈ y)
-- I was unable to write this nicely using Data.Product.map...
-- hence it is moved here to a toplevel where it can pattern-match
-- on the product of proofs
transform-resp : ∀ {a ℓ₁ ℓ₂ ℓ} {A : Set a} {≈ : Rel A ℓ₁} {≈′ : Rel A ℓ₂} {< : Rel A ℓ} →
≈ ⇔₂ ≈′ →
< Respects₂ ≈ → < Respects₂ ≈′
transform-resp (to , from) = λ { (resp₁ , resp₂) → (resp₁ ∘ from , resp₂ ∘ from) }
transform-isSTO : ∀ {a ℓ₁ ℓ₂ ℓ} {A : Set a} {≈ : Rel A ℓ₁} {≈′ : Rel A ℓ₂} {< : Rel A ℓ} →
≈ ⇔₂ ≈′ →
IsStrictTotalOrder ≈ < → IsStrictTotalOrder ≈′ <
transform-isSTO {≈′ = ≈′} {< = <} (to , from) isSTO = record
{ isEquivalence = let open IsEquivalence (IsStrictTotalOrder.isEquivalence isSTO)
in record { refl = to refl
; sym = to ∘ sym ∘ from
; trans = λ x y → to (trans (from x) (from y))
}
; trans = IsStrictTotalOrder.trans isSTO
; compare = compare
; <-resp-≈ = transform-resp (to , from) (IsStrictTotalOrder.<-resp-≈ isSTO)
}
where
compare : Trichotomous ≈′ <
compare x y with IsStrictTotalOrder.compare isSTO x y
compare x y | tri< a ¬b ¬c = tri< a (¬b ∘ from) ¬c
compare x y | tri≈ ¬a b ¬c = tri≈ ¬a (to b) ¬c
compare x y | tri> ¬a ¬b c = tri> ¬a (¬b ∘ from) c
Then we can use this to solve your original problem:
ℕ×ℕ-isSTO′ : IsStrictTotalOrder _≡_ (×-Lex _≡_ _<_ _<_)
ℕ×ℕ-isSTO′ = transform-isSTO (to , from) ℕ×ℕ-isSTO
where
open import Function using (_⟨_⟩_)
open import Relation.Binary.PropositionalEquality
to : ∀ {a b} {A : Set a} {B : Set b}
{x x′ : A} {y y′ : B} → (x , y) ⟨ _≡_ ×-Rel _≡_ ⟩ (x′ , y′) → (x , y) ≡ (x′ , y′)
to (refl , refl) = refl
from : ∀ {a b} {A : Set a} {B : Set b}
{x x′ : A} {y y′ : B} → (x , y) ≡ (x′ , y′) → (x , y) ⟨ _≡_ ×-Rel _≡_ ⟩ (x′ , y′)
from refl = refl , refl

Implicit arguments and applying a function to the tail-part of fixed-size-vectors

I wrote an Agda-function applyPrefix to apply a fixed-size-vector-function to the initial part of a longer vector where the vector-sizes m, n and k may stay implicit. Here's the definition together with a helper-function split:
split : ∀ {A m n} → Vec A (n + m) → (Vec A n) × (Vec A m)
split {_} {_} {zero} xs = ( [] , xs )
split {_} {_} {suc _} (x ∷ xs) with split xs
... | ( ys , zs ) = ( (x ∷ ys) , zs )
applyPrefix : ∀ {A n m k} → (Vec A n → Vec A m) → Vec A (n + k) → Vec A (m + k)
applyPrefix f xs with split xs
... | ( ys , zs ) = f ys ++ zs
I need a symmetric function applyPostfix which applies a fixed-size-vector-function to the tail-part of a longer vector.
applyPostfix ∀ {A n m k} → (Vec A n → Vec A m) → Vec A (k + n) → Vec A (k + m)
applyPostfix {k = k} f xs with split {_} {_} {k} xs
... | ( ys , zs ) = ys ++ (f zs)
As the definition of applyPrefix already shows, the k-Argument cannot stay implicit when applyPostfix is used. For example:
change2 : {A : Set} → Vec A 2 → Vec A 2
change2 ( x ∷ y ∷ [] ) = (y ∷ x ∷ [] )
changeNpre : {A : Set}{n : ℕ} → Vec A (2 + n) → Vec A (2 + n)
changeNpre = applyPrefix change2
changeNpost : {A : Set}{n : ℕ} → Vec A (n + 2) → Vec A (n + 2)
changeNpost = applyPost change2 -- does not work; n has to be provided
Does anyone know a technique, how to implement applyPostfix so that the k-argument may stay implicit when using applyPostfix?
What I did is proofing / programming:
lem-plus-comm : (n m : ℕ) → (n + m) ≡ (m + n)
and use that lemma when defining applyPostfix:
postfixApp2 : ∀ {A}{n m k : ℕ} → (Vec A n → Vec A m) → Vec A (k + n) → Vec A (k + m)
postfixApp2 {A} {n} {m} {k} f xs rewrite lem-plus-comm n k | lem-plus-comm k n | lem-plus-comm k m | lem-plus-comm m k = reverse (drop {n = n} (reverse xs)) ++ f (reverse (take {n = n} (reverse xs)))
Unfortunately, this didnt help, since I use the k-Parameter for calling the lemma :-(
Any better ideas how to avoid k to be explicit? Maybe I should use a snoc-View on Vectors?
What you can do is to give postfixApp2 the same type as applyPrefix.
The source of the problem is that a natural number n can be unified with p + q only if p is known. This is because + is defined via induction on the first argument.
So this one works (I'm using the standard-library version of commutativity on +):
+-comm = comm
where
open IsCommutativeSemiring isCommutativeSemiring
open IsCommutativeMonoid +-isCommutativeMonoid
postfixApp2 : {A : Set} {n m k : ℕ}
→ (Vec A n → Vec A m)
→ Vec A (n + k) → Vec A (m + k)
postfixApp2 {A} {n} {m} {k} f xs rewrite +-comm n k | +-comm m k =
applyPostfix {k = k} f xs
Yes, I'm reusing the original applyPostfix here and just give it a different type by rewriting twice.
And testing:
changeNpost : {A : Set} {n : ℕ} → Vec A (2 + n) → Vec A (2 + n)
changeNpost = postfixApp2 change2
test : changeNpost (1 ∷ 2 ∷ 3 ∷ 4 ∷ []) ≡ 1 ∷ 2 ∷ 4 ∷ 3 ∷ []
test = refl

Resources