Functional induction
Note
The functional induction (FunInd) plugin is legacy functionality. For new code and new projects, we recommend Equations, a more powerful plugin that provides most of FunInd's features. It can be installed through the Coq Platform. Refer to the Equations documentation to learn more. FunInd is not deprecated and not planned for removal yet because porting code from FunInd to Equations can be difficult (due to differences in the generated induction principles).
Note
The tactics described in this chapter require the Stdlib library.
Advanced recursive functions
The following command is available when the FunInd library has been loaded via From Stdlib Require Import FunInd:
- Command Function fix_definition with fix_definition*
- This command is a generalization of - Fixpoint. It is a wrapper for several ways of defining a function and other useful related objects, namely: an induction principle that reflects the recursive structure of the function (see- functional induction) and its fixpoint equality. This defines a function similar to those defined by- Fixpoint. As in- Fixpoint, the decreasing argument must be given (unless the function is not recursive), but it might not necessarily be structurally decreasing. Use the- fixannotclause to name the decreasing argument and to describe which kind of decreasing criteria to use to ensure termination of recursive calls.- Functionalso supports the- withclause to create mutually recursive definitions, however this feature is limited to structurally recursive functions (i.e. when- fixannotis a- structclause).- See - functional inductionand- Functional Schemefor how to use the induction principle to reason easily about the function.- The form of the - fixannotclause determines which definition mechanism- Functionuses. (Note that references to- identbelow refer to the name of the function being defined.):- If - fixannotis not specified,- Functiondefines the nonrecursive function- identas if it was declared with- Definition. In addition, the following are defined:
- If - { struct ... }is specified,- Functiondefines the structural recursive function- identas if it was declared with- Fixpoint. In addition, the following are defined:
- If - { measure ... }or- { wf ... }are specified,- Functiondefines a recursive function by well-founded recursion. The module- Recdefof the standard library must be loaded for this feature.- {measure one_term1 ident? one_term2? }: where- identis the decreasing argument and- one_term1is a function from the type of- identto- natfor which the decreasing argument decreases (for the- ltorder on- nat) for each recursive call of the function. The parameters of the function are bound in- one_term1.
- {wf one_term ident }: where- identis the decreasing argument and- one_termis an ordering relation on the type of- ident(i.e. of type- T\(_{\sf ident}\) →- T\(_{\sf ident}\) →- Prop) for which the decreasing argument decreases for each recursive call of the function. The order must be well-founded. The parameters of the function are bound in- one_term.
 - If the clause is - measureor- wf, the user is left with some proof obligations that will be used to define the function. These proofs are: proofs that each recursive call is actually decreasing with respect to the given criteria, and (if the criteria is- wf) a proof that the ordering relation is well-founded. Once proof obligations are discharged, the following objects are defined:- The way this recursive function is defined is the subject of several papers by Yves Bertot and Antonia Balaa on the one hand, and Gilles Barthe, Julien Forest, David Pichardie, and Vlad Rusu on the other hand. 
 
Note
To obtain the right principle, it is better to put rigid parameters of the function as first arguments. For example it is better to define plus like this:
- From Stdlib Require Import FunInd.
- [Loading ML file rocq-runtime.plugins.extraction ... done] [Loading ML file rocq-runtime.plugins.funind ... done]
- Function plus (m n : nat) {struct n} : nat := match n with | 0 => m | S p => S (plus m p) end.
- plus is defined plus is recursively defined (guarded on 2nd argument) plus_equation is defined plus_rect is defined plus_ind is defined plus_rec is defined R_plus_correct is defined R_plus_complete is defined
than like this:
- From Stdlib Require Import FunInd.
- Function plus (n m : nat) {struct n} : nat := match n with | 0 => m | S p => S (plus p m) end.
- plus is defined plus is recursively defined (guarded on 1st argument) plus_equation is defined plus_rect is defined plus_ind is defined plus_rec is defined R_plus_correct is defined R_plus_complete is defined
Limitations
term must be built as a pure pattern matching tree (match … with)
with applications only at the end of each branch.
Function does not support partial application of the function being
defined. Thus, the following example cannot be accepted due to the
presence of partial application of wrong in the body of wrong:
- From Stdlib Require List.
- Import List.ListNotations.
- Function wrong (C:nat) : nat := List.hd 0 (List.map wrong (C::nil)).
- Toplevel input, characters 0-70: > Function wrong (C:nat) : nat := List.hd 0 (List.map wrong (C::nil)). > ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ Error: Recursive definition of wrong is ill-formed. In environment wrong : nat -> nat C : nat A := nat : Type B := nat : Type l := [C]%list : list A Recursive call to wrong has principal argument equal to "C" instead of a subterm of "C". Recursive definition is: "fun C : nat => List.hd 0 (ListDef.map wrong [C])".
For now, dependent cases are not treated for non-structurally terminating functions.
- Error The recursive argument must be specified.
- Error Cannot use mutual definition with well-founded recursion or measure.
- Warning Cannot define graph for ident.
- The generation of the graph relation ( - R_ident) used to compute the induction scheme of ident raised a typing error. Only- identis defined; the induction scheme will not be generated. This error happens generally when:- the definition uses pattern matching on dependent types, which - Functioncannot deal with yet.
- the definition is not a pattern matching tree as explained above. 
 
- Warning Cannot define principle(s) for ident.
- The generation of the graph relation ( - R_ident) succeeded but the induction principle could not be built. Only- identis defined. Please report.
- Warning Cannot build functional inversion principle.
- functional inversionwill not be available for the function.
Tactics
- Tactic functional induction term using one_term_with_bindings? as simple_intropattern?
- Performs case analysis and induction following the definition of a function - qualid, which must be fully applied to its arguments as part of- term. It uses a principle generated by- Functionor- Functional Scheme. Note that this tactic is only available after a- From Stdlib Require Import FunInd. See the- Functioncommand.- using one_term
- Specifies the induction principle (aka elimination scheme). 
- with bindings
- Specifies the arguments of the induction principle. 
- as simple_intropattern
- Provides names for the introduced variables. 
 - Example - From Stdlib Require Import FunInd.
- Functional Scheme minus_ind := Induction for minus Sort Prop.
- sub_equation is defined minus_ind is defined
- Check minus_ind.
- minus_ind : forall P : nat -> nat -> nat -> Prop, (forall n m : nat, n = 0 -> P 0 m n) -> (forall n m k : nat, n = S k -> m = 0 -> P (S k) 0 n) -> (forall n m k : nat, n = S k -> forall l : nat, m = S l -> P k l (k - l) -> P (S k) (S l) (k - l)) -> forall n m : nat, P n m (n - m)
- Lemma le_minus (n m:nat) : n - m <= n.
- 1 goal n, m : nat ============================ n - m <= n
- functional induction (minus n m) using minus_ind; simpl; auto.
- No more goals.
- Qed.
 - Note - functional induction (f x1 x2 x3)is actually a wrapper for- induction x1, x2, x3, (f x1 x2 x3) using qualidfollowed by a cleaning phase, where- qualidis the induction principle registered for- f(by the- Functionor- Functional Schemecommand) corresponding to the sort of the goal. Therefore- functional inductionmay fail if the induction scheme- qualidis not defined.- Note - There is a difference between obtaining an induction scheme for a function by using - Functionand by using- Functional Schemeafter a normal definition using- Fixpointor- Definition.- Error Not the right number of induction arguments.
 
- Tactic soft functional induction one_term+ using one_term_with_bindings? as simple_intropattern?
- Tactic functional inversion identnatural qualid?
- Performs inversion on hypothesis - identof the form- qualid term+ = termor- term = qualid term+when- qualidis defined using- Function. Note that this tactic is only available after a- From Stdlib Require Import FunInd.- natural
- Does the same thing as - intros until naturalfollowed by- functional inversion identwhere- identis the identifier for the last introduced hypothesis.
- qualid
- If the hypothesis - ident(or- natural) has a type of the form- qualid1 termi+ = qualid2 termj+where- qualid1and- qualid2are valid candidates to functional inversion, this variant allows choosing which- qualidis inverted.
 
Generation of induction principles with Functional Scheme
- Command Functional Scheme func_scheme_def with func_scheme_def*
- func_scheme_def::=ident := Induction for qualid Sort sort_familyAn experimental high-level tool that automatically generates induction principles corresponding to functions that may be mutually recursive. The command generates an induction principle named identfor each given function namedqualid. Thequalids must be given in the same order as when they were defined.Note the command must be made available via From StdlibRequire ImportFunInd.
Warning
There is a difference between induction schemes generated by the command
Functional Scheme and these generated by the Function. Indeed,
Function generally produces smaller principles that are closer to how
a user would implement them. See Advanced recursive functions for details.
Example
Induction scheme for div2.
We define the function div2 as follows:
- From Stdlib Require Import FunInd.
- From Stdlib Require Import Arith.
- [Loading ML file rocq-runtime.plugins.ring ... done]
- Fixpoint div2 (n:nat) : nat := match n with | O => 0 | S O => 0 | S (S n') => S (div2 n') end.
- div2 is defined div2 is recursively defined (guarded on 1st argument)
The definition of a principle of induction corresponding to the
recursive structure of div2 is defined by the command:
- 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:
- 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:
- Lemma div2_le' : forall n:nat, div2 n <= n.
- 1 goal ============================ forall n : nat, div2 n <= n
- intro n.
- 1 goal n : nat ============================ div2 n <= n
- pattern n, (div2 n).
- 1 goal n : nat ============================ (fun n0 n1 : nat => n1 <= n0) n (div2 n)
- apply div2_ind; intros.
- 3 goals n, n0 : nat e : n0 = 0 ============================ 0 <= 0 goal 2 is: 0 <= 1 goal 3 is: S (div2 n') <= S (S n')
- auto with arith.
- 2 goals n, n0, n1 : nat e : n0 = S n1 e0 : n1 = 0 ============================ 0 <= 1 goal 2 is: S (div2 n') <= S (S n')
- auto with arith.
- 1 goal n, n0, n1 : nat e : n0 = S n1 n' : nat e0 : n1 = S n' H : div2 n' <= n' ============================ S (div2 n') <= S (S n')
- simpl; auto with arith.
- No more goals.
- Qed.
We can use directly the functional induction (functional induction) tactic instead
of the pattern/apply trick:
- Reset div2_le'.
- Lemma div2_le : forall n:nat, div2 n <= n.
- 1 goal ============================ forall n : nat, div2 n <= n
- intro n.
- 1 goal n : nat ============================ div2 n <= n
- functional induction (div2 n).
- 3 goals ============================ 0 <= 0 goal 2 is: 0 <= 1 goal 3 is: S (div2 n') <= S (S n')
- auto with arith.
- 2 goals ============================ 0 <= 1 goal 2 is: S (div2 n') <= S (S n')
- auto with arith.
- 1 goal n' : nat IHn0 : div2 n' <= n' ============================ S (div2 n') <= S (S n')
- auto with arith.
- No more goals.
- Qed.
Example
Induction scheme for tree_size.
We define trees by the following mutual inductive type:
- Axiom A : Set.
- A is declared
- Inductive tree : Set := node : A -> forest -> tree with forest : Set := | empty : 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
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.
- From Stdlib Require Import FunInd.
- Function tree_size (t:tree) : nat := match t with | node A f => S (forest_size f) end with forest_size (f:forest) : nat := match f with | empty => 0 | cons t f' => (tree_size t + forest_size f') end.
- tree_size is defined forest_size is defined tree_size, forest_size are recursively defined (guarded respectively on 1st, 1st arguments) tree_size_equation is defined tree_size_rect is defined tree_size_ind is defined tree_size_rec is defined forest_size_equation is defined forest_size_rect is defined forest_size_ind is defined forest_size_rec is defined R_tree_size_correct is defined R_forest_size_correct is defined R_tree_size_complete is defined R_forest_size_complete is defined
Notice that the induction principles tree_size_ind and forest_size_ind
generated by Function are not mutual.
- 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)
Mutual induction principles following the recursive structure of tree_size
and forest_size can be generated by the following command:
- Functional Scheme tree_size_ind2 := Induction for tree_size Sort Prop with forest_size_ind2 := Induction for forest_size Sort Prop.
- tree_size_ind2 is defined forest_size_ind2 is defined
You may now look at the type of tree_size_ind2:
- 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 f : forest, f = empty -> P0 empty 0) -> (forall (f : forest) (t : tree) (f' : forest), f = 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)
- Command Functional Case func_scheme_def
- Command Generate graph for qualid
- Internal debugging commands.