How can I automate counting within proofs in Coq? - coq

I have a function count that counts how many times a given predicate is provable when applied to elements of a list. It is defined as follows:
Parameter T : Type.
Parameter dec: forall (p: T -> Prop) (w: T), {p w} + {~ (p w)}.
Fixpoint count (p: T -> Prop) (l: list T) := match l with
| nil => 0
| (cons head tail) => if (dec p head) then (1 + (count p tail)) else (count p tail)
end.
I then use this function to state lemmas like the following:
Parameter a b c: T.
Parameter q: T -> Prop.
Axiom Aa: (q a).
Axiom Ab: (q b).
Axiom Ac: ~ (q c).
Lemma example: (count q (cons a (cons b (cons c nil)))) = 2.
My proofs of such lemmas tend to be quite tedious:
Lemma example: (count q (cons a (cons b (cons c nil)))) = 2.
Proof.
unfold count.
assert (q a); [apply Aa| auto].
assert (q b); [apply Ab| auto].
assert (~ (q c)); [apply Ac| auto].
destruct (dec q a); [auto | contradiction].
destruct (dec q b); [auto | contradiction].
destruct (dec q c); [contradiction | auto].
Qed.
What can I do to automate such tedious proofs that involve computation with my count function?

This is typically the kind of cases where you are better off proving things by reflection. See how things go smoothly (of course I modified a bit your example to avoid all these axioms):
Require Import List.
Import ListNotations.
Fixpoint count {T : Type} (p : T -> bool) (l : list T) :=
match l with
| [] => 0
| h :: t => if p h then S (count p t) else (count p t)
end.
Inductive T := a | b | c.
Definition q x :=
match x with
| a => true
| b => true
| c => false
end.
Lemma example: (count q [a; b; c]) = 2.
Proof.
reflexivity.
Qed.
I realize that your definition of count was taking a propositional predicate on type T (but with the assumption that all predicates on type T are decidable) and instead I propose to define count to take a boolean predicate. But you may realize that having a decidable propositional predicate or having a boolean predicate is actually equivalent.
E.g. from your axioms, I can define a function which transform any propositional predicate into a boolean one:
Parameter T : Type.
Parameter dec: forall (p: T -> Prop) (w: T), {p w} + {~ (p w)}.
Definition prop_to_bool_predicate (p : T -> Prop) (x : T) : bool :=
if dec p x then true else false.
Of course, because there are axioms involved in your example, it won't actually be possible to compute with the boolean predicate. But I'm assuming that you put all these axioms for the purpose of the example and that your actual application doesn't have them.
Answer to your comment
As I told you, as soon as you have defined some function in terms of an axiom (or of a Parameter since this is the same thing), there is no way you can compute with it anymore.
However, here is a solution where the decidability of propositional predicate p is a lemma instead. I ended the proof of the lemma with Defined instead of Qed to allow computing with it (otherwise, it wouldn't be any better than an axiom). As you can see I also redefined the count function to take a predicate and a proof of its decidability. The proof by reflection still works in that case. There is no bool but it is strictly equivalent.
Require Import List.
Import ListNotations.
Fixpoint count {T : Type}
(p : T -> Prop) (dec : forall (w: T), {p w} + {~ (p w)}) (l : list T) :=
match l with
| [] => 0
| h :: t => if dec h then S (count p dec t) else (count p dec t)
end.
Inductive T := a | b | c.
Definition p x := match x with | a => True | b => True | c => False end.
Lemma dec_p: forall (w: T), {p w} + {~ (p w)}.
Proof.
intros []; simpl; auto.
Defined.
Lemma example2: (count p dec_p [a; b; c]) = 2. Proof. reflexivity. Qed.

Let's create our custom hint database and add your axioms there:
Hint Resolve Aa : axiom_db.
Hint Resolve Ab : axiom_db.
Hint Resolve Ac : axiom_db.
Now, the firstorder tactic can make use of the hint database:
Lemma example: count q (cons a (cons b (cons c nil))) = 2.
Proof.
unfold count.
destruct (dec q a), (dec q b), (dec q c); firstorder with axiom_db.
Qed.
We can automate our solution using the following piece of Ltac:
Ltac solve_the_probem :=
match goal with
|- context [if dec ?q ?x then _ else _] =>
destruct (dec q x);
firstorder with axioms_db;
solve_the_probem
end.
Then, unfold count; solve_the_probem. will be able to prove the lemma.

Related

Proving that s-expressions printing is injective

I defined a type of s-expressions and it's printing functions.
Inductive sexp : Set :=
K : string -> (list sexp) -> sexp
.
Fixpoint sexpprint (s:sexp) : list string :=
match s with
K n l => ["("%string]++[n]++(concat (map sexpprint l))++[")"%string]
end.
(Yes, I understand it can be just string, not the list of strings, but Coq have small amount of theorems for working with strings, but a big amount for working with lists.)
(* more usual function
Fixpoint sexpprint (s:sexp) :string :=
match s with
K n l => ("(":string)++n++(String.concat "" (map sexpprint l))++")"
end.
*)
And I've got stuck trying to prove this theorem:
Theorem sexpprint_inj s1 s2:
sexpprint s1 = sexpprint s2 -> s1 = s2.
Maybe there are some sources which can help me to plan the theorem's proof? (books/articles/codes) How to prove it?
(Maybe I need a special kind of inductive principle, could you formulate its statement?)
Also I defined depth function, it may somehow help
Fixpoint depth (s:sexp) : nat :=
match s with
K n l =>
(match l with
nil => 0
| _ => S (list_max (map depth l))
end)
end.
Thanks!
p.s. some additional thoughts:
Theorem depth_decr n l s m:
depth (K n l) = m
->
In s l
->
depth s < m
.
Proof.
Admitted.
Theorem step_lem (m:nat) :
(forall s1 s2,
depth s1 < m ->
depth s2 < m ->
sexpprint s1 = sexpprint s2 -> s1 = s2
) ->
(forall s1 s2,
depth s1 = m ->
depth s2 = m ->
sexpprint s1 = sexpprint s2 -> s1 = s2
).
Proof.
intros H s1 s2 Q1 Q2 E.
destruct s1 as [n1 l1], s2 as [n2 l2].
simpl in E.
inversion E as [E1].
apply (app_inv_tail) in H0.
Search "concat".
cut (l1=l2).
intros []; reflexivity.
Search "In".
induction l1, l2.
+ trivial.
+ simpl in H0.
destruct s.
unfold sexpprint in H0.
simpl in H0.
inversion H0.
+ simpl in H0.
destruct a.
unfold sexpprint in H0.
simpl in H0.
inversion H0.
+ admit.
Admitted.
p.p.s. I feel like the main obstacle is performing induction on two lists.
The type sexp is an example of a nested inductive type, where one of the recursive occurrences appears inside of another induction. Such types are hard to work with in Coq, because the induction principles that it generates by default are not useful. However, you can fix this issue by writing down your own induction principle by hand. Here is one possibility:
Require Import Coq.Lists.List Coq.Strings.String.
Import ListNotations.
Unset Elimination Schemes.
Inductive sexp : Type :=
| K : string -> list sexp -> sexp.
Set Elimination Schemes.
Definition tuple (T : sexp -> Type) (es : list sexp) :=
fold_right (fun e R => T e * R)%type unit es.
Definition sexp_rect
(T : sexp -> Type)
(H : forall s es, tuple T es -> T (K s es)) :
forall e, T e :=
fix outer (e : sexp) : T e :=
match e with
| K s es =>
let fix inner (es : list sexp) : tuple T es :=
match es return tuple T es with
| [] => tt
| e :: es => (outer e, inner es)
end in
H s es (inner es)
end.
Definition sexp_ind (T : sexp -> Prop) := sexp_rect T.
With this induction principle, it is now possible to prove your lemma (exercise!), but you will need to generalize its statement a bit.
For a deeper discussion about these nested inductives, you can have a look at CPDT.

Can't prove trivial lemma about function with non-standard recursion

I'm having a great difficulty trying to prove even very simple lemmas about a function I defined. This is my definition:
Require Import List.
Require Export Omega.
Require Export FunInd.
Require Export Recdef.
Notation "A :: B" := (cons A B).
Notation "[]" := nil.
Notation "[[ A ]]" := (A :: nil).
Inductive tm :=
| E: nat -> tm
| L: list tm -> tm.
Definition T := list tm.
Fixpoint add_list (l: list nat) : nat :=
match l with
| [] => 0
| n :: l' => n + (add_list l')
end.
Fixpoint depth (t: tm) : nat :=
match t with
| E _ => 1
| L l => 1 + (add_list (map depth l))
end.
Definition sum_depth (l: T) := add_list (map depth l).
Function sum_total (l: T) {measure sum_depth l} : nat :=
match l with
| [] => 0
| [[E n]] => n
| [[L li]] => sum_total li
| E n :: l' => n + (sum_total l')
| L li :: l' => (sum_total li) + (sum_total l')
end.
Proof.
- auto.
- intros; unfold sum_depth; subst. simpl; omega.
- intros; subst; unfold sum_depth; simpl; omega.
- intros; subst; unfold sum_depth; simpl; omega.
Defined.
The inductive type can't be changed.
I can prove simple propositions like Lemma test : forall n, sum_total [[E n]] = n. using the compute tactic, but another trivial lemma like Lemma test2 : forall l, sum_total [[L l]] = sum_total l. hangs.
First, it seems OK that the compute tactic "hangs" on the goal you mention (because when using the Function … Proof. … Defined. definition methodology, your function sum_total incorporates some proof terms, which are not intended to be computed − all the more on an arbitrary argument l; maybe a tactic such as simpl or cbn would be more suitable in this context).
Independently of my comment on list notations, I had a closer look on your formalization and it seems the Function command is unneeded in your case, because sum_total is essentially structural, so you could use a mere Fixpoint, provided the inductive type you are looking at is slightly rephrased to be defined in one go as a mutually-defined inductive type (see the corresponding doc of the Inductive command in Coq's refman which gives a similar, typical example of "tree / forest").
To elaborate on your example, you may want to adapt your definition (if it is possible for your use case) like this:
Inductive tm :=
| E: nat -> tm
| L: T -> tm
with T :=
Nil : T
| Cons : forall (e : tm) (l : T), T.
Notation "[[ A ]]" := (Cons A Nil).
Fixpoint sum_total (l: T) {struct l} : nat :=
match l with
| Nil => 0
| [[E n]] => n
| [[L li]] => sum_total li
| Cons (E n) l' => n + (sum_total l')
| Cons (L li) l' => (sum_total li) + (sum_total l')
end.
(* and the lemma you were talking about is immediate *)
Lemma test2 : forall l, sum_total [[L l]] = sum_total l.
reflexivity.
Qed.
Otherwise (if you cannot rephrase your tm inductive like this), another solution would be to use another strategy than Function to define your sum_total function, e.g. Program Fixpoint, or the Equations plugin (which are much more flexible and robust than Function when dealing with non-structural recursion / dependently-typed pattern matching).
Edit: as the OP mentions the inductive type itself can't be changed, there is a direct solution, even when using the mere Function machinery: relying on the "equation lemma" that is automatically generated by the definition.
To be more precise, if you take your script as is, then you get the following lemma "for free":
Search sum_total "equation".
(*
sum_total_equation:
forall l : T,
sum_total l =
match l with
| [] => 0
| [[E n]] => n
| E n :: (_ :: _) as l' => n + sum_total l'
| [[L li]] => sum_total li
| L li :: (_ :: _) as l' => sum_total li + sum_total l'
end
*)
So you could easily state and prove the lemma you are interested in by doing:
Lemma test2 : forall l, sum_total [[L l]] = sum_total l.
intros l.
rewrite sum_total_equation.
reflexivity.
Qed.
Here is an answer that doesn't require changing the inductive type.
There is a simple definition of sum_total that is both comparatively easy to understand and gives (almost) the lemma you are looking for by compute.
Fixpoint sum_tm (t : tm) : nat :=
match t with
| E n => n
| L li => list_sum (map sum_tm li)
end.
Definition sum_total (l : T) : nat := list_sum (map sum_tm l).
Lemma test2 : forall l, sum_total [[L l]] = sum_total l + 0.
reflexivity.
Qed.
(list_sum comes from the List module.)
Notice how the definition of sum_tm and sum_total exactly follows the structure of the definition of term and T, with list_sum (composed with map) corresponding to the use of list. This pattern is in general effective for these problems with nested inductives.
If you want to get rid of the + 0, you can define a different version of list_sum that includes a case for the singleton list (and you can fuse this with map if you want, though it is not necessary).
That would look like replacing list_sum with list_sum_alt defined as
Fixpoint list_sum_alt (l : list nat) : nat :=
match l with
| [] => 0
| [[n]] => n
| n :: li => n + list_sum_alt li
end.
With this definition, test2 holds by compute.

Applying a Program Definition fails with "unable to unify Prop with [goal]"

In Coq, I showed the associativity of append on vectors using:
Require Import Coq.Vectors.VectorDef Omega.
Program Definition t_app_assoc v p q r (a : t v p) (b : t v q) (c : t v r) :=
append (append a b) c = append a (append b c).
Next Obligation. omega. Qed.
I now want to apply this equality in a proof. Below is the easiest goal that I would expect to be provable with t_app_assoc. Of course it can be proven by simpl - this is just an example.
Goal (append (append (nil nat) (nil _)) (nil _)
= append (nil _) (append (nil _) (nil _))).
apply t_app_assoc.
This apply fails with:
Error: Unable to unify "Prop" with
"append (append (nil nat) (nil nat)) (nil nat) =
append (nil nat) (append (nil nat) (nil nat))".
How can I apply t_app_assoc? Or is there a better way to define it? I thought I needed a Program Definition, because simply using a Lemma leads to a type error because t v (p + (q + r)) and t v (p + q + r) are not the same to Coq.
Prologue
I guess what you want to to is to prove that the vector concatenation is associative and then use that fact as a lemma.
But t_app_assoc as you define it has the following type:
t_app_assoc
: forall (v : Type) (p q r : nat), t v p -> t v q -> t v r -> Prop
You basically want to use : instead of := as follows.
From Coq Require Import Vector Arith.
Import VectorNotations.
Import EqNotations. (* rew notation, see below *)
Section Append.
Variable A : Type.
Variable p q r : nat.
Variables (a : t A p) (b : t A q) (c : t A r).
Fail Lemma t_app_assoc :
append (append a b) c = append a (append b c).
Unfortunately, we cannot even state a lemma like this using the usual homogeneous equality.
The left-hand side has the following type:
Check append (append a b) c : t A (p + q + r).
whereas the right-hand side is of type
Check append a (append b c) : t A (p + (q + r)).
Since t A (p + q + r) is not the same as t A (p + (q + r)) we cannot use = to state the above lemma.
Let me describe some ways of working around this issue:
rew notation
Lemma t_app_assoc_rew :
append (append a b) c = rew (plus_assoc _ _ _) in
append a (append b c).
Admitted.
Here we just use the law of associativity of addition for natural numbers to cast the type of RHS to t A (p + q + r).
To make it work one needs to Import EqNotations. before.
cast function
This is a common problem, so the authors of the Vector library decided to provide a cast function with the following type:
cast :
forall (A : Type) (m : nat),
Vector.t A m -> forall n : nat, m = n -> Vector.t A n
Let me show how one can use it to prove the law of associativity for vectors. But let's prove the following auxiliary lemma first:
Lemma uncast {X n} {v : Vector.t X n} e :
cast v e = v.
Proof. induction v as [|??? IH]; simpl; rewrite ?IH; reflexivity. Qed.
Now we are all set:
Lemma t_app_assoc_cast (a : t A p) (b : t A q) (c : t A r) :
append (append a b) c = cast (append a (append b c)) (plus_assoc _ _ _).
Proof.
generalize (Nat.add_assoc p q r).
induction a as [|h p' a' IH]; intros e.
- now rewrite uncast.
- simpl; f_equal. apply IH.
Qed.
Heterogeneous equality (a.k.a. John Major equality)
Lemma t_app_assoc_jmeq :
append (append a b) c ~= append a (append b c).
Admitted.
End Append.
If you compare the definition of the homogeneous equality
Inductive eq (A : Type) (x : A) : A -> Prop :=
eq_refl : x = x.
and the definition of heterogeneous equality
Inductive JMeq (A : Type) (x : A) : forall B : Type, B -> Prop :=
JMeq_refl : x ~= x.
you will see that with JMeq the LHS and RHS don't have to be of the same type and this is why the statement of t_app_assoc_jmeq looks a bit simpler than the previous ones.
Other approaches to vectors
See e.g. this question
and this one;
I also find this answer
very useful too.

coq induction with passing in equality

I have a list with a known value and want to induct on it, keeping track of what the original list was, and referring to it by element. That is, I need to refer to it by l[i] with varying i instead of just having (a :: l).
I tried to make an induction principle to allow me to do that. Here is a program with all of the unnecessary Theorems replaced with Admitted, using a simplified example. The objective is to prove allLE_countDown using countDown_nth, and have list_nth_rect in a convenient form. (The theorem is easy to prove directly without any of those.)
Require Import Arith.
Require Import List.
Definition countDown1 := fix f a i := match i with
| 0 => nil
| S i0 => (a + i0) :: f a i0
end.
(* countDown from a number to another, excluding greatest. *)
Definition countDown a b := countDown1 b (a - b).
Theorem countDown_nth a b i d (boundi : i < length (countDown a b))
: nth i (countDown a b) d = a - i - 1.
Admitted.
Definition allLE := fix f l m := match l with
| nil => true
| a :: l0 => if Nat.leb a m then f l0 m else false
end.
Definition drop {A} := fix f (l : list A) n := match n with
| 0 => l
| S a => match l with
| nil => nil
| _ :: l2 => f l2 a
end
end.
Theorem list_nth_rect_aux {A : Type} (P : list A -> list A -> nat -> Type)
(Pnil : forall l, P l nil (length l))
(Pcons : forall i s l d (boundi : i < length l), P l s (S i) -> P l ((nth i l d) :: s) i)
l s i (size : length l = i + length s) (sub : s = drop l i) : P l s i.
Admitted.
Theorem list_nth_rect {A : Type} (P : list A -> list A -> nat -> Type)
(Pnil : forall l, P l nil (length l))
(Pcons : forall i s l d (boundi : i < length l), P l s (S i) -> P l ((nth i l d) :: s) i)
l s (leqs : l = s): P l s 0.
Admitted.
Theorem allLE_countDown a b : allLE (countDown a b) a = true.
remember (countDown a b) as l.
refine (list_nth_rect (fun l s _ => l = countDown a b -> allLE s a = true) _ _ l l eq_refl Heql);
intros; subst; [ apply eq_refl | ].
rewrite countDown_nth; [ | apply boundi ].
pose proof (Nat.le_sub_l a (i + 1)).
rewrite Nat.sub_add_distr in H0.
apply leb_correct in H0.
simpl; rewrite H0; clear H0.
apply (H eq_refl).
Qed.
So, I have list_nth_rect and was able to use it with refine to prove the theorem by referring to the nth element, as desired. However, I had to construct the Proposition P myself. Normally, you'd like to use induction.
This requires distinguishing which elements are the original list l vs. the sublist s that is inducted on. So, I can use remember.
Theorem allLE_countDown a b : allLE (countDown a b) a = true.
remember (countDown a b) as s.
remember s as l.
rewrite Heql.
This puts me at
a, b : nat
s, l : list nat
Heql : l = s
Heqs : l = countDown a b
============================
allLE s a = true
However, I can't seem to pass the equality as I just did above. When I try
induction l, s, Heql using list_nth_rect.
I get the error
Error: Abstracting over the terms "l", "s" and "0" leads to a term
fun (l0 : list ?X133#{__:=a; __:=b; __:=s; __:=l; __:=Heql; __:=Heqs})
(s0 : list ?X133#{__:=a; __:=b; __:=s; __:=l0; __:=Heql; __:=Heqs})
(_ : nat) =>
(fun (l1 l2 : list nat) (_ : l1 = l2) =>
l1 = countDown a b -> allLE l2 a = true) l0 s0 Heql
which is ill-typed.
Reason is: Illegal application:
The term
"fun (l l0 : list nat) (_ : l = l0) =>
l = countDown a b -> allLE l0 a = true" of type
"forall l l0 : list nat, l = l0 -> Prop"
cannot be applied to the terms
"l0" : "list nat"
"s0" : "list nat"
"Heql" : "l = s"
The 3rd term has type "l = s" which should be coercible to
"l0 = s0".
So, how can I change the induction principle
such that it works with the induction tactic?
It looks like it's getting confused between
the outer variables and the ones inside the
function. But, I don't have a way to talk
about the inner variables that aren't in scope.
It's very strange, since invoking it with
refine works without issues.
I know for match, there's as clauses, but
I can't figure out how to apply that here.
Or, is there a way to make list_nth_rect use
P l l 0 and still indicate which variables correspond to l and s?
First, you can prove this result much more easily by reusing more basic ones. Here's a version based on definitions of the ssreflect library:
From mathcomp
Require Import ssreflect ssrfun ssrbool ssrnat eqtype seq.
Definition countDown n m := rev (iota m (n - m)).
Lemma allLE_countDown n m : all (fun k => k <= n) (countDown n m).
Proof.
rewrite /countDown all_rev; apply/allP=> k; rewrite mem_iota.
have [mn|/ltnW] := leqP m n.
by rewrite subnKC //; case/andP => _; apply/leqW.
by rewrite -subn_eq0 => /eqP ->; rewrite addn0 ltnNge andbN.
Qed.
Here, iota n m is the list of m elements that counts starting from n, and all is a generic version of your allLE. Similar functions and results exist in the standard library.
Back to your original question, it is true that sometimes we need to induct on a list while remembering the entire list we started with. I don't know if there is a way to get what you want with the standard induction tactic; I didn't even know that it had a multi-argument variant. When I want to prove P l using this strategy, I usually proceed as follows:
Find a predicate Q : nat -> Prop such that Q (length l) implies P l. Typically, Q n will have the form n <= length l -> R (take n l) (drop n l), where R : list A -> list A -> Prop.
Prove Q n for all n by induction.
I do not know if this answers your question, but induction seems to accept with clauses. Thus, you can write the following.
Theorem allLE_countDown a b : allLE (countDown a b) a = true.
remember (countDown a b) as s.
remember s as l.
rewrite Heql.
induction l, s, Heql using list_nth_rect
with (P:=fun l s _ => l = countDown a b -> allLE s a = true).
But the benefit is quite limited w.r.t. the refine version, since you need to specify manually the predicate.
Now, here is how I would have proved such a result using objects from the standard library.
Require Import List. Import ListNotations.
Require Import Omega.
Definition countDown1 := fix f a i := match i with
| 0 => nil
| S i0 => (a + i0) :: f a i0
end.
(* countDown from a number to another, excluding greatest. *)
Definition countDown a b := countDown1 b (a - b).
Theorem countDown1_nth a i k d (boundi : k < i) :
nth k (countDown1 a i) d = a + i -k - 1.
Proof.
revert k boundi.
induction i; intros.
- inversion boundi.
- simpl. destruct k.
+ omega.
+ rewrite IHi; omega.
Qed.
Lemma countDown1_length a i : length (countDown1 a i) = i.
Proof.
induction i.
- reflexivity.
- simpl. rewrite IHi. reflexivity.
Qed.
Theorem countDown_nth a b i d (boundi : i < length (countDown a b))
: nth i (countDown a b) d = a - i - 1.
Proof.
unfold countDown in *.
rewrite countDown1_length in boundi.
rewrite countDown1_nth.
replace (b+(a-b)) with a by omega. reflexivity. assumption.
Qed.
Theorem allLE_countDown a b : Forall (ge a) (countDown a b).
Proof.
apply Forall_forall. intros.
apply In_nth with (d:=0) in H.
destruct H as (n & H & H0).
rewrite countDown_nth in H0 by assumption. omega.
Qed.
EDIT:
You can state an helper lemma to make an even more concise proof.
Lemma Forall_nth : forall {A} (P:A->Prop) l,
(forall d i, i < length l -> P (nth i l d)) ->
Forall P l.
Proof.
intros. apply Forall_forall.
intros. apply In_nth with (d:=x) in H0.
destruct H0 as (n & H0 & H1).
rewrite <- H1. apply H. assumption.
Qed.
Theorem allLE_countDown a b : Forall (ge a) (countDown a b).
Proof.
apply Forall_nth.
intros. rewrite countDown_nth. omega. assumption.
Qed.
The issue is that, for better or for worse, induction seems to assume that its arguments are independent. The solution, then, is to let induction automatically infer l and s from Heql:
Theorem list_nth_rect {A : Type} {l s : list A} (P : list A -> list A -> nat -> Type)
(Pnil : P l nil (length l))
(Pcons : forall i s d (boundi : i < length l), P l s (S i) -> P l ((nth i l d) :: s) i)
(leqs : l = s): P l s 0.
Admitted.
Theorem allLE_countDown a b : allLE (countDown a b) a = true.
remember (countDown a b) as s.
remember s as l.
rewrite Heql.
induction Heql using list_nth_rect;
intros; subst; [ apply eq_refl | ].
rewrite countDown_nth; [ | apply boundi ].
pose proof (Nat.le_sub_l a (i + 1)).
rewrite Nat.sub_add_distr in H.
apply leb_correct in H.
simpl; rewrite H; clear H.
assumption.
Qed.
I had to change around the type of list_nth_rect a bit; I hope I haven't made it false.

Coq rewriting using lambda arguments

We have a function that inserts an element into a specific index of a list.
Fixpoint inject_into {A} (x : A) (l : list A) (n : nat) : option (list A) :=
match n, l with
| 0, _ => Some (x :: l)
| S k, [] => None
| S k, h :: t => let kwa := inject_into x t k
in match kwa with
| None => None
| Some l' => Some (h :: l')
end
end.
The following property of the aforementioned function is of relevance to the problem (proof omitted, straightforward induction on l with n not being fixed):
Theorem inject_correct_index : forall A x (l : list A) n,
n <= length l -> exists l', inject_into x l n = Some l'.
And we have a computational definition of permutations, with iota k being a list of nats [0...k]:
Fixpoint permute {A} (l : list A) : list (list A) :=
match l with
| [] => [[]]
| h :: t => flat_map (
fun x => map (
fun y => match inject_into h x y with
| None => []
| Some permutations => permutations
end
) (iota (length t))) (permute t)
end.
The theorem we're trying to prove:
Theorem num_permutations : forall A (l : list A) k,
length l = k -> length (permute l) = factorial k.
By induction on l we can (eventually) get to following goal: length (permute (a :: l)) = S (length l) * length (permute l). If we now simply cbn, the resulting goal is stated as follows:
length
(flat_map
(fun x : list A =>
map
(fun y : nat =>
match inject_into a x y with
| Some permutations => permutations
| None => []
end) (iota (length l))) (permute l)) =
length (permute l) + length l * length (permute l)
Here I would like to proceed by destruct (inject_into a x y), which is impossible considering x and y are lambda arguments. Please note that we will never get the None branch as a result of the lemma inject_correct_index.
How does one proceed from this proof state? (Please do note that I am not trying to simply complete the proof of the theorem, that's completely irrelevant.)
There is a way to rewrite under binders: the setoid_rewrite tactic (see §27.3.1 of the Coq Reference manual).
However, direct rewriting under lambdas is not possible without assuming an axiom as powerful as the axiom of functional extensionality (functional_extensionality).
Otherwise, we could have proved:
(* classical example *)
Goal (fun n => n + 0) = (fun n => n).
Fail setoid_rewrite <- plus_n_O.
Abort.
See here for more detail.
Nevertheless, if you are willing to accept such axiom, then you can use the approach described by Matthieu Sozeau in this Coq Club post to rewrite under lambdas like so:
Require Import Coq.Logic.FunctionalExtensionality.
Require Import Coq.Setoids.Setoid.
Require Import Coq.Classes.Morphisms.
Generalizable All Variables.
Instance pointwise_eq_ext {A B : Type} `(sb : subrelation B RB eq)
: subrelation (pointwise_relation A RB) eq.
Proof. intros f g Hfg. apply functional_extensionality. intro x; apply sb, (Hfg x). Qed.
Goal (fun n => n + 0) = (fun n => n).
setoid_rewrite <- plus_n_O.
reflexivity.
Qed.