I'm trying to define binary numbers in agda but agda cant see that ⟦_⇑⟧ is terminating. I really dont want to have to break out accessibility relations. How can I show agda that n gets smaller?
module Binary where
open import Data.Nat using (ℕ; _+_; +-suc)
open ℕ
2* : ℕ → ℕ
2* n = n + n
data Parity : ℕ → Set where
𝕖 : ∀ k → Parity (2* k)
𝕠 : ∀ k → Parity (1 + 2* k)
parity : ∀ n → Parity n
parity zero = 𝕖 0
parity (suc n) with parity n
parity (suc .(k + k)) | 𝕖 k = 𝕠 k
parity (suc .(suc (k + k))) | 𝕠 k rewrite sym (+-suc k k) = 𝕖 (suc k)
data 𝔹 : Set where
𝕫 : 𝔹
𝕖 : 𝔹 → 𝔹
𝕠 : 𝔹 → 𝔹
⟦_⇓⟧ : 𝔹 → ℕ
⟦ 𝕫 ⇓⟧ = 0
⟦ 𝕖 b ⇓⟧ = 2* ⟦ b ⇓⟧
⟦ 𝕠 b ⇓⟧ = 1 + 2* ⟦ b ⇓⟧
⟦_⇑⟧ : ℕ → 𝔹
⟦ zero ⇑⟧ = 𝕫
⟦ suc n ⇑⟧ with parity n
⟦ suc .(k + k) ⇑⟧ | 𝕖 k = 𝕠 ⟦ k ⇑⟧
⟦ suc .(suc (k + k)) ⇑⟧ | 𝕠 k = 𝕖 ⟦ k ⇑⟧
Error:
Termination checking failed for the following functions:
⟦_⇑⟧
Problematic calls:
⟦ k ⇑⟧
Termination checking fails for good reasons, since the following function is not structurally recursive over its input:
⟦_⇑⟧ : ℕ → 𝔹
⟦ zero ⇑⟧ = 𝕫
⟦ suc n ⇑⟧ with parity n
⟦ suc .(k + k) ⇑⟧ | 𝕖 k = 𝕠 ⟦ k ⇑⟧
⟦ suc .(suc (k + k)) ⇑⟧ | 𝕠 k = 𝕖 ⟦ k ⇑⟧
Agda even tells you what is the problematic call : ⟦ k ⇑⟧. (in this case, there are two of these ill-formed calls).
While I agree it might look that the function is called on a structurally smaller argument it's not the case. To understand why, you have to see that k is used as input for a function call, _+_, and that only the result of this function is structurally smaller than n, not k itself. And agda has no way of knowing the following property about _+_ :
∀ {n} → n < suc (n + n)
Should you provide a proof of such lemma, you could use the fact that _<_ is well founded to make your function structurally recursive over Acc, but it seems you are reluctant to do so.
An easy way to understand why Agda cannot know that this terminates is the following: imagine you replace suc .(k + k) by suc .(a ∸ b) and recursively call your function over a. From agda's point of view, both are the same cases, and in this case, it does not usually terminate unless b happens to be 0.
Here is the module corrected in terms of termination :
module Binary where
open import Data.Nat
open import Data.Nat.Properties
open import Relation.Binary.PropositionalEquality
open import Data.Product
open import Induction.Nat
open import Induction.WellFounded
2* : ℕ → ℕ
2* n = n + n
data Parity : ℕ → Set where
𝕖 : ∀ k → Parity (2* k)
𝕠 : ∀ k → Parity (1 + 2* k)
parity : ∀ n → Parity n
parity zero = 𝕖 0
parity (suc n) with parity n
parity (suc .(k + k)) | 𝕖 k = 𝕠 k
parity (suc .(suc (k + k))) | 𝕠 k rewrite sym (+-suc k k) = 𝕖 (suc k)
data 𝔹 : Set where
𝕫 : 𝔹
𝕖 : 𝔹 → 𝔹
𝕠 : 𝔹 → 𝔹
⟦_⇓⟧ : 𝔹 → ℕ
⟦ 𝕫 ⇓⟧ = 0
⟦ 𝕖 b ⇓⟧ = 2* ⟦ b ⇓⟧
⟦ 𝕠 b ⇓⟧ = 1 + 2* ⟦ b ⇓⟧
decr : ∀ {n} → n < suc (n + n)
decr {n} = s≤s (m≤m+n n n)
helper : (n : ℕ) → Acc _<_ n → 𝔹
helper zero a = 𝕫
helper (suc n) a with parity n
helper (suc .(k + k)) (acc rs) | 𝕖 k = 𝕠 (helper k (rs _ decr))
helper (suc .(suc (k + k))) (acc rs) | 𝕠 k = 𝕖 (helper k (rs _ (s≤s (<⇒≤ decr))))
⟦_⇑⟧ : ℕ → 𝔹
⟦ n ⇑⟧ = helper n (<-wellFounded n)
I was also skeptical about the correctness of your definitions and it turns out I was right, for instance, considering the following definition:
test : ℕ
test = ⟦ ⟦ 124 ⇑⟧ ⇓⟧
Agda returns 2 when evaluating test.
Considering the definition:
test₁ : ℕ
test₁ = ⟦ ⟦ 16 ⇑⟧ ⇓⟧
Agda returns 14 when evaluating test₁
In order to correct your definitions, you can take inspiration from what is done in the standard library, either in the module Data.Bin (deprecated) or in the module Data.Nat.Binary depending on which version of the stdlib you have.
Related
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
From what I understand these two definitions a̶r̶e̶ ̶e̶q̶u̶i̶v̶a̶l̶e̶n̶t̶ represent the same behaviour:
data _<_ : ℕ → ℕ → Set where
lt-zero : {n : ℕ} → zero < suc n
lt-suc : {m n : ℕ} → m < n → (suc m) < (suc n)
lt : ℕ → ℕ → Bool
lt _ zero = false
lt zero (suc n) = true
lt (suc m) (suc n) = m < n
Except that _<_ can be more easily used to prove things about itself, and lt, from what I have found, is easier to use for programming other behaviours. For example I can see how I can easily define a min function using lt:
min : ℕ → ℕ → ℕ
min x y where lt x y
... | true = x
... | false = y
Is there a way for me to define min and other similar functions using _<_? From what I have found there is no way to pattern match x < y if y is less than x. Is there a different method to use _<_ in these cases?
EDIT: Would adding a not-true case to _<_ be a sensible idea?
̶E̶D̶I̶T̶ ̶2̶:̶ ̶O̶r̶ ̶i̶s̶ ̶t̶h̶i̶s̶ ̶t̶h̶e̶ ̶'̶c̶o̶r̶r̶e̶c̶t̶'̶ ̶w̶a̶y̶ ̶t̶o̶ ̶d̶o̶ ̶i̶t̶?̶
EDIT 3: I have changed the below definition of min, in order to be able to omit proofs for x <? y when using min.
data _<?_ (n m : ℕ) : Set where
isLT : n < m → n <? m
notLT : m ≤ n → n <? m
min : (x y : ℕ) → {p : x <? y} → ℕ
min x y (isLT _) = x
min x y (notLT _) = y
This did not work as expected. If I true to evaluate min 1 2, using C-c C-n, it returns min 1 2. I can get it to return the min, if I give it a proof, however if I evaluate min 2 1 (isLT _) it returns 2, instead of giving me an error message. Is there a way for me to define min using _<_, so that Agda could evaluate min 1 2?
There is a way to eliminate impossible patterns in Agda. Using your definition of _<_
you can try to prove transitivity to illustrate this.
le-trans : ∀ {m n k} → m < n → n < k → m < k
le-trans {k = k} lt-zero b = {!!}
le-trans (lt-suc a) (lt-suc b) = lt-suc (le-trans a b)
Goal: 0 < k
b : suc .n < k
We pattern match on the first argument, the step case is a simple application of the hypothesis. In the base case, we have to pattern match on k since the base constructor of your data type says zero < suc n, but we don't know anything about k yet. Upon pattern matching on k we see two cases
le-trans : ∀ {m n k} → m < n → n < k → m < k
le-trans {k = zero} lt-zero b = {!!}
le-trans {k = suc k} lt-zero b = {!!}
le-trans (lt-suc a) (lt-suc b) = lt-suc (le-trans a b)
In the first one, we see something impossible, i.e, Goal: 0 < zero and we have an element of type b : suc .n < zero, which cannot occur. Then, you can pattern match on b and Agda will see that you cannot construct such thing and will eliminate this case le-trans {k = zero} lt-zero (). Whereas in the other case, you can prove it with the base constructor.
le-trans : ∀ {m n k} → m < n → n < k → m < k
le-trans {k = zero} lt-zero ()
le-trans {k = suc k} lt-zero b = lt-zero
le-trans (lt-suc a) (lt-suc b) = lt-suc (le-trans a b)
Therefore, defining a not-true case in datatypes is not appropriate. You define how elements are constructed. Your last definition of _<?_ makes sense and can be, in fact, used.
Edit on min
Once you have the inductive relation for _<?_, you can work as follows. Define a function that gives you m <? n and then for the min function, do a with abstraction calling that function.
data _≥_ : ℕ → ℕ → Set where
get-z : ∀ {n} → n ≥ zero
get-s : ∀ {m n} → m ≥ n → (suc m) ≥ (suc n)
data _<?_ (n m : ℕ) : Set where
y-< : n < m → n <? m
n-< : n ≥ m → n <? m
f<? : (m n : ℕ) → m <? n
f<? zero zero = n-< get-z
f<? zero (suc n) = y-< lt-zero
f<? (suc m) zero = n-< get-z
f<? (suc m) (suc n) with f<? m n
f<? (suc m) (suc n) | y-< x = y-< (lt-suc x)
f<? (suc m) (suc n) | n-< x = n-< (get-s x)
min : ℕ → ℕ → ℕ
min x y with f<? x y
min x y | y-< _ = x
min x y | n-< _ = y
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
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.
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