Detailed examples of tactics¶
This chapter presents detailed examples of certain tactics, to illustrate their behavior.
dependent induction¶
The tactics dependent induction
and dependent destruction
are another
solution for inverting inductive predicate instances and potentially
doing induction at the same time. It is based on the BasicElim
tactic
of Conor McBride which works by abstracting each argument of an
inductive instance by a variable and constraining it by equalities
afterwards. This way, the usual induction and destruct tactics can be
applied to the abstracted instance and after simplification of the
equalities we get the expected goals.
The abstracting tactic is called generalize_eqs and it takes as argument a hypothesis to generalize. It uses the JMeq datatype defined in Coq.Logic.JMeq, hence we need to require it before. For example, revisiting the first example of the inversion documentation:
Require Import Coq.Logic.JMeq. Inductive Le : nat -> nat -> Set := | LeO : forall n:nat, Le 0 n | LeS : forall n m:nat, Le n m -> Le (S n) (S m). Parameter P : nat -> nat -> Prop.forall n m : nat, Le (S n) m -> P n mn, m:natH:Le (S n) mP n mn, m, gen_x:natH:Le gen_x mgen_x = S n -> P n m
The index S n
gets abstracted by a variable here, but a corresponding
equality is added under the abstract instance so that no information
is actually lost. The goal is now almost amenable to do induction or
case analysis. One should indeed first move n
into the goal to
strengthen it before doing induction, or n
will be fixed in the
inductive hypotheses (this does not matter for case analysis). As a
rule of thumb, all the variables that appear inside constructors in
the indices of the hypothesis should be generalized. This is exactly
what the generalize_eqs_vars
variant does:
n, m, gen_x:natH:Le gen_x mgen_x = S n -> P n mn, n0:nat0 = S n -> P n n0n, n0, m:natH:Le n0 mIHLe:n0 = S n -> P n mS n0 = S n -> P n (S m)
As the hypothesis itself did not appear in the goal, we did not need to use an heterogeneous equality to relate the new hypothesis to the old one (which just disappeared here). However, the tactic works just as well in this case, e.g.:
Parameter Q : forall (n m : nat), Le n m -> Prop.forall (n m : nat) (p : Le (S n) m), Q (S n) m pn, m:natp:Le (S n) mQ (S n) m pm, gen_x:natp:Le gen_x mforall (n : nat) (p0 : Le (S n) m), gen_x = S n -> p ~= p0 -> Q (S n) m p0
One drawback of this approach is that in the branches one will have to
substitute the equalities back into the instance to get the right
assumptions. Sometimes injection of constructors will also be needed
to recover the needed equalities. Also, some subgoals should be
directly solved because of inconsistent contexts arising from the
constraints on indexes. The nice thing is that we can make a tactic
based on discriminate, injection and variants of substitution to
automatically do such simplifications (which may involve the axiom K).
This is what the simplify_dep_elim
tactic from Coq.Program.Equality
does. For example, we might simplify the previous goals considerably:
n, m:natp:Le n mIHp:forall (n0 : nat) (p0 : Le (S n0) m),
n = S n0 -> p ~= p0 -> Q (S n0) m p0Q (S n) (S m) (LeS n m p)
The higher-order tactic do_depind
defined in Coq.Program.Equality
takes a tactic and combines the building blocks we have seen with it:
generalizing by equalities calling the given tactic with the
generalized induction hypothesis as argument and cleaning the subgoals
with respect to equalities. Its most important instantiations
are dependent induction
and dependent destruction
that do induction or
simply case analysis on the generalized hypothesis. For example we can
redo what we’ve done manually with dependent destruction:
forall n m : nat, Le (S n) m -> P n mn, m:natH:Le (S n) mP n mn, m:natH:Le n mP n (S m)
This gives essentially the same result as inversion. Now if the destructed hypothesis actually appeared in the goal, the tactic would still be able to invert it, contrary to dependent inversion. Consider the following example on vectors:
Set Implicit Arguments. Parameter A : Set.
Inductive vector : nat -> Type :=
| vnil : vector 0
| vcons : A -> forall n, vector n -> vector (S n).
forall (n : nat) (v : vector (S n)), exists (v' : vector n) (a : A), v = vcons a v'n:natv:vector (S n)exists (v' : vector n) (a : A), v = vcons a v'n:nata:Av:vector nexists (v' : vector n) (a0 : A), vcons a v = vcons a0 v'
In this case, the v
variable can be replaced in the goal by the
generalized hypothesis only when it has a type of the form vector (S n)
,
that is only in the second case of the destruct. The first one is
dismissed because S n <> 0
.
A larger example¶
Let’s see how the technique works with induction on inductive predicates on a real example. We will develop an example application to the theory of simply-typed lambda-calculus formalized in a dependently-typed style:
Inductive type : Type :=
| base : type
| arrow : type -> type -> type.
Notation " t --> t' " := (arrow t t') (at level 20, t' at next level).
Inductive ctx : Type :=
| empty : ctx
| snoc : ctx -> type -> ctx.
Notation " G , tau " := (snoc G tau) (at level 20, tau at next level).
Fixpoint conc (G D : ctx) : ctx :=
match D with
| empty => G
| snoc D' x => snoc (conc G D') x
end.
Notation " G ; D " := (conc G D) (at level 20).
Inductive term : ctx -> type -> Type :=
| ax : forall G tau, term (G, tau) tau
| weak : forall G tau,
term G tau -> forall tau', term (G, tau') tau
| abs : forall G tau tau',
term (G , tau) tau' -> term G (tau --> tau')
| app : forall G tau tau',
term G (tau --> tau') -> term G tau -> term G tau'.
We have defined types and contexts which are snoc-lists of types. We
also have a conc
operation that concatenates two contexts. The term
datatype represents in fact the possible typing derivations of the
calculus, which are isomorphic to the well-typed terms, hence the
name. A term is either an application of:
the axiom rule to type a reference to the first variable in a context
the weakening rule to type an object in a larger context
the abstraction or lambda rule to type a function
the application to type an application of a function to an argument
Once we have this datatype we want to do proofs on it, like weakening:
forall (G D : ctx) (tau : type),
term (G; D) tau ->
forall tau' : type, term ((G, tau'); D) tau
The problem here is that we can’t just use induction on the typing
derivation because it will forget about the G ; D
constraint appearing
in the instance. A solution would be to rewrite the goal as:
forall (G' : ctx) (tau : type),
term G' tau ->
forall G D : ctx,
G; D = G' ->
forall tau' : type, term ((G, tau'); D) tau
With this proper separation of the index from the instance and the
right induction loading (putting G
and D
after the inducted-on
hypothesis), the proof will go through, but it is a very tedious
process. One is also forced to make a wrapper lemma to get back the
more natural statement. The dependent induction
tactic alleviates this
trouble by doing all of this plumbing of generalizing and substituting
back automatically. Indeed we can simply write:
Require Import Coq.Program.Tactics. Require Import Coq.Program.Equality.forall (G D : ctx) (tau : type), term (G; D) tau -> forall tau' : type, term ((G, tau'); D) tauforall (G D : ctx) (tau : type), term (G; D) tau -> forall tau' : type, term ((G, tau'); D) tauG, D:ctxtau:typeH:term (G; D) tauforall tau' : type, term ((G, tau'); D) tauG0:ctxtau:typeG, D:ctxx:G0, tau = G; Dtau':typeterm ((G, tau'); D) tauG0:ctxtau:typeH:term G0 tautau':typeIHterm:forall G1 D0 : ctx, G0 = G1; D0 -> forall tau'1 : type, term ((G1, tau'1); D0) tauG, D:ctxx:G0, tau' = G; Dtau'0:typeterm ((G, tau'0); D) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'
This call to dependent induction has an additional arguments which is a list of variables appearing in the instance that should be generalized in the goal, so that they can vary in the induction hypotheses. By default, all variables appearing inside constructors (except in a parameter position) of the instantiated hypothesis will be generalized automatically but one can always give the list explicitly.
The simpl_depind
tactic includes an automatic tactic that tries to
simplify equalities appearing at the beginning of induction
hypotheses, generally using trivial applications of reflexivity
. In
cases where the equality is not between constructor forms though, one
must help the automation by giving some arguments, using the
specialize
tactic for example.
G0:ctxtau, tau':typeterm ((G0, tau), tau') tauG, D:ctxt, tau':typeterm (((G, tau'); D), t) tG0:ctxtau:typeH:term G0 tautau':typeIHterm:forall G1 D0 : ctx, G0 = G1; D0 -> forall tau'1 : type, term ((G1, tau'1); D0) tauG, D:ctxx:G0, tau' = G; Dtau'0:typeterm ((G, tau'0); D) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'G, D:ctxt, tau':typeterm (((G, tau'); D), t) tG0:ctxtau:typeH:term G0 tautau':typeIHterm:forall G1 D0 : ctx, G0 = G1; D0 -> forall tau'1 : type, term ((G1, tau'1); D0) tauG, D:ctxx:G0, tau' = G; Dtau'0:typeterm ((G, tau'0); D) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'G0:ctxtau:typeH:term G0 tautau':typeIHterm:forall G1 D0 : ctx, G0 = G1; D0 -> forall tau'1 : type, term ((G1, tau'1); D0) tauG, D:ctxx:G0, tau' = G; Dtau'0:typeterm ((G, tau'0); D) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'G0:ctxtau:typeH:term G0 tautau':typeIHterm:forall G D : ctx, G0 = G; D -> forall tau'1 : type, term ((G, tau'1); D) tautau'0:typeterm ((G0, tau'), tau'0) tautau:typeG, D:ctxIHterm:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau' : type, term ((G0, tau'); D0) tauH:term (G; D) taut, tau'0:typeterm (((G, tau'0); D), t) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'G0:ctxtau:typeH:term G0 tautau':typeIHterm:forall tau'1 : type, term ((G0, tau'1); empty) tautau'0:typeterm ((G0, tau'), tau'0) tautau:typeG, D:ctxIHterm:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau' : type, term ((G0, tau'); D0) tauH:term (G; D) taut, tau'0:typeterm (((G, tau'0); D), t) tauG, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx, (G; D), tau = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx, G; D = G0; D0 -> forall tau'1 : type, term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'
Once the induction hypothesis has been narrowed to the right equality, it can be used directly.
tau:typeG, D:ctxIHterm:forall G0 D0 : ctx,
G; D = G0; D0 ->
forall tau' : type, term ((G0, tau'); D0) tauH:term (G; D) taut, tau'0:typeterm (((G, tau'0); D), t) tau
G, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx,
(G; D), tau = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')
G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx,
G; D = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx,
G; D = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'
Now concluding this subgoal is easy.
G, D:ctxtau, tau':typeH:term ((G; D), tau) tau'IHterm:forall G0 D0 : ctx,
(G; D), tau = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) tau'tau'0:typeterm ((G, tau'0); D) (tau --> tau')
G, D:ctxtau, tau':typeH:term (G; D) (tau --> tau')H0:term (G; D) tauIHterm1:forall G0 D0 : ctx,
G; D = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) (tau --> tau')IHterm2:forall G0 D0 : ctx,
G; D = G0; D0 ->
forall tau'1 : type,
term ((G0, tau'1); D0) tautau'0:typeterm ((G, tau'0); D) tau'
autorewrite¶
Here are two examples of autorewrite
use. The first one ( Ackermann
function) shows actually a quite basic use where there is no
conditional rewriting. The second one ( Mac Carthy function)
involves conditional rewritings and shows how to deal with them using
the optional tactic of the Hint Rewrite
command.
Example: Ackermann function
Require Import Arith. Parameter Ack : nat -> nat -> nat. Axiom Ack0 : forall m:nat, Ack 0 m = S m. Axiom Ack1 : forall n:nat, Ack (S n) 0 = Ack n 1. Axiom Ack2 : forall n m:nat, Ack (S n) (S m) = Ack n (Ack (S n) m). Hint Rewrite Ack0 Ack1 Ack2 : base0.autorewrite with base0 using try reflexivity. Qed.Ack 3 2 = 29
Example: MacCarthy function
Require Import Omega. Parameter g : nat -> nat -> nat. Axiom g0 : forall m:nat, g 0 m = m. Axiom g1 : forall n m:nat, (n > 0) -> (m > 100) -> g n m = g (pred n) (m - 10). Axiom g2 : forall n m:nat, (n > 0) -> (m <= 100) -> g n m = g (S n) (m + 11). Hint Rewrite g0 g1 g2 using omega : base1.autorewrite with base1 using reflexivity || simpl. Qed.g 1 110 = 100autorewrite with base1 using reflexivity || simpl. Qed.g 1 95 = 91