Proof schemes¶
Generation of induction principles with Scheme
¶
-
Command
Scheme ident :=? scheme_kind with ident :=? scheme_kind*
¶ - scheme_kind
::=
Equality for reference|
InductionMinimalityEliminationCase for reference Sort sort_familysort_family
::=
Set|
Prop|
SProp|
TypeA high-level tool for automatically generating (possibly mutual) induction principles for given types and sorts. Each
reference
is a different inductive type identifier belonging to the same package of mutual inductive definitions. The command generates theident
s as mutually recursive definitions. Each termident
proves a general principle of mutual induction for objects in typereference
.ident
The name of the scheme. If not provided, the scheme name will be determined automatically from the sorts involved.
Minimality for reference Sort sort_family
Defines a non-dependent elimination principle more natural for inductively defined relations.
Equality for reference
Tries to generate a Boolean equality and a proof of the decidability of the usual equality. If
reference
involves other inductive types, their equality has to be defined first.
Example
Induction scheme for tree and forest.
A mutual induction principle for tree and forest in sort
Set
can be defined using the command
- Axiom A : Set.
- A is declared
- Axiom B : Set.
- B is declared
- Inductive tree : Set := node : A -> forest -> tree with forest : Set := leaf : B -> forest | cons : tree -> forest -> forest.
- tree, forest are defined tree_rect is defined tree_ind is defined tree_rec is defined tree_sind is defined forest_rect is defined forest_ind is defined forest_rec is defined forest_sind is defined
- Scheme tree_forest_rec := Induction for tree Sort Set with forest_tree_rec := Induction for forest Sort Set.
- forest_tree_rec is defined tree_forest_rec is defined tree_forest_rec, forest_tree_rec are recursively defined
You may now look at the type of tree_forest_rec:
- 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 andforests; 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.
Example
Predicates odd and even on naturals.
Let odd and even be inductively defined as:
- Inductive odd : nat -> Prop := oddS : forall n:nat, even n -> odd (S n) with even : nat -> Prop := | evenO : even 0 | evenS : forall n:nat, odd n -> even (S n).
- odd, even are defined odd_ind is defined odd_sind is defined even_ind is defined even_sind is defined
The following command generates a powerful elimination principle:
- Scheme odd_even := Minimality for odd Sort Prop with even_odd := Minimality for even Sort Prop.
- even_odd is defined odd_even is defined odd_even, even_odd are recursively defined
The type of odd_even for instance will be:
- 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)->(P0 n)
.
Automatic declaration of schemes¶
-
Flag
Elimination Schemes
¶ Enables automatic declaration of induction principles when defining a new inductive type. Defaults to on.
-
Flag
Nonrecursive Elimination Schemes
¶ Enables automatic declaration of induction principles for types declared with the
Variant
andRecord
commands. Defaults to off.
-
Flag
Case Analysis Schemes
¶ This flag governs the generation of case analysis lemmas for inductive types, i.e. corresponding to the pattern matching term alone and without fixpoint.
-
Flag
Boolean Equality Schemes
¶ -
Flag
Decidable Equality Schemes
¶ These flags control the automatic declaration of those Boolean equalities (see the second variant of
Scheme
).
Warning
You have to be careful with these flags since Coq may now reject well-defined inductive types because it cannot compute a Boolean equality for them.
-
Flag
Rewriting Schemes
¶ This flag governs generation of equality-related schemes such as congruence.
Combined Scheme¶
-
Command
Combined Scheme identdef from ident+,
¶ This command is a tool for combining induction principles generated by the
Scheme
command. Eachident
is a different inductive principle that must belong to the same package of mutual inductive principle definitions. This command generatesidentdef
as the conjunction of the principles: it is built from the common premises of the principles and concluded by the conjunction of their conclusions. In the case where all the inductive principles used are in sortProp
, the propositional conjunctionand
is used, otherwise the simple productprod
is used instead.
Example
We can define the induction principles for trees and forests using:
- Scheme tree_forest_ind := Induction for tree Sort Prop with forest_tree_ind := Induction for forest Sort Prop.
- forest_tree_ind is defined tree_forest_ind is defined 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:
- Combined Scheme tree_forest_mutind from tree_forest_ind,forest_tree_ind.
- tree_forest_mutind is defined tree_forest_mutind is recursively defined
The type of tree_forest_mutind will be:
- 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)
Example
We can also combine schemes at sort
Type
:
- Scheme tree_forest_rect := Induction for tree Sort Type with forest_tree_rect := Induction for forest Sort Type.
- forest_tree_rect is defined tree_forest_rect is defined tree_forest_rect, forest_tree_rect are recursively defined
- Combined Scheme tree_forest_mutrect from tree_forest_rect, forest_tree_rect.
- tree_forest_mutrect is defined tree_forest_mutrect is recursively defined
- Check tree_forest_mutrect.
- tree_forest_mutrect : forall (P : tree -> Type) (P0 : forest -> Type), (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)
Generation of inversion principles with Derive
Inversion
¶
-
Command
Derive Inversion ident with one_term Sort sort_family?
¶ Generates an inversion lemma for the
inversion ... using ...
tactic.ident
is the name of the generated lemma.one_term
should be in the formqualid
or(forall binder+, qualid term)
wherequalid
is the name of an inductive predicate andbinder+
binds the variables occurring in the termterm
. The lemma is generated for the sortsort_family
corresponding toone_term
. Applying the lemma is equivalent to inverting the instance with theinversion
tactic.
-
Command
Derive Inversion_clear ident with one_term Sort sort_family?
¶ When applied, it is equivalent to having inverted the instance with the tactic inversion replaced by the tactic
inversion_clear
.
-
Command
Derive Dependent Inversion ident with one_term Sort sort_family
¶ When applied, it is equivalent to having inverted the instance with the tactic
dependent inversion
.
-
Command
Derive Dependent Inversion_clear ident with one_term Sort sort_family
¶ When applied, it is equivalent to having inverted the instance with the tactic
dependent inversion_clear
.
Example
Consider the relation Le
over natural numbers and the following
parameter P
:
- 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).
- Le is defined Le_rect is defined Le_ind is defined Le_rec is defined Le_sind is defined
- Parameter P : nat -> nat -> Prop.
- P is declared
To generate the inversion lemma for the instance (Le (S n) m)
and the
sort Prop
, we do:
- Derive Inversion_clear leminv with (forall n m:nat, Le (S n) m) Sort Prop.
- leminv is defined
- 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:
- Goal forall (n m : nat) (H : Le (S n) m), P n m.
- 1 subgoal ============================ forall n m : nat, Le (S n) m -> P n m
- intros.
- 1 subgoal n, m : nat H : Le (S n) m ============================ P n m
- Show.
- 1 subgoal n, m : nat H : Le (S n) m ============================ P n m
- inversion H using leminv.
- 1 subgoal n, m : nat H : Le (S n) m ============================ forall m0 : nat, Le n m0 -> P n (S m0)