Compute with a recursive function defined by well-defined induction

半腔热情 提交于 2019-12-01 15:47:00

I've managed to pin-point the place that causes troubles: it's inversion H2. in lemma1. It turns out we don't need that case-analysis and intuition can finish the proof (it doesn't pattern-match on H2):

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. intuition.
  - apply le_n.
Defined.

If we use lemma1 with this proof, the computation of log2 10 results in 3.

By the way, here is my version of lt_wf2 (it lets us compute as well):

Lemma lt_wf2 : well_founded lt.
Proof.
  unfold well_founded; intros n.
  induction n; constructor; intros k Hk.
  - inversion Hk.
  - constructor; intros m Hm.
    apply IHn; omega.
 (* OR: apply IHn, Nat.lt_le_trans with (m := k); auto with arith. *)
Defined.

I believe the Using Coq's evaluation mechanisms in anger blog post by Xavier Leroy explains this kind of behavior.

it eliminates the proof of equality between the heads before recursing over the tails and finally deciding whether to produce a left or a right. This makes the left/right data part of the final result dependent on a proof term, which in general does not reduce!

In our case we eliminate the proof of inequality (inversion H2.) in the proof of lemma1 and the Function 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 for n = 0 and n = 1, log2 n is defined within the match expression as base cases.

To illustrate this point, let me show an example when we can prevent evaluation of log2 n on any values n and n + 1 of our choice, where n > 1 and nowhere else!

Lemma lt_wf2' : well_founded lt.
Proof.
  unfold well_founded; intros n.
  induction n; constructor; intros k Hk.
  - inversion Hk.          (* n = 0 *)
  - destruct n. intuition. (* n = 1 *)
    destruct n. intuition. (* n = 2 *)
    destruct n. intuition. (* n = 3 *)
    destruct n. inversion Hk; intuition. (* n = 4 and n = 5 - won't evaluate *)
    (* n > 5 *)
    constructor; intros m Hm; apply IHn; omega.
Defined.

If you use this modified lemma in the definition of log2 you'll see that it computes everywhere except n = 4 and n = 5. Well, almost everywhere -- computations with large nats can result in stack overflow or segmentation fault, as Coq warns us:

Warning: Stack overflow or segmentation fault happens when working with large numbers in nat (observed threshold may vary from 5000 to 70000 depending on your system limits and on the command executed).

And to make log2 work for n = 4 and n = 5 even for the above "flawed" proof, we could amend log2 like this

Function log2 n {wf lt n} :=
  match n with
  | 0 => 0
  | 1 => 0
  | 4 => 2
  | 5 => 2
  | n => S (log2 (Nat.div2 n))
  end.

adding 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 the lt_wf lemma to compute the decreasing termination guard. If we look at the term produced by Eval (in a case where evaluation fails to produce a nat), we'll see something along these lines:
fix Ffix (x : nat) (x0 : Acc (fun x0 x1 : nat => S x0 <= x1) x) {struct x0}

It's easy to see that x0 : Prop, so it gets erased when extracting the functional program log2 into, say OCaml, but Coq's internal evaluation mechanism have to use it to ensure termination.

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 over Function, because the latter is much older and (I think) less well maintained. Thus, you would end up with a definition like this:

Require Import Coq.Numbers.Natural.Peano.NPeano.
Require Import Coq.Program.Wf.
Require Import Coq.Program.Tactics.

Program Fixpoint log2 n {wf lt n} :=
  match n with
  | 0 => 0
  | 1 => 0
  | n => S (log2 (Nat.div2 n))
  end.

Next Obligation.
admit.
Qed.

Now, you just need to use the program_simpl tactic to simplify calls to log2. Here's an example:

Lemma foo : log2 4 = 2.
Proof.
  program_simpl.
Qed.
标签
易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!