How to prove that division is cancellative? - agda

I want to prove that (c * a) / (c * b) = a / b in agda using the division function defined in the standard library. The proofs keep coming back to this thing div-helper that is very difficult to work with and reason about.
open import Data.Nat.DivMod using (_/_)
open import Relation.Binary.PropositionalEquality using (_≡_)
open import Data.Nat using (ℕ; suc; zero)
/-cancelˡ : ∀ c a b {b≢0} {b*c≢0} → ((c * a) / (c * b)) {b*c≢0} ≡ (a / b) {b≢0}
/-cancelˡ (suc c) a (suc b) {b≢0} {b*c≢0} = ?
That hole simplifies to:
div-helper 0 (b + c * suc b) (a + c * a) (b + c * suc b) ≡ div-helper 0 b a b

I could come up with two different proofs.
Proof 1 is completely agnostic of div-helper's structure. It is based on lemmas about division from the standard library. It proves a few additional lemmas about division, which it then uses to prove the property we're after. It's a little involved, but I think that the additional lemmas are useful in their own right.
Proof 2 considers the structure of div-helper and proves two invariants from which the property follows trivially. It's way more concise.
This is the property we're after:
∀ (a b n : ℕ) → divₕ 0 (b + n + b * n) (a + n * a) (b + n + b * n) ≡ divₕ 0 b a b
What makes this look a little weird is that the division helper takes n (as opposed to suc n) for a / suc n. That's where b + n + b * n comes from - suc (b + n + b * n) is equal to suc b * suc n.
Stated in terms of / instead of the division helper, the proof thus says that (a * suc n) / (suc b * suc n) is equal to a / suc b.
This is proof 1:
module Answer where
open import Relation.Binary.PropositionalEquality as P using (_≡_; cong; refl; subst; sym)
open import Data.Nat
open import Data.Nat.DivMod
open import Data.Nat.DivMod.Core
open import Data.Nat.Properties
open import Agda.Builtin.Nat using () renaming (div-helper to divₕ; mod-helper to modₕ)
divₕ′ : ℕ → ℕ → ℕ
divₕ′ a m = divₕ 0 m a m
modₕ′ : ℕ → ℕ → ℕ
modₕ′ a m = modₕ 0 m a m
a<n⇒a[divₕ]n≡0 : ∀ (a n n′ : ℕ) → a ≤ n → divₕ 0 n′ a n ≡ 0
a<n⇒a[divₕ]n≡0 zero _ _ _ = refl
a<n⇒a[divₕ]n≡0 (suc a) (suc n) n′ (s≤s a≤n) = a<n⇒a[divₕ]n≡0 a n n′ a≤n
a*n[modₕ]n≡0 : ∀ (a n : ℕ) → modₕ 0 n (a * suc n) n ≡ 0
a*n[modₕ]n≡0 zero n = refl
a*n[modₕ]n≡0 (suc a) n rewrite +-comm (suc n) (a * suc n) | a+n[modₕ]n≡a[modₕ]n 0 (a * suc n) n = a*n[modₕ]n≡0 a n
a<n⇒a+b*n[divₕ]n≡b : ∀ (a b n : ℕ) → a ≤ n → divₕ 0 n (a + b * suc n) n ≡ b
a<n⇒a+b*n[divₕ]n≡b a b n a≤n =
begin
divₕ′ (a + b * suc n) n
≡⟨ +-distrib-divₕ 0 0 a (b * suc n) n lem₁ ⟩
divₕ′ a n + divₕ′ (b * suc n) n
≡⟨ cong (_+ divₕ′ (b * suc n) n) (a<n⇒a[divₕ]n≡0 a n n a≤n) ⟩
0 + divₕ′ (b * suc n) n
≡⟨ +-identityˡ (divₕ 0 n (b * suc n) n) ⟩
divₕ′ (b * suc n) n
≡⟨ m*n/n≡m b (suc n) ⟩
b
∎
where
open P.≡-Reasoning
≤₁ = a[modₕ]n<n 0 a n
≤₂ = ≤-reflexive (a*n[modₕ]n≡0 b n)
<₃ : n + 0 < suc n
<₃ = s≤s (≤-reflexive (+-identityʳ n))
lem₁ : modₕ′ a n + modₕ′ (b * suc n) n < suc n
lem₁ = <-transʳ (+-mono-≤ ≤₁ ≤₂) <₃
a[divₕ]m*n≡a[divₕ]m[divₕ]n : ∀ (a m n : ℕ) → divₕ 0 (m + n + m * n) a (m + n + m * n) ≡ divₕ 0 n (divₕ 0 m a m) n
a[divₕ]m*n≡a[divₕ]m[divₕ]n a m n =
begin
divₕ′ a mn
≡⟨ cong (λ # → divₕ′ # mn) (div-mod-lemma 0 0 a m) ⟩
divₕ′ (modₕ′ a m + divₕ′ a m * suc m) mn
≡⟨ cong (λ # → divₕ′ (modₕ′ a m + # * suc m) mn) (div-mod-lemma 0 0 (divₕ′ a m) n) ⟩
divₕ′ (modₕ′ a m + (modₕ′ (divₕ′ a m) n + divₕ′ (divₕ′ a m) n * suc n) * suc m) mn
≡⟨ cong (λ # → divₕ′ (modₕ′ a m + #) mn) (*-distribʳ-+ (suc m) (modₕ′ (divₕ′ a m) n) (divₕ′ (divₕ′ a m) n * suc n)) ⟩
divₕ′ (modₕ′ a m + (modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * suc n * suc m)) mn
≡⟨ cong (λ # → divₕ′ # mn) (sym (+-assoc (modₕ′ a m) (modₕ′ (divₕ′ a m) n * suc m) (divₕ′ (divₕ′ a m) n * suc n * suc m))) ⟩
divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * suc n * suc m) mn
≡⟨ cong (λ # → divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + #) mn) (*-assoc (divₕ′ (divₕ′ a m) n) (suc n) (suc m))⟩
divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * (suc n * suc m)) mn
≡⟨ cong (λ # → divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * #) mn) (*-comm (suc n) (suc m)) ⟩
divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * (suc m * suc n)) mn
≡⟨ cong (λ # → divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * #) mn) lem₁ ⟩
divₕ′ (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m + divₕ′ (divₕ′ a m) n * suc mn) mn
≡⟨ a<n⇒a+b*n[divₕ]n≡b (modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m) (divₕ′ (divₕ′ a m) n) mn lem₄ ⟩
divₕ 0 n (divₕ′ a m) n
∎
where
open P.≡-Reasoning
mn = m + n + m * n
lem₁ : suc m * suc n ≡ suc mn
lem₁ rewrite +-comm m n | *-comm m (suc n) | *-comm m n | +-assoc n m (n * m) = refl
lem₂ : m + n * suc m ≡ mn
lem₂ rewrite *-comm n (suc m) | +-assoc m n (m * n) = refl
≤₁ = a[modₕ]n<n 0 a m
≤₂ = a[modₕ]n<n 0 (divₕ 0 m a m) n
≤₃ = ≤-refl {suc m}
lem₃ = +-mono-≤ ≤₁ (*-mono-≤ ≤₂ ≤₃)
lem₄ : modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m ≤ mn
lem₄ = subst (λ # → modₕ′ a m + modₕ′ (divₕ′ a m) n * suc m ≤ #) lem₂ lem₃
a*n[divₕ]b*n≡a[divₕ]b : ∀ (a b n : ℕ) → divₕ 0 (b + n + b * n) (a + n * a) (b + n + b * n) ≡ divₕ 0 b a b
a*n[divₕ]b*n≡a[divₕ]b a b n =
begin
divₕ′ na nb
≡⟨ cong (λ # → divₕ′ na #) nb≡bn ⟩
divₕ′ na bn
≡⟨ a[divₕ]m*n≡a[divₕ]m[divₕ]n na n b ⟩
divₕ′ (divₕ′ na n) b
≡⟨ cong (λ # → divₕ′ (divₕ′ # n) b) lem₁ ⟩
divₕ′ (divₕ′ (a * suc n) n) b
≡⟨ cong (λ # → divₕ′ # b) (a*n[divₕ]n≡a 0 a n) ⟩
divₕ 0 b a b
∎
where
open P.≡-Reasoning
na = a + n * a
nb = b + n + b * n
bn = n + b + n * b
nb≡bn : nb ≡ bn
nb≡bn rewrite +-comm n b | *-comm n b = refl
lem₁ : na ≡ a * suc n
lem₁ rewrite *-comm a (suc n) = refl
Now let's look at proof 2. Simply append it to the above code, so that I don't have to repeat the imports:
a*n[divₕ]b*n≡a[divₕ]b′ : ∀ (a b n : ℕ) → divₕ 0 (b + n + b * n) (a + n * a) (b + n + b * n) ≡ divₕ 0 b a b
a*n[divₕ]b*n≡a[divₕ]b′ a b n rewrite +-comm b n | *-comm b n = lem₂ n 0 b a b
where
lem₁ : ∀ (a k m n j : ℕ) → divₕ k m (a + n) (a + j) ≡ divₕ k m n j
lem₁ zero k m n j = refl
lem₁ (suc a) k m n j = lem₁ a k m n j
lem₂ : ∀ (a k m n j : ℕ) → divₕ k (a + m + a * m) (suc a * n) (a + j + a * j) ≡ divₕ k m n j
lem₂ a k m zero j rewrite *-zeroʳ a = refl
lem₂ a k m (suc n) zero
rewrite *-zeroʳ a
| +-identityʳ a
| *-suc a n
| P.sym (+-assoc n a (a * n))
| +-comm n a
| +-assoc a n (a * n)
| P.sym (+-suc a (n + a * n))
| lem₁ a k (a + m + a * m) (suc (n + a * n)) 0
= lem₂ a (suc k) m n m
lem₂ a k m (suc n) (suc j)
rewrite *-suc a n
| P.sym (+-assoc n a (a * n))
| +-comm n a
| +-assoc a n (a * n)
| P.sym (+-suc a (n + a * n))
| +-assoc a (suc j) (a * suc j)
| lem₁ a k (a + m + a * m) (suc (n + a * n)) (suc j + a * suc j)
| *-suc a j
| P.sym (+-assoc j a (a * j))
| +-comm j a
= lem₂ a k m n j
The first lemma is an invariant that says that we can add an identical value a to the last two arguments of div-helper, n and j. It's used by the second lemma, which is the more interesting one.
The second lemma is an invariant that says that you can multiply the last three arguments, m, n, and j, with an identical value a. Its a proof by induction on the last two arguments, n and j. This makes it closely follow the structure of div-helper.
The proof thus covers the same three different cases that show up in the three defining equations of div-helper:
n is zero
n is non-zero, but j is zero
n and j are both non-zero
The property that we're after then follows directly from the second lemma. This is basically the idea outlined in the other response, where the second lemma is the more general version of the property that we're after.

Note that the invariants div-helper respects are spelt out in the builtin file Agda.Builtin.Nat. You should state a more general version of your lemma satisfied by div-helper.

Related

How do multiple rewrites expand into with?

I am a beginner in PLFA, when I read the Induction section, I accidentally wrote a +-swap proof that I can't understand:
+-suc': ∀ (m n: ℕ) → m + suc n ≡ suc (m + n)
+-suc' zero n = refl
+-suc' (suc m) n rewrite +-suc' m n = refl
+-swap: ∀ (m n p: ℕ) → m + (n + p) ≡ n + (m + p)
+-swap zero n p = refl
+-swap (suc m) n p rewrite +-suc' n (m + p) | +-swap m n p = refl
I don't know why this proof is right, so I try to prove it by a chain of equations (which is wrong):
+-swap (suc m) n p =
begin
(suc m) + (n + p)
≡⟨⟩ n + (suc (m + p))
≡⟨ +-suc' n (m + p)⟩
suc (n + (m + p))
≡⟨ cong suc (+-swap m n p)⟩
n + ((suc m) + p)
∎
I know I really don't understand how rewrite works. I learn from the following document that rewrite will expand into with:
https://agda.readthedocs.io/en/v2.6.2/language/with-abstraction.html#with-rewrite
But I don't find how rewrite containing | expands in the document. I guess the | in rewrite is also a kind of syntactic sugar:
+-swap (suc m) n p rewrite +-suc' n (m + p) | +-swap m n p = refl
will expand into:
+-swap (suc m) n p rewrite +-suc' n (m + p) rewrite +-swap m n p = refl
I tried to replace the second rewrite with with, no problem:
+-swap (suc m) n p rewrite +-suc' n (m + p)
with m + (n + p) | +-swap m n p
... | .(n + (m + p)) | refl = refl
But if I replaced the first rewrite with with, it gives an error:
+-swap (suc m) n p with n + (suc (m + p)) | +-suc' n (m + p)
... | .(suc (n + m + p)) | refl
rewrite +-swap m n p = refl
+-swap (suc m) n p with n + (suc (m + p)) | +-suc' n (m + p)
... | .(suc (n + m + p)) | refl
with m + (n + p) | +-swap m n p
... | .(n + (m + p)) | refl = refl
Error message:
n + m != n of type ℕ
when checking that the given dot pattern suc (n + m + p) matches
the inferred value suc (n + (m + p))
How do multiple rewrites expand into with? How can this proof be accomplished with an equation chain?

Termination checking failed to prove ∃-even′ : ∀ {n : ℕ} → ∃[ m ] ( 2 * m ≡ n) → even n

The PLFA exercise: what if we write the arithmetic more "naturally" in Quantifiers chapter (https://plfa.github.io/Quantifiers/) ?
∃-even′ : ∀ {n : ℕ} → ∃[ m ] ( 2 * m ≡ n) → even n
∃-odd′ : ∀ {n : ℕ} → ∃[ m ] (2 * m + 1 ≡ n) → odd n
I have make the type right. But have got Termination checking failed for the following functions:
dbl≡2* : ∀ n → n + n ≡ 2 * n
dbl≡2* n = cong (n +_) (sym (+-identityʳ n))
+-suc1 : ∀ (m : ℕ) → m + 1 ≡ suc m
+-suc1 m =
begin
m + 1
≡⟨⟩
m + (suc zero)
≡⟨ +-suc m zero ⟩
suc (m + zero)
≡⟨ cong suc (+-identityʳ m) ⟩
suc m
∎
help1 : ∀ m → 2 * m + 1 ≡ suc (m + m)
help1 m =
begin
2 * m + 1
≡⟨ sym ( cong (_+ 1) (dbl≡2* m) ) ⟩
m + m + 1 -- must use every rule
≡⟨ +-assoc m m 1 ⟩
m + (m + 1)
≡⟨ cong (m +_) (+-suc1 m) ⟩
m + suc m
≡⟨ +-suc m m ⟩
suc (m + m)
∎
∃-even′ ⟨ zero , refl ⟩ = even-zero
∃-even′ ⟨ suc m , refl ⟩ rewrite +-identityʳ m
| +-suc m m
= even-suc (∃-odd′ ⟨ (m) , help1 m ⟩)
∃-odd′ ⟨ m , refl ⟩ rewrite +-suc (2 * m) 0
| +-identityʳ m
| +-identityʳ (m + m)
| dbl≡2* m
= odd-suc (∃-even′ ⟨ m , refl ⟩)
For the normal version, the same mutually-recursive define can work fine.
∃-even : ∀ {n : ℕ} → ∃[ m ] ( m * 2 ≡ n) → even n
∃-odd : ∀ {n : ℕ} → ∃[ m ] (1 + m * 2 ≡ n) → odd n
∃-even ⟨ zero , refl ⟩ = even-zero
∃-even ⟨ suc x , refl ⟩ = even-suc (∃-odd ⟨ x , refl ⟩)
∃-odd ⟨ x , refl ⟩ = odd-suc (∃-even ⟨ x , refl ⟩)
∃-even′ ⟨ zero , refl ⟩ = even-zero
∃-even′ ⟨ suc m , refl ⟩ rewrite +-identityʳ m
| +-suc m m
= even-suc (∃-odd′ ⟨ m , help1 m ⟩)
∃-odd′ ⟨ m , refl ⟩ rewrite +-suc (2 * m) 0
| +-identityʳ m
| +-identityʳ (m + m)
| dbl≡2* m
= odd-suc (∃-even′ ⟨ m , refl ⟩)
Your recursive calls are:
∃-even′ ⟨ suc m , refl ⟩ -> ∃-odd′ ⟨ m , help1 m ⟩
∃-odd′ ⟨ m , refl ⟩ -> ∃-even′ ⟨ m , refl ⟩
In the first one, suc m -> m decreases, but refl -> help1 m (on its surface) increases. If you passed refl as the second argument to ∃-odd′, then the termination checker would accept it, since it means the second argument stays the same, while the first one strictly monotonically decreases over a complete chain of two calls.
So how can we change that first recursive call to ∃-odd′ ⟨ m , refl ⟩? By rewriting by sym (help1 m):
∃-even′ ( suc m , refl ) rewrite +-identityʳ m
| +-suc m m
| sym (help1 m)
= even-suc (∃-odd′ (m , refl))
This code is then accepted by the termination checker.

Why does `sym` need to be used in this case when using `rewriting`?

Given the Peano definition of natural numbers:
data ℕ : Set where
zero : ℕ
suc : ℕ → ℕ
_+_ : ℕ → ℕ → ℕ
zero + n = n
(suc m) + n = suc (m + n)
We can prove by different methods the property ∀ (m : ℕ) → zero + m ≡ m + zero.
For example:
comm-+₀ : ∀ (m : ℕ) → zero + m ≡ m + zero
comm-+₀ zero = refl
comm-+₀ (suc n) =
begin
zero + suc n
≡⟨⟩
zero + suc (zero + n)
≡⟨⟩
suc (zero + n)
≡⟨ cong suc (comm-+₀ n) ⟩
suc (n + zero)
≡⟨⟩
suc n + zero
∎
And more compactly:
comm-+₀ : ∀ (m : ℕ) → zero + m ≡ m + zero
comm-+₀ zero = refl
comm-+₀ (suc n) = cong suc (comm-+₀ n)
If we want, we can even use rewrite and forgo cong:
comm-+₀ : ∀ (m : ℕ) → zero + m ≡ m + zero
comm-+₀ zero = refl
comm-+₀ (suc n) rewrite comm-+₀ n = refl
But wait! That doesn't work. Agda will tell us that the expression is wrong because it can't prove the following:
suc (n + 0) ≡ suc (n + 0 + 0)
If we present Agda the symmetrical rewrite of the property, sym (comm-+₀ n), it will type check without errors.
So, my question is: why do we need sym in this case? The proof worked perfectly fine without it with the other strategies. Does rewrite work on both sides simultaneously and not just the left side?
In every cases, the goal when m is of the form suc n is:
suc n ≡ suc (n + 0)
To solve this goal by providing a correctly typed term, the right way is, as you noticed:
cong suc (comm-+₀ n)
However, when using rewrite with an equality a ≡ b you modify directly the goal by substituting all occurences of a by b In your case, using rewrite on the quantity comm-+₀ n whose type is n ≡ n + 0 leads to the replacing of every occurence of n by n + 0, thus transforming the goal from
suc n ≡ suc (n + 0)
to
suc (n + 0) ≡ suc (n + 0 + 0)
which is not what you want to do. Since rewriting replaces all occurences of the left side by the right side, reversing the equality using sym will instead replace the only occurence of n + 0 by n thus transforming the goal from
suc n ≡ suc (n + 0)
to
suc n ≡ suc n
which is your expected behaviour and let you conclude using refl direcly. This explains why you need to use sym.
To summarize :
rewrite interacts directly with the type of the goal.
rewrite rewrites from left to right.
rewrite rewrites all occurences it finds in the type of the goal.
More on rewrite can be found here:
https://agda.readthedocs.io/en/v2.6.0.1/language/with-abstraction.html#with-rewrite

Dependent type involving nat addition

I need to define two versions of an operation with a slightly different definition. It is a series of compositions with Nat indices involved.
open import Data.Nat
data Hom : ℕ → ℕ → Set where
id : (m : ℕ) → Hom m m
_∘_ : ∀ {m n k} → Hom n k → Hom m n → Hom m k
p : (n : ℕ) → Hom (suc n) n
p1 : (m n : ℕ) → Hom (m + n) n
p1 zero n = id n
p1 (suc m) n = p1 m n ∘ p (m + n)
p2 : (m n : ℕ) → Hom (m + n) n
p2 zero n = id n
p2 (suc m) n = {!!} -- p n ∘ p2 m (1 + n)
-- Goal: Hom (suc (m + n)) n
-- Have: Hom (m + suc n) n
I would like to define both p1 and p2 and be able to use them interchangeably. Is this doable?
You can define p2 by direct recursion (no subst or rewriting) over _+_ using the trick described here. Looks like this:
record Homable (H : ℕ → ℕ → Set) : Set where
field
id-able : (m : ℕ) → H m m
_∘-able_ : ∀ {m n k} → H n k → H m n → H m k
p-able : (n : ℕ) → H (suc n) n
suc-homable : ∀ {H} → Homable H → Homable (λ m n -> H (suc m) (suc n))
suc-homable homable = record
{ id-able = λ m → id-able (suc m)
; _∘-able_ = _∘-able_
; p-able = λ m → p-able (suc m)
} where open Homable homable
p2-go : ∀ {H} → Homable H → (m : ℕ) → H m 0
p2-go homable zero = id-able 0 where
open Homable homable
p2-go homable (suc m) = p-able 0 ∘-able p2-go (suc-homable homable) m where
open Homable homable
plus-homable-hom : ∀ k → Homable (λ m n → Hom (m + k) (n + k))
plus-homable-hom k = record
{ id-able = λ n → id (n + k)
; _∘-able_ = _∘_
; p-able = λ n → p (n + k)
}
p2 : (m n : ℕ) → Hom (m + n) n
p2 m n = p2-go (plus-homable-hom n) m
The cost is that you need to maintain those Homable records which is somewhat tedious, but to my experience proving things about functions defined this way is simpler than about functions defined in terms of subst or over _+′_ (unless you never want to coerce _+′_ to _+_, of course).
Well, the value you provide has a type that is equal to the type of the hole, but Agda does not see this fact. More formally, the two types are propositionally equal but not judgementally equal. The problem is caused by the index m + suc n, which is propositionally but not judgementally equal to suc m + n because of how addition is defined. One way to solve your problem is to manually explain to Agda that the two types are equal:
open import Data.Nat
open import Data.Nat.Properties
open import Relation.Binary.PropositionalEquality
data Hom : ℕ → ℕ → Set where
id : (m : ℕ) → Hom m m
_∘_ : ∀ {m n k} → Hom n k → Hom m n → Hom m k
p : (n : ℕ) → Hom (suc n) n
p1 : (m n : ℕ) → Hom (m + n) n
p1 zero n = id n
p1 (suc m) n = p1 m n ∘ p (m + n)
p2 : (m n : ℕ) → Hom (m + n) n
p2 zero n = id n
p2 (suc m) n = subst (λ k → Hom k n) (+-suc m n) (p n ∘ p2 m (suc n))
However, this approach is not without downsides, as p2 (suc m) n is now not judgementally equal to your intended definition but to the expression above involving subst.
The problem seems essentially linked to what you're trying to do: IIUC, p1 and p2 are actually provably equal but defined using a different recursion structure. That's fine, but then the indices of your result type should follow the same recursion structure, i.e. you should define p2 using a different version of + that recurses in the appropriate way for p2:
_+′_ : ℕ → ℕ → ℕ
zero +′ n = n
suc m +′ n = m +′ suc n
p2′ : (m n : ℕ) → Hom (m +′ n) n
p2′ zero n = id n
p2′ (suc m) n = p n ∘ p2′ m (suc n)
However, this has another downside that the type of p1 and p2′ are no longer judgementally equal (but still propositionally equal though).
Another thing you can try is to use Agda's rewrite rules to give _+_ satisfy additional judgemental equalities, but this is dangerous as it may break some of Agda's desirable qualities as a logic. In this case, I suspect it's fine, but I'd have to check.
In summary, there are a number of things you can try but none is without downsides. Which is your best option depends on what you're trying to use this for.

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