coq-club AT inria.fr
Subject: The Coq mailing list
List archive
- From: Hendrik Tews <tews AT os.inf.tu-dresden.de>
- To: coq-club AT inria.fr
- Subject: [Coq-Club] rewrite and subgoals for dependent premises
- Date: Fri, 9 Sep 2011 14:22:51 +0200
Hi,
I am looking for a variants of rewrite and apply that generate
subgoals for dependent premises whose type is in Prop, if the
unification does not yield an instance.
Alternatively I am looking for a variant of erewrite and eapply
that turn existential variables of a type in Prop into subgoals.
For a small, artificial example, consider
Require Import Arith.
Parameter f : forall(n : nat)(pos : n > 0), nat.
(* f is simply a partial function, whose value
* does not depend on the argument pos.
*)
Axiom fvalue_tcc : forall(n : nat), n > 1 -> n - 1 > 0.
Axiom fvalue : forall(n : nat)(pos : n > 0)(gt1 : n > 1),
f n pos = f (n - 1) (fvalue_tcc n gt1).
(* This is a rewrite rule for f. Here pos is universally
* quantified, because the rule should apply to all witnesses
* of n > 0. n > 1 is a simple precondition for the rule.
* It only gets a name here, because I have to construct
* the witness for the f on the right-hand side.
*)
Lemma lemma : f 2 (gt_Sn_O 1) = 0.
Proof.
(* Here I would like to rewrite with fvalue, but in such a way
* that I get an additional subgoal 2 > 1, whose proof will
* provide the instantiation for argument gt1.
* erewrite does not really help here, because the created
* existential will most likely disappear without getting
* instantiated.
*)
Thanks in advance,
Hendrik
- [Coq-Club] rewrite and subgoals for dependent premises, Hendrik Tews
- Re: [Coq-Club] rewrite and subgoals for dependent premises, Arnaud Spiwack
Archive powered by MhonArc 2.6.16.