When I use Function
to define a non-structurally recursive function in Coq, the resulting object behaves strangely when a specific computation is asked. Indeed, instead of giving directly the result, the Eval compute in ...
directive return a rather long (typically 170 000 lines) expression. It seems that Coq cannot evaluate everything, and therefore returns a simplified (but long) expression instead of just a value.
The problem seems to come from the way I prove the obligations generated by Function
. First, I thought the problem came from the opaque terms I used, and I converted all the lemmas to transparent constants. By the way, is there a way to list the opaque terms appearing in a term ? Or any other way to turn opaque lemmas into transparent ones ?
I then remarked that the problem came more precisely from the proof that the order used is well-founded. But I got strange results.
For example, I define log2
on the natural numbers by repeatedly applying div2
. Here is the definition:
Function log2 n {wf lt n} :=
match n with
| 0 => 0
| 1 => 0
| n => S (log2 (Nat.div2 n))
end.
I get two proof obligations. The first one checks that n
respects the relation lt
in the recursive calls and can be proved easily.
forall n n0 n1 : nat, n0 = S n1 -> n = S (S n1) -> Nat.div2 (S (S n1)) < S (S n1)
intros. apply Nat.lt_div2. apply le_n_S. apply le_0_n.
The second one checks that lt
is a well-founded order. This is already proved in the standard library. The corresponding lemma is Coq.Arith.Wf_nat.lt_wf
.
If I use this proof, the resulting function behaves normally. Eval compute in log24 10.
returns 3
.
But if I want to do the proof myself, I do not always get this behaviour. First, if I end the proof with Qed
instead of Defined
, the result of the computation (even on small numbers) is a complex expression and not a single number. So I use Defined
and try to use only transparent lemmas.
Lemma lt_wf2 : well_founded lt.
Proof.
unfold well_founded. intros n.
apply (lemma1 n). clear n.
intros. constructor. apply H.
Defined.
Here, lemma1 is a proof of the well-founded induction on the natural numbers. Here again, I can use already existing lemmas, such as lt_wf_ind
, lt_wf_rec
, lt_wf_rec1
located in Coq.Arith.Wf_nat
, or even well_founded_ind lt_wf
. The first one does not work, it seems this is because it is opaque. The three others work.
I tried to prove it directly using the standard induction on the natural numbers, nat_ind
. This gives:
Lemma lemma1 : forall n (P:nat -> Prop),
(forall n, (forall p, p < n -> P p) -> P n) -> P n.
Proof.
intros n P H. pose proof (nat_ind (fun n => forall p, p < n -> P p)).
simpl in H0. apply H0 with (n:=S n).
- intros. inversion H1.
- intros. inversion H2.
+ apply H. exact H1.
+ apply H1. assumption.
- apply le_n.
Defined.
With this proof (and some variants of it), log2
has the same strange behaviour. And this proof seems to use only transparent objects, so maybe the problem is not there.
How can I define a Function
that returns understandable results on specific values ?
The reduction behavior of functions defined by well-founded recursion in Coq is generally not very good, even when you declare your proofs to be transparent. The reason for this is that arguments of well-foundedness usually need to be done with complicated proof terms. Since these proofs terms end up appearing in well-founded recursive definitions, "simplifying" your function will make all of those proof terms appear, as you noticed.
It is easier to rely on custom tactics and lemmas to reduce functions defined this way. First, I would recommend favoring
Program Fixpoint
overFunction
, because the latter is much older and (I think) less well maintained. Thus, you would end up with a definition like this:Now, you just need to use the
program_simpl
tactic to simplify calls tolog2
. Here's an example:I've managed to pin-point the place that causes troubles: it's
inversion H2.
inlemma1
. It turns out we don't need that case-analysis andintuition
can finish the proof (it doesn't pattern-match onH2
):If we use
lemma1
with this proof, the computation oflog2 10
results in3
.By the way, here is my version of
lt_wf2
(it lets us compute as well):I believe the Using Coq's evaluation mechanisms in anger blog post by Xavier Leroy explains this kind of behavior.
In our case we eliminate the proof of inequality (
inversion H2.
) in the proof oflemma1
and theFunction
mechanism makes our computations depend on a proof term. Hence, the evaluator can't proceed when n > 1.And the reason
inversion H1.
in the body of the lemma doesn't influence computations is that forn = 0
andn = 1
,log2 n
is defined within thematch
expression as base cases.To illustrate this point, let me show an example when we can prevent evaluation of
log2 n
on any valuesn
andn + 1
of our choice, wheren > 1
and nowhere else!If you use this modified lemma in the definition of
log2
you'll see that it computes everywhere exceptn = 4
andn = 5
. Well, almost everywhere -- computations with largenat
s can result in stack overflow or segmentation fault, as Coq warns us:And to make
log2
work forn = 4
andn = 5
even for the above "flawed" proof, we could amendlog2
like thisadding the necessary proofs at the end.
The "well-founded" proof must be transparent and can't rely on pattern-matching on proof objects because the
Function
mechanism actually uses thelt_wf
lemma to compute the decreasing termination guard. If we look at the term produced byEval
(in a case where evaluation fails to produce anat
), we'll see something along these lines:It's easy to see that
x0 : Prop
, so it gets erased when extracting the functional programlog2
into, say OCaml, but Coq's internal evaluation mechanism have to use it to ensure termination.