Why does this 'with' block spoil the totality of this function? - idris

I'm trying to compute parity together with the floor of the half, over natural numbers:
data IsEven : Nat -> Nat -> Type where
Times2 : (n : Nat) -> IsEven (n + n) n
data IsOdd : Nat -> Nat -> Type where
Times2Plus1 : (n : Nat) -> IsOdd (S (n + n)) n
parity : (n : Nat) -> Either (Exists (IsEven n)) (Exists (IsOdd n))
I tried going with the obvious implementation of parity:
parity Z = Left $ Evidence _ $ Times2 0
parity (S Z) = Right $ Evidence _ $ Times2Plus1 0
parity (S (S n)) with (parity n)
parity (S (S (k + k))) | Left (Evidence _ (Times2 k)) =
Left $ rewrite plusSuccRightSucc k k in Evidence _ $ Times2 (S k)
parity (S (S (S ((k + k))))) | Right (Evidence _ (Times2Plus1 k)) =
Right $ rewrite plusSuccRightSucc k k in Evidence _ $ Times2Plus1 (S k)
This typechecks and works as expected. However, if I try to mark parity as total, Idris starts complaining:
parity is possibly not total due to: with block in parity
The only with block I see in parity is the one with the recursive call from parity (S (S n)) to parity n, but clearly that is well-founded, since n is structurally smaller than S (S n).
How do I convince Idris that parity is total?

It looks like a bug to me, because the following solution based on case passes the totality checker:
total
parity : (n : Nat) -> Either (Exists (IsEven n)) (Exists (IsOdd n))
parity Z = Left $ Evidence _ $ Times2 0
parity (S Z) = Right $ Evidence _ $ Times2Plus1 0
parity (S (S k)) =
case (parity k) of
Left (Evidence k (Times2 k)) =>
Left $ rewrite plusSuccRightSucc k k in Evidence _ $ Times2 (S k)
Right (Evidence k (Times2Plus1 k)) =>
Right $ rewrite plusSuccRightSucc k k in Evidence _ $ Times2Plus1 (S k)

Related

Is this a bug in Idris's type-checker (or unification)?

The following snippet compiles fine in Idris 2,
divMod2 : Nat -> (Nat, Bool)
divMod2 Z = (Z, False)
divMod2 (S Z) = (Z, True)
divMod2 (S (S n)) = case divMod2 n of (k, r) => (S k, r)
lm : {q:_} -> divMod2 (q+q) = (q, False)
lm {q=0} = Refl
lm {q=S q} = rewrite sym $ plusSuccRightSucc q q in
rewrite lm {q} in Refl
but it caused Idris 1.3.4 to fail:
|
14 | rewrite lm {q} in Refl
| ~~~~~~
When checking right hand side of lm with expected type
divMod2 (S q + S q) = (S q, False)
When checking an application of function rewrite__impl:
Type mismatch between
(q1, False) = (q1, False) (Type of Refl)
and
case divMod2 (plus q q) of (k, r) => (S k, r) = (q1, False) (Expected type)
Specifically:
Type mismatch between
(q1, False)
and
case divMod2 (plus q q) of
(k, r) => (S k, r)
I know Idris 1 has become obsolete, but am still curious about the underlying reason behind this observation. Could someone help to explain this?
I've done a bit more dig here and found surprisingly that Idris1 failed to compile the following even simpler code snippet, which compiles correctly by Idris2:
f : Nat -> (Nat, Bool)
f Z = (Z, False)
f (S n) = case f n of (k, r) => (S k, r)
h2 : {q:_} -> f (S (q+1)) = case f (q+1) of (k, r) => (S k, r)
h2 = Refl
It spit error message:
|
19 | h2 = Refl
| ~~~~
When checking right hand side of h2 with expected type
f (S (q + 1)) = case f (q + 1) of (k, r) => (S k, r)
Type mismatch between
case f (plus q 1) of (k, r) => (S k, r) = case f (plus q 1) of (k, r) => (S k, r) (Type of Refl)
and
case f (plus q 1) of (k, r) => (S k, r) = case f (plus q 1) of (k, r) => (S k, r) (Expected type)
Specifically:
Type mismatch between
case f (plus q 1) of
(k, r) => (S k, r)
and
case f (plus q 1) of
(k, r) => (S k, r)
Aren't the two mismatching terms identical? Why couldn't be constructed with Refl?

Why can (and must) the variable x appear twice on the left side of the function definition? And what's the meaning?

I'm a green hand with Idris,and get confused with this definition, as I don't understand how it works.
The definitionare as follows.
sameS : (k : Nat)->(j : Nat)->(k = j)->((S k) = (S j))
sameS x x Refl=Refl
Let us start by breaking down the type signature:
sameS : (k : Nat) -> (j : Nat) -> (k = j) -> ((S k) = (S j))
sameS is a function.
sameS take the following arguments:
(k : Nat) a parameter k of type Nat
(j : Nat) a parameter j of type Nat
(k = j) A proof that k and j are equal
sameS returns:
((S k) = (S j)) proof that S k and S j are equal.
Now let us breakdown the definition:
sameS x x Refl = Refl
The type of Refl is a = a.
x is both the first and second argument because both are identical. We know this because the 3rd argument is Refl.
Refl is returned because S x = S x.

Proving equality of types depending on (further) proofs

Suppose we'd like to have a "proper" minus on Nats, requiring m <= n for n `minus` m to make sense:
%hide minus
minus : (n, m : Nat) -> { auto prf : m `LTE` n } -> Nat
minus { prf = LTEZero } n Z = n
minus { prf = LTESucc prevPrf } (S n) (S m) = minus n m
Now let's try to prove the following lemma, stating that (n + (1 + m)) - k = ((1 + n) + m) - k, assuming both sides are valid:
minusPlusTossS : (n, m, k : Nat) ->
{ auto prf1 : k `LTE` n + S m } ->
{ auto prf2 : k `LTE` S n + m } ->
minus (n + S m) k = minus (S n + m) k
The goal suggests the following sublemma might help:
plusTossS : (n, m : Nat) -> n + S m = S n + m
plusTossS Z m = Refl
plusTossS (S n) m = cong $ plusTossS n m
so we try to use it:
minusPlusTossS n m k =
let tossPrf = plusTossS n m
in rewrite tossPrf in ?rhs
And here we fail:
When checking right hand side of minusPlusTossS with expected type
minus (n + S m) k = minus (S n + m) k
When checking argument prf to function Main.minus:
Type mismatch between
LTE k (S n + m) (Type of prf2)
and
LTE k replaced (Expected type)
Specifically:
Type mismatch between
S (plus n m)
and
replaced
If I understand this error correctly, it just means that it tries to rewrite the RHS of the target equality (which is minus { prf = prf2 } (S n + m) k) to minus { prf = prf2 } (n + S m) k and fails. Rightfully, of course, since prf is a proof for a different inequality! And while replace could be used to produce a proof of (S n + m) k (or prf1 would do as well), it does not look like it's possible to simultaneously rewrite and change the proof object so that it matches the rewrite.
How do I work around this? Or, more generally, how do I prove this lemma?
Ok, I guess I solved it. Bottom line: if you don't know what to do, do a lemma!
So we have a proof of two minuends n1, n2 being equal, and we need to produce a proof of n1 `minus` m = n2 `minus` m. Let's write this down!
minusReflLeft : { n1, n2, m : Nat } -> (prf : n1 = n2) -> (prf_n1 : m `LTE` n1) -> (prf_n2 : m `LTE` n2) -> n1 `minus` m = n2 `minus` m
minusReflLeft Refl LTEZero LTEZero = Refl
minusReflLeft Refl (LTESucc prev1) (LTESucc prev2) = minusReflLeft Refl prev1 prev2
I don't even need plusTossS anymore, which can be replaced by a more directly applicable lemma:
plusRightS : (n, m : Nat) -> n + S m = S (n + m)
plusRightS Z m = Refl
plusRightS (S n) m = cong $ plusRightS n m
After that, the original one becomes trivial:
minusPlusTossS : (n, m, k : Nat) ->
{ auto prf1 : k `LTE` n + S m } ->
{ auto prf2 : k `LTE` S n + m } ->
minus (n + S m) k = minus (S n + m) k
minusPlusTossS {prf1} {prf2} n m k = minusReflLeft (plusRightS n m) prf1 prf2

An Idris proof about `mod`

I was trying to write a proof in Idris regarding the following subtraction-based mod operator:
mod : (x, y : Nat) -> Not (y = Z) -> Nat
mod x Z p = void (p Refl)
mod x (S k) _ = if lt x (S k) then x else helper x (minus x (S k)) (S k)
where total
helper : Nat -> Nat -> Nat -> Nat
helper Z x y = x
helper (S k) x y = if lt x y then x else helper k (minus x y) y
The theorem I wanted to prove is that the remainder as produced by "mod" above is always smaller than the divider. Namely,
mod_prop : (x, y : Nat) -> (p : Not (y=0))-> LT (mod x y p) y
I constructed a proof but was stuck by a final hole. My full code is pasted below
mod : (x, y : Nat) -> Not (y = Z) -> Nat
mod x Z p = void (p Refl)
mod x (S k) _ = if lt x (S k) then x else helper x (minus x (S k)) (S k)
where total
helper : Nat -> Nat -> Nat -> Nat
helper Z x y = x
helper (S k) x y = if lt x y then x else helper k (minus x y) y
lteZK : LTE Z k
lteZK {k = Z} = LTEZero
lteZK {k = (S k)} = let ih = lteZK {k=k} in
lteSuccRight {n=Z} {m=k} ih
lte2LTE_True : True = lte a b -> LTE a b
lte2LTE_True {a = Z} prf = lteZK
lte2LTE_True {a = (S _)} {b = Z} Refl impossible
lte2LTE_True {a = (S k)} {b = (S j)} prf =
let ih = lte2LTE_True {a=k} {b=j} prf in LTESucc ih
lte2LTE_False : False = lte a b -> GT a b
lte2LTE_False {a = Z} Refl impossible
lte2LTE_False {a = (S k)} {b = Z} prf = LTESucc lteZK
lte2LTE_False {a = (S k)} {b = (S j)} prf =
let ih = lte2LTE_False {a=k} {b=j} prf in (LTESucc ih)
total
mod_prop : (x, y : Nat) -> (p : Not (y=0))-> LT (mod x y p) y
mod_prop x Z p = void (p Refl)
mod_prop x (S k) p with (lte x k) proof lxk
mod_prop x (S k) p | True = LTESucc (lte2LTE_True lxk)
mod_prop Z (S k) p | False = LTESucc lteZK
mod_prop (S x) (S k) p | False with (lte (minus x k) k) proof lxk'
mod_prop (S x) (S k) p | False | True = LTESucc (lte2LTE_True lxk')
mod_prop (S x) (S Z) p | False | False = LTESucc ?hole
Once I run the type checker, the hole is described as follows:
- + Main.hole [P]
`-- x : Nat
p : (1 = 0) -> Void
lxk : False = lte (S x) 0
lxk' : False = lte (minus x 0) 0
--------------------------------------------------------------------------
Main.hole : LTE (Main.mod, helper (S x) 0 p x (minus (minus x 0) 1) 1) 0
I am not familiar with the syntax of Main.mod, helper (S x) 0 p x (minus (minus x 0) 1) 1 given in the idris-holes window. I guess (S x) 0 p are the three parameters of "mod" while (minus (minus x 0) 1) 1 are the three parameters of the local "helper" function of "mod"?
It seems that it's time to leverage an induction hypothesis. But how can I finish up the proof using induction?
(Main.mod, helper (S x) 0 p x (minus (minus x 0) 1) 1)
can be read as
Main.mod, helper - a qualified name for helper function, which is defined in the where clause of the mod function (Main is a module name);
Arguments of mod which are also passed to helper - (S x), 0 and p (see docs):
Any names which are visible in the outer scope are also visible in the
where clause (unless they have been redefined, such as xs here). A
name which appears only in the type will be in scope in the where
clause if it is a parameter to one of the types, i.e. it is fixed
across the entire structure.
Arguments of helper itself - x, (minus (minus x 0) 1) and 1.
Also below is another implementation of mod which uses Fin n type for the remainder in division by n. It turns out to be easier to reason about, since any value of Fin n is always less than n:
import Data.Fin
%default total
mod' : (x, y : Nat) -> {auto ok: GT y Z} -> Fin y
mod' Z (S _) = FZ
mod' (S x) (S y) with (strengthen $ mod' x (S y))
| Left _ = FZ
| Right rem = FS rem
mod : (x, y : Nat) -> {auto ok: GT y Z} -> Nat
mod x y = finToNat $ mod' x y
finLessThanBound : (f : Fin n) -> LT (finToNat f) n
finLessThanBound FZ = LTESucc LTEZero
finLessThanBound (FS f) = LTESucc (finLessThanBound f)
mod_prop : (x, y : Nat) -> {auto ok: GT y Z} -> LT (mod x y) y
mod_prop x y = finLessThanBound (mod' x y)
Here for convenience I used auto implicits for proofs that y > 0.

Why doesn't equality involving “minus” typecheck in Idris?

Why won't the following typecheck:
minusReduces : (n : Nat) -> n `minus` Z = n
minusReduces n = Refl
Yet this will typecheck fine:
plusReduces : (n : Nat) -> Z `plus` n = n
plusReduces n = Refl
minus n doesn't reduce because minus is defined with pattern matching on the first argument:
total minus : Nat -> Nat -> Nat
minus Z right = Z
minus left Z = left
minus (S left) (S right) = minus left right
So you'll need to split your Z and S n cases as well:
minusReduces : (n : Nat) -> n `minus` Z = n
minusReduces Z = Refl
minusReduces (S k) = Refl