Chapter 10  Detailed examples of tactics

This chapter presents detailed examples of certain tactics, to illustrate their behavior.

10.1  refine

This tactic applies to any goal. It behaves like exact with a big difference : the user can leave some holes (denoted by _ or (_:type)) in the term. refine will generate as many subgoals as they are holes in the term. The type of holes must be either synthesized by the system or declared by an explicit cast like (\_:nat->Prop). This low-level tactic can be useful to advanced users.


Example:

Coq < Inductive Option : Set :=
Coq <   | Fail : Option
Coq <   | Ok : bool -> Option.

Coq < Definition get : forall x:Option, x <> Fail -> bool.
1 subgoal
  
  ============================
   forall x : Option, x <> Fail -> bool

Coq < refine
Coq <  (fun x:Option =>
Coq <     match x return x <> Fail -> bool with
Coq <     | Fail => _
Coq <     | Ok b => fun _ => b
Coq <     end).
1 subgoal
  
  x : Option
  ============================
   Fail <> Fail -> bool

Coq < intros; absurd (Fail = Fail); trivial.
Proof completed.

Coq < Defined.

10.2  eapply


Example: Assume we have a relation on nat which is transitive:

Coq < Variable R : nat -> nat -> Prop.

Coq < Hypothesis Rtrans : forall x y z:nat, R x y -> R y z -> R x z.

Coq < Variables n m p : nat.

Coq < Hypothesis Rnm : R n m.

Coq < Hypothesis Rmp : R m p.

Consider the goal (R n p) provable using the transitivity of R:

Coq < Goal R n p.

The direct application of Rtrans with apply fails because no value for y in Rtrans is found by apply:

Coq < apply Rtrans.
Unnamed_thm < Unnamed_thm < Toplevel input, characters 144-156:
> apply Rtrans.
> ^^^^^^^^^^^^
Error: Unable to find an instance for the variable y.

A solution is to rather apply (Rtrans n m p).

Coq < apply (Rtrans n m p).
2 subgoals
  
  ============================
   R n m
subgoal 2 is:
 R m p

More elegantly, apply Rtrans with (y:=m) allows to only mention the unknown m:

Coq < 
Coq <   apply Rtrans with (y := m).
2 subgoals
  
  ============================
   R n m
subgoal 2 is:
 R m p

Another solution is to mention the proof of (R x y) in Rtrans...

Coq < 
Coq <   apply Rtrans with (1 := Rnm).
1 subgoal
  
  ============================
   R m p

... or the proof of (R y z):

Coq < 
Coq <   apply Rtrans with (2 := Rmp).
1 subgoal
  
  ============================
   R n m

On the opposite, one can use eapply which postpone the problem of finding m. Then one can apply the hypotheses Rnm and Rmp. This instantiates the existential variable and completes the proof.

Coq < eapply Rtrans.
2 subgoals
  
  ============================
   R n ?9
subgoal 2 is:
 R ?9 p

Coq < apply Rnm.
1 subgoal
  
  ============================
   R m p

Coq < apply Rmp.
Proof completed.

10.3  Scheme


Example 1: Induction scheme for tree and forest

The definition of principle of mutual induction for tree and forest over the sort Set is defined by the command:

Coq < Inductive tree : Set :=
Coq <     node : A -> forest -> tree
Coq < with forest : Set :=
Coq <   | leaf : B -> forest
Coq <   | cons : tree -> forest -> forest.

Coq < 
Coq < Scheme tree_forest_rec := Induction for tree Sort Set
Coq <   with forest_tree_rec := Induction for forest Sort Set.

You may now look at the type of tree_forest_rec:

Coq < Check tree_forest_rec.
tree_forest_rec
     : forall (P : tree -> Set) (P0 : forest -> Set),
       (forall (a : A) (f : forest), P0 f -> P (node a f)) ->
       (forall b : B, P0 (leaf b)) ->
       (forall t : tree, P t -> forall f1 : forest, P0 f1 -> P0 (cons t f1)) ->
       forall t : tree, P t

This principle involves two different predicates for trees and forests; it also has three premises each one corresponding to a constructor of one of the inductive definitions.

The principle forest_tree_rec shares exactly the same premises, only the conclusion now refers to the property of forests.

Coq < Check forest_tree_rec.
forest_tree_rec
     : forall (P : tree -> Set) (P0 : forest -> Set),
       (forall (a : A) (f : forest), P0 f -> P (node a f)) ->
       (forall b : B, P0 (leaf b)) ->
       (forall t : tree, P t -> forall f1 : forest, P0 f1 -> P0 (cons t f1)) ->
       forall f2 : forest, P0 f2


Example 2: Predicates odd and even on naturals

Let odd and even be inductively defined as:

Coq < Inductive odd : nat -> Prop :=
Coq <     oddS : forall n:nat, even n -> odd (S n)
Coq < with even : nat -> Prop :=
Coq <   | evenO : even 0
Coq <   | evenS : forall n:nat, odd n -> even (S n).

The following command generates a powerful elimination principle:

Coq < Scheme odd_even := Minimality for   odd Sort Prop
Coq <   with even_odd := Minimality for even Sort Prop.
odd_even, even_odd are recursively defined

The type of odd_even for instance will be:

Coq < Check odd_even.
odd_even
     : forall P P0 : nat -> Prop,
       (forall n : nat, even n -> P0 n -> P (S n)) ->
       P0 0 ->
       (forall n : nat, odd n -> P n -> P0 (S n)) ->
       forall n : nat, odd n -> P n

The type of even_odd shares the same premises but the conclusion is (n:nat)(even n)->(Q n).

10.3.1  Combined Scheme

We can define the induction principles for trees and forests using:

Coq < Scheme tree_forest_ind := Induction for tree Sort Prop
Coq <   with forest_tree_ind := Induction for forest Sort Prop.
tree_forest_ind, forest_tree_ind are recursively defined

Then we can build the combined induction principle which gives the conjunction of the conclusions of each individual principle:

Coq < Combined Scheme tree_forest_mutind from tree_forest_ind, forest_tree_ind.
tree_forest_mutind is recursively defined

The type of tree_forest_mutrec will be:

Coq < Check tree_forest_mutind.
tree_forest_mutind
     : forall (P : tree -> Prop) (P0 : forest -> Prop),
       (forall (a : A) (f : forest), P0 f -> P (node a f)) ->
       (forall b : B, P0 (leaf b)) ->
       (forall t : tree, P t -> forall f1 : forest, P0 f1 -> P0 (cons t f1)) ->
       (forall t : tree, P t) /\ (forall f2 : forest, P0 f2)

10.4  Functional Scheme and functional induction


Example 1: Induction scheme for div2

We define the function div2 as follows:

Coq < Require Import Arith.

Coq < Fixpoint div2 (n:nat) : nat :=
Coq <   match n with
Coq <   | O => 0
Coq <   | S O => 0
Coq <   | S (S n’) => S (div2 n’)
Coq <   end.

The definition of a principle of induction corresponding to the recursive structure of div2 is defined by the command:

Coq < Functional Scheme div2_ind := Induction for div2 Sort Prop.
div2_equation is defined
div2_ind is defined

You may now look at the type of div2_ind:

Coq < Check div2_ind.
div2_ind
     : forall P : nat -> nat -> Prop,
       (forall n : nat, n = 0 -> P 0 0) ->
       (forall n n0 : nat, n = S n0 -> n0 = 0 -> P 1 0) ->
       (forall n n0 : nat,
        n = S n0 ->
        forall n’ : nat,
        n0 = S n’ -> P n’ (div2 n’) -> P (S (S n’)) (S (div2 n’))) ->
       forall n : nat, P n (div2 n)

We can now prove the following lemma using this principle:

Coq < Lemma div2_le’ : forall n:nat, div2 n <= n.

Coq < intro n.

Coq <  pattern n , (div2 n).
Coq < apply div2_ind; intros.
3 subgoals
  
  n : nat
  n0 : nat
  e : n0 = 0
  ============================
   0 <= 0
subgoal 2 is:
 0 <= 1
subgoal 3 is:
 S (div2 n’) <= S (S n’)
Coq < auto with arith.

Coq < auto with arith.

Coq < simpl; auto with arith.

Coq < Qed.

We can use directly the functional induction (8.7.7) tactic instead of the pattern/apply trick:

Coq < Reset div2_le’.

Coq < Lemma div2_le : forall n:nat, div2 n <= n.

Coq < intro n.
Coq < functional induction (div2 n).
3 subgoals
  
  ============================
   0 <= 0
subgoal 2 is:
 0 <= 1
subgoal 3 is:
 S (div2 n’) <= S (S n’)
Coq < auto with arith.

Coq < auto with arith.

Coq < auto with arith.

Coq < Qed.


Remark: There is a difference between obtaining an induction scheme for a function by using Function (see Section 2.3) and by using Functional Scheme after a normal definition using Fixpoint or Definition. See 2.3 for details.


Example 2: Induction scheme for tree_size

We define trees by the following mutual inductive type:

Coq < Variable A : Set.

Coq < Inductive tree : Set :=
Coq <     node : A -> forest -> tree
Coq < with forest : Set :=
Coq <   | empty : forest
Coq <   | cons : tree -> forest -> forest.

We define the function tree_size that computes the size of a tree or a forest. Note that we use Function which generally produces better principles.

Coq < Function tree_size (t:tree) : nat :=
Coq <   match t with
Coq <   | node A f => S (forest_size f)
Coq <   end
Coq <  with forest_size (f:forest) : nat :=
Coq <   match f with
Coq <   | empty => 0
Coq <   | cons t f’ => (tree_size t + forest_size f’)
Coq <   end.

Remark: Function generates itself non mutual induction principles tree_size_ind and forest_size_ind:

Coq < Check tree_size_ind.
tree_size_ind
     : forall P : tree -> nat -> Prop,
       (forall (t : tree) (A : A) (f : forest),
        t = node A f -> P (node A f) (S (forest_size f))) ->
       forall t : tree, P t (tree_size t)

The definition of mutual induction principles following the recursive structure of tree_size and forest_size is defined by the command:

Coq < Functional Scheme tree_size_ind2 := Induction for tree_size Sort Prop
Coq < with forest_size_ind2 := Induction for forest_size Sort Prop.

You may now look at the type of tree_size_ind2:

Coq < Check tree_size_ind2.
tree_size_ind2
     : forall (P : tree -> nat -> Prop) (P0 : forest -> nat -> Prop),
       (forall (t : tree) (A : A) (f : forest),
        t = node A f ->
        P0 f (forest_size f) -> P (node A f) (S (forest_size f))) ->
       (forall f0 : forest, f0 = empty -> P0 empty 0) ->
       (forall (f1 : forest) (t : tree) (f’ : forest),
        f1 = cons t f’ ->
        P t (tree_size t) ->
        P0 f’ (forest_size f’) ->
        P0 (cons t f’) (tree_size t + forest_size f’)) ->
       forall t : tree, P t (tree_size t)

10.5  inversion

Generalities about inversion

When working with (co)inductive predicates, we are very often faced to some of these situations:

The inversion tactics are very useful to simplify the work in these cases. Inversion tools can be classified in three groups:

  1. tactics for inverting an instance without stocking the inversion lemma in the context; this includes the tactics (dependent) inversion and (dependent) inversion_clear.
  2. commands for generating and stocking in the context the inversion lemma corresponding to an instance; this includes Derive (Dependent) Inversion and Derive (Dependent) Inversion_clear.
  3. tactics for inverting an instance using an already defined inversion lemma; this includes the tactic inversion …using.

As inversion proofs may be large in size, we recommend the user to stock the lemmas whenever the same instance needs to be inverted several times.


Example 1: Non-dependent inversion

Let’s consider the relation Le over natural numbers and the following variables:

Coq < Inductive Le : nat -> nat -> Set :=
Coq <   | LeO : forall n:nat, Le 0 n
Coq <   | LeS : forall n m:nat, Le n m -> Le (S n) (S m).

Coq < Variable P : nat -> nat -> Prop.

Coq < Variable Q : forall n m:nat, Le n m -> Prop.

For example, consider the goal:

Coq < Show.
1 subgoal
  
  n : nat
  m : nat
  H : Le (S n) m
  ============================
   P n m

To prove the goal we may need to reason by cases on H and to derive that m is necessarily of the form (S m0) for certain m0 and that (Le n m0). Deriving these conditions corresponds to prove that the only possible constructor of (Le (S n) m) is LeS and that we can invert the -> in the type of LeS. This inversion is possible because Le is the smallest set closed by the constructors LeO and LeS.

Coq < inversion_clear H.
1 subgoal
  
  n : nat
  m : nat
  m0 : nat
  H0 : Le n m0
  ============================
   P n (S m0)

Note that m has been substituted in the goal for (S m0) and that the hypothesis (Le n m0) has been added to the context.

Sometimes it is interesting to have the equality m=(S m0) in the context to use it after. In that case we can use inversion that does not clear the equalities:

Coq < Undo.
Coq < inversion H.
1 subgoal
  
  n : nat
  m : nat
  H : Le (S n) m
  n0 : nat
  m0 : nat
  H1 : Le n m0
  H0 : n0 = n
  H2 : S m0 = m
  ============================
   P n (S m0)


Example 2: Dependent Inversion

Let us consider the following goal:

Coq < Show.
1 subgoal
  
  n : nat
  m : nat
  H : Le (S n) m
  ============================
   Q (S n) m H

As H occurs in the goal, we may want to reason by cases on its structure and so, we would like inversion tactics to substitute H by the corresponding term in constructor form. Neither Inversion nor Inversion_clear make such a substitution. To have such a behavior we use the dependent inversion tactics:

Coq < dependent inversion_clear H.
1 subgoal
  
  n : nat
  m : nat
  m0 : nat
  l : Le n m0
  ============================
   Q (S n) (S m0) (LeS n m0 l)

Note that H has been substituted by (LeS n m0 l) and m by (S m0).


Example 3: using already defined inversion lemmas

For example, to generate the inversion lemma for the instance (Le (S n) m) and the sort Prop we do:

Coq < Derive Inversion_clear leminv with (forall n m:nat, Le (S n) m) Sort
Coq <  Prop.
Coq < Check leminv.
leminv
     : forall (n m : nat) (P : nat -> nat -> Prop),
       (forall m0 : nat, Le n m0 -> P n (S m0)) -> Le (S n) m -> P n m

Then we can use the proven inversion lemma:

Coq < Show.
1 subgoal
  
  n : nat
  m : nat
  H : Le (S n) m
  ============================
   P n m
Coq < inversion H using leminv.
1 subgoal
  
  n : nat
  m : nat
  H : Le (S n) m
  ============================
   forall m0 : nat, Le n m0 -> P n (S m0)

10.6  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 an 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 above:

Coq < Require Import Coq.Logic.JMeq.
Coq < Goal forall n m:nat, Le (S n) m -> P n m.

Coq < intros n m H.

Coq < generalize_eqs H.
1 subgoal
  
  n : nat
  m : nat
  gen_x : nat
  H : Le gen_x m
  ============================
   gen_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 ammenable 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:

Coq < generalize_eqs_vars H.
1 subgoal
  
  m : nat
  gen_x : nat
  H : Le gen_x m
  ============================
   forall n : nat, gen_x = S n -> P n m

Coq < induction H.
2 subgoals
  
  n : nat
  ============================
   forall n0 : nat, 0 = S n0 -> P n0 n
subgoal 2 is:
 forall n0 : nat, S n = S n0 -> P n0 (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 a well in this case, e.g.:

Coq < Goal forall n m (p : Le (S n) m), Q (S n) m p.
1 subgoal
  
  ============================
   forall (n m : nat) (p : Le (S n) m), Q (S n) m p

Coq < intros n m p ; generalize_eqs_vars p.
1 subgoal
  
  m : nat
  gen_x : nat
  p : Le gen_x m
  ============================
   forall (n : nat) (p0 : Le (S n) m),
   gen_x = S n -> [p : (Le gen_x m)]= [p0 : (Le (S n) m)] -> 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 indices. 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 K axiom). This is what the simplify_dep_elim tactic from Coq.Program.Equality does. For example, we might simplify the previous goals considerably:

Coq < induction p ; simplify_dep_elim.
1 subgoal
  
  n0 : nat
  m : nat
  p : Le n0 m
  IHp : forall (n : nat) (p0 : Le (S n) m),
        n0 = S n -> [p : (Le n0 m)]= [p0 : (Le (S n) m)] -> Q (S n) m p0
  ============================
   Q (S n0) (S m) (LeS n0 m p)

The higher-order tactic do_depind defined in Coq.Program.Equality takes a tactic and combines the building blocks we’ve 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 :

Coq < Require Import Coq.Program.Equality.

Coq < Lemma ex : forall n m:nat, Le (S n) m -> P n m.

Coq < intros n m H.

Coq < dependent destruction H.
1 subgoal
  
  n : nat
  m : nat
  H : Le n m
  ============================
   P 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:

Coq < Require Import Coq.Program.Equality.

Coq < Set Implicit Arguments.

Coq < Variable A : Set.

Coq < Inductive vector : nat -> Type := 
Coq < | vnil : vector 0 
Coq < | vcons : A -> forall n, vector n -> vector (S n).

Coq < Goal forall n, forall v : vector (S n), 
Coq <   exists v’ : vector n, exists a : A, v = vcons a v’.

Coq <   intros n v.

Coq <   dependent destruction v.
1 subgoal
  
  n : nat
  a : A
  v : vector n
  ============================
   exists v’ : vector n, exists 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.

10.6.1  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:

Coq < Inductive type : Type :=
Coq < | base : type
Coq < | arrow : type -> type -> type.

Coq < Notation " t --> t’ " := (arrow t t’) (at level 20, t’ at next level).

Coq < Inductive ctx : Type :=
Coq < | empty : ctx
Coq < | snoc : ctx -> type -> ctx.

Coq < Notation " G , tau " := (snoc G tau) (at level 20, t at next level).

Coq < Fixpoint conc (G D : ctx) : ctx :=
Coq <   match D with
Coq <     | empty => G
Coq <     | snoc D’ x => snoc (conc G D’) x
Coq <   end.

Coq < Notation " G ; D " := (conc G D) (at level 20).

Coq < Inductive term : ctx -> type -> Type :=
Coq < | ax : forall G tau, term (G, tau) tau
Coq < | weak : forall G tau, 
Coq <   term G tau -> forall tau’, term (G, tau’) tau
Coq < | abs : forall G tau tau’, 
Coq <   term (G , tau) tau’ -> term G (tau --> tau’)
Coq < | app : forall G tau tau’, 
Coq <   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:

Once we have this datatype we want to do proofs on it, like weakening:

Coq < Lemma weakening : forall G D tau, term (G ; D) tau -> 
Coq <   forall tau’, 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:

Coq < Lemma weakening’ : forall G’ tau, term G’ tau -> 
Coq <   forall G D, (G ; D) = G’ ->
Coq <   forall tau’, 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 aleviates this trouble by doing all of this plumbing of generalizing and substituting back automatically. Indeed we can simply write:

Coq < Require Import Coq.Program.Tactics.

Coq < Lemma weakening : forall G D tau, term (G ; D) tau -> 
Coq <   forall tau’, term (G , tau’ ; D) tau.

Coq < Proof with simpl in * ; simpl_depind ; auto.

Coq <   intros G D tau H. dependent induction H generalizing G D ; intros.

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 explicitely.

Coq <   Show.
4 subgoals
  
  G : ctx
  tau : type
  G0 : ctx
  D : ctx
  H : G, tau = G0; D
  tau’ : type
  ============================
   term ((G0, tau’); D) tau
subgoal 2 is:
 term ((G0, tau’0); D) tau
subgoal 3 is:
 term ((G0, tau’0); D) (tau –> tau’)
subgoal 4 is:
 term ((G0, tau’0); D) tau’

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 reflexiviy. In cases where the equality is not between constructor forms though, one must help the automation by giving some arguments, using the specialize tactic.

Coq < destruct D... apply weak ; apply ax. apply ax.

Coq < destruct D...

Coq < Show.
4 subgoals
  
  G : ctx
  tau : type
  H : term G tau
  tau’ : type
  IHterm : forall G0 D : ctx,
           G = G0; D -> forall tau’ : type, term ((G0, tau’); D) tau
  tau’0 : type
  ============================
   term ((G, tau’), tau’0) tau
subgoal 2 is:
 term (((G0, tau’0); D), t) tau
subgoal 3 is:
 term ((G0, tau’0); D) (tau –> tau’)
subgoal 4 is:
 term ((G0, tau’0); D) tau’

Coq <   specialize (IHterm G empty).
4 subgoals
  
  G : ctx
  tau : type
  H : term G tau
  tau’ : type
  IHterm : G = G; empty -> forall tau’ : type, term ((G, tau’); empty) tau
  tau’0 : type
  ============================
   term ((G, tau’), tau’0) tau
subgoal 2 is:
 term (((G0, tau’0); D), t) tau
subgoal 3 is:
 term ((G0, tau’0); D) (tau –> tau’)
subgoal 4 is:
 term ((G0, tau’0); D) tau’

Then the automation can find the needed equality G = G to narrow the induction hypothesis further. This concludes our example.

Coq <   simpl_depind.
4 subgoals
  
  G : ctx
  tau : type
  H : term G tau
  tau’ : type
  tau’0 : type
  IHterm : forall tau’ : type, term ((G, tau’); empty) tau
  ============================
   term ((G, tau’), tau’0) tau
subgoal 2 is:
 term (((G0, tau’0); D), t) tau
subgoal 3 is:
 term ((G0, tau’0); D) (tau –> tau’)
subgoal 4 is:
 term ((G0, tau’0); D) tau’


See also: The induction 11, case 9 and inversion 8.10 tactics.

10.7  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 1: Ackermann function

Coq < Require Import Arith.

Coq < Variable Ack : 
Coq <            nat -> nat -> nat.

Coq < Axiom Ack0 : 
Coq <         forall m:nat, Ack 0 m = S m.

Coq < Axiom Ack1 : forall n:nat, Ack (S n) 0 = Ack n 1.

Coq < Axiom Ack2 : forall n m:nat, Ack (S n) (S m) = Ack n (Ack (S n) m).
Coq < Hint Rewrite Ack0 Ack1 Ack2 : base0.
Toplevel input, characters 13-17:
> Hint Rewrite Ack0 Ack1 Ack2 : base0.
>              ^^^^
Error: The reference Ack0 was not found in the current environment.

Coq < Lemma ResAck0 : 
Coq <  Ack 3 2 = 29.
Toplevel input, characters 18-21:
>  Ack 3 2 = 29.
>  ^^^
Error: The reference Ack was not found in the current environment.

Coq < autorewrite with base0 using try reflexivity.
Toplevel input, characters 0-44:
> autorewrite with base0 using try reflexivity.
> ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
Error: Rewriting base base0 does not exist.


Example 2: Mac Carthy function

Coq < Require Import Omega.

Coq < Variable g :   
Coq <            nat -> nat -> nat.

Coq < Axiom g0 : 
Coq <         forall m:nat, g 0 m = m.

Coq < Axiom
Coq <   g1 :
Coq <     forall n m:nat,
Coq <       (n > 0) -> (m > 100) -> g n m = g (pred n) (m - 10).

Coq < Axiom
Coq <   g2 :
Coq <     forall n m:nat,
Coq <       (n > 0) -> (m <= 100) -> g n m = g (S n) (m + 11).
Coq < Hint Rewrite g0 g1 g2 using omega : base1.

Coq < Lemma Resg0 : 
Coq <  g 1 110 = 100.
1 subgoal
  
  ============================
   g 1 110 = 100

Coq < autorewrite with base1 using reflexivity || simpl.
Proof completed.
Coq < Lemma Resg1 : g 1 95 = 91.
1 subgoal
  
  ============================
   g 1 95 = 91

Coq < autorewrite with base1 using reflexivity || simpl.
Proof completed.

10.8  quote

The tactic quote allows to use Barendregt’s so-called 2-level approach without writing any ML code. Suppose you have a language L of ’abstract terms’ and a type A of ’concrete terms’ and a function f : L -> A. If L is a simple inductive datatype and f a simple fixpoint, quote f will replace the head of current goal by a convertible term of the form (f t). L must have a constructor of type: A -> L.

Here is an example:

Coq < Require Import Quote.

Coq < Parameters A B C : Prop.
A is assumed
B is assumed
C is assumed

Coq < Inductive formula : Type :=
Coq <   | f_and : formula -> formula -> formula (* binary constructor *)
Coq <   | f_or : formula -> formula -> formula
Coq <   | f_not : formula -> formula (* unary constructor *)
Coq <   | f_true : formula (* 0-ary constructor *)
Coq <   | f_const : Prop -> formula (* contructor for constants *).
formula is defined
formula_rect is defined
formula_ind is defined
formula_rec is defined

Coq < Fixpoint interp_f (f:
Coq <                    formula) : Prop :=
Coq <   match f with
Coq <   | f_and f1 f2 => interp_f f1 /\ interp_f f2
Coq <   | f_or f1 f2 => interp_f f1 \/ interp_f f2
Coq <   | f_not f1 => ~ interp_f f1
Coq <   | f_true => True
Coq <   | f_const c => c
Coq <   end.
interp_f is recursively defined (decreasing on 1st argument)

Coq < Goal A /\ (A \/ True) /\ ~ B /\ (A <-> A).
1 subgoal
  
  ============================
   A /\ (A \/ True) /\ ~ B /\ (A <-> A)

Coq < quote interp_f.
1 subgoal
  
  ============================
   interp_f
     (f_and (f_const A)
        (f_and (f_or (f_const A) f_true)
           (f_and (f_not (f_const B)) (f_const (A <-> A)))))

The algorithm to perform this inversion is: try to match the term with right-hand sides expression of f. If there is a match, apply the corresponding left-hand side and call yourself recursively on sub-terms. If there is no match, we are at a leaf: return the corresponding constructor (here f_const) applied to the term.


Error messages:

  1. quote: not a simple fixpoint
    Happens when quote is not able to perform inversion properly.

10.8.1  Introducing variables map

The normal use of quote is to make proofs by reflection: one defines a function simplify : formula -> formula and proves a theorem simplify_ok: (f:formula)(interp_f (simplify f)) -> (interp_f f). Then, one can simplify formulas by doing:

   quote interp_f.
   apply simplify_ok.
   compute.

But there is a problem with leafs: in the example above one cannot write a function that implements, for example, the logical simplifications AAA or A ∧ ¬ AFalse. This is because the Prop is impredicative.

It is better to use that type of formulas:

Coq < Inductive formula : Set :=
Coq <   | f_and : formula -> formula -> formula
Coq <   | f_or : formula -> formula -> formula
Coq <   | f_not : formula -> formula
Coq <   | f_true : formula
Coq <   | f_atom : index -> formula.
formula is defined
formula_rect is defined
formula_ind is defined
formula_rec is defined

index is defined in module quote. Equality on that type is decidable so we are able to simplify AA into A at the abstract level.

When there are variables, there are bindings, and quote provides also a type (varmap A) of bindings from index to any set A, and a function varmap_find to search in such maps. The interpretation function has now another argument, a variables map:

Coq < Fixpoint interp_f (vm:
Coq <                     varmap Prop) (f:formula) {struct f} : Prop :=
Coq <   match f with
Coq <   | f_and f1 f2 => interp_f vm f1 /\ interp_f vm f2
Coq <   | f_or f1 f2 => interp_f vm f1 \/ interp_f vm f2
Coq <   | f_not f1 => ~ interp_f vm f1
Coq <   | f_true => True
Coq <   | f_atom i => varmap_find True i vm
Coq <   end.
interp_f is recursively defined (decreasing on 2nd argument)

quote handles this second case properly:

Coq < Goal A /\ (B \/ A) /\ (A \~ B).
1 subgoal
  
  ============================
   A /\ (B \/ A) /\ (A \~ B)

Coq < quote interp_f.
1 subgoal
  
  ============================
   interp_f
     (Node_vm B (Node_vm A (Empty_vm Prop) (Empty_vm Prop)) (Empty_vm Prop))
     (f_and (f_atom (Left_idx End_idx))
        (f_and (f_or (f_atom End_idx) (f_atom (Left_idx End_idx)))
           (f_or (f_atom (Left_idx End_idx)) (f_not (f_atom End_idx)))))

It builds vm and t such that (f vm t) is convertible with the conclusion of current goal.

10.8.2  Combining variables and constants

One can have both variables and constants in abstracts terms; that is the case, for example, for the ring tactic (chapter 23). Then one must provide to quote a list of constructors of constants. For example, if the list is [O S] then closed natural numbers will be considered as constants and other terms as variables.

Example:

Coq < Inductive formula : Type :=
Coq <   | f_and : formula -> formula -> formula
Coq <   | f_or : formula -> formula -> formula
Coq <   | f_not : formula -> formula
Coq <   | f_true : formula
Coq <   | f_const : Prop -> formula (* constructor for constants *)
Coq <   | f_atom : index -> formula.

Coq < Fixpoint interp_f
Coq <  (vm:            (* constructor for variables *)
Coq <   varmap Prop) (f:formula) {struct f} : Prop :=
Coq <   match f with
Coq <   | f_and f1 f2 => interp_f vm f1 /\ interp_f vm f2
Coq <   | f_or f1 f2 => interp_f vm f1 \/ interp_f vm f2
Coq <   | f_not f1 => ~ interp_f vm f1
Coq <   | f_true => True
Coq <   | f_const c => c
Coq <   | f_atom i => varmap_find True i vm
Coq <   end.

Coq < Goal 
Coq < A /\ (A \/ True) /\ ~ B /\ (C <-> C).
Coq < quote interp_f [ A B ].
1 subgoal
  
  ============================
   interp_f (Node_vm (C <-> C) (Empty_vm Prop) (Empty_vm Prop))
     (f_and (f_const A)
        (f_and (f_or (f_const A) f_true)
           (f_and (f_not (f_const B)) (f_atom End_idx))))

Coq < Undo.
1 subgoal
  
  ============================
   A /\ (A \/ True) /\ ~ B /\ (C <-> C)

Coq <   quote interp_f [ B C iff ].
1 subgoal
  
  ============================
   interp_f (Node_vm A (Empty_vm Prop) (Empty_vm Prop))
     (f_and (f_atom End_idx)
        (f_and (f_or (f_atom End_idx) f_true)
           (f_and (f_not (f_const B)) (f_const (C <-> C)))))


Warning: Since function inversion is undecidable in general case, don’t expect miracles from it!


See also: comments of source file tactics/contrib/polynom/quote.ml


See also: the ring tactic (Chapter 23)

10.9  Using the tactical language

10.9.1  About the cardinality of the set of natural numbers

A first example which shows how to use the pattern matching over the proof contexts is the proof that natural numbers have more than two elements. The proof of such a lemma can be done as follows:

Coq < Lemma card_nat :
Coq <  ~ (exists x : nat, exists y : nat, forall z:nat, x = z \/ y = z).

Coq < Proof.

Coq < red; intros (x, (y, Hy)).

Coq < elim (Hy 0); elim (Hy 1); elim (Hy 2); intros;
Coq <  match goal with
Coq <  | [_:(?a = ?b),_:(?a = ?c) |- _ ] =>
Coq <      cut (b = c); [ discriminate | apply trans_equal with a; auto ]
Coq <  end.

Coq < Qed.

We can notice that all the (very similar) cases coming from the three eliminations (with three distinct natural numbers) are successfully solved by a match goal structure and, in particular, with only one pattern (use of non-linear matching).

10.9.2  Permutation on closed lists

Another more complex example is the problem of permutation on closed lists. The aim is to show that a closed list is a permutation of another one.

First, we define the permutation predicate as shown in table 10.1.


Coq < Section Sort.

Coq < Variable A : Set.

Coq < Inductive permut : list A -> list A -> Prop :=
Coq <   | permut_refl   : forall l, permut l l
Coq <   | permut_cons   :
Coq <       forall a l0 l1, permut l0 l1 -> permut (a :: l0) (a :: l1)
Coq <   | permut_append : forall a l, permut (a :: l) (l ++ a :: nil)
Coq <   | permut_trans  :
Coq <       forall l0 l1 l2, permut l0 l1 -> permut l1 l2 -> permut l0 l2.

Coq < End Sort.
Figure 10.1: Definition of the permutation predicate

A more complex example is the problem of permutation on closed lists. The aim is to show that a closed list is a permutation of another one. First, we define the permutation predicate as shown on Figure 10.1.


Coq < Ltac Permut n :=
Coq <   match goal with
Coq <   | |- (permut _ ?l ?l) => apply permut_refl
Coq <   | |- (permut _ (?a :: ?l1) (?a :: ?l2)) =>
Coq <       let newn := eval compute in (length l1) in
Coq <       (apply permut_cons; Permut newn)
Coq <   | |- (permut ?A (?a :: ?l1) ?l2) =>
Coq <       match eval compute in n with
Coq <       | 1 => fail
Coq <       | _ =>
Coq <           let l1’ := constr:(l1 ++ a :: nil) in
Coq <           (apply (permut_trans A (a :: l1) l1’ l2);
Coq <             [ apply permut_append | compute; Permut (pred n) ])
Coq <       end
Coq <   end.
Permut is defined

Coq < Ltac PermutProve :=
Coq <   match goal with
Coq <   | |- (permut _ ?l1 ?l2) =>
Coq <       match eval compute in (length l1 = length l2) with
Coq <       | (?n = ?n) => Permut n
Coq <       end
Coq <   end.
PermutProve is defined
Figure 10.2: Permutation tactic

Next, we can write naturally the tactic and the result can be seen on Figure 10.2. We can notice that we use two toplevel definitions PermutProve and Permut. The function to be called is PermutProve which computes the lengths of the two lists and calls Permut with the length if the two lists have the same length. Permut works as expected. If the two lists are equal, it concludes. Otherwise, if the lists have identical first elements, it applies Permut on the tail of the lists. Finally, if the lists have different first elements, it puts the first element of one of the lists (here the second one which appears in the permut predicate) at the end if that is possible, i.e., if the new first element has been at this place previously. To verify that all rotations have been done for a list, we use the length of the list as an argument for Permut and this length is decremented for each rotation down to, but not including, 1 because for a list of length n, we can make exactly n−1 rotations to generate at most n distinct lists. Here, it must be noticed that we use the natural numbers of Coq for the rotation counter. On Figure 9.1, we can see that it is possible to use usual natural numbers but they are only used as arguments for primitive tactics and they cannot be handled, in particular, we cannot make computations with them. So, a natural choice is to use Coq data structures so that Coq makes the computations (reductions) by eval compute in and we can get the terms back by match.

With PermutProve, we can now prove lemmas as follows:

Coq < Lemma permut_ex1 :
Coq <   permut nat (1 :: 2 :: 3 :: nil) (3 :: 2 :: 1 :: nil).

Coq < Proof. PermutProve. Qed.

Coq < Lemma permut_ex2 :
Coq <   permut nat
Coq <     (0 :: 1 :: 2 :: 3 :: 4 :: 5 :: 6 :: 7 :: 8 :: 9 :: nil)
Coq <     (0 :: 2 :: 4 :: 6 :: 8 :: 9 :: 7 :: 5 :: 3 :: 1 :: nil).

Coq < Proof. PermutProve. Qed.

10.9.3  Deciding intuitionistic propositional logic


Coq < Ltac Axioms :=
Coq <   match goal with
Coq <   | |- True => trivial
Coq <   | _:False |- _  => elimtype False; assumption
Coq <   | _:?A |- ?A  => auto
Coq <   end.
Axioms is defined
Figure 10.3: Deciding intuitionistic propositions (1)


Coq < Ltac DSimplif :=
Coq <   repeat
Coq <    (intros;
Coq <     match goal with
Coq <      | id:(~ _) |- _ => red in id
Coq <      | id:(_ /\ _) |- _ =>
Coq <          elim id; do 2 intro; clear id
Coq <      | id:(_ \/ _) |- _ =>
Coq <          elim id; intro; clear id
Coq <      | id:(?A /\ ?B -> ?C) |- _ =>
Coq <          cut (A -> B -> C);
Coq <           [ intro | intros; apply id; split; assumption ]
Coq <      | id:(?A \/ ?B -> ?C) |- _ =>
Coq <          cut (B -> C);
Coq <           [ cut (A -> C);
Coq <              [ intros; clear id
Coq <              | intro; apply id; left; assumption ]
Coq <           | intro; apply id; right; assumption ]
Coq <      | id0:(?A -> ?B),id1:?A |- _ =>
Coq <          cut B; [ intro; clear id0 | apply id0; assumption ]
Coq <      | |- (_ /\ _) => split
Coq <      | |- (~ _) => red
Coq <      end).
DSimplif is defined

Coq < Ltac TautoProp :=
Coq <   DSimplif;
Coq <    Axioms ||
Coq <      match goal with
Coq <      | id:((?A -> ?B) -> ?C) |- _ =>
Coq <           cut (B -> C);
Coq <           [ intro; cut (A -> B);
Coq <              [ intro; cut C;
Coq <                 [ intro; clear id | apply id; assumption ]
Coq <              | clear id ]
Coq <           | intro; apply id; intro; assumption ]; TautoProp
Coq <      | id:(~ ?A -> ?B) |- _ =>
Coq <          cut (False -> B);
Coq <           [ intro; cut (A -> False);
Coq <              [ intro; cut B;
Coq <                 [ intro; clear id | apply id; assumption ]
Coq <              | clear id ]
Coq <           | intro; apply id; red; intro; assumption ]; TautoProp
Coq <      | |- (_ \/ _) => (left; TautoProp) || (right; TautoProp)
Coq <      end.
TautoProp is defined
Figure 10.4: Deciding intuitionistic propositions (2)

The pattern matching on goals allows a complete and so a powerful backtracking when returning tactic values. An interesting application is the problem of deciding intuitionistic propositional logic. Considering the contraction-free sequent calculi LJT* of Roy Dyckhoff ([54]), it is quite natural to code such a tactic using the tactic language as shown on Figures 10.3 and 10.4. The tactic Axioms tries to conclude using usual axioms. The tactic DSimplif applies all the reversible rules of Dyckhoff’s system. Finally, the tactic TautoProp (the main tactic to be called) simplifies with DSimplif, tries to conclude with Axioms and tries several paths using the backtracking rules (one of the four Dyckhoff’s rules for the left implication to get rid of the contraction and the right or).

For example, with TautoProp, we can prove tautologies like those:

Coq < Lemma tauto_ex1 : forall A B:Prop, A /\ B -> A \/ B.

Coq < Proof. TautoProp. Qed.

Coq < Lemma tauto_ex2 :
Coq <    forall A B:Prop, (~ ~ B -> B) -> (A -> B) -> ~ ~ A -> B.

Coq < Proof. TautoProp. Qed.

10.9.4  Deciding type isomorphisms

A more tricky problem is to decide equalities between types and modulo isomorphisms. Here, we choose to use the isomorphisms of the simply typed λ-calculus with Cartesian product and unit type (see, for example, [43]). The axioms of this λ-calculus are given by table 10.5.


Coq < Open Scope type_scope.

Coq < Section Iso_axioms.

Coq < Variables A B C : Set.

Coq < Axiom Com : A * B = B * A.

Coq < Axiom Ass : A * (B * C) = A * B * C.

Coq < Axiom Cur : (A * B -> C) = (A -> B -> C).

Coq < Axiom Dis : (A -> B * C) = (A -> B) * (A -> C).

Coq < Axiom P_unit : A * unit = A.

Coq < Axiom AR_unit : (A -> unit) = unit.

Coq < Axiom AL_unit : (unit -> A) = A.

Coq < Lemma Cons : B = C -> A * B = A * C.

Coq < Proof.

Coq < intro Heq; rewrite Heq; apply refl_equal.

Coq < Qed.

Coq < End Iso_axioms.
Figure 10.5: Type isomorphism axioms

A more tricky problem is to decide equalities between types and modulo isomorphisms. Here, we choose to use the isomorphisms of the simply typed λ-calculus with Cartesian product and unit type (see, for example, [43]). The axioms of this λ-calculus are given on Figure 10.5.


Coq < Ltac DSimplif trm :=
Coq <   match trm with
Coq <   | (?A * ?B * ?C) =>
Coq <       rewrite <- (Ass A B C); try MainSimplif
Coq <   | (?A * ?B -> ?C) =>
Coq <       rewrite (Cur A B C); try MainSimplif
Coq <   | (?A -> ?B * ?C) =>
Coq <       rewrite (Dis A B C); try MainSimplif
Coq <   | (?A * unit) =>
Coq <       rewrite (P_unit A); try MainSimplif
Coq <   | (unit * ?B) =>
Coq <       rewrite (Com unit B); try MainSimplif
Coq <   | (?A -> unit) =>
Coq <       rewrite (AR_unit A); try MainSimplif
Coq <   | (unit -> ?B) =>
Coq <       rewrite (AL_unit B); try MainSimplif
Coq <   | (?A * ?B) =>
Coq <       (DSimplif A; try MainSimplif) || (DSimplif B; try MainSimplif)
Coq <   | (?A -> ?B) =>
Coq <       (DSimplif A; try MainSimplif) || (DSimplif B; try MainSimplif)
Coq <   end
Coq <  with MainSimplif :=
Coq <   match goal with
Coq <   | |- (?A = ?B) => try DSimplif A; try DSimplif B
Coq <   end.
DSimplif is defined
MainSimplif is defined

Coq < Ltac Length trm :=
Coq <   match trm with
Coq <   | (_ * ?B) => let succ := Length B in constr:(S succ)
Coq <   | _ => constr:1
Coq <   end.
Length is defined

Coq < Ltac assoc := repeat rewrite <- Ass.
assoc is defined
Figure 10.6: Type isomorphism tactic (1)


Coq < Ltac DoCompare n :=
Coq <   match goal with
Coq <   | [ |- (?A = ?A) ] => apply refl_equal
Coq <   | [ |- (?A * ?B = ?A * ?C) ] =>
Coq <       apply Cons; let newn := Length B in
Coq <                   DoCompare newn
Coq <   | [ |- (?A * ?B = ?C) ] =>
Coq <       match eval compute in n with
Coq <       | 1 => fail
Coq <       | _ =>
Coq <           pattern (A * B) at 1; rewrite Com; assoc; DoCompare (pred n)
Coq <       end
Coq <   end.
DoCompare is defined

Coq < Ltac CompareStruct :=
Coq <   match goal with
Coq <   | [ |- (?A = ?B) ] =>
Coq <       let l1 := Length A
Coq <       with l2 := Length B in
Coq <       match eval compute in (l1 = l2) with
Coq <       | (?n = ?n) => DoCompare n
Coq <       end
Coq <   end.
CompareStruct is defined

Coq < Ltac IsoProve := MainSimplif; CompareStruct.
IsoProve is defined
Figure 10.7: Type isomorphism tactic (2)

The tactic to judge equalities modulo this axiomatization can be written as shown on Figures 10.6 and 10.7. The algorithm is quite simple. Types are reduced using axioms that can be oriented (this done by MainSimplif). The normal forms are sequences of Cartesian products without Cartesian product in the left component. These normal forms are then compared modulo permutation of the components (this is done by CompareStruct). The main tactic to be called and realizing this algorithm is IsoProve.

Here are examples of what can be solved by IsoProve.

Coq < Lemma isos_ex1 : 
Coq <   forall A B:Set, A * unit * B = B * (unit * A).

Coq < Proof.

Coq < intros; IsoProve.

Coq < Qed.

Coq < 
Coq < Lemma isos_ex2 :
Coq <   forall A B C:Set,
Coq <     (A * unit -> B * (C * unit)) =
Coq <     (A * unit -> (C -> unit) * C) * (unit -> A -> B).

Coq < Proof.

Coq < intros; IsoProve.

Coq < Qed.