coq-club AT inria.fr
Subject: The Coq mailing list
List archive
- From: Xavier Leroy <xavier.leroy AT college-de-france.fr>
- To: coq-club AT inria.fr
- Subject: Re: [Coq-Club] Straight reasoning with existential variables
- Date: Fri, 2 Aug 2019 20:08:52 +0200
- Authentication-results: mail2-smtp-roc.national.inria.fr; spf=None smtp.pra=xavier.leroy AT college-de-france.fr; spf=Pass smtp.mailfrom=xavier.leroy AT college-de-france.fr; spf=None smtp.helo=postmaster AT smtpout01-ext1.partage.renater.fr
On Fri, Aug 2, 2019 at 8:02 PM Joshua Gancher <jrg358 AT cornell.edu> wrote:
I am not sure if this is exactly what you want, but you might be interested in the Procrastination library: https://github.com/fakusb/coq-procrastination
Giving credit to the original author, this should be https://github.com/Armael/coq-procrastination .
Here is a pesentation that explains Procrastination in more details: https://easychair.org/smart-slide/slide/5pGk#
Enjoy,
- Xavier Leroy
Using this library, I can indeed do (something equivalent to) eexists where you want it:Require Import Procrastination.
Inductive sum : nat -> nat -> nat -> Prop :=
| Sum0 : forall a, sum 0 a a
| SumS : forall a b c, sum a (S b) c -> sum (S a) b c.
Theorem sum_exists : forall a b, exists c, sum a b c.
induction a; intro.
- exists b; constructor.
- begin defer assuming c.
+ exists c.
constructor.
defer.
+ end defer.
apply (IHa (S b)).
Qed.JoshOn Fri, Aug 2, 2019 at 10:40 AM Ilia Zaichuk <zoickx AT ztld.org> wrote:Hello everyone,Consider the following example:Inductive sum : nat -> nat -> nat -> Prop :=| SumO : forall a, sum 0 a a| SumS : forall a b c, sum a (S b) c -> sum (S a) b c.Theorem sum_exists :forall a b, exists c, sum a b c.Proof.induction a; intro.- exists b. constructor.- eexists.constructor.specialize (IHa (S b)).destruct IHa as [c IHa].Fail apply IHa.Abort.The problem here (as suggested by [apply]) is that the shelved [?c] does nothave [c] in its scope. Fair enough. This problem has been coming up in a few ofmy proofs lately, and I have learned to deal with it by moving the [destruct] -the instantiation of [c] - above the [eexists].However, this comes with a problem - it breaks the "flow of thought" in a proof.You either have to know in advance what variables you need instantiated, or godown to the point where you need them and move the instantiation above [eexists].That would be fine for this small example, but in more complex proofs, it quicklybecomes unwieldy.Without [eexists], I cannot apply the [SumS] constructor.It is possible to add a lemma:Lemma SumS_existential (a b : nat) :(exists c, sum a (S b) c) -> exists c, sum (S a) b c.Proof.intros.destruct H as [c H].exists c. constructor.assumption.Qed., with which the proof goes smoothly. But this, once again, is unsustainablefor more complex problems. I have tried writing a generic lemma for"existential implication", but was unsuccessful.Can such a lemma exist? (Or is it a tactic I am looking for?)I have also tried manipulating the contexts of the main goal and the existentialvariables separately (via [Unshelve]), and failed at proving the lemma this way at all.Is it at all possible to cleanly do that in this case?Overall, how can I prove this lemma without getting the reasoning tangled?Thanks and regards,Ilia.
- [Coq-Club] Straight reasoning with existential variables, Ilia Zaichuk, 08/02/2019
- Re: [Coq-Club] Straight reasoning with existential variables, Joshua Gancher, 08/02/2019
- Re: [Coq-Club] Straight reasoning with existential variables, Xavier Leroy, 08/02/2019
- Re: [Coq-Club] Straight reasoning with existential variables, Joshua Gancher, 08/02/2019
Archive powered by MHonArc 2.6.18.